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.
- to_add_submonoid_class : add_submonoid_class S R
- to_one_mem_class : one_mem_class S R
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
Equations
- add_submonoid_with_one_class.to_add_monoid_with_one s = {nat_cast := λ (n : ℕ), ⟨↑n, _⟩, add := add_monoid.add (add_submonoid_class.to_add_monoid s), add_assoc := _, zero := add_monoid.zero (add_submonoid_class.to_add_monoid s), zero_add := _, add_zero := _, nsmul := add_monoid.nsmul (add_submonoid_class.to_add_monoid s), nsmul_zero' := _, nsmul_succ' := _, one := ⟨1, _⟩, nat_cast_zero := _, nat_cast_succ := _}
- to_submonoid_class : submonoid_class S R
- to_add_submonoid_class : add_submonoid_class S R
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
A subsemiring of a non_assoc_semiring
inherits a non_assoc_semiring
structure
Equations
The natural ring hom from a subsemiring of semiring R
to R
.
Equations
- subsemiring_class.subtype s = {to_fun := coe coe_to_lift, map_one' := _, map_mul' := _, map_zero' := _, map_add' := _}
A subsemiring of a semiring
is a semiring
.
Equations
- subsemiring_class.to_semiring s = function.injective.semiring coe _ _ _ _ _ _ _ _
A subsemiring of a comm_semiring
is a comm_semiring
.
Equations
- subsemiring_class.to_comm_semiring s = function.injective.comm_semiring coe _ _ _ _ _ _ _ _
A subsemiring of an ordered_semiring
is an ordered_semiring
.
Equations
- subsemiring_class.to_ordered_semiring s = function.injective.ordered_semiring coe _ _ _ _ _ _ _ _
A subsemiring of an strict_ordered_semiring
is an strict_ordered_semiring
.
Equations
A subsemiring of an ordered_comm_semiring
is an ordered_comm_semiring
.
Equations
A subsemiring of an strict_ordered_comm_semiring
is an strict_ordered_comm_semiring
.
Equations
A subsemiring of a linear_ordered_semiring
is a linear_ordered_semiring
.
Equations
- subsemiring_class.to_linear_ordered_semiring s = function.injective.linear_ordered_semiring coe _ _ _ _ _ _ _ _ _ _
A subsemiring of a linear_ordered_comm_semiring
is a linear_ordered_comm_semiring
.
Equations
- subsemiring_class.to_linear_ordered_comm_semiring s = function.injective.linear_ordered_comm_semiring coe _ _ _ _ _ _ _ _ _ _
Reinterpret a subsemiring
as an add_submonoid
.
Reinterpret a subsemiring
as a submonoid
.
- carrier : set R
- mul_mem' : ∀ {a b : R}, a ∈ self.carrier → b ∈ self.carrier → a * b ∈ self.carrier
- one_mem' : 1 ∈ self.carrier
- add_mem' : ∀ {a b : R}, a ∈ self.carrier → b ∈ self.carrier → a + b ∈ self.carrier
- zero_mem' : 0 ∈ self.carrier
A subsemiring of a semiring R
is a subset s
that is both a multiplicative and an additive
submonoid.
Instances for subsemiring
Equations
- subsemiring.set_like = {coe := subsemiring.carrier _inst_1, coe_injective' := _}
Two subsemirings are equal if they have the same elements.
Copy of a subsemiring with a new carrier
equal to the old one. Useful to fix definitional
equalities.
Construct a subsemiring R
from a set s
, a submonoid sm
, and an additive
submonoid sa
such that x ∈ s ↔ x ∈ sm ↔ x ∈ sa
.
A subsemiring contains the semiring's 1.
A subsemiring contains the semiring's 0.
A subsemiring is closed under multiplication.
A subsemiring is closed under addition.
Product of a list of elements in a subsemiring
is in the subsemiring
.
Sum of a list of elements in a subsemiring
is in the subsemiring
.
Product of a multiset of elements in a subsemiring
of a comm_semiring
is in the subsemiring
.
Sum of a multiset of elements in a subsemiring
of a semiring
is
in the add_subsemiring
.
Product of elements of a subsemiring of a comm_semiring
indexed by a finset
is in the
subsemiring.
Sum of elements in an subsemiring
of an semiring
indexed by a finset
is in the add_subsemiring
.
A subsemiring of a non_assoc_semiring
inherits a non_assoc_semiring
structure
Equations
- s.to_non_assoc_semiring = {add := add_comm_monoid.add s.to_add_submonoid.to_add_comm_monoid, add_assoc := _, zero := add_comm_monoid.zero s.to_add_submonoid.to_add_comm_monoid, zero_add := _, add_zero := _, nsmul := add_comm_monoid.nsmul s.to_add_submonoid.to_add_comm_monoid, nsmul_zero' := _, nsmul_succ' := _, add_comm := _, mul := mul_one_class.mul s.to_submonoid.to_mul_one_class, left_distrib := _, right_distrib := _, zero_mul := _, mul_zero := _, one := mul_one_class.one s.to_submonoid.to_mul_one_class, one_mul := _, mul_one := _, nat_cast := λ (n : ℕ), ⟨↑n, _⟩, nat_cast_zero := _, nat_cast_succ := _}
A subsemiring of a semiring
is a semiring
.
Equations
- s.to_semiring = {add := non_assoc_semiring.add s.to_non_assoc_semiring, add_assoc := _, zero := non_assoc_semiring.zero s.to_non_assoc_semiring, zero_add := _, add_zero := _, nsmul := non_assoc_semiring.nsmul s.to_non_assoc_semiring, nsmul_zero' := _, nsmul_succ' := _, add_comm := _, mul := non_assoc_semiring.mul s.to_non_assoc_semiring, left_distrib := _, right_distrib := _, zero_mul := _, mul_zero := _, mul_assoc := _, one := non_assoc_semiring.one s.to_non_assoc_semiring, one_mul := _, mul_one := _, nat_cast := non_assoc_semiring.nat_cast s.to_non_assoc_semiring, nat_cast_zero := _, nat_cast_succ := _, npow := monoid.npow s.to_submonoid.to_monoid, npow_zero' := _, npow_succ' := _}
A subsemiring of a comm_semiring
is a comm_semiring
.
Equations
- s.to_comm_semiring = {add := semiring.add s.to_semiring, add_assoc := _, zero := semiring.zero s.to_semiring, zero_add := _, add_zero := _, nsmul := semiring.nsmul s.to_semiring, nsmul_zero' := _, nsmul_succ' := _, add_comm := _, mul := semiring.mul s.to_semiring, left_distrib := _, right_distrib := _, zero_mul := _, mul_zero := _, mul_assoc := _, one := semiring.one s.to_semiring, one_mul := _, mul_one := _, nat_cast := semiring.nat_cast s.to_semiring, nat_cast_zero := _, nat_cast_succ := _, npow := semiring.npow s.to_semiring, npow_zero' := _, npow_succ' := _, mul_comm := _}
The natural ring hom from a subsemiring of semiring R
to R
.
A subsemiring of an ordered_semiring
is an ordered_semiring
.
Equations
- s.to_ordered_semiring = function.injective.ordered_semiring coe _ _ _ _ _ _ _ _
A subsemiring of a strict_ordered_semiring
is a strict_ordered_semiring
.
Equations
- s.to_strict_ordered_semiring = function.injective.strict_ordered_semiring coe _ _ _ _ _ _ _ _
A subsemiring of an ordered_comm_semiring
is an ordered_comm_semiring
.
Equations
- s.to_ordered_comm_semiring = function.injective.ordered_comm_semiring coe _ _ _ _ _ _ _ _
A subsemiring of a strict_ordered_comm_semiring
is a strict_ordered_comm_semiring
.
Equations
- s.to_strict_ordered_comm_semiring = function.injective.strict_ordered_comm_semiring coe _ _ _ _ _ _ _ _
A subsemiring of a linear_ordered_semiring
is a linear_ordered_semiring
.
Equations
- s.to_linear_ordered_semiring = function.injective.linear_ordered_semiring coe _ _ _ _ _ _ _ _ _ _
A subsemiring of a linear_ordered_comm_semiring
is a linear_ordered_comm_semiring
.
Equations
- s.to_linear_ordered_comm_semiring = function.injective.linear_ordered_comm_semiring coe _ _ _ _ _ _ _ _ _ _
The subsemiring R
of the semiring R
.
The preimage of a subsemiring along a ring homomorphism is a subsemiring.
The image of a subsemiring along a ring homomorphism is a subsemiring.
A subsemiring is isomorphic to its image under an injective function
The range of a ring homomorphism is a subsemiring. See Note [range copy pattern].
Instances for ↥ring_hom.srange
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
- subsemiring.has_bot = {bot := (nat.cast_ring_hom R).srange}
Equations
The inf of two subsemirings is their intersection.
Equations
- subsemiring.has_Inf = {Inf := λ (s : set (subsemiring R)), subsemiring.mk' (⋂ (t : subsemiring R) (H : t ∈ s), ↑t) (⨅ (t : subsemiring R) (H : t ∈ s), t.to_submonoid) _ (⨅ (t : subsemiring R) (H : t ∈ s), t.to_add_submonoid) _}
Subsemirings of a semiring form a complete lattice.
Equations
- subsemiring.complete_lattice = {sup := complete_lattice.sup (complete_lattice_of_Inf (subsemiring R) subsemiring.complete_lattice._proof_1), le := complete_lattice.le (complete_lattice_of_Inf (subsemiring R) subsemiring.complete_lattice._proof_1), lt := complete_lattice.lt (complete_lattice_of_Inf (subsemiring R) subsemiring.complete_lattice._proof_1), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := has_inf.inf subsemiring.has_inf, inf_le_left := _, inf_le_right := _, le_inf := _, Sup := complete_lattice.Sup (complete_lattice_of_Inf (subsemiring R) subsemiring.complete_lattice._proof_1), le_Sup := _, Sup_le := _, Inf := complete_lattice.Inf (complete_lattice_of_Inf (subsemiring R) subsemiring.complete_lattice._proof_1), Inf_le := _, le_Inf := _, top := ⊤, bot := ⊥, le_top := _, bot_le := _}
The center of a semiring R
is the set of elements that commute with everything in R
Equations
- subsemiring.center R = {carrier := set.center R (distrib.to_has_mul R), mul_mem' := _, one_mem' := _, add_mem' := _, zero_mem' := _}
Instances for ↥subsemiring.center
Equations
- subsemiring.decidable_mem_center = λ (_x : R), decidable_of_iff' (∀ (g : R), g * _x = _x * g) subsemiring.mem_center_iff
The center is commutative.
Equations
- subsemiring.center.comm_semiring = {add := semiring.add (subsemiring.center R).to_semiring, add_assoc := _, zero := semiring.zero (subsemiring.center R).to_semiring, zero_add := _, add_zero := _, nsmul := semiring.nsmul (subsemiring.center R).to_semiring, nsmul_zero' := _, nsmul_succ' := _, add_comm := _, mul := comm_monoid.mul submonoid.center.comm_monoid, left_distrib := _, right_distrib := _, zero_mul := _, mul_zero := _, mul_assoc := _, one := comm_monoid.one submonoid.center.comm_monoid, one_mul := _, mul_one := _, nat_cast := semiring.nat_cast (subsemiring.center R).to_semiring, nat_cast_zero := _, nat_cast_succ := _, npow := comm_monoid.npow submonoid.center.comm_monoid, npow_zero' := _, npow_succ' := _, mul_comm := _}
The centralizer of a set as subsemiring.
Equations
- subsemiring.centralizer s = {carrier := s.centralizer (distrib.to_has_mul R), mul_mem' := _, one_mem' := _, add_mem' := _, zero_mem' := _}
The subsemiring
generated by a set.
Equations
- subsemiring.closure s = has_Inf.Inf {S : subsemiring R | s ⊆ ↑S}
The subsemiring generated by a set includes the set.
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
.
The additive closure of a submonoid is a subsemiring.
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
.
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
.
An induction principle for closure membership for predicates with two arguments.
closure
forms a Galois insertion with the coercion to set.
Equations
- subsemiring.gi R = {choice := λ (s : set R) (_x : ↑(subsemiring.closure s) ≤ s), subsemiring.closure s, gc := _, le_l_u := _, choice_eq := _}
Closure of a subsemiring S
equals S
.
Given subsemiring
s s
, t
of semirings R
, S
respectively, s.prod t
is s × t
as a subsemiring of R × S
.
Product of subsemirings is isomorphic to their product as monoids.
Restriction of a ring homomorphism to a subsemiring of the domain.
Equations
- f.dom_restrict s = f.comp (subsemiring_class.subtype s)
Restriction of a ring homomorphism to a subsemiring of the codomain.
The ring homomorphism from the preimage of s
to s
.
Equations
- f.restrict s' s h = (f.dom_restrict s').cod_restrict s _
Restriction of a ring homomorphism to its range interpreted as a subsemiring.
This is the bundled version of set.range_factorization
.
Equations
- f.srange_restrict = f.cod_restrict f.srange _
The range of a surjective ring homomorphism is the whole of the codomain.
The subsemiring of elements x : R
such that f x = g x
If two ring homomorphisms are equal on a set, then they are equal on its subsemiring closure.
The image under a ring homomorphism of the subsemiring generated by a set equals the subsemiring generated by the image of the set.
The ring homomorphism associated to an inclusion of subsemirings.
Equations
- subsemiring.inclusion h = S.subtype.cod_restrict T _
Makes the identity isomorphism from a proof two subsemirings of a multiplicative monoid are equal.
Equations
- ring_equiv.subsemiring_congr h = {to_fun := (equiv.set_congr _).to_fun, inv_fun := (equiv.set_congr _).inv_fun, left_inv := _, right_inv := _, map_mul' := _, map_add' := _}
Restrict a ring homomorphism with a left inverse to a ring isomorphism to its
ring_hom.srange
.
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
- e.subsemiring_map s = {to_fun := (e.to_add_equiv.add_submonoid_map s.to_add_submonoid).to_fun, inv_fun := (e.to_add_equiv.add_submonoid_map s.to_add_submonoid).inv_fun, left_inv := _, right_inv := _, map_mul' := _, map_add' := _}
Actions by subsemiring
s #
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.
The action by a subsemiring is the action by the underlying semiring.
Equations
Note that this provides is_scalar_tower S R R
which is needed by smul_mul_assoc
.
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
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
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
- S.module = {to_distrib_mul_action := {to_mul_action := {to_has_smul := {smul := has_smul.smul S.has_smul}, one_smul := _, mul_smul := _}, smul_zero := _, smul_add := _}, add_smul := _, zero_smul := _}
The action by a subsemiring is the action by the underlying semiring.
Equations
The center of a semiring acts commutatively on that semiring.
The center of a semiring acts commutatively on that semiring.
If all the elements of a set s
commute, then closure s
is a commutative monoid.
Equations
- subsemiring.closure_comm_semiring_of_comm hcomm = {add := semiring.add (subsemiring.closure s).to_semiring, add_assoc := _, zero := semiring.zero (subsemiring.closure s).to_semiring, zero_add := _, add_zero := _, nsmul := semiring.nsmul (subsemiring.closure s).to_semiring, nsmul_zero' := _, nsmul_succ' := _, add_comm := _, mul := semiring.mul (subsemiring.closure s).to_semiring, left_distrib := _, right_distrib := _, zero_mul := _, mul_zero := _, mul_assoc := _, one := semiring.one (subsemiring.closure s).to_semiring, one_mul := _, mul_one := _, nat_cast := semiring.nat_cast (subsemiring.closure s).to_semiring, nat_cast_zero := _, nat_cast_succ := _, npow := semiring.npow (subsemiring.closure s).to_semiring, npow_zero' := _, npow_succ' := _, mul_comm := _}
Submonoid of positive elements of an ordered semiring.