Documentation

Mathlib.Algebra.Group.Defs

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

For basic lemmas about these classes see Algebra.Group.Basic.

We register the following instances:

Notation #

def leftMul {G : Type u_1} [Mul G] :
GGG

leftMul g denotes left multiplication by g

Equations
def leftAdd {G : Type u_1} [Add G] :
GGG

leftAdd g denotes left addition by g

Equations
def rightMul {G : Type u_1} [Mul G] :
GGG

rightMul g denotes right multiplication by g

Equations
def rightAdd {G : Type u_1} [Add G] :
GGG

rightAdd g denotes right addition by g

Equations
class IsLeftCancelMul (G : Type u) [Mul G] :

A mixin for left cancellative multiplication.

  • mul_left_cancel (a b c : G) : a * b = a * cb = c

    Multiplication is left cancellative.

Instances
class IsRightCancelMul (G : Type u) [Mul G] :

A mixin for right cancellative multiplication.

  • mul_right_cancel (a b c : G) : a * b = c * ba = c

    Multiplication is right cancellative.

Instances
class IsCancelMul (G : Type u) [Mul G] extends IsLeftCancelMul G, IsRightCancelMul G :

A mixin for cancellative multiplication.

Instances
class IsLeftCancelAdd (G : Type u) [Add G] :

A mixin for left cancellative addition.

  • add_left_cancel (a b c : G) : a + b = a + cb = c

    Addition is left cancellative.

Instances
class IsRightCancelAdd (G : Type u) [Add G] :

A mixin for right cancellative addition.

  • add_right_cancel (a b c : G) : a + b = c + ba = c

    Addition is right cancellative.

Instances
class IsCancelAdd (G : Type u) [Add G] extends IsLeftCancelAdd G, IsRightCancelAdd G :

A mixin for cancellative addition.

Instances
theorem mul_left_cancel {G : Type u_1} [Mul G] [IsLeftCancelMul G] {a b c : G} :
a * b = a * cb = c
theorem add_left_cancel {G : Type u_1} [Add G] [IsLeftCancelAdd G] {a b c : G} :
a + b = a + cb = c
theorem mul_left_cancel_iff {G : Type u_1} [Mul G] [IsLeftCancelMul G] {a b c : G} :
a * b = a * c b = c
theorem add_left_cancel_iff {G : Type u_1} [Add G] [IsLeftCancelAdd G] {a b c : G} :
a + b = a + c b = c
theorem mul_right_cancel {G : Type u_1} [Mul G] [IsRightCancelMul G] {a b c : G} :
a * b = c * ba = c
theorem add_right_cancel {G : Type u_1} [Add G] [IsRightCancelAdd G] {a b c : G} :
a + b = c + ba = c
theorem mul_right_cancel_iff {G : Type u_1} [Mul G] [IsRightCancelMul G] {a b c : G} :
b * a = c * a b = c
theorem add_right_cancel_iff {G : Type u_1} [Add G] [IsRightCancelAdd G] {a b c : G} :
b + a = c + a b = c
theorem Semigroup.ext {G : Type u} {x y : Semigroup G} (mul : Mul.mul = Mul.mul) :
x = y
theorem AddSemigroup.ext {G : Type u} {x y : AddSemigroup G} (add : Add.add = Add.add) :
x = y
theorem mul_assoc {G : Type u_1} [Semigroup G] (a b c : G) :
a * b * c = a * (b * c)
theorem add_assoc {G : Type u_1} [AddSemigroup G] (a b c : G) :
a + b + c = a + (b + c)
class AddCommMagma (G : Type u) extends Add G :

A commutative additive magma is a type with an addition which commutes.

  • add : GGG
  • add_comm (a b : G) : a + b = b + a

    Addition is commutative in an commutative additive magma.

Instances
    theorem AddCommMagma.ext {G : Type u} {x y : AddCommMagma G} (add : Add.add = Add.add) :
    x = y
    class CommMagma (G : Type u) extends Mul G :

    A commutative multiplicative magma is a type with a multiplication which commutes.

    • mul : GGG
    • mul_comm (a b : G) : a * b = b * a

      Multiplication is commutative in a commutative multiplicative magma.

    Instances
      theorem CommMagma.ext {G : Type u} {x y : CommMagma G} (mul : Mul.mul = Mul.mul) :
      x = y
      theorem CommSemigroup.ext {G : Type u} {x y : CommSemigroup G} (mul : Mul.mul = Mul.mul) :
      x = y
      theorem AddCommSemigroup.ext {G : Type u} {x y : AddCommSemigroup G} (add : Add.add = Add.add) :
      x = y
      theorem mul_comm {G : Type u_1} [CommMagma G] (a b : G) :
      a * b = b * a
      theorem add_comm {G : Type u_1} [AddCommMagma G] (a b : G) :
      a + b = b + a
      class MulOneClass (M : Type u) extends One M, Mul M :

      Typeclass for expressing that a type M with multiplication and a one satisfies 1 * a = a and a * 1 = a for all a : M.

      • one : M
      • mul : MMM
      • one_mul (a : M) : 1 * a = a

        One is a left neutral element for multiplication

      • mul_one (a : M) : a * 1 = a

        One is a right neutral element for multiplication

      Instances
      class AddZeroClass (M : Type u) extends Zero M, Add M :

      Typeclass for expressing that a type M with addition and a zero satisfies 0 + a = a and a + 0 = a for all a : M.

      • zero : M
      • add : MMM
      • zero_add (a : M) : 0 + a = a

        Zero is a left neutral element for addition

      • add_zero (a : M) : a + 0 = a

        Zero is a right neutral element for addition

      Instances
      theorem AddZeroClass.ext {M : Type u} ⦃m₁ m₂ : AddZeroClass M :
      Add.add = Add.addm₁ = m₂
      theorem MulOneClass.ext {M : Type u} ⦃m₁ m₂ : MulOneClass M :
      Mul.mul = Mul.mulm₁ = m₂
      @[simp]
      theorem one_mul {M : Type u} [MulOneClass M] (a : M) :
      1 * a = a
      @[simp]
      theorem zero_add {M : Type u} [AddZeroClass M] (a : M) :
      0 + a = a
      @[simp]
      theorem mul_one {M : Type u} [MulOneClass M] (a : M) :
      a * 1 = a
      @[simp]
      theorem add_zero {M : Type u} [AddZeroClass M] (a : M) :
      a + 0 = a
      def npowRec {M : Type u} [One M] [Mul M] :
      MM

      The fundamental power operation in a monoid. npowRec n a = a*a*...*a n times. Use instead a ^ n, which has better definitional behavior.

      Equations
      def nsmulRec {M : Type u} [Zero M] [Add M] :
      MM

      The fundamental scalar multiplication in an additive monoid. nsmulRec n a = a+a+...+a n times. Use instead n • a, which has better definitional behavior.

      Equations
      theorem npowRec_add {M : Type u} [One M] [Semigroup M] (m n : ) (hn : n 0) (a : M) (ha : 1 * a = a) :
      npowRec (m + n) a = npowRec m a * npowRec n a
      theorem nsmulRec_add {M : Type u} [Zero M] [AddSemigroup M] (m n : ) (hn : n 0) (a : M) (ha : 0 + a = a) :
      nsmulRec (m + n) a = nsmulRec m a + nsmulRec n a
      theorem npowRec_succ {M : Type u} [One M] [Semigroup M] (n : ) (hn : n 0) (a : M) (ha : 1 * a = a) :
      npowRec (n + 1) a = a * npowRec n a
      theorem nsmulRec_succ {M : Type u} [Zero M] [AddSemigroup M] (n : ) (hn : n 0) (a : M) (ha : 0 + a = a) :
      nsmulRec (n + 1) a = a + nsmulRec n a

      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.

      def npowBinRec {M : Type u_2} [One M] [Mul M] (k : ) :
      MM

      Exponentiation by repeated squaring.

      Equations
      def nsmulBinRec {M : Type u_2} [Zero M] [Add M] (k : ) :
      MM

      Scalar multiplication by repeated self-addition, the additive version of exponentiation by repeated squaring.

      Equations
      def nsmulBinRec.go {M : Type u_2} [Add M] (k : ) :
      MMM

      Auxiliary tail-recursive implementation for nsmulBinRec.

      Equations
      def npowBinRec.go {M : Type u_2} [Mul M] (k : ) :
      MMM

      Auxiliary tail-recursive implementation for npowBinRec.

      Equations
      • npowBinRec.go k = Nat.binaryRec (motive := fun (x : ) => MMM) (fun (y x : M) => y) (fun (bn : Bool) (_n : ) (fn : MMM) (y x : M) => fn (bif bn then y * x else y) (x * x)) k
      def npowRec' {M : Type u_2} [One M] [Mul M] :
      MM

      A variant of npowRec which is a semigroup homomorphisms from ℕ₊ to M.

      Equations
      def nsmulRec' {M : Type u_2} [Zero M] [Add M] :
      MM

      A variant of nsmulRec which is a semigroup homomorphisms from ℕ₊ to M.

      Equations
      theorem npowRec'_succ {M : Type u_2} [Semigroup M] [One M] {k : } :
      k 0∀ (m : M), npowRec' (k + 1) m = npowRec' k m * m
      theorem nsmulRec'_succ {M : Type u_2} [AddSemigroup M] [Zero M] {k : } :
      k 0∀ (m : M), nsmulRec' (k + 1) m = nsmulRec' k m + m
      theorem npowRec'_two_mul {M : Type u_2} [Semigroup M] [One M] (k : ) (m : M) :
      npowRec' (2 * k) m = npowRec' k (m * m)
      theorem nsmulRec'_two_add {M : Type u_2} [AddSemigroup M] [Zero M] (k : ) (m : M) :
      nsmulRec' (2 * k) m = nsmulRec' k (m + m)
      theorem npowRec'_mul_comm {M : Type u_2} [Semigroup M] [One M] {k : } (k0 : k 0) (m : M) :
      m * npowRec' k m = npowRec' k m * m
      theorem nsmulRec'_add_comm {M : Type u_2} [AddSemigroup M] [Zero M] {k : } (k0 : k 0) (m : M) :
      m + nsmulRec' k m = nsmulRec' k m + m
      theorem npowRec_eq {M : Type u_2} [Semigroup M] [One M] (k : ) (m : M) :
      npowRec (k + 1) m = 1 * npowRec' (k + 1) m
      theorem nsmulRec_eq {M : Type u_2} [AddSemigroup M] [Zero M] (k : ) (m : M) :
      nsmulRec (k + 1) m = 0 + nsmulRec' (k + 1) m
      theorem npowBinRec.go_spec {M : Type u_2} [Semigroup M] [One M] (k : ) (m n : M) :
      go (k + 1) m n = m * npowRec' (k + 1) n
      theorem nsmulBinRec.go_spec {M : Type u_2} [AddSemigroup M] [Zero M] (k : ) (m n : M) :
      go (k + 1) m n = m + nsmulRec' (k + 1) n
      @[reducible, inline]
      abbrev npowRecAuto {M : Type u_2} [Semigroup M] [One M] (k : ) (m : M) :
      M

      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
      @[reducible, inline]
      abbrev nsmulRecAuto {M : Type u_2} [AddSemigroup M] [Zero M] (k : ) (m : M) :
      M

      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
      @[reducible, inline]
      abbrev npowBinRecAuto {M : Type u_2} [Semigroup M] [One M] (k : ) (m : M) :
      M

      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
      @[reducible, inline]
      abbrev nsmulBinRecAuto {M : Type u_2} [AddSemigroup M] [Zero M] (k : ) (m : M) :
      M

      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
      class AddMonoid (M : Type u) extends AddSemigroup M, AddZeroClass M :

      An AddMonoid is an AddSemigroup with an element 0 such that 0 + a = a + 0 = a.

      Instances
      class Monoid (M : Type u) extends Semigroup M, MulOneClass M :

      A Monoid is a Semigroup with an element 1 such that 1 * a = a * 1 = a.

      Instances
      @[defaultInstance 10000]
      instance Monoid.toNatPow {M : Type u_2} [Monoid M] :
      Equations
      @[simp]
      theorem npow_eq_pow {M : Type u_2} [Monoid M] (n : ) (x : M) :
      Monoid.npow n x = x ^ n
      @[simp]
      theorem nsmul_eq_smul {M : Type u_2} [AddMonoid M] (n : ) (x : M) :
      theorem left_inv_eq_right_inv {M : Type u_2} [Monoid M] {a b c : M} (hba : b * a = 1) (hac : a * c = 1) :
      b = c
      theorem left_neg_eq_right_neg {M : Type u_2} [AddMonoid M] {a b c : M} (hba : b + a = 0) (hac : a + c = 0) :
      b = c
      @[simp]
      theorem pow_zero {M : Type u_2} [Monoid M] (a : M) :
      a ^ 0 = 1
      theorem zero_nsmul {M : Type u_2} [AddMonoid M] (a : M) :
      0 a = 0
      theorem pow_succ {M : Type u_2} [Monoid M] (a : M) (n : ) :
      a ^ (n + 1) = a ^ n * a
      theorem succ_nsmul {M : Type u_2} [AddMonoid M] (a : M) (n : ) :
      (n + 1) a = n a + a
      @[simp]
      theorem pow_one {M : Type u_2} [Monoid M] (a : M) :
      a ^ 1 = a
      @[simp]
      theorem one_nsmul {M : Type u_2} [AddMonoid M] (a : M) :
      1 a = a
      theorem pow_succ' {M : Type u_2} [Monoid M] (a : M) (n : ) :
      a ^ (n + 1) = a * a ^ n
      theorem succ_nsmul' {M : Type u_2} [AddMonoid M] (a : M) (n : ) :
      (n + 1) a = a + n a
      theorem pow_mul_comm' {M : Type u_2} [Monoid M] (a : M) (n : ) :
      a ^ n * a = a * a ^ n
      theorem nsmul_add_comm' {M : Type u_2} [AddMonoid M] (a : M) (n : ) :
      n a + a = a + n a
      theorem pow_two {M : Type u_2} [Monoid M] (a : M) :
      a ^ 2 = a * a

      Note that most of the lemmas about powers of two refer to it as sq.

      theorem two_nsmul {M : Type u_2} [AddMonoid M] (a : M) :
      2 a = a + a
      theorem sq {M : Type u_2} [Monoid M] (a : M) :
      a ^ 2 = a * a

      Alias of pow_two.


      Note that most of the lemmas about powers of two refer to it as sq.

      theorem pow_three' {M : Type u_2} [Monoid M] (a : M) :
      a ^ 3 = a * a * a
      theorem three'_nsmul {M : Type u_2} [AddMonoid M] (a : M) :
      3 a = a + a + a
      theorem pow_three {M : Type u_2} [Monoid M] (a : M) :
      a ^ 3 = a * (a * a)
      theorem three_nsmul {M : Type u_2} [AddMonoid M] (a : M) :
      3 a = a + (a + a)
      @[simp]
      theorem one_pow {M : Type u_2} [Monoid M] (n : ) :
      1 ^ n = 1
      theorem nsmul_zero {M : Type u_2} [AddMonoid M] (n : ) :
      n 0 = 0
      theorem pow_add {M : Type u_2} [Monoid M] (a : M) (m n : ) :
      a ^ (m + n) = a ^ m * a ^ n
      theorem add_nsmul {M : Type u_2} [AddMonoid M] (a : M) (m n : ) :
      (m + n) a = m a + n a
      theorem pow_mul_comm {M : Type u_2} [Monoid M] (a : M) (m n : ) :
      a ^ m * a ^ n = a ^ n * a ^ m
      theorem nsmul_add_comm {M : Type u_2} [AddMonoid M] (a : M) (m n : ) :
      m a + n a = n a + m a
      theorem pow_mul {M : Type u_2} [Monoid M] (a : M) (m n : ) :
      a ^ (m * n) = (a ^ m) ^ n
      theorem mul_nsmul {M : Type u_2} [AddMonoid M] (a : M) (m n : ) :
      (m * n) a = n m a
      theorem pow_mul' {M : Type u_2} [Monoid M] (a : M) (m n : ) :
      a ^ (m * n) = (a ^ n) ^ m
      theorem mul_nsmul' {M : Type u_2} [AddMonoid M] (a : M) (m n : ) :
      (m * n) a = m n a
      theorem pow_right_comm {M : Type u_2} [Monoid M] (a : M) (m n : ) :
      (a ^ m) ^ n = (a ^ n) ^ m
      theorem nsmul_left_comm {M : Type u_2} [AddMonoid M] (a : M) (m n : ) :
      n m a = m n a

      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
      class LeftCancelMonoid (M : Type u) extends Monoid M, LeftCancelSemigroup M :

      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
      class RightCancelMonoid (M : Type u) extends Monoid M, RightCancelSemigroup M :

      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
      class CancelCommMonoid (M : Type u) extends CommMonoid M, LeftCancelMonoid M :

      Commutative version of CancelMonoid.

      Instances
      @[instance 100]

      Any CancelMonoid G satisfies IsCancelMul G.

      @[instance 100]

      Any AddCancelMonoid G satisfies IsCancelAdd G.

      def zpowRec {G : Type u_1} [One G] [Mul G] [Inv G] (npow : GG := npowRec) :
      GG

      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
      def zsmulRec {G : Type u_1} [Zero G] [Add G] [Neg G] (nsmul : GG := nsmulRec) :
      GG

      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
      @[simp]
      theorem inv_inv {G : Type u_1} [InvolutiveInv G] (a : G) :
      @[simp]
      theorem neg_neg {G : Type u_1} [InvolutiveNeg G] (a : G) :
      - -a = a

      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:

      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:

      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.

      def DivInvMonoid.div' {G : Type u} [Monoid G] [Inv G] (a b : G) :
      G

      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
      class DivInvMonoid (G : Type u) extends Monoid G, Inv G, Div G :

      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.

      Instances
      def SubNegMonoid.sub' {G : Type u} [AddMonoid G] [Neg G] (a b : G) :
      G

      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
      class SubNegMonoid (G : Type u) extends AddMonoid G, Neg G, Sub G :

      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.

      Instances
      instance DivInvMonoid.toZPow {M : Type u_2} [DivInvMonoid M] :
      Equations
      class IsAddCyclic (G : Type u) [SMul G] :

      A group is called cyclic if it is generated by a single element.

      Instances
        class IsCyclic (G : Type u) [Pow G ] :

        A group is called cyclic if it is generated by a single element.

        Instances
          theorem exists_zpow_surjective (G : Type u_2) [Pow G ] [IsCyclic G] :
          ∃ (g : G), Function.Surjective fun (x : ) => g ^ x
          theorem exists_zsmul_surjective (G : Type u_2) [SMul G] [IsAddCyclic G] :
          ∃ (g : G), Function.Surjective fun (x : ) => x g
          @[simp]
          theorem zpow_eq_pow {G : Type u_1} [DivInvMonoid G] (n : ) (x : G) :
          @[simp]
          theorem zsmul_eq_smul {G : Type u_1} [SubNegMonoid G] (n : ) (x : G) :
          @[simp]
          theorem zpow_zero {G : Type u_1} [DivInvMonoid G] (a : G) :
          a ^ 0 = 1
          @[simp]
          theorem zero_zsmul {G : Type u_1} [SubNegMonoid G] (a : G) :
          0 a = 0
          @[simp]
          theorem zpow_natCast {G : Type u_1} [DivInvMonoid G] (a : G) (n : ) :
          a ^ n = a ^ n
          @[simp]
          theorem natCast_zsmul {G : Type u_1} [SubNegMonoid G] (a : G) (n : ) :
          n a = n a
          theorem zpow_ofNat {G : Type u_1} [DivInvMonoid G] (a : G) (n : ) :
          theorem ofNat_zsmul {G : Type u_1} [SubNegMonoid G] (a : G) (n : ) :
          @[simp]
          theorem zpow_negSucc {G : Type u_1} [DivInvMonoid G] (a : G) (n : ) :
          a ^ Int.negSucc n = (a ^ (n + 1))⁻¹
          @[simp]
          theorem negSucc_zsmul {G : Type u_2} [SubNegMonoid G] (a : G) (n : ) :
          Int.negSucc n a = -((n + 1) a)
          theorem div_eq_mul_inv {G : Type u_1} [DivInvMonoid G] (a b : G) :
          a / b = a * b⁻¹

          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.

          theorem sub_eq_add_neg {G : Type u_1} [SubNegMonoid G] (a b : G) :
          a - b = a + -b

          Subtracting an element is the same as adding by its negative. This is a duplicate of SubNegMonoid.sub_eq_mul_neg ensuring that the types unfold better.

          theorem division_def {G : Type u_1} [DivInvMonoid G] (a b : G) :
          a / b = a * b⁻¹

          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.

          @[simp]
          theorem zpow_one {G : Type u_1} [DivInvMonoid G] (a : G) :
          a ^ 1 = a
          @[simp]
          theorem one_zsmul {G : Type u_1} [SubNegMonoid G] (a : G) :
          1 a = a
          theorem zpow_two {G : Type u_1} [DivInvMonoid G] (a : G) :
          a ^ 2 = a * a
          theorem two_zsmul {G : Type u_1} [SubNegMonoid G] (a : G) :
          2 a = a + a
          theorem zpow_neg_one {G : Type u_1} [DivInvMonoid G] (x : G) :
          x ^ (-1) = x⁻¹
          theorem neg_one_zsmul {G : Type u_1} [SubNegMonoid G] (x : G) :
          -1 x = -x
          theorem zpow_neg_coe_of_pos {G : Type u_1} [DivInvMonoid G] (a : G) {n : } :
          0 < na ^ (-n) = (a ^ n)⁻¹
          theorem zsmul_neg_coe_of_pos {G : Type u_1} [SubNegMonoid G] (a : G) {n : } :
          0 < n-n a = -(n a)
          class NegZeroClass (G : Type u_2) extends Zero G, Neg G :
          Type u_2

          Typeclass for expressing that -0 = 0.

          Instances
          class SubNegZeroMonoid (G : Type u_2) extends SubNegMonoid G, NegZeroClass G :
          Type u_2

          A SubNegMonoid where -0 = 0.

          Instances
          class InvOneClass (G : Type u_2) extends One G, Inv G :
          Type u_2

          Typeclass for expressing that 1⁻¹ = 1.

          Instances
          class DivInvOneMonoid (G : Type u_2) extends DivInvMonoid G, InvOneClass G :
          Type u_2

          A DivInvMonoid where 1⁻¹ = 1.

          Instances
          @[simp]
          theorem inv_one {G : Type u_1} [InvOneClass G] :
          1⁻¹ = 1
          @[simp]
          theorem neg_zero {G : Type u_1} [NegZeroClass G] :
          -0 = 0
          class SubtractionMonoid (G : Type u) extends SubNegMonoid G, InvolutiveNeg G :

          A SubtractionMonoid is a SubNegMonoid with involutive negation and such that -(a + b) = -b + -a and a + b = 0 → -a = b.

          Instances
          class DivisionMonoid (G : Type u) extends DivInvMonoid G, InvolutiveInv G :

          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.

          Instances
          @[simp]
          theorem mul_inv_rev {G : Type u_1} [DivisionMonoid G] (a b : G) :
          (a * b)⁻¹ = b⁻¹ * a⁻¹
          @[simp]
          theorem neg_add_rev {G : Type u_1} [SubtractionMonoid G] (a b : G) :
          -(a + b) = -b + -a
          theorem inv_eq_of_mul_eq_one_right {G : Type u_1} [DivisionMonoid G] {a b : G} :
          a * b = 1a⁻¹ = b
          theorem neg_eq_of_add_eq_zero_right {G : Type u_1} [SubtractionMonoid G] {a b : G} :
          a + b = 0-a = b
          theorem inv_eq_of_mul_eq_one_left {G : Type u_1} [DivisionMonoid G] {a b : G} (h : a * b = 1) :
          b⁻¹ = a
          theorem neg_eq_of_add_eq_zero_left {G : Type u_1} [SubtractionMonoid G] {a b : G} (h : a + b = 0) :
          -b = a
          theorem eq_inv_of_mul_eq_one_left {G : Type u_1} [DivisionMonoid G] {a b : G} (h : a * b = 1) :
          a = b⁻¹
          theorem eq_neg_of_add_eq_zero_left {G : Type u_1} [SubtractionMonoid G] {a b : G} (h : a + b = 0) :
          a = -b

          Commutative SubtractionMonoid.

          Instances
          class DivisionCommMonoid (G : Type u) extends DivisionMonoid G, CommMonoid G :

          Commutative DivisionMonoid.

          This is the immediate common ancestor of CommGroup and CommGroupWithZero.

          Instances
          class Group (G : Type u) extends DivInvMonoid G :

          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
          class AddGroup (A : Type u) extends SubNegMonoid A :

          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
          @[simp]
          theorem inv_mul_cancel {G : Type u_1} [Group G] (a : G) :
          a⁻¹ * a = 1
          @[simp]
          theorem neg_add_cancel {G : Type u_1} [AddGroup G] (a : G) :
          -a + a = 0
          @[simp]
          theorem mul_inv_cancel {G : Type u_1} [Group G] (a : G) :
          a * a⁻¹ = 1
          @[simp]
          theorem add_neg_cancel {G : Type u_1} [AddGroup G] (a : G) :
          a + -a = 0
          @[deprecated inv_mul_cancel (since := "2024-08-12")]
          theorem mul_left_inv {G : Type u_1} [Group G] (a : G) :
          a⁻¹ * a = 1

          Alias of inv_mul_cancel.

          @[deprecated mul_inv_cancel (since := "2024-08-12")]
          theorem mul_right_inv {G : Type u_1} [Group G] (a : G) :
          a * a⁻¹ = 1

          Alias of mul_inv_cancel.

          @[deprecated neg_add_cancel (since := "2024-08-12")]
          theorem add_left_neg {G : Type u_1} [AddGroup G] (a : G) :
          -a + a = 0

          Alias of neg_add_cancel.

          @[deprecated add_neg_cancel (since := "2024-08-12")]
          theorem add_right_neg {G : Type u_1} [AddGroup G] (a : G) :
          a + -a = 0

          Alias of add_neg_cancel.

          @[deprecated inv_mul_cancel (since := "2024-08-12")]
          theorem inv_mul_self {G : Type u_1} [Group G] (a : G) :
          a⁻¹ * a = 1

          Alias of inv_mul_cancel.

          @[deprecated mul_inv_cancel (since := "2024-08-12")]
          theorem mul_inv_self {G : Type u_1} [Group G] (a : G) :
          a * a⁻¹ = 1

          Alias of mul_inv_cancel.

          @[deprecated neg_add_cancel (since := "2024-08-12")]
          theorem neg_add_self {G : Type u_1} [AddGroup G] (a : G) :
          -a + a = 0

          Alias of neg_add_cancel.

          @[deprecated add_neg_cancel (since := "2024-08-12")]
          theorem add_right_self {G : Type u_1} [AddGroup G] (a : G) :
          a + -a = 0

          Alias of add_neg_cancel.

          @[simp]
          theorem inv_mul_cancel_left {G : Type u_1} [Group G] (a b : G) :
          a⁻¹ * (a * b) = b
          @[simp]
          theorem neg_add_cancel_left {G : Type u_1} [AddGroup G] (a b : G) :
          -a + (a + b) = b
          @[simp]
          theorem mul_inv_cancel_left {G : Type u_1} [Group G] (a b : G) :
          a * (a⁻¹ * b) = b
          @[simp]
          theorem add_neg_cancel_left {G : Type u_1} [AddGroup G] (a b : G) :
          a + (-a + b) = b
          @[simp]
          theorem mul_inv_cancel_right {G : Type u_1} [Group G] (a b : G) :
          a * b * b⁻¹ = a
          @[simp]
          theorem add_neg_cancel_right {G : Type u_1} [AddGroup G] (a b : G) :
          a + b + -b = a
          @[simp]
          theorem inv_mul_cancel_right {G : Type u_1} [Group G] (a b : G) :
          a * b⁻¹ * b = a
          @[simp]
          theorem neg_add_cancel_right {G : Type u_1} [AddGroup G] (a b : G) :
          a + -b + b = a
          @[instance 100]
          Equations
          @[instance 100]
          instance Group.toCancelMonoid {G : Type u_1} [Group G] :
          Equations
          class AddCommGroup (G : Type u) extends AddGroup G, AddCommMonoid G :

          An additive commutative group is an additive group with commutative (+).

          Instances
          class CommGroup (G : Type u) extends Group G, CommMonoid G :

          A commutative group is a group with commutative (*).

          Instances
          @[simp]
          theorem inv_mul_cancel_comm {G : Type u_1} [CommGroup G] (a b : G) :
          a⁻¹ * b * a = b
          @[simp]
          theorem neg_add_cancel_comm {G : Type u_1} [AddCommGroup G] (a b : G) :
          -a + b + a = b
          @[simp]
          theorem mul_inv_cancel_comm {G : Type u_1} [CommGroup G] (a b : G) :
          a * b * a⁻¹ = b
          @[simp]
          theorem add_neg_cancel_comm {G : Type u_1} [AddCommGroup G] (a b : G) :
          a + b + -a = b
          @[simp]
          theorem inv_mul_cancel_comm_assoc {G : Type u_1} [CommGroup G] (a b : G) :
          a⁻¹ * (b * a) = b
          @[simp]
          theorem neg_add_cancel_comm_assoc {G : Type u_1} [AddCommGroup G] (a b : G) :
          -a + (b + a) = b
          @[simp]
          theorem mul_inv_cancel_comm_assoc {G : Type u_1} [CommGroup G] (a b : G) :
          a * (b * a⁻¹) = b
          @[simp]
          theorem add_neg_cancel_comm_assoc {G : Type u_1} [AddCommGroup G] (a b : G) :
          a + (b + -a) = b

          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.