Typeclasses for (semi)groups and monoids #
In this file we define typeclasses for algebraic structures with one binary operation. The classes are named (Add)?(Comm)?(Semigroup|Monoid|Group), where Add means that the class uses additive notation and Comm means that the class assumes that the binary operation is commutative.
The file does not contain any lemmas except for
- axioms of typeclasses restated in the root namespace;
- lemmas required for instances.
For basic lemmas about these classes see Algebra.Group.Basic.
We register the following instances:
- Pow M ℕ, for monoids- M, and- Pow G ℤfor groups- G;
- SMul ℕ Mfor additive monoids- M, and- SMul ℤ Gfor additive groups- G.
Notation #
A mixin for left cancellative multiplication.
- mul_left_cancel (a : G) : IsLeftRegular aMultiplication is left cancellative (i.e. left regular). 
Instances
A mixin for right cancellative multiplication.
- mul_right_cancel (a : G) : IsRightRegular aMultiplication is right cancellative (i.e. right regular). 
Instances
A mixin for cancellative multiplication.
Instances
A mixin for left cancellative addition.
- add_left_cancel (a : G) : IsAddLeftRegular aAddition is left cancellative (i.e. left regular). 
Instances
A mixin for right cancellative addition.
- add_right_cancel (a : G) : IsAddRightRegular aAddition is right cancellative (i.e. right regular). 
Instances
A mixin for cancellative addition.
Instances
If all multiplications cancel on the left then every element is left-regular.
If all additions cancel on the left then every element is add-left-regular.
If all multiplications cancel on the right then every element is right-regular.
If all additions cancel on the right then every element is add-right-regular.
If all multiplications cancel then every element is regular.
If all additions cancel then every element is add-regular.
Any CommMagma G that satisfies IsRightCancelMul G also satisfies IsLeftCancelMul G.
Any AddCommMagma G that satisfies IsRightCancelAdd G also satisfies IsLeftCancelAdd G.
Any CommMagma G that satisfies IsLeftCancelMul G also satisfies IsRightCancelMul G.
Any AddCommMagma G that satisfies IsLeftCancelAdd G also satisfies IsRightCancelAdd G.
Any CommMagma G that satisfies IsLeftCancelMul G also satisfies IsCancelMul G.
Any AddCommMagma G that satisfies IsLeftCancelAdd G also satisfies IsCancelAdd G.
Any CommMagma G that satisfies IsRightCancelMul G also satisfies IsCancelMul G.
Any AddCommMagma G that satisfies IsRightCancelAdd G also satisfies IsCancelAdd G.
We lower the priority of inheriting from cancellative structures. This attempts to avoid expensive checks involving bundling and unbundling with the IsDomain class. since IsDomain already depends on Semiring, we can synthesize that one first. Zulip discussion: https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/Why.20is.20.60simpNF.60.20complaining.20here.3F
Instances For
An AddLeftCancelSemigroup is an additive semigroup such that a + b = a + c implies b = c.
- add : G → G → G
Instances
A RightCancelSemigroup is a semigroup such that a * b = c * b implies a = c.
- mul : G → G → G
Instances
An AddRightCancelSemigroup is an additive semigroup such that a + b = c + b implies a = c.
- add : G → G → G
Instances
Suppose that one can put two mathematical structures on a type, a rich one R and a poor one P, and that one can deduce the poor structure from the rich structure through a map F (called a forgetful functor) (think R = MetricSpace and P = TopologicalSpace). A possible implementation would be to have a type class rich containing a field R, a type class poor containing a field P, and an instance from rich to poor. However, this creates diamond problems, and a better approach is to let rich extend poor and have a field saying that F R = P.
To illustrate this, consider the pair MetricSpace / TopologicalSpace. Consider the topology on a product of two metric spaces. With the first approach, it could be obtained by going first from each metric space to its topology, and then taking the product topology. But it could also be obtained by considering the product metric space (with its sup distance) and then the topology coming from this distance. These would be the same topology, but not definitionally, which means that from the point of view of Lean's kernel, there would be two different TopologicalSpace instances on the product. This is not compatible with the way instances are designed and used: there should be at most one instance of a kind on each type. This approach has created an instance diamond that does not commute definitionally.
The second approach solves this issue. Now, a metric space contains both a distance, a topology, and a proof that the topology coincides with the one coming from the distance. When one defines the product of two metric spaces, one uses the sup distance and the product topology, and one has to give the proof that the sup distance induces the product topology. Following both sides of the instance diamond then gives rise (definitionally) to the product topology on the product space.
Another approach would be to have the rich type class take the poor type class as an instance parameter. It would solve the diamond problem, but it would lead to a blow up of the number of type classes one would need to declare to work with complicated classes, say a real inner product space, and would create exponential complexity when working with products of such complicated spaces, that are avoided by bundling things carefully as above.
Note that this description of this specific case of the product of metric spaces is oversimplified compared to mathlib, as there is an intermediate typeclass between MetricSpace and TopologicalSpace called UniformSpace. The above scheme is used at both levels, embedding a topology in the uniform space structure, and a uniform structure in the metric space structure.
Note also that, when P is a proposition, there is no such issue as any two proofs of P are definitionally equivalent in Lean.
To avoid boilerplate, there are some designs that can automatically fill the poor fields when creating a rich structure if one doesn't want to do something special about them. For instance, in the definition of metric spaces, default tactics fill the uniform space fields if they are not given explicitly. One can also have a helper function creating the rich structure from a structure with fewer fields, where the helper function fills the remaining fields. See for instance UniformSpace.ofCore or RealInnerProduct.ofCore.
For more details on this question, called the forgetful inheritance pattern, see Competing inheritance paths in dependent type theory: a case study in functional analysis.
Instances For
Design note on AddMonoid and Monoid #
An AddMonoid has a natural ℕ-action, defined by n • a = a + ... + a, that we want to declare as an instance as it makes it possible to use the language of linear algebra. However, there are often other natural ℕ-actions. For instance, for any semiring R, the space of polynomials Polynomial R has a natural R-action defined by multiplication on the coefficients. This means that Polynomial ℕ would have two natural ℕ-actions, which are equal but not defeq. The same goes for linear maps, tensor products, and so on (and even for ℕ itself).
To solve this issue, we embed an ℕ-action in the definition of an AddMonoid (which is by default equal to the naive action a + ... + a, but can be adjusted when needed), and declare a SMul ℕ α instance using this action. See Note [forgetful inheritance] for more explanations on this pattern.
For example, when we define Polynomial R, then we declare the ℕ-action to be by multiplication on each coefficient (using the ℕ-action on R that comes from the fact that R is an AddMonoid). In this way, the two natural SMul ℕ (Polynomial ℕ) instances are defeq.
The tactic to_additive transfers definitions and results from multiplicative monoids to additive monoids. To work, it has to map fields to fields. This means that we should also add corresponding fields to the multiplicative structure Monoid, which could solve defeq problems for powers if needed. These problems do not come up in practice, so most of the time we will not need to adjust the npow field when defining multiplicative objects.
Scalar multiplication by repeated self-addition, the additive version of exponentiation by repeated squaring.
Equations
- nsmulBinRec k = nsmulBinRec.go k 0
Instances For
Auxiliary tail-recursive implementation for nsmulBinRec.
Equations
- nsmulBinRec.go k = Nat.binaryRec (motive := fun (x : ℕ) => M → M → M) (fun (y x : M) => y) (fun (bn : Bool) (_n : ℕ) (fn : M → M → M) (y x : M) => fn (bif bn then y + x else y) (x + x)) k
Instances For
Auxiliary tail-recursive implementation for npowBinRec.
Equations
- npowBinRec.go k = Nat.binaryRec (motive := fun (x : ℕ) => M → M → M) (fun (y x : M) => y) (fun (bn : Bool) (_n : ℕ) (fn : M → M → M) (y x : M) => fn (bif bn then y * x else y) (x * x)) k
Instances For
An abbreviation for npowRec with an additional typeclass assumption on associativity so that we can use @[csimp] to replace it with an implementation by repeated squaring in compiled code.
Equations
- npowRecAuto k m = npowRec k m
Instances For
An abbreviation for nsmulRec with an additional typeclass assumptions on associativity so that we can use @[csimp] to replace it with an implementation by repeated doubling in compiled code as an automatic parameter.
Equations
- nsmulRecAuto k m = nsmulRec k m
Instances For
An abbreviation for npowBinRec with an additional typeclass assumption on associativity so that we can use it in @[csimp] for more performant code generation.
Equations
- npowBinRecAuto k m = npowBinRec k m
Instances For
An abbreviation for nsmulBinRec with an additional typeclass assumption on associativity so that we can use it in @[csimp] for more performant code generation as an automatic parameter.
Equations
- nsmulBinRecAuto k m = nsmulBinRec k m
Instances For
An AddMonoid is an AddSemigroup with an element 0 such that 0 + a = a + 0 = a.
- add : M → M → M
- zero : M
- nsmul : ℕ → M → MMultiplication by a natural number. Set this to nsmulRecunlessModulediamonds are possible.
- Multiplication by - (0 : ℕ)gives- 0.
- Multiplication by - (n + 1 : ℕ)behaves as expected.
Instances
Equations
- Monoid.toNatPow = { pow := fun (x : M) (n : ℕ) => Monoid.npow n x }
Equations
- AddMonoid.toNatSMul = { smul := AddMonoid.nsmul }
An additive monoid is torsion-free if scalar multiplication by every non-zero element n : ℕ is injective.
Instances
A monoid is torsion-free if power by every non-zero element n : ℕ is injective.
Instances
An additive commutative monoid is an additive monoid with commutative (+).
Instances
A commutative monoid is a monoid with commutative (*).
Instances
An additive monoid in which addition is left-cancellative. Main examples are ℕ and groups. This is the right typeclass for many sum lemmas, as having a zero is useful to define the sum over the empty set, so AddLeftCancelSemigroup is not enough.
Instances
A monoid in which multiplication is left-cancellative.
Instances
An additive monoid in which addition is right-cancellative. Main examples are ℕ and groups. This is the right typeclass for many sum lemmas, as having a zero is useful to define the sum over the empty set, so AddRightCancelSemigroup is not enough.
Instances
A monoid in which multiplication is right-cancellative.
Instances
An additive monoid in which addition is cancellative on both sides. Main examples are ℕ and groups. This is the right typeclass for many sum lemmas, as having a zero is useful to define the sum over the empty set, so AddRightCancelMonoid is not enough.
Instances
A monoid in which multiplication is cancellative.
Instances
Commutative version of AddCancelMonoid.
Instances
Commutative version of CancelMonoid.
Instances
Equations
- CancelCommMonoid.toCancelMonoid M = { toLeftCancelMonoid := inst✝.toLeftCancelMonoid, toIsRightCancelMul := ⋯ }
Equations
- AddCancelCommMonoid.toAddCancelMonoid M = { toAddLeftCancelMonoid := inst✝.toAddLeftCancelMonoid, toIsRightCancelAdd := ⋯ }
Any CancelMonoid G satisfies IsCancelMul G.
Any AddCancelMonoid G satisfies IsCancelAdd G.
The fundamental power operation in a group. zpowRec n a = a*a*...*a n times, for integer n. Use instead a ^ n, which has better definitional behavior.
Equations
Instances For
The fundamental scalar multiplication in an additive group. zpowRec n a = a+a+...+a n times, for integer n. Use instead n • a, which has better definitional behavior.
Equations
Instances For
Design note on DivInvMonoid/SubNegMonoid and DivisionMonoid/SubtractionMonoid #
Those two pairs of made-up classes fulfill slightly different roles.
DivInvMonoid/SubNegMonoid provides the minimum amount of information to define the ℤ action (zpow or zsmul). Further, it provides a div field, matching the forgetful inheritance pattern. This is useful to shorten extension clauses of stronger structures (Group, GroupWithZero, DivisionRing, Field) and for a few structures with a rather weak pseudo-inverse (Matrix).
DivisionMonoid/SubtractionMonoid is targeted at structures with stronger pseudo-inverses. It is an ad hoc collection of axioms that are mainly respected by three things:
- Groups
- Groups with zero
- The pointwise monoids Set α,Finset α,Filter α
It acts as a middle ground for structures with an inversion operator that plays well with multiplication, except for the fact that it might not be a true inverse (a / a ≠ 1 in general). The axioms are pretty arbitrary (many other combinations are equivalent to it), but they are independent:
- Without DivisionMonoid.div_eq_mul_inv, you can define/arbitrarily.
- Without DivisionMonoid.inv_inv, you can considerWithTop Unitwitha⁻¹ = ⊤for alla.
- Without DivisionMonoid.mul_inv_rev, you can considerWithTop αwitha⁻¹ = afor allawhereαnoncommutative.
- Without DivisionMonoid.inv_eq_of_mul, you can consider anyCommMonoidwitha⁻¹ = afor alla.
As a consequence, a few natural structures do not fit in this framework. For example, ENNReal respects everything except for the fact that (0 * ∞)⁻¹ = 0⁻¹ = ∞ while ∞⁻¹ * 0⁻¹ = 0 * ∞ = 0.
In a class equipped with instances of both Monoid and Inv, this definition records what the default definition for Div would be: a * b⁻¹. This is later provided as the default value for the Div instance in DivInvMonoid.
We keep it as a separate definition rather than inlining it in DivInvMonoid so that the Div field of individual DivInvMonoids constructed using that default value will not be unfolded at .instance transparency.
Equations
- DivInvMonoid.div' a b = a * b⁻¹
Instances For
A DivInvMonoid is a Monoid with operations / and ⁻¹ satisfying div_eq_mul_inv : ∀ a b, a / b = a * b⁻¹.
This deduplicates the name div_eq_mul_inv. The default for div is such that a / b = a * b⁻¹ holds by definition.
Adding div as a field rather than defining a / b := a * b⁻¹ allows us to avoid certain classes of unification failures, for example: Let Foo X be a type with a ∀ X, Div (Foo X) instance but no ∀ X, Inv (Foo X), e.g. when Foo X is a EuclideanDomain. Suppose we also have an instance ∀ X [Cromulent X], GroupWithZero (Foo X). Then the (/) coming from GroupWithZero.div cannot be definitionally equal to the (/) coming from Foo.Div.
In the same way, adding a zpow field makes it possible to avoid definitional failures in diamonds. See the definition of Monoid and Note [forgetful inheritance] for more explanations on this.
- mul : G → G → G
- one : G
- inv : G → G
- div : G → G → G
- a / b := a * b⁻¹
- zpow : ℤ → G → GThe power operation: a ^ n = a * ··· * a;a ^ (-n) = a⁻¹ * ··· a⁻¹(ntimes)
- a ^ 0 = 1
- a ^ (n + 1) = a ^ n * a
- a ^ -(n + 1) = (a ^ (n + 1))⁻¹
Instances
In a class equipped with instances of both AddMonoid and Neg, this definition records what the default definition for Sub would be: a + -b. This is later provided as the default value for the Sub instance in SubNegMonoid.
We keep it as a separate definition rather than inlining it in SubNegMonoid so that the Sub field of individual SubNegMonoids constructed using that default value will not be unfolded at .instance transparency.
Equations
- SubNegMonoid.sub' a b = a + -b
Instances For
A SubNegMonoid is an AddMonoid with unary - and binary - operations satisfying sub_eq_add_neg : ∀ a b, a - b = a + -b.
The default for sub is such that a - b = a + -b holds by definition.
Adding sub as a field rather than defining a - b := a + -b allows us to avoid certain classes of unification failures, for example: Let foo X be a type with a ∀ X, Sub (Foo X) instance but no ∀ X, Neg (Foo X). Suppose we also have an instance ∀ X [Cromulent X], AddGroup (Foo X). Then the (-) coming from AddGroup.sub cannot be definitionally equal to the (-) coming from Foo.Sub.
In the same way, adding a zsmul field makes it possible to avoid definitional failures in diamonds. See the definition of AddMonoid and Note [forgetful inheritance] for more explanations on this.
- add : G → G → G
- zero : G
- neg : G → G
- sub : G → G → G
- zsmul : ℤ → G → GMultiplication by an integer. Set this to zsmulRecunlessModulediamonds are possible.
Instances
Equations
- DivInvMonoid.toZPow = { pow := fun (x : M) (n : ℤ) => DivInvMonoid.zpow n x }
Equations
- SubNegMonoid.toZSMul = { smul := SubNegMonoid.zsmul }
Dividing by an element is the same as multiplying by its inverse.
This is a duplicate of DivInvMonoid.div_eq_mul_inv ensuring that the types unfold better.
Subtracting an element is the same as adding by its negative. This is a duplicate of SubNegMonoid.sub_eq_add_neg ensuring that the types unfold better.
Alias of div_eq_mul_inv.
Dividing by an element is the same as multiplying by its inverse.
This is a duplicate of DivInvMonoid.div_eq_mul_inv ensuring that the types unfold better.
A SubNegMonoid where -0 = 0.
Instances
A DivInvMonoid where 1⁻¹ = 1.
Instances
A SubtractionMonoid is a SubNegMonoid with involutive negation and such that -(a + b) = -b + -a and a + b = 0 → -a = b.
- add : G → G → G
- zero : G
- neg : G → G
- sub : G → G → G
- Despite the asymmetry of - neg_eq_of_add, the symmetric version is true thanks to the involutivity of negation.
Instances
A DivisionMonoid is a DivInvMonoid with involutive inversion and such that (a * b)⁻¹ = b⁻¹ * a⁻¹ and a * b = 1 → a⁻¹ = b.
This is the immediate common ancestor of Group and GroupWithZero.
- mul : G → G → G
- one : G
- inv : G → G
- div : G → G → G
- Despite the asymmetry of - inv_eq_of_mul, the symmetric version is true thanks to the involutivity of inversion.
Instances
Commutative SubtractionMonoid.
Instances
Commutative DivisionMonoid.
This is the immediate common ancestor of CommGroup and CommGroupWithZero.
Instances
A Group is a Monoid with an operation ⁻¹ satisfying a⁻¹ * a = 1.
There is also a division operation / such that a / b = a * b⁻¹, with a default so that a / b = a * b⁻¹ holds by definition.
Use Group.ofLeftAxioms or Group.ofRightAxioms to define a group structure on a type with the minimum proof obligations.
Instances
An AddGroup is an AddMonoid with a unary - satisfying -a + a = 0.
There is also a binary operation - such that a - b = a + -b, with a default so that a - b = a + -b holds by definition.
Use AddGroup.ofLeftAxioms or AddGroup.ofRightAxioms to define an additive group structure on a type with the minimum proof obligations.
Instances
Equations
- Group.toDivisionMonoid = { toDivInvMonoid := inst✝.toDivInvMonoid, inv_inv := ⋯, mul_inv_rev := ⋯, inv_eq_of_mul := ⋯ }
Equations
- AddGroup.toSubtractionMonoid = { toSubNegMonoid := inst✝.toSubNegMonoid, neg_neg := ⋯, neg_add_rev := ⋯, neg_eq_of_add := ⋯ }
Equations
- Group.toCancelMonoid = { toMonoid := inst✝.toMonoid, toIsLeftCancelMul := ⋯, toIsRightCancelMul := ⋯ }
Equations
- AddGroup.toAddCancelMonoid = { toAddMonoid := inst✝.toAddMonoid, toIsLeftCancelAdd := ⋯, toIsRightCancelAdd := ⋯ }
An additive commutative group is an additive group with commutative (+).
Instances
A commutative group is a group with commutative (*).
Instances
Equations
- CommGroup.toCancelCommMonoid = { toMonoid := inst✝.toMonoid, mul_comm := ⋯, toIsLeftCancelMul := ⋯ }
Equations
- AddCommGroup.toAddCancelCommMonoid = { toAddMonoid := inst✝.toAddMonoid, add_comm := ⋯, toIsLeftCancelAdd := ⋯ }
Equations
- CommGroup.toDivisionCommMonoid = { toDivInvMonoid := inst✝.toDivInvMonoid, inv_inv := ⋯, mul_inv_rev := ⋯, inv_eq_of_mul := ⋯, mul_comm := ⋯ }
Equations
- AddCommGroup.toDivisionAddCommMonoid = { toSubNegMonoid := inst✝.toSubNegMonoid, neg_neg := ⋯, neg_add_rev := ⋯, neg_eq_of_add := ⋯, add_comm := ⋯ }
A Prop stating that the addition is commutative.
- is_comm : Std.Commutative fun (x1 x2 : M) => x1 + x2
Instances
A Prop stating that the multiplication is commutative.
- is_comm : Std.Commutative fun (x1 x2 : M) => x1 * x2
Instances
Equations
- CommMonoid.ofIsMulCommutative = { toMonoid := inst✝¹, mul_comm := ⋯ }
Equations
- AddCommMonoid.ofIsAddCommutative = { toAddMonoid := inst✝¹, add_comm := ⋯ }
Equations
- CommGroup.ofIsMulCommutative = { toGroup := inst✝¹, mul_comm := ⋯ }
Equations
- AddCommGroup.ofIsAddCommutative = { toAddGroup := inst✝¹, add_comm := ⋯ }
We initialize all projections for @[simps] here, so that we don't have to do it in later files.
Note: the lemmas generated for the npow/zpow projections will not apply to x ^ y, since the argument order of these projections doesn't match the argument order of ^. The nsmul/zsmul lemmas will be correct.