Submonoids: membership criteria #
In this file we prove various facts about membership in a submonoid:
pow_mem
,nsmul_mem
: ifx ∈ S
whereS
is a multiplicative (resp., additive) submonoid andn
is a natural number, thenx^n
(resp.,n • x
) belongs toS
;mem_iSup_of_directed
,coe_iSup_of_directed
,mem_sSup_of_directedOn
,coe_sSup_of_directedOn
: the supremum of a directed collection of submonoid is their union.sup_eq_range
,mem_sup
: supremum of two submonoidsS
,T
of a commutative monoid is the set of products;closure_singleton_eq
,mem_closure_singleton
,mem_closure_pair
: the multiplicative (resp., additive) closure of{x}
consists of powers (resp., natural multiples) ofx
, and a similar result holds for the closure of{x, y}
.
Tags #
submonoid, submonoids
An induction principle for elements of ⨆ i, S i
.
If C
holds for 1
and all elements of S i
for all i
, and is preserved under multiplication,
then it holds for all elements of the supremum of S
.
An induction principle for elements of ⨆ i, S i
.
If C
holds for 0
and all elements of S i
for all i
, and is preserved under addition,
then it holds for all elements of the supremum of S
.
A dependent version of Submonoid.iSup_induction
.
A dependent version of AddSubmonoid.iSup_induction
.
The submonoid generated by an element.
Equations
- Submonoid.powers n = (MonoidHom.mrange ((powersHom M) n)).copy (Set.range fun (x : ℕ) => n ^ x) ⋯
Instances For
Equations
- Submonoid.decidableMemPowers = Classical.decPred fun (x : M) => x ∈ Submonoid.powers a
Equations
Exponentiation map from natural numbers to powers.
Equations
- Submonoid.pow n m = ((powersHom M) n).mrangeRestrict (Multiplicative.ofAdd m)
Instances For
The exponentiation map is an isomorphism from the additive monoid on natural numbers to powers when it is injective. The inverse is given by the logarithms.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The additive submonoid generated by an element.
Equations
- AddSubmonoid.multiples x = (AddMonoidHom.mrange ((multiplesHom A) x)).copy (Set.range fun (i : ℕ) => i • x) ⋯
Instances For
Equations
- AddSubmonoid.decidableMemMultiples = Classical.decPred fun (x : M) => x ∈ AddSubmonoid.multiples a
Equations
The additive submonoid generated by an element is an additive group if that element has finite order.
Equations
- AddSubmonoid.addGroupMultiples hpos hx = AddGroup.mk ⋯
Instances For
Lemmas about additive closures of Subsemigroup
.
The product of an element of the additive closure of a multiplicative subsemigroup M
and an element of M
is contained in the additive closure of M
.
The product of two elements of the additive closure of a submonoid M
is an element of the
additive closure of M
.
The product of an element of S
and an element of the additive closure of a multiplicative
submonoid S
is contained in the additive closure of S
.
The submonoid of primal elements in a cancellative commutative monoid with zero.
Equations
- Submonoid.isPrimal α = { carrier := {a : α | IsPrimal a}, mul_mem' := ⋯, one_mem' := ⋯ }