Documentation

Mathlib.Algebra.Ring.Subsemiring.Basic

Bundled subsemirings #

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

instance SubsemiringClass.instCharZero {R : Type u} {S : Type v} [NonAssocSemiring R] [SetLike S R] [hSR : SubsemiringClass S R] (s : S) [CharZero R] :
theorem Subsemiring.list_prod_mem {R : Type u_1} [Semiring R] (s : Subsemiring R) {l : List R} :
(∀ xl, x s)l.prod 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)l.sum 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)m.prod 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)m.sum s

Sum of a multiset of elements in a Subsemiring of a NonAssocSemiring is in the 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) :
it, 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) :
it, f i s

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

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

Equations
  • Subsemiring.topEquiv = { toFun := fun (r : ) => r, invFun := fun (r : R) => r, , left_inv := , right_inv := , map_mul' := , map_add' := }
@[simp]
@[simp]
theorem Subsemiring.topEquiv_apply {R : Type u} [NonAssocSemiring R] (r : ) :
topEquiv r = r
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
  • Subsemiring.comap f s = { carrier := f ⁻¹' s, mul_mem' := , one_mem' := , add_mem' := , zero_mem' := }
@[simp]
theorem Subsemiring.coe_comap {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (s : Subsemiring S) :
(comap f s) = f ⁻¹' s
@[simp]
theorem Subsemiring.comap_toSubmonoid {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (s : Subsemiring S) :
(comap f s).toSubmonoid = { carrier := f ⁻¹' s, mul_mem' := , one_mem' := }
@[simp]
theorem Subsemiring.mem_comap {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {s : Subsemiring S} {f : R →+* S} {x : R} :
x comap f s f x s
theorem Subsemiring.comap_comap {R : Type u} {S : Type v} {T : Type w} [NonAssocSemiring R] [NonAssocSemiring S] [NonAssocSemiring T] (s : Subsemiring T) (g : S →+* T) (f : R →+* S) :
comap f (comap g s) = comap (g.comp f) s
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
  • Subsemiring.map f s = { carrier := f '' s, mul_mem' := , one_mem' := , add_mem' := , zero_mem' := }
@[simp]
theorem Subsemiring.map_toSubmonoid {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (s : Subsemiring R) :
(map f s).toSubmonoid = { carrier := f '' s, mul_mem' := , one_mem' := }
@[simp]
theorem Subsemiring.coe_map {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (s : Subsemiring R) :
(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 map f s xs, f x = y
@[simp]
theorem Subsemiring.map_id {R : Type u} [NonAssocSemiring R] (s : Subsemiring R) :
map (RingHom.id R) s = s
theorem Subsemiring.map_map {R : Type u} {S : Type v} {T : Type w} [NonAssocSemiring R] [NonAssocSemiring S] [NonAssocSemiring T] (s : Subsemiring R) (g : S →+* T) (f : R →+* S) :
map g (map f s) = map (g.comp f) s
theorem Subsemiring.map_le_iff_le_comap {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {f : R →+* S} {s : Subsemiring R} {t : Subsemiring S} :
map f s t s comap f t
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 ≃+* (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) :
((s.equivMapOfInjective 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) :
f.rangeS = Set.range f
@[simp]
theorem RingHom.rangeS_toSubmonoid {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) :
f.rangeS.toSubmonoid = { carrier := Set.range f, mul_mem' := , one_mem' := }
@[simp]
theorem RingHom.mem_rangeS {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {f : R →+* S} {y : S} :
y f.rangeS ∃ (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) :
f x f.rangeS
theorem RingHom.map_rangeS {R : Type u} {S : Type v} {T : Type w} [NonAssocSemiring R] [NonAssocSemiring S] [NonAssocSemiring T] (g : S →+* T) (f : R →+* S) :
instance RingHom.fintypeRangeS {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] [Fintype R] [DecidableEq S] (f : R →+* S) :

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
theorem Subsemiring.mem_bot {R : Type u} [NonAssocSemiring R] {x : R} :
x ∃ (n : ), n = x
Equations
@[simp]
theorem Subsemiring.coe_sInf {R : Type u} [NonAssocSemiring R] (S : Set (Subsemiring R)) :
(sInf S) = sS, 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.coe_iInf {R : Type u} [NonAssocSemiring R] {ι : Sort u_1} {S : ιSubsemiring R} :
(⨅ (i : ι), S i) = ⋂ (i : ι), (S i)
theorem Subsemiring.mem_iInf {R : Type u} [NonAssocSemiring R] {ι : Sort u_1} {S : ιSubsemiring R} {x : R} :
x ⨅ (i : ι), S i ∀ (i : ι), x S i
@[simp]

Subsemirings of a semiring form a complete lattice.

Equations
  • One or more equations did not get rendered due to their size.
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
@[simp]
theorem Subsemiring.center_toSubmonoid (R : Type u) [NonAssocSemiring R] :
(center R).toSubmonoid = { carrier := (NonUnitalSubsemiring.center R).carrier, mul_mem' := , one_mem' := }
@[reducible, inline]

The center is commutative and associative.

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

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

The center of isomorphic (not necessarily associative) semirings are isomorphic.

Equations
@[simp]
theorem Subsemiring.centerCongr_apply_coe {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (e : R ≃+* S) (r : (Subsemigroup.center R)) :
((centerCongr e) r) = e r
@[simp]
theorem Subsemiring.centerCongr_symm_apply_coe {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (e : R ≃+* S) (s : (Subsemigroup.center S)) :
((centerCongr e).symm s) = (↑e).symm s

The center of a (not necessarily associative) semiring is isomorphic to the center of its opposite.

Equations

The center is commutative.

Equations
  • One or more equations did not get rendered due to their size.
theorem Subsemiring.mem_center_iff {R : Type u_1} [Semiring R] {z : R} :
z center R ∀ (g : R), g * z = z * g
instance Subsemiring.decidableMemCenter {R : Type u_1} [Semiring R] [DecidableEq R] [Fintype R] :
DecidablePred fun (x : R) => x center R
Equations
@[simp]
def Subsemiring.centralizer {R : Type u_1} [Semiring R] (s : Set R) :

The centralizer of a set as subsemiring.

Equations
@[simp]
theorem Subsemiring.coe_centralizer {R : Type u_1} [Semiring R] (s : Set R) :
theorem Subsemiring.mem_centralizer_iff {R : Type u_1} [Semiring R] {s : Set R} {z : R} :
z centralizer s gs, g * z = z * g
theorem Subsemiring.centralizer_le {R : Type u_1} [Semiring R] (s t : Set R) (h : s t) :

The Subsemiring generated by a set.

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

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 : Pclosure s) :
Ps
@[simp]
theorem Subsemiring.closure_le {R : Type u} [NonAssocSemiring R] {s : Set R} {t : Subsemiring R} :
closure s t s t

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

theorem Subsemiring.closure_mono {R : Type u} [NonAssocSemiring R] s 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 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} :
x map (↑f) K f.symm x K
theorem Subsemiring.map_equiv_eq_comap_symm {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R ≃+* S) (K : Subsemiring R) :
map (↑f) K = comap (↑f.symm) K
theorem Subsemiring.comap_equiv_eq_map_symm {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R ≃+* S) (K : Subsemiring S) :
comap (↑f) K = map (↑f.symm) K

The additive closure of a submonoid is a subsemiring.

Equations

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 : (x : R) → x closure sProp} (mem : ∀ (x : R) (hx : x s), p x ) (zero : p 0 ) (one : p 1 ) (add : ∀ (x y : R) (hx : x closure s) (hy : y closure s), p x hxp y hyp (x + y) ) (mul : ∀ (x y : R) (hx : x closure s) (hy : y closure s), p x hxp y hyp (x * y) ) {x : R} (hx : x closure s) :
p x hx

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 y : R) → x closure sy closure sProp} (mem_mem : ∀ (x y : R) (hx : x s) (hy : y s), p x y ) (zero_left : ∀ (x : R) (hx : x closure s), p 0 x hx) (zero_right : ∀ (x : R) (hx : x closure s), p x 0 hx ) (one_left : ∀ (x : R) (hx : x closure s), p 1 x hx) (one_right : ∀ (x : R) (hx : x closure s), p x 1 hx ) (add_left : ∀ (x y z : R) (hx : x closure s) (hy : y closure s) (hz : z closure s), p x z hx hzp y z hy hzp (x + y) z hz) (add_right : ∀ (x y z : R) (hx : x closure s) (hy : y closure s) (hz : z closure s), p x y hx hyp x z hx hzp x (y + z) hx ) (mul_left : ∀ (x y z : R) (hx : x closure s) (hy : y closure s) (hz : z closure s), p x z hx hzp y z hy hzp (x * y) z hz) (mul_right : ∀ (x y z : R) (hx : x closure s) (hy : y closure s) (hz : z closure s), p x y hx hyp x z hx hzp x (y * z) hx ) {x y : R} (hx : x closure s) (hy : y closure s) :
p x y hx hy

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 closure s ∃ (L : List (List R)), (∀ tL, yt, y s) (List.map List.prod L).sum = x

closure forms a Galois insertion with the coercion to set.

Equations
@[simp]
theorem Subsemiring.closure_eq {R : Type u} [NonAssocSemiring R] (s : Subsemiring R) :
closure s = s

Closure of a subsemiring S equals S.

theorem Subsemiring.closure_union {R : Type u} [NonAssocSemiring R] (s t : Set R) :
closure (s t) = closure sclosure t
theorem Subsemiring.closure_iUnion {R : Type u} [NonAssocSemiring R] {ι : Sort u_1} (s : ιSet R) :
closure (⋃ (i : ι), s i) = ⨆ (i : ι), closure (s i)
theorem Subsemiring.closure_sUnion {R : Type u} [NonAssocSemiring R] (s : Set (Set R)) :
closure (⋃₀ s) = ts, closure t
@[simp]
theorem Subsemiring.closure_insert_natCast {R : Type u} [NonAssocSemiring R] (n : ) (s : Set R) :
closure (insert (↑n) s) = closure s
theorem Subsemiring.map_sup {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (s t : Subsemiring R) (f : R →+* S) :
map f (st) = map f smap f t
theorem Subsemiring.map_iSup {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {ι : Sort u_1} (f : R →+* S) (s : ιSubsemiring R) :
map f (iSup s) = ⨆ (i : ι), map f (s i)
theorem Subsemiring.map_inf {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (s t : Subsemiring R) (f : R →+* S) (hf : Function.Injective f) :
map f (st) = map f smap f t
theorem Subsemiring.map_iInf {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {ι : Sort u_1} [Nonempty ι] (f : R →+* S) (hf : Function.Injective f) (s : ιSubsemiring R) :
map f (iInf s) = ⨅ (i : ι), map f (s i)
theorem Subsemiring.comap_inf {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (s t : Subsemiring S) (f : R →+* S) :
comap f (st) = comap f scomap f t
theorem Subsemiring.comap_iInf {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {ι : Sort u_1} (f : R →+* S) (s : ιSubsemiring S) :
comap f (iInf s) = ⨅ (i : ι), comap f (s i)
@[simp]
theorem Subsemiring.map_bot {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) :
@[simp]
theorem Subsemiring.comap_top {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) :

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

Equations
  • s.prod t = { carrier := s ×ˢ t, mul_mem' := , one_mem' := , add_mem' := , zero_mem' := }
theorem Subsemiring.coe_prod {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (s : Subsemiring R) (t : Subsemiring S) :
(s.prod 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 s.prod t p.1 s p.2 t
theorem Subsemiring.prod_mono {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] s₁ s₂ : Subsemiring R (hs : s₁ s₂) t₁ t₂ : Subsemiring S (ht : t₁ t₂) :
s₁.prod t₁ s₂.prod t₂
def Subsemiring.prodEquiv {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (s : Subsemiring R) (t : Subsemiring S) :
(s.prod 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} [ : Nonempty ι] {S : ιSubsemiring R} (hS : Directed (fun (x1 x2 : Subsemiring R) => x1 x2) 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} [ : Nonempty ι] {S : ιSubsemiring R} (hS : Directed (fun (x1 x2 : Subsemiring R) => x1 x2) S) :
(⨆ (i : ι), S i) = ⋃ (i : ι), (S i)
theorem Subsemiring.mem_sSup_of_directedOn {R : Type u} [NonAssocSemiring R] {S : Set (Subsemiring R)} (Sne : S.Nonempty) (hS : DirectedOn (fun (x1 x2 : Subsemiring R) => x1 x2) S) {x : R} :
x sSup S sS, x s
theorem Subsemiring.coe_sSup_of_directedOn {R : Type u} [NonAssocSemiring R] {S : Set (Subsemiring R)} (Sne : S.Nonempty) (hS : DirectedOn (fun (x1 x2 : Subsemiring R) => x1 x2) S) :
(sSup S) = sS, s
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
  • f.codRestrict s h = { toFun := fun (n : R) => f n, , map_one' := , map_mul' := , map_zero' := , map_add' := }
@[simp]
theorem RingHom.codRestrict_apply {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) (x : R) :
((f.codRestrict s h) x) = f x
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') :
((f.restrict 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) :
def RingHom.rangeSRestrict {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) :
R →+* f.rangeS

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) :
(f.rangeSRestrict x) = f x
@[simp]

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

theorem RingHom.eqOn_sclosure {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {f 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 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 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 T : Subsemiring R} (h : S T) :
S →+* T

The ring homomorphism associated to an inclusion of subsemirings.

Equations
@[simp]
def RingEquiv.subsemiringCongr {R : Type u} [NonAssocSemiring R] {s 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) :
R ≃+* f.rangeS

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) :
((ofLeftInverseS h) x) = f x
@[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 : f.rangeS) :
(ofLeftInverseS h).symm x = g x
def RingEquiv.subsemiringMap {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (e : R ≃+* S) (s : Subsemiring R) :
s ≃+* (Subsemiring.map (↑e) s)

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

Equations
@[simp]
theorem RingEquiv.subsemiringMap_apply_coe {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (e : R ≃+* S) (s : Subsemiring R) (x : s) :
((e.subsemiringMap s) x) = e x
@[simp]
theorem RingEquiv.subsemiringMap_symm_apply_coe {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (e : R ≃+* S) (s : Subsemiring R) (x : (Subsemiring.map e.toRingHom s)) :
((e.subsemiringMap s).symm x) = e.symm x

Actions by Subsemirings #

These are just copies of the definitions about Submonoid starting from Submonoid.mulAction. 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) α β
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) β
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.

@[instance 100]
instance Subsemiring.instFaithfulSMulSubtypeMem {M' : Type u_5} {α : Type u_6} [SMul M' α] {S' : Type u_7} [SetLike S' M'] (s : S') [FaithfulSMul M' α] :
FaithfulSMul (↥s) α
instance Subsemiring.faithfulSMul {R' : Type u_1} {α : Type u_2} [NonAssocSemiring R'] [SMul R' α] [FaithfulSMul R' α] (S : Subsemiring R') :
FaithfulSMul (↥S) α
@[instance 100]
instance Subsemiring.instSMulWithZeroSubtypeMem {R' : Type u_1} {α : Type u_2} [NonAssocSemiring R'] {S' : Type u_5} [SetLike S' R'] [SubsemiringClass S' R'] (s : S') [Zero α] [SMulWithZero R' α] :
SMulWithZero (↥s) α
Equations
instance Subsemiring.instSMulWithZeroSubtypeMem_1 {R' : Type u_1} {α : Type u_2} [NonAssocSemiring R'] [Zero α] [SMulWithZero R' α] (S : Subsemiring R') :
SMulWithZero (↥S) α

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

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 100]
instance Subsemiring.instModuleSubtypeMem {R' : Type u_1} {α : Type u_2} [Semiring R'] [AddCommMonoid α] [Module R' α] {S' : Type u_5} [SetLike S' R'] [SubsemiringClass S' R'] (s : S') :
Module (↥s) α
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 100]
instance Subsemiring.instMulActionWithZeroSubtypeMem {R' : Type u_1} {α : Type u_2} [Semiring R'] {S' : Type u_5} [SetLike S' R'] [SubsemiringClass S' R'] (s : S') [Zero α] [MulActionWithZero R' α] :
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 100]
instance Subsemiring.instModuleSubtypeMem_1 {R' : Type u_1} {α : Type u_2} [Semiring R'] [AddCommMonoid α] [Module R' α] {S' : Type u_5} [SetLike S' R'] [SubsemiringClass S' R'] (s : S') :
Module (↥s) α
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
instance Subsemiring.instMulSemiringActionSubtypeMem {R' : Type u_1} {α : Type u_2} [Semiring R'] [Semiring α] [MulSemiringAction R' α] (S : Subsemiring R') :

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

Equations
instance Subsemiring.center.smulCommClass_left {R' : Type u_1} [Semiring R'] :
SMulCommClass (↥(center R')) R' R'

The center of a semiring acts commutatively on that semiring.

instance Subsemiring.center.smulCommClass_right {R' : Type u_1} [Semiring R'] :
SMulCommClass R' (↥(center R')) R'

The center of a semiring acts commutatively on that semiring.

@[reducible, inline]
abbrev Subsemiring.closureCommSemiringOfComm {R' : Type u_1} [Semiring R'] {s : Set R'} (hcomm : xs, ys, x * y = y * x) :

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

Equations
theorem Subsemiring.map_comap_eq {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (t : Subsemiring S) :
map f (comap f t) = tf.rangeS
theorem Subsemiring.map_comap_eq_self {R : Type u} {S : Type v} [NonAssocSemiring R] [NonAssocSemiring S] {f : R →+* S} {t : Subsemiring S} (h : t f.rangeS) :
map f (comap f t) = t