mathlib3 documentation

ring_theory.subsemiring.basic

Bundled subsemirings #

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.

We define bundled subsemirings and some standard constructions: complete_lattice structure, subtype and inclusion ring homomorphisms, subsemiring map, comap and range (srange) of a ring_hom etc.

@[class]
structure add_submonoid_with_one_class (S : Type u_1) (R : Type u_2) [add_monoid_with_one R] [set_like S R] :
Prop

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

Instances of this typeclass
theorem nat_cast_mem {S : Type u_1} {R : Type u_2} [add_monoid_with_one R] [set_like S R] (s : S) [add_submonoid_with_one_class S R] (n : ) :
n s
@[class]
structure subsemiring_class (S : Type u_1) (R : Type u) [non_assoc_semiring R] [set_like S R] :
Prop

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

Instances of this typeclass
theorem coe_nat_mem {R : Type u} {S : Type v} [non_assoc_semiring R] [set_like S R] [hSR : subsemiring_class S R] (s : S) (n : ) :
n s
@[protected, instance]

A subsemiring of a non_assoc_semiring inherits a non_assoc_semiring structure

Equations
@[protected, instance]
def subsemiring_class.nontrivial {R : Type u} {S : Type v} [non_assoc_semiring R] [set_like S R] [hSR : subsemiring_class S R] (s : S) [nontrivial R] :
@[protected, instance]
def subsemiring_class.subtype {R : Type u} {S : Type v} [non_assoc_semiring R] [set_like S R] [hSR : subsemiring_class S R] (s : S) :

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

Equations
@[simp]
@[protected, instance]
def subsemiring_class.to_semiring {S : Type v} (s : S) {R : Type u_1} [semiring R] [set_like S R] [subsemiring_class S R] :

A subsemiring of a semiring is a semiring.

Equations
@[simp, norm_cast]
theorem subsemiring_class.coe_pow {S : Type v} (s : S) {R : Type u_1} [semiring R] [set_like S R] [subsemiring_class S R] (x : s) (n : ) :
(x ^ n) = x ^ n
@[protected, instance]

A subsemiring of a comm_semiring is a comm_semiring.

Equations
structure subsemiring (R : Type u) [non_assoc_semiring R] :

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

Instances for subsemiring
@[protected, instance]
Equations
@[simp]
theorem subsemiring.mem_carrier {R : Type u} [non_assoc_semiring R] {s : subsemiring R} {x : R} :
x s.carrier x s
@[ext]
theorem subsemiring.ext {R : Type u} [non_assoc_semiring R] {S T : subsemiring R} (h : (x : R), x S x T) :
S = T

Two subsemirings are equal if they have the same elements.

@[protected]
def subsemiring.copy {R : Type u} [non_assoc_semiring 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
@[simp]
theorem subsemiring.coe_copy {R : Type u} [non_assoc_semiring R] (S : subsemiring R) (s : set R) (hs : s = S) :
(S.copy s hs) = s
theorem subsemiring.copy_eq {R : Type u} [non_assoc_semiring R] (S : subsemiring R) (s : set R) (hs : s = S) :
S.copy s hs = S
@[protected]
def subsemiring.mk' {R : Type u} [non_assoc_semiring R] (s : set R) (sm : submonoid R) (hm : sm = s) (sa : add_submonoid 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
@[simp]
theorem subsemiring.coe_mk' {R : Type u} [non_assoc_semiring R] {s : set R} {sm : submonoid R} (hm : sm = s) {sa : add_submonoid R} (ha : sa = s) :
(subsemiring.mk' s sm hm sa ha) = s
@[simp]
theorem subsemiring.mem_mk' {R : Type u} [non_assoc_semiring R] {s : set R} {sm : submonoid R} (hm : sm = s) {sa : add_submonoid R} (ha : sa = s) {x : R} :
x subsemiring.mk' s sm hm sa ha x s
@[simp]
theorem subsemiring.mk'_to_submonoid {R : Type u} [non_assoc_semiring R] {s : set R} {sm : submonoid R} (hm : sm = s) {sa : add_submonoid R} (ha : sa = s) :
(subsemiring.mk' s sm hm sa ha).to_submonoid = sm
@[simp]
theorem subsemiring.mk'_to_add_submonoid {R : Type u} [non_assoc_semiring R] {s : set R} {sm : submonoid R} (hm : sm = s) {sa : add_submonoid R} (ha : sa = s) :
@[protected]
theorem subsemiring.one_mem {R : Type u} [non_assoc_semiring R] (s : subsemiring R) :
1 s

A subsemiring contains the semiring's 1.

@[protected]
theorem subsemiring.zero_mem {R : Type u} [non_assoc_semiring R] (s : subsemiring R) :
0 s

A subsemiring contains the semiring's 0.

@[protected]
theorem subsemiring.mul_mem {R : Type u} [non_assoc_semiring R] (s : subsemiring R) {x y : R} :
x s y s x * y s

A subsemiring is closed under multiplication.

@[protected]
theorem subsemiring.add_mem {R : Type u} [non_assoc_semiring R] (s : subsemiring R) {x y : R} :
x s y s x + 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} :
( (x : R), x l x s) l.prod s

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

@[protected]
theorem subsemiring.list_sum_mem {R : Type u} [non_assoc_semiring R] (s : subsemiring R) {l : list R} :
( (x : R), x l x s) l.sum s

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

@[protected]
theorem subsemiring.multiset_prod_mem {R : Type u_1} [comm_semiring R] (s : subsemiring R) (m : multiset R) :
( (a : R), a m a s) m.prod s

Product of a multiset of elements in a subsemiring of a comm_semiring is in the subsemiring.

@[protected]
theorem subsemiring.multiset_sum_mem {R : Type u} [non_assoc_semiring R] (s : subsemiring R) (m : multiset R) :
( (a : R), a m a s) m.sum s

Sum of a multiset of elements in a subsemiring of a semiring is in the add_subsemiring.

@[protected]
theorem subsemiring.prod_mem {R : Type u_1} [comm_semiring R] (s : subsemiring R) {ι : Type u_2} {t : finset ι} {f : ι R} (h : (c : ι), c t f c s) :
t.prod (λ (i : ι), f i) s

Product of elements of a subsemiring of a comm_semiring indexed by a finset is in the subsemiring.

@[protected]
theorem subsemiring.sum_mem {R : Type u} [non_assoc_semiring R] (s : subsemiring R) {ι : Type u_1} {t : finset ι} {f : ι R} (h : (c : ι), c t f c s) :
t.sum (λ (i : ι), f i) s

Sum of elements in an subsemiring of an semiring indexed by a finset is in the add_subsemiring.

@[simp, norm_cast]
theorem subsemiring.coe_one {R : Type u} [non_assoc_semiring R] (s : subsemiring R) :
1 = 1
@[simp, norm_cast]
theorem subsemiring.coe_zero {R : Type u} [non_assoc_semiring R] (s : subsemiring R) :
0 = 0
@[simp, norm_cast]
theorem subsemiring.coe_add {R : Type u} [non_assoc_semiring R] (s : subsemiring R) (x y : s) :
(x + y) = x + y
@[simp, norm_cast]
theorem subsemiring.coe_mul {R : Type u} [non_assoc_semiring R] (s : subsemiring R) (x y : s) :
(x * y) = x * y
@[protected, instance]
@[protected]
theorem subsemiring.pow_mem {R : Type u_1} [semiring R] (s : subsemiring R) {x : R} (hx : x s) (n : ) :
x ^ n s
@[simp, norm_cast]
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
@[simp]
@[protected, instance]

A subsemiring of an ordered_semiring is an ordered_semiring.

Equations
@[protected]
theorem subsemiring.nsmul_mem {R : Type u} [non_assoc_semiring R] (s : subsemiring R) {x : R} (hx : x s) (n : ) :
n x s
@[simp]
theorem subsemiring.mem_to_submonoid {R : Type u} [non_assoc_semiring R] {s : subsemiring R} {x : R} :
@[simp]
@[protected, instance]

The subsemiring R of the semiring R.

Equations
@[simp]
theorem subsemiring.mem_top {R : Type u} [non_assoc_semiring R] (x : R) :

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

Equations
def subsemiring.comap {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] (f : R →+* S) (s : subsemiring S) :

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

Equations
@[simp]
@[simp]
theorem subsemiring.mem_comap {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] {s : subsemiring S} {f : R →+* S} {x : R} :
def subsemiring.map {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] (f : R →+* S) (s : subsemiring R) :

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

Equations
@[simp]
theorem subsemiring.coe_map {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] (f : R →+* S) (s : subsemiring R) :
@[simp]
theorem subsemiring.mem_map {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] {f : R →+* S} {s : subsemiring R} {y : S} :
y subsemiring.map f s (x : R) (H : x s), f x = y
@[simp]
theorem subsemiring.map_map {R : Type u} {S : Type v} {T : Type w} [non_assoc_semiring R] [non_assoc_semiring S] [non_assoc_semiring T] (s : subsemiring R) (g : S →+* T) (f : R →+* S) :

A subsemiring is isomorphic to its image under an injective function

Equations
def ring_hom.srange {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] (f : R →+* S) :

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

Equations
Instances for ring_hom.srange
@[simp]
theorem ring_hom.coe_srange {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] (f : R →+* S) :
@[simp]
theorem ring_hom.mem_srange {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] {f : R →+* S} {y : S} :
y f.srange (x : R), f x = y
theorem ring_hom.mem_srange_self {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] (f : R →+* S) (x : R) :
theorem ring_hom.map_srange {R : Type u} {S : Type v} {T : Type w} [non_assoc_semiring R] [non_assoc_semiring S] [non_assoc_semiring T] (g : S →+* T) (f : R →+* S) :
@[protected, instance]

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
@[protected, instance]
Equations
@[protected, instance]
Equations
theorem subsemiring.mem_bot {R : Type u} [non_assoc_semiring R] {x : R} :
x (n : ), n = x
@[protected, instance]

The inf of two subsemirings is their intersection.

Equations
@[simp]
theorem subsemiring.coe_inf {R : Type u} [non_assoc_semiring R] (p p' : subsemiring R) :
(p p') = p p'
@[simp]
theorem subsemiring.mem_inf {R : Type u} [non_assoc_semiring R] {p p' : subsemiring R} {x : R} :
x p p' x p x p'
@[protected, instance]
Equations
@[simp, norm_cast]
theorem subsemiring.coe_Inf {R : Type u} [non_assoc_semiring R] (S : set (subsemiring R)) :
(has_Inf.Inf S) = (s : subsemiring R) (H : s S), s
theorem subsemiring.mem_Inf {R : Type u} [non_assoc_semiring R] {S : set (subsemiring R)} {x : R} :
x has_Inf.Inf S (p : subsemiring R), p S x p
@[protected, instance]

Subsemirings of a semiring form a complete lattice.

Equations
theorem subsemiring.eq_top_iff' {R : Type u} [non_assoc_semiring R] (A : subsemiring R) :
A = (x : R), x A
def subsemiring.center (R : Type u_1) [semiring R] :

The center of a semiring R is the set of elements that commute with everything in R

Equations
Instances for subsemiring.center
theorem subsemiring.mem_center_iff {R : Type u_1} [semiring R] {z : R} :
z subsemiring.center R (g : R), g * z = z * g
@[protected, instance]
Equations
def subsemiring.centralizer {R : Type u_1} [semiring R] (s : set R) :

The centralizer of a set as subsemiring.

Equations
@[simp, norm_cast]
theorem subsemiring.mem_centralizer_iff {R : Type u_1} [semiring R] {s : set R} {z : R} :
z subsemiring.centralizer s (g : R), g s g * z = z * g

The subsemiring generated by a set.

Equations
theorem subsemiring.mem_closure {R : Type u} [non_assoc_semiring R] {x : R} {s : set R} :
@[simp]

The subsemiring generated by a set includes the set.

theorem subsemiring.not_mem_of_not_mem_closure {R : Type u} [non_assoc_semiring R] {s : set R} {P : R} (hP : P subsemiring.closure s) :
P s
@[simp]

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

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} [non_assoc_semiring 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} [non_assoc_semiring R] [non_assoc_semiring S] {f : R ≃+* S} {K : subsemiring R} {x : S} :

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} [non_assoc_semiring R] {s : set R} {p : R Prop} {x : R} (h : x subsemiring.closure s) (Hs : (x : R), x s p x) (H0 : p 0) (H1 : p 1) (Hadd : (x y : R), p x p y p (x + y)) (Hmul : (x y : R), p x p y p (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} [non_assoc_semiring R] {s : set R} {p : R R Prop} {x y : R} (hx : x subsemiring.closure s) (hy : y subsemiring.closure s) (Hs : (x : R), x s (y : R), y s 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₁ y p x₂ y p (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₁ y p x₂ y p (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)), ( (t : list R), t L (y : R), y t y s) (list.map list.prod L).sum = x
@[protected]

closure forms a Galois insertion with the coercion to set.

Equations

Closure of a subsemiring S equals S.

theorem subsemiring.closure_Union {R : Type u} [non_assoc_semiring R] {ι : Sort u_1} (s : ι set R) :
subsemiring.closure ( (i : ι), s i) = (i : ι), subsemiring.closure (s i)
theorem subsemiring.map_supr {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] {ι : Sort u_1} (f : R →+* S) (s : ι subsemiring R) :
subsemiring.map f (supr s) = (i : ι), subsemiring.map f (s i)
theorem subsemiring.comap_infi {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] {ι : Sort u_1} (f : R →+* S) (s : ι subsemiring S) :
@[simp]
@[simp]

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

Equations
@[norm_cast]
theorem subsemiring.coe_prod {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] (s : subsemiring R) (t : subsemiring S) :
(s.prod t) = s ×ˢ t
theorem subsemiring.mem_prod {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] {s : subsemiring R} {t : subsemiring S} {p : R × S} :
p s.prod t p.fst s p.snd t
theorem subsemiring.prod_mono {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] ⦃s₁ s₂ : subsemiring R⦄ (hs : s₁ s₂) ⦃t₁ t₂ : subsemiring S⦄ (ht : t₁ t₂) :
s₁.prod t₁ s₂.prod t₂

Product of subsemirings is isomorphic to their product as monoids.

Equations
theorem subsemiring.mem_supr_of_directed {R : Type u} [non_assoc_semiring R] {ι : Sort u_1} [hι : nonempty ι] {S : ι subsemiring R} (hS : directed has_le.le S) {x : R} :
(x (i : ι), S i) (i : ι), x S i
theorem subsemiring.coe_supr_of_directed {R : Type u} [non_assoc_semiring R] {ι : Sort u_1} [hι : nonempty ι] {S : ι subsemiring R} (hS : directed has_le.le S) :
( (i : ι), S i) = (i : ι), (S i)
theorem subsemiring.mem_Sup_of_directed_on {R : Type u} [non_assoc_semiring R] {S : set (subsemiring R)} (Sne : S.nonempty) (hS : directed_on has_le.le S) {x : R} :
x has_Sup.Sup S (s : subsemiring R) (H : s S), x s
def ring_hom.dom_restrict {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] {σR : Type u_1} [set_like σR R] [subsemiring_class σR R] (f : R →+* S) (s : σR) :

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

Equations
@[simp]
theorem ring_hom.restrict_apply {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] {σR : Type u_1} [set_like σR R] [subsemiring_class σR R] (f : R →+* S) {s : σR} (x : s) :
def ring_hom.cod_restrict {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] {σS : Type u_2} [set_like σS S] [subsemiring_class σS S] (f : R →+* S) (s : σS) (h : (x : R), f x s) :

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

Equations
def ring_hom.restrict {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] {σR : Type u_1} {σS : Type u_2} [set_like σR R] [set_like σS S] [subsemiring_class σR R] [subsemiring_class σS S] (f : R →+* S) (s' : σR) (s : σS) (h : (x : R), x s' f x s) :

The ring homomorphism from the preimage of s to s.

Equations
@[simp]
theorem ring_hom.coe_restrict_apply {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] {σR : Type u_1} {σS : Type u_2} [set_like σR R] [set_like σS S] [subsemiring_class σR R] [subsemiring_class σS S] (f : R →+* S) (s' : σR) (s : σS) (h : (x : R), x s' f x s) (x : s') :
((f.restrict s' s h) x) = f x
@[simp]
theorem ring_hom.comp_restrict {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] {σR : Type u_1} {σS : Type u_2} [set_like σR R] [set_like σS S] [subsemiring_class σR R] [subsemiring_class σS S] (f : R →+* S) (s' : σR) (s : σS) (h : (x : R), x s' f x s) :

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

This is the bundled version of set.range_factorization.

Equations
@[simp]
theorem ring_hom.coe_srange_restrict {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] (f : R →+* S) (x : R) :

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

def ring_hom.eq_slocus {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] (f g : R →+* S) :

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

Equations
@[simp]
theorem ring_hom.eq_slocus_same {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] (f : R →+* S) :
theorem ring_hom.eq_on_sclosure {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] {f g : R →+* S} {s : set R} (h : set.eq_on f g s) :

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

theorem ring_hom.eq_of_eq_on_stop {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] {f g : R →+* S} (h : set.eq_on f g ) :
f = g
theorem ring_hom.eq_of_eq_on_sdense {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] {s : set R} (hs : subsemiring.closure s = ) {f g : R →+* S} (h : set.eq_on 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} [non_assoc_semiring R] {S T : subsemiring R} (h : S T) :

The ring homomorphism associated to an inclusion of subsemirings.

Equations
def ring_equiv.subsemiring_congr {R : Type u} [non_assoc_semiring R] {s t : subsemiring R} (h : s = t) :

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

Equations
def ring_equiv.sof_left_inverse {R : Type u} {S : Type v} [non_assoc_semiring R] [non_assoc_semiring S] {g : S R} {f : R →+* S} (h : function.left_inverse g f) :

Restrict a ring homomorphism with a left inverse to a ring isomorphism to its ring_hom.srange.

Equations
@[simp]

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

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.of_subsemiring provides a stronger result than those found in this file, which uses the same scalar action.

@[protected, instance]
def subsemiring.has_smul {R' : Type u_1} {α : Type u_2} [non_assoc_semiring R'] [has_smul R' α] (S : subsemiring R') :

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

Equations
theorem subsemiring.smul_def {R' : Type u_1} {α : Type u_2} [non_assoc_semiring R'] [has_smul R' α] {S : subsemiring R'} (g : S) (m : α) :
g m = g m
@[protected, instance]
def subsemiring.smul_comm_class_left {R' : Type u_1} {α : Type u_2} {β : Type u_3} [non_assoc_semiring R'] [has_smul R' β] [has_smul α β] [smul_comm_class R' α β] (S : subsemiring R') :
@[protected, instance]
def subsemiring.smul_comm_class_right {R' : Type u_1} {α : Type u_2} {β : Type u_3} [non_assoc_semiring R'] [has_smul α β] [has_smul R' β] [smul_comm_class α R' β] (S : subsemiring R') :
@[protected, instance]
def subsemiring.is_scalar_tower {R' : Type u_1} {α : Type u_2} {β : Type u_3} [non_assoc_semiring R'] [has_smul α β] [has_smul R' α] [has_smul R' β] [is_scalar_tower R' α β] (S : subsemiring R') :

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

@[protected, instance]
@[protected, instance]
def subsemiring.smul_with_zero {R' : Type u_1} {α : Type u_2} [non_assoc_semiring R'] [has_zero α] [smul_with_zero R' α] (S : subsemiring R') :

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

Equations
@[protected, instance]
def subsemiring.mul_action {R' : Type u_1} {α : Type u_2} [semiring R'] [mul_action R' α] (S : subsemiring R') :

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

Equations
@[protected, instance]

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

Equations
@[protected, instance]

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

Equations
@[protected, instance]

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

Equations
@[protected, instance]
def subsemiring.module {R' : Type u_1} {α : Type u_2} [semiring R'] [add_comm_monoid α] [module R' α] (S : subsemiring R') :

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

Equations
@[protected, instance]

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

Equations
@[protected, instance]

The center of a semiring acts commutatively on that semiring.

@[protected, instance]

The center of a semiring acts commutatively on that semiring.

Submonoid of positive elements of an ordered semiring.

Equations
@[simp]
theorem mem_pos_monoid {R : Type u_1} [strict_ordered_semiring R] (u : Rˣ) :