Documentation

Mathlib.RingTheory.Subsemiring.Basic

Bundled subsemirings #

We define bundled subsemirings and some standard constructions: CompleteLattice structure, Subtype and inclusion ring homomorphisms, subsemiring map, comap and range (rangeS) of a RingHom etc.

AddSubmonoidWithOneClass S R says S is a type of subsets s ≤ R that contain 0, 1, and are closed under (+)

    Instances
    theorem natCast_mem {S : Type u_1} {R : Type u_2} [AddMonoidWithOne R] [SetLike S R] (s : S) [AddSubmonoidWithOneClass S R] (n : ) :
    n s
    theorem ofNat_mem {S : Type u_1} {R : Type u_2} [AddMonoidWithOne R] [SetLike S R] [AddSubmonoidWithOneClass S R] (s : S) (n : ) [Nat.AtLeastTwo n] :

    SubsemiringClass S R states that S is a type of subsets s ⊆ R that are both a multiplicative and an additive submonoid.

      Instances
      theorem coe_nat_mem {R : Type u} {S : Type v} [NonAssocSemiring R] [SetLike S R] [hSR : SubsemiringClass S R] (s : S) (n : ) :
      n s
      instance SubsemiringClass.toNonAssocSemiring {R : Type u} {S : Type v} [NonAssocSemiring R] [SetLike S R] [hSR : SubsemiringClass S R] (s : S) :

      A subsemiring of a NonAssocSemiring inherits a NonAssocSemiring structure

      Equations
      instance SubsemiringClass.nontrivial {R : Type u} {S : Type v} [NonAssocSemiring R] [SetLike S R] [hSR : SubsemiringClass S R] (s : S) [Nontrivial R] :
      Equations
      • =
      instance SubsemiringClass.noZeroDivisors {R : Type u} {S : Type v} [NonAssocSemiring R] [SetLike S R] [hSR : SubsemiringClass S R] (s : S) [NoZeroDivisors R] :
      Equations
      • =
      def SubsemiringClass.subtype {R : Type u} {S : Type v} [NonAssocSemiring R] [SetLike S R] [hSR : SubsemiringClass S R] (s : S) :
      s →+* R

      The natural ring hom from a subsemiring of semiring R to R.

      Equations
      • One or more equations did not get rendered due to their size.
      @[simp]
      theorem SubsemiringClass.coe_subtype {R : Type u} {S : Type v} [NonAssocSemiring R] [SetLike S R] [hSR : SubsemiringClass S R] (s : S) :
      (SubsemiringClass.subtype s) = Subtype.val
      instance SubsemiringClass.toSemiring {S : Type v} (s : S) {R : Type u_1} [Semiring R] [SetLike S R] [SubsemiringClass S R] :

      A subsemiring of a Semiring is a Semiring.

      Equations
      @[simp]
      theorem SubsemiringClass.coe_pow {S : Type v} (s : S) {R : Type u_1} [Semiring R] [SetLike S R] [SubsemiringClass S R] (x : s) (n : ) :
      (x ^ n) = x ^ n
      instance SubsemiringClass.toCommSemiring {S : Type v} (s : S) {R : Type u_1} [CommSemiring R] [SetLike S R] [SubsemiringClass S R] :

      A subsemiring of a CommSemiring is a CommSemiring.

      Equations
      instance SubsemiringClass.instCharZero {R : Type u} {S : Type v} [NonAssocSemiring R] [SetLike S R] [hSR : SubsemiringClass S R] (s : S) [CharZero R] :
      Equations
      • =
      structure Subsemiring (R : Type u) [NonAssocSemiring R] extends Submonoid :

      A subsemiring of a semiring R is a subset s that is both a multiplicative and an additive submonoid.

      • carrier : Set R
      • mul_mem' : ∀ {a b : R}, a self.carrierb self.carriera * b self.carrier
      • one_mem' : 1 self.carrier
      • add_mem' : ∀ {a b : R}, a self.carrierb self.carriera + b self.carrier

        The sum of two elements of an additive subsemigroup belongs to the subsemigroup.

      • zero_mem' : 0 self.carrier

        An additive submonoid contains 0.

      Instances For
      @[reducible]

      Reinterpret a Subsemiring as an AddSubmonoid.

      Equations
      Equations
      • Subsemiring.instSetLikeSubsemiring = { coe := fun (s : Subsemiring R) => s.carrier, coe_injective' := }
      @[simp]
      theorem Subsemiring.mem_toSubmonoid {R : Type u} [NonAssocSemiring R] {s : Subsemiring R} {x : R} :
      x s.toSubmonoid x s
      theorem Subsemiring.mem_carrier {R : Type u} [NonAssocSemiring R] {s : Subsemiring R} {x : R} :
      x s.carrier x s
      theorem Subsemiring.ext {R : Type u} [NonAssocSemiring R] {S : Subsemiring R} {T : Subsemiring R} (h : ∀ (x : R), x S x T) :
      S = T

      Two subsemirings are equal if they have the same elements.

      def Subsemiring.copy {R : Type u} [NonAssocSemiring R] (S : Subsemiring R) (s : Set R) (hs : s = S) :

      Copy of a subsemiring with a new carrier equal to the old one. Useful to fix definitional equalities.

      Equations
      • One or more equations did not get rendered due to their size.
      @[simp]
      theorem Subsemiring.coe_copy {R : Type u} [NonAssocSemiring R] (S : Subsemiring R) (s : Set R) (hs : s = S) :
      (Subsemiring.copy S s hs) = s
      theorem Subsemiring.copy_eq {R : Type u} [NonAssocSemiring R] (S : Subsemiring R) (s : Set R) (hs : s = S) :
      theorem Subsemiring.toSubmonoid_strictMono {R : Type u} [NonAssocSemiring R] :
      StrictMono Subsemiring.toSubmonoid
      theorem Subsemiring.toSubmonoid_mono {R : Type u} [NonAssocSemiring R] :
      Monotone Subsemiring.toSubmonoid
      theorem Subsemiring.toAddSubmonoid_strictMono {R : Type u} [NonAssocSemiring R] :
      StrictMono Subsemiring.toAddSubmonoid
      theorem Subsemiring.toAddSubmonoid_mono {R : Type u} [NonAssocSemiring R] :
      Monotone Subsemiring.toAddSubmonoid
      def Subsemiring.mk' {R : Type u} [NonAssocSemiring R] (s : Set R) (sm : Submonoid R) (hm : sm = s) (sa : AddSubmonoid R) (ha : sa = s) :

      Construct a Subsemiring R from a set s, a submonoid sm, and an additive submonoid sa such that x ∈ s ↔ x ∈ sm ↔ x ∈ sa.

      Equations
      • Subsemiring.mk' s sm hm sa ha = { toSubmonoid := { toSubsemigroup := { carrier := s, mul_mem' := }, one_mem' := }, add_mem' := , zero_mem' := }
      @[simp]
      theorem Subsemiring.coe_mk' {R : Type u} [NonAssocSemiring R] {s : Set R} {sm : Submonoid R} (hm : sm = s) {sa : AddSubmonoid R} (ha : sa = s) :
      (Subsemiring.mk' s sm hm sa ha) = s
      @[simp]
      theorem Subsemiring.mem_mk' {R : Type u} [NonAssocSemiring R] {s : Set R} {sm : Submonoid R} (hm : sm = s) {sa : AddSubmonoid R} (ha : sa = s) {x : R} :
      x Subsemiring.mk' s sm hm sa ha x s
      @[simp]
      theorem Subsemiring.mk'_toSubmonoid {R : Type u} [NonAssocSemiring R] {s : Set R} {sm : Submonoid R} (hm : sm = s) {sa : AddSubmonoid R} (ha : sa = s) :
      (Subsemiring.mk' s sm hm sa ha).toSubmonoid = sm
      @[simp]
      theorem Subsemiring.mk'_toAddSubmonoid {R : Type u} [NonAssocSemiring R] {s : Set R} {sm : Submonoid R} (hm : sm = s) {sa : AddSubmonoid R} (ha : sa = s) :
      theorem Subsemiring.one_mem {R : Type u} [NonAssocSemiring R] (s : Subsemiring R) :
      1 s

      A subsemiring contains the semiring's 1.

      theorem Subsemiring.zero_mem {R : Type u} [NonAssocSemiring R] (s : Subsemiring R) :
      0 s

      A subsemiring contains the semiring's 0.

      theorem Subsemiring.mul_mem {R : Type u} [NonAssocSemiring R] (s : Subsemiring R) {x : R} {y : R} :
      x sy sx * y s

      A subsemiring is closed under multiplication.

      theorem Subsemiring.add_mem {R : Type u} [NonAssocSemiring R] (s : Subsemiring R) {x : R} {y : R} :
      x sy sx + y s

      A subsemiring is closed under addition.

      theorem Subsemiring.list_prod_mem {R : Type u_1} [Semiring R] (s : Subsemiring R) {l : List R} :
      (xl, x s)List.prod l s

      Product of a list of elements in a Subsemiring is in the Subsemiring.

      theorem Subsemiring.list_sum_mem {R : Type u} [NonAssocSemiring R] (s : Subsemiring R) {l : List R} :
      (xl, x s)List.sum l s

      Sum of a list of elements in a Subsemiring is in the Subsemiring.

      theorem Subsemiring.multiset_prod_mem {R : Type u_1} [CommSemiring R] (s : Subsemiring R) (m : Multiset R) :
      (am, a s)Multiset.prod m s

      Product of a multiset of elements in a Subsemiring of a CommSemiring is in the Subsemiring.

      theorem Subsemiring.multiset_sum_mem {R : Type u} [NonAssocSemiring R] (s : Subsemiring R) (m : Multiset R) :
      (am, a s)Multiset.sum m s

      Sum of a multiset of elements in a Subsemiring of a Semiring is in the add_subsemiring.

      theorem Subsemiring.prod_mem {R : Type u_1} [CommSemiring R] (s : Subsemiring R) {ι : Type u_2} {t : Finset ι} {f : ιR} (h : ct, f c s) :
      (Finset.prod t fun (i : ι) => f i) s

      Product of elements of a subsemiring of a CommSemiring indexed by a Finset is in the subsemiring.

      theorem Subsemiring.sum_mem {R : Type u} [NonAssocSemiring R] (s : Subsemiring R) {ι : Type u_1} {t : Finset ι} {f : ιR} (h : ct, f c s) :
      (Finset.sum t fun (i : ι) => f i) s

      Sum of elements in a Subsemiring of a Semiring indexed by a Finset is in the add_subsemiring.

      @[simp]
      theorem Subsemiring.coe_one {R : Type u} [NonAssocSemiring R] (s : Subsemiring R) :
      1 = 1
      @[simp]
      theorem Subsemiring.coe_zero {R : Type u} [NonAssocSemiring R] (s : Subsemiring R) :
      0 = 0
      @[simp]
      theorem Subsemiring.coe_add {R : Type u} [NonAssocSemiring R] (s : Subsemiring R) (x : s) (y : s) :
      (x + y) = x + y
      @[simp]
      theorem Subsemiring.coe_mul {R : Type u} [NonAssocSemiring R] (s : Subsemiring R) (x : s) (y : s) :
      (x * y) = x * y
      Equations
      • =
      theorem Subsemiring.pow_mem {R : Type u_1} [Semiring R] (s : Subsemiring R) {x : R} (hx : x s) (n : ) :
      x ^ n s
      instance Subsemiring.toSemiring {R : Type u_1} [Semiring R] (s : Subsemiring R) :

      A subsemiring of a Semiring is a Semiring.

      Equations
      @[simp]
      theorem Subsemiring.coe_pow {R : Type u_1} [Semiring R] (s : Subsemiring R) (x : s) (n : ) :
      (x ^ n) = x ^ n

      The natural ring hom from a subsemiring of semiring R to R.

      Equations
      • One or more equations did not get rendered due to their size.
      @[simp]
      theorem Subsemiring.coe_subtype {R : Type u} [NonAssocSemiring R] (s : Subsemiring R) :
      (Subsemiring.subtype s) = Subtype.val
      theorem Subsemiring.nsmul_mem {R : Type u} [NonAssocSemiring R] (s : Subsemiring R) {x : R} (hx : x s) (n : ) :
      n x s
      @[simp]
      theorem Subsemiring.coe_toSubmonoid {R : Type u} [NonAssocSemiring R] (s : Subsemiring R) :
      s.toSubmonoid = s
      @[simp]
      theorem Subsemiring.coe_carrier_toSubmonoid {R : Type u} [NonAssocSemiring R] (s : Subsemiring R) :
      s.carrier = s

      The subsemiring R of the semiring R.

      Equations
      • Subsemiring.instTopSubsemiring = { top := let __src := ; let __src_1 := ; { toSubmonoid := __src, add_mem' := , zero_mem' := } }
      @[simp]
      theorem Subsemiring.mem_top {R : Type u} [NonAssocSemiring R] (x : R) :
      @[simp]
      theorem Subsemiring.coe_top {R : Type u} [NonAssocSemiring R] :
      = Set.univ
      @[simp]
      theorem Subsemiring.topEquiv_symm_apply_coe {R : Type u} [NonAssocSemiring R] (r : R) :
      ((RingEquiv.symm Subsemiring.topEquiv) r) = r
      @[simp]
      theorem Subsemiring.topEquiv_apply {R : Type u} [NonAssocSemiring R] (r : ) :
      Subsemiring.topEquiv r = r

      The ring equiv between the top element of Subsemiring R and R.

      Equations
      • One or more equations did not get rendered due to their size.
      def Subsemiring.comap {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (s : Subsemiring S) :

      The preimage of a subsemiring along a ring homomorphism is a subsemiring.

      Equations
      • One or more equations did not get rendered due to their size.
      @[simp]
      theorem Subsemiring.coe_comap {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (s : Subsemiring S) (f : R →+* S) :
      (Subsemiring.comap f s) = f ⁻¹' s
      @[simp]
      theorem Subsemiring.mem_comap {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {s : Subsemiring S} {f : R →+* S} {x : R} :
      def Subsemiring.map {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (s : Subsemiring R) :

      The image of a subsemiring along a ring homomorphism is a subsemiring.

      Equations
      • One or more equations did not get rendered due to their size.
      @[simp]
      theorem Subsemiring.coe_map {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (s : Subsemiring R) :
      (Subsemiring.map f s) = f '' s
      @[simp]
      theorem Subsemiring.mem_map {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {f : R →+* S} {s : Subsemiring R} {y : S} :
      y Subsemiring.map f s ∃ x ∈ s, f x = y
      noncomputable def Subsemiring.equivMapOfInjective {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (s : Subsemiring R) (f : R →+* S) (hf : Function.Injective f) :
      s ≃+* (Subsemiring.map f s)

      A subsemiring is isomorphic to its image under an injective function

      Equations
      @[simp]
      theorem Subsemiring.coe_equivMapOfInjective_apply {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (s : Subsemiring R) (f : R →+* S) (hf : Function.Injective f) (x : s) :
      ((Subsemiring.equivMapOfInjective s f hf) x) = f x
      def RingHom.rangeS {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) :

      The range of a ring homomorphism is a subsemiring. See Note [range copy pattern].

      Equations
      @[simp]
      theorem RingHom.coe_rangeS {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) :
      @[simp]
      theorem RingHom.mem_rangeS {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {f : R →+* S} {y : S} :
      y RingHom.rangeS f ∃ (x : R), f x = y
      theorem RingHom.mem_rangeS_self {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (x : R) :

      The range of a morphism of semirings is a fintype, if the domain is a fintype. Note: this instance can form a diamond with Subtype.fintype in the presence of Fintype S.

      Equations
      Equations
      Equations
      • Subsemiring.instInhabitedSubsemiring = { default := }
      theorem Subsemiring.mem_bot {R : Type u} [NonAssocSemiring R] {x : R} :
      x ∃ (n : ), n = x

      The inf of two subsemirings is their intersection.

      Equations
      • One or more equations did not get rendered due to their size.
      @[simp]
      theorem Subsemiring.coe_inf {R : Type u} [NonAssocSemiring R] (p : Subsemiring R) (p' : Subsemiring R) :
      (p p') = p p'
      @[simp]
      theorem Subsemiring.mem_inf {R : Type u} [NonAssocSemiring R] {p : Subsemiring R} {p' : Subsemiring R} {x : R} :
      x p p' x p x p'
      Equations
      @[simp]
      theorem Subsemiring.coe_sInf {R : Type u} [NonAssocSemiring R] (S : Set (Subsemiring R)) :
      (sInf S) = ⋂ s ∈ S, s
      theorem Subsemiring.mem_sInf {R : Type u} [NonAssocSemiring R] {S : Set (Subsemiring R)} {x : R} :
      x sInf S pS, x p
      @[simp]
      theorem Subsemiring.sInf_toSubmonoid {R : Type u} [NonAssocSemiring R] (s : Set (Subsemiring R)) :
      (sInf s).toSubmonoid = ⨅ t ∈ s, t.toSubmonoid

      Subsemirings of a semiring form a complete lattice.

      Equations
      theorem Subsemiring.eq_top_iff' {R : Type u} [NonAssocSemiring R] (A : Subsemiring R) :
      A = ∀ (x : R), x A

      The center of a non-associative semiring R is the set of elements that commute and associate with everything in R

      Equations
      • One or more equations did not get rendered due to their size.
      @[inline, reducible]

      The center is commutative and associative.

      This is not an instance as it forms a non-defeq diamond with NonUnitalSubringClass.tNonUnitalring in the npow field.

      Equations

      The center is commutative.

      Equations
      theorem Subsemiring.mem_center_iff {R : Type u_1} [Semiring R] {z : R} :
      z Subsemiring.center R ∀ (g : R), g * z = z * g
      Equations
      def Subsemiring.centralizer {R : Type u_1} [Semiring R] (s : Set R) :

      The centralizer of a set as subsemiring.

      Equations
      • One or more equations did not get rendered due to their size.
      theorem Subsemiring.mem_centralizer_iff {R : Type u_1} [Semiring R] {s : Set R} {z : R} :
      z Subsemiring.centralizer s gs, g * z = z * g

      The Subsemiring generated by a set.

      Equations
      theorem Subsemiring.mem_closure {R : Type u} [NonAssocSemiring R] {x : R} {s : Set R} :
      x Subsemiring.closure s ∀ (S : Subsemiring R), s Sx S
      @[simp]

      The subsemiring generated by a set includes the set.

      theorem Subsemiring.not_mem_of_not_mem_closure {R : Type u} [NonAssocSemiring R] {s : Set R} {P : R} (hP : PSubsemiring.closure s) :
      Ps
      @[simp]
      theorem Subsemiring.closure_le {R : Type u} [NonAssocSemiring R] {s : Set R} {t : Subsemiring R} :

      A subsemiring S includes closure s if and only if it includes s.

      theorem Subsemiring.closure_mono {R : Type u} [NonAssocSemiring R] ⦃s : Set R ⦃t : Set R (h : s t) :

      Subsemiring closure of a set is monotone in its argument: if s ⊆ t, then closure s ≤ closure t.

      theorem Subsemiring.closure_eq_of_le {R : Type u} [NonAssocSemiring R] {s : Set R} {t : Subsemiring R} (h₁ : s t) (h₂ : t Subsemiring.closure s) :
      theorem Subsemiring.mem_map_equiv {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {f : R ≃+* S} {K : Subsemiring R} {x : S} :

      The additive closure of a submonoid is a subsemiring.

      Equations
      • One or more equations did not get rendered due to their size.

      The Subsemiring generated by a multiplicative submonoid coincides with the Subsemiring.closure of the submonoid itself .

      The elements of the subsemiring closure of M are exactly the elements of the additive closure of a multiplicative submonoid M.

      theorem Subsemiring.closure_induction {R : Type u} [NonAssocSemiring R] {s : Set R} {p : RProp} {x : R} (h : x Subsemiring.closure s) (mem : xs, p x) (zero : p 0) (one : p 1) (add : ∀ (x y : R), p xp yp (x + y)) (mul : ∀ (x y : R), p xp yp (x * y)) :
      p x

      An induction principle for closure membership. If p holds for 0, 1, and all elements of s, and is preserved under addition and multiplication, then p holds for all elements of the closure of s.

      theorem Subsemiring.closure_induction' {R : Type u} [NonAssocSemiring R] {s : Set R} {p : (x : R) → x Subsemiring.closure sProp} (mem : ∀ (x : R) (h : x s), p x ) (zero : p 0 ) (one : p 1 ) (add : ∀ (x : R) (hx : x Subsemiring.closure s) (y : R) (hy : y Subsemiring.closure s), p x hxp y hyp (x + y) ) (mul : ∀ (x : R) (hx : x Subsemiring.closure s) (y : R) (hy : y Subsemiring.closure s), p x hxp y hyp (x * y) ) {a : R} (ha : a Subsemiring.closure s) :
      p a ha
      theorem Subsemiring.closure_induction₂ {R : Type u} [NonAssocSemiring R] {s : Set R} {p : RRProp} {x : R} {y : R} (hx : x Subsemiring.closure s) (hy : y Subsemiring.closure s) (Hs : xs, ys, p x y) (H0_left : ∀ (x : R), p 0 x) (H0_right : ∀ (x : R), p x 0) (H1_left : ∀ (x : R), p 1 x) (H1_right : ∀ (x : R), p x 1) (Hadd_left : ∀ (x₁ x₂ y : R), p x₁ yp x₂ yp (x₁ + x₂) y) (Hadd_right : ∀ (x y₁ y₂ : R), p x y₁p x y₂p x (y₁ + y₂)) (Hmul_left : ∀ (x₁ x₂ y : R), p x₁ yp x₂ yp (x₁ * x₂) y) (Hmul_right : ∀ (x y₁ y₂ : R), p x y₁p x y₂p x (y₁ * y₂)) :
      p x y

      An induction principle for closure membership for predicates with two arguments.

      theorem Subsemiring.mem_closure_iff_exists_list {R : Type u_1} [Semiring R] {s : Set R} {x : R} :
      x Subsemiring.closure s ∃ (L : List (List R)), (tL, yt, y s) List.sum (List.map List.prod L) = x
      def Subsemiring.gi (R : Type u) [NonAssocSemiring R] :
      GaloisInsertion Subsemiring.closure SetLike.coe

      closure forms a Galois insertion with the coercion to set.

      Equations

      Closure of a subsemiring S equals S.

      theorem Subsemiring.closure_iUnion {R : Type u} [NonAssocSemiring R] {ι : Sort u_1} (s : ιSet R) :
      Subsemiring.closure (⋃ (i : ι), s i) = ⨆ (i : ι), Subsemiring.closure (s i)
      theorem Subsemiring.map_iSup {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {ι : Sort u_1} (f : R →+* S) (s : ιSubsemiring R) :
      Subsemiring.map f (iSup s) = ⨆ (i : ι), Subsemiring.map f (s i)
      theorem Subsemiring.comap_iInf {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {ι : Sort u_1} (f : R →+* S) (s : ιSubsemiring S) :
      Subsemiring.comap f (iInf s) = ⨅ (i : ι), Subsemiring.comap f (s i)
      @[simp]

      Given Subsemirings s, t of semirings R, S respectively, s.prod t is s × t as a subsemiring of R × S.

      Equations
      • One or more equations did not get rendered due to their size.
      theorem Subsemiring.coe_prod {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (s : Subsemiring R) (t : Subsemiring S) :
      (Subsemiring.prod s t) = s ×ˢ t
      theorem Subsemiring.mem_prod {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {s : Subsemiring R} {t : Subsemiring S} {p : R × S} :
      p Subsemiring.prod s t p.1 s p.2 t
      theorem Subsemiring.prod_mono {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] ⦃s₁ : Subsemiring R ⦃s₂ : Subsemiring R (hs : s₁ s₂) ⦃t₁ : Subsemiring S ⦃t₂ : Subsemiring S (ht : t₁ t₂) :
      def Subsemiring.prodEquiv {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (s : Subsemiring R) (t : Subsemiring S) :
      (Subsemiring.prod s t) ≃+* s × t

      Product of subsemirings is isomorphic to their product as monoids.

      Equations
      theorem Subsemiring.mem_iSup_of_directed {R : Type u} [NonAssocSemiring R] {ι : Sort u_1} [hι : Nonempty ι] {S : ιSubsemiring R} (hS : Directed (fun (x x_1 : Subsemiring R) => x x_1) S) {x : R} :
      x ⨆ (i : ι), S i ∃ (i : ι), x S i
      theorem Subsemiring.coe_iSup_of_directed {R : Type u} [NonAssocSemiring R] {ι : Sort u_1} [hι : Nonempty ι] {S : ιSubsemiring R} (hS : Directed (fun (x x_1 : Subsemiring R) => x x_1) S) :
      (⨆ (i : ι), S i) = ⋃ (i : ι), (S i)
      theorem Subsemiring.mem_sSup_of_directedOn {R : Type u} [NonAssocSemiring R] {S : Set (Subsemiring R)} (Sne : Set.Nonempty S) (hS : DirectedOn (fun (x x_1 : Subsemiring R) => x x_1) S) {x : R} :
      x sSup S ∃ s ∈ S, x s
      theorem Subsemiring.coe_sSup_of_directedOn {R : Type u} [NonAssocSemiring R] {S : Set (Subsemiring R)} (Sne : Set.Nonempty S) (hS : DirectedOn (fun (x x_1 : Subsemiring R) => x x_1) S) :
      (sSup S) = ⋃ s ∈ S, s
      def RingHom.domRestrict {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {σR : Type u_1} [SetLike σR R] [SubsemiringClass σR R] (f : R →+* S) (s : σR) :
      s →+* S

      Restriction of a ring homomorphism to a subsemiring of the domain.

      Equations
      @[simp]
      theorem RingHom.restrict_apply {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {σR : Type u_1} [SetLike σR R] [SubsemiringClass σR R] (f : R →+* S) {s : σR} (x : s) :
      (RingHom.domRestrict f s) x = f x
      def RingHom.codRestrict {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {σS : Type u_2} [SetLike σS S] [SubsemiringClass σS S] (f : R →+* S) (s : σS) (h : ∀ (x : R), f x s) :
      R →+* s

      Restriction of a ring homomorphism to a subsemiring of the codomain.

      Equations
      • One or more equations did not get rendered due to their size.
      def RingHom.restrict {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {σR : Type u_1} {σS : Type u_2} [SetLike σR R] [SetLike σS S] [SubsemiringClass σR R] [SubsemiringClass σS S] (f : R →+* S) (s' : σR) (s : σS) (h : xs', f x s) :
      s' →+* s

      The ring homomorphism from the preimage of s to s.

      Equations
      @[simp]
      theorem RingHom.coe_restrict_apply {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {σR : Type u_1} {σS : Type u_2} [SetLike σR R] [SetLike σS S] [SubsemiringClass σR R] [SubsemiringClass σS S] (f : R →+* S) (s' : σR) (s : σS) (h : xs', f x s) (x : s') :
      ((RingHom.restrict f s' s h) x) = f x
      @[simp]
      theorem RingHom.comp_restrict {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {σR : Type u_1} {σS : Type u_2} [SetLike σR R] [SetLike σS S] [SubsemiringClass σR R] [SubsemiringClass σS S] (f : R →+* S) (s' : σR) (s : σS) (h : xs', f x s) :

      Restriction of a ring homomorphism to its range interpreted as a subsemiring.

      This is the bundled version of Set.rangeFactorization.

      Equations
      @[simp]
      theorem RingHom.coe_rangeSRestrict {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (x : R) :
      ((RingHom.rangeSRestrict f) x) = f x
      @[simp]

      The range of a surjective ring homomorphism is the whole of the codomain.

      def RingHom.eqLocusS {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (g : R →+* S) :

      The subsemiring of elements x : R such that f x = g x

      Equations
      • One or more equations did not get rendered due to their size.
      @[simp]
      theorem RingHom.eqOn_sclosure {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {f : R →+* S} {g : R →+* S} {s : Set R} (h : Set.EqOn (f) (g) s) :
      Set.EqOn f g (Subsemiring.closure s)

      If two ring homomorphisms are equal on a set, then they are equal on its subsemiring closure.

      theorem RingHom.eq_of_eqOn_stop {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {f : R →+* S} {g : R →+* S} (h : Set.EqOn f g ) :
      f = g
      theorem RingHom.eq_of_eqOn_sdense {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {s : Set R} (hs : Subsemiring.closure s = ) {f : R →+* S} {g : R →+* S} (h : Set.EqOn (f) (g) s) :
      f = g

      The image under a ring homomorphism of the subsemiring generated by a set equals the subsemiring generated by the image of the set.

      def Subsemiring.inclusion {R : Type u} [NonAssocSemiring R] {S : Subsemiring R} {T : Subsemiring R} (h : S T) :
      S →+* T

      The ring homomorphism associated to an inclusion of subsemirings.

      Equations
      def RingEquiv.subsemiringCongr {R : Type u} [NonAssocSemiring R] {s : Subsemiring R} {t : Subsemiring R} (h : s = t) :
      s ≃+* t

      Makes the identity isomorphism from a proof two subsemirings of a multiplicative monoid are equal.

      Equations
      def RingEquiv.ofLeftInverseS {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {g : SR} {f : R →+* S} (h : Function.LeftInverse g f) :

      Restrict a ring homomorphism with a left inverse to a ring isomorphism to its RingHom.rangeS.

      Equations
      • One or more equations did not get rendered due to their size.
      @[simp]
      theorem RingEquiv.ofLeftInverseS_apply {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {g : SR} {f : R →+* S} (h : Function.LeftInverse g f) (x : R) :
      @[simp]
      theorem RingEquiv.ofLeftInverseS_symm_apply {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {g : SR} {f : R →+* S} (h : Function.LeftInverse g f) (x : (RingHom.rangeS f)) :
      @[simp]
      theorem RingEquiv.subsemiringMap_apply_coe {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (e : R ≃+* S) (s : Subsemiring R) (x : (Subsemiring.toAddSubmonoid s)) :
      ((RingEquiv.subsemiringMap e s) x) = e x

      Given an equivalence e : R ≃+* S of semirings and a subsemiring s of R, subsemiring_map e s is the induced equivalence between s and s.map e

      Equations
      • One or more equations did not get rendered due to their size.

      Actions by Subsemirings #

      These are just copies of the definitions about Submonoid starting from submonoid.mul_action. The only new result is subsemiring.module.

      When R is commutative, Algebra.ofSubsemiring provides a stronger result than those found in this file, which uses the same scalar action.

      instance Subsemiring.smul {R' : Type u_1} {α : Type u_2} [NonAssocSemiring R'] [SMul R' α] (S : Subsemiring R') :
      SMul (S) α

      The action by a subsemiring is the action by the underlying semiring.

      Equations
      theorem Subsemiring.smul_def {R' : Type u_1} {α : Type u_2} [NonAssocSemiring R'] [SMul R' α] {S : Subsemiring R'} (g : S) (m : α) :
      g m = g m
      instance Subsemiring.smulCommClass_left {R' : Type u_1} {α : Type u_2} {β : Type u_3} [NonAssocSemiring R'] [SMul R' β] [SMul α β] [SMulCommClass R' α β] (S : Subsemiring R') :
      SMulCommClass (S) α β
      Equations
      • =
      instance Subsemiring.smulCommClass_right {R' : Type u_1} {α : Type u_2} {β : Type u_3} [NonAssocSemiring R'] [SMul α β] [SMul R' β] [SMulCommClass α R' β] (S : Subsemiring R') :
      SMulCommClass α (S) β
      Equations
      • =
      instance Subsemiring.isScalarTower {R' : Type u_1} {α : Type u_2} {β : Type u_3} [NonAssocSemiring R'] [SMul α β] [SMul R' α] [SMul R' β] [IsScalarTower R' α β] (S : Subsemiring R') :
      IsScalarTower (S) α β

      Note that this provides IsScalarTower S R R which is needed by smul_mul_assoc.

      Equations
      • =
      instance Subsemiring.faithfulSMul {R' : Type u_1} {α : Type u_2} [NonAssocSemiring R'] [SMul R' α] [FaithfulSMul R' α] (S : Subsemiring R') :
      FaithfulSMul (S) α
      Equations
      • =
      instance Subsemiring.mulAction {R' : Type u_1} {α : Type u_2} [Semiring R'] [MulAction R' α] (S : Subsemiring R') :
      MulAction (S) α

      The action by a subsemiring is the action by the underlying semiring.

      Equations
      instance Subsemiring.distribMulAction {R' : Type u_1} {α : Type u_2} [Semiring R'] [AddMonoid α] [DistribMulAction R' α] (S : Subsemiring R') :

      The action by a subsemiring is the action by the underlying semiring.

      Equations
      instance Subsemiring.mulDistribMulAction {R' : Type u_1} {α : Type u_2} [Semiring R'] [Monoid α] [MulDistribMulAction R' α] (S : Subsemiring R') :

      The action by a subsemiring is the action by the underlying semiring.

      Equations
      instance Subsemiring.mulActionWithZero {R' : Type u_1} {α : Type u_2} [Semiring R'] [Zero α] [MulActionWithZero R' α] (S : Subsemiring R') :

      The action by a subsemiring is the action by the underlying semiring.

      Equations
      instance Subsemiring.module {R' : Type u_1} {α : Type u_2} [Semiring R'] [AddCommMonoid α] [Module R' α] (S : Subsemiring R') :
      Module (S) α

      The action by a subsemiring is the action by the underlying semiring.

      Equations

      The action by a subsemiring is the action by the underlying semiring.

      Equations
      • One or more equations did not get rendered due to their size.

      The center of a semiring acts commutatively on that semiring.

      Equations
      • =

      The center of a semiring acts commutatively on that semiring.

      Equations
      • =
      def Subsemiring.closureCommSemiringOfComm {R' : Type u_1} [Semiring R'] {s : Set R'} (hcomm : as, bs, a * b = b * a) :

      If all the elements of a set s commute, then closure s is a commutative monoid.

      Equations