Documentation

Mathlib.GroupTheory.Divisible

Divisible Group and rootable group #

In this file, we define a divisible add monoid and a rootable monoid with some basic properties.

Main definition #

Main results #

For additive monoids and groups:

and their multiplicative counterparts:

TODO: Show that divisibility implies injectivity in the category of AddCommGroup.

class DivisibleBy (A : Type u_1) (α : Type u_2) [AddMonoid A] [SMul α A] [Zero α] :
Type (max u_1 u_2)

An AddMonoid A is α-divisible iff n • x = a has a solution for all n ≠ 0 ∈ α and a ∈ A. Here we adopt a constructive approach where we ask an explicit div : A → α → A function such that

  • div a 0 = 0 for all a ∈ A
  • n • div a n = a for all n ≠ 0 ∈ α and a ∈ A.
  • div : AαA

    The division function

  • div_zero (a : A) : div a 0 = 0
  • div_cancel {n : α} (a : A) : n 0n div a n = a
Instances
class RootableBy (A : Type u_1) (α : Type u_2) [Monoid A] [Pow A α] [Zero α] :
Type (max u_1 u_2)

A Monoid A is α-rootable iff xⁿ = a has a solution for all n ≠ 0 ∈ α and a ∈ A. Here we adopt a constructive approach where we ask an explicit root : A → α → A function such that

  • root a 0 = 1 for all a ∈ A
  • (root a n)ⁿ = a for all n ≠ 0 ∈ α and a ∈ A.
  • root : AαA

    The root function

  • root_zero (a : A) : root a 0 = 1
  • root_cancel {n : α} (a : A) : n 0root a n ^ n = a
Instances
theorem pow_left_surj_of_rootableBy (A : Type u_1) (α : Type u_2) [Monoid A] [Pow A α] [Zero α] [RootableBy A α] {n : α} (hn : n 0) :
Function.Surjective fun (a : A) => a ^ n
theorem smul_right_surj_of_divisibleBy (A : Type u_1) (α : Type u_2) [AddMonoid A] [SMul α A] [Zero α] [DivisibleBy A α] {n : α} (hn : n 0) :
Function.Surjective fun (a : A) => n a
noncomputable def rootableByOfPowLeftSurj (A : Type u_1) (α : Type u_2) [Monoid A] [Pow A α] [Zero α] (H : ∀ {n : α}, n 0Function.Surjective fun (a : A) => a ^ n) :

A Monoid A is α-rootable iff the pow _ n function is surjective, i.e. the constructive version implies the textbook approach.

Equations
noncomputable def divisibleByOfSMulRightSurj (A : Type u_1) (α : Type u_2) [AddMonoid A] [SMul α A] [Zero α] (H : ∀ {n : α}, n 0Function.Surjective fun (a : A) => n a) :

An AddMonoid A is α-divisible iff n • _ is a surjective function, i.e. the constructive version implies the textbook approach.

Equations
instance Pi.rootableBy {ι : Type u_3} {β : Type u_4} (B : ιType u_5) [(i : ι) → Pow (B i) β] [Zero β] [(i : ι) → Monoid (B i)] [(i : ι) → RootableBy (B i) β] :
RootableBy ((i : ι) → B i) β
Equations
instance Pi.divisibleBy {ι : Type u_3} {β : Type u_4} (B : ιType u_5) [(i : ι) → SMul β (B i)] [Zero β] [(i : ι) → AddMonoid (B i)] [(i : ι) → DivisibleBy (B i) β] :
DivisibleBy ((i : ι) → B i) β
Equations
instance Prod.rootableBy {β : Type u_3} {B : Type u_4} {B' : Type u_5} [Pow B β] [Pow B' β] [Zero β] [Monoid B] [Monoid B'] [RootableBy B β] [RootableBy B' β] :
RootableBy (B × B') β
Equations
instance Prod.divisibleBy {β : Type u_3} {B : Type u_4} {B' : Type u_5} [SMul β B] [SMul β B'] [Zero β] [AddMonoid B] [AddMonoid B'] [DivisibleBy B β] [DivisibleBy B' β] :
DivisibleBy (B × B') β
Equations
instance ULift.instRootableBy (A : Type u_1) (α : Type u_2) [Monoid A] [Pow A α] [Zero α] [RootableBy A α] :
Equations
instance ULift.instDivisibleBy (A : Type u_1) (α : Type u_2) [AddMonoid A] [SMul α A] [Zero α] [DivisibleBy A α] :
Equations
noncomputable def AddCommGroup.divisibleByIntOfSMulTopEqTop (A : Type u_1) [AddCommGroup A] (H : ∀ {n : }, n 0n = ) :

If for all n ≠ 0 ∈ ℤ, n • A = A, then A is divisible.

Equations
@[instance 100]
instance divisibleByIntOfCharZero {𝕜 : Type u_1} [DivisionRing 𝕜] [CharZero 𝕜] :
Equations

A group is -rootable if it is -rootable.

Equations
  • One or more equations did not get rendered due to their size.

An additive group is -divisible if it is -divisible.

Equations
  • One or more equations did not get rendered due to their size.

A group is -rootable if it is -rootable

Equations

An additive group is -divisible if it -divisible.

Equations
noncomputable def Function.Surjective.rootableBy {A : Type u_1} {B : Type u_2} {α : Type u_3} [Zero α] [Monoid A] [Monoid B] [Pow A α] [Pow B α] [RootableBy A α] (f : AB) (hf : Surjective f) (hpow : ∀ (a : A) (n : α), f (a ^ n) = f a ^ n) :

If f : A → B is a surjective homomorphism and A is α-rootable, then B is also α-rootable.

Equations
noncomputable def Function.Surjective.divisibleBy {A : Type u_1} {B : Type u_2} {α : Type u_3} [Zero α] [AddMonoid A] [AddMonoid B] [SMul α A] [SMul α B] [DivisibleBy A α] (f : AB) (hf : Surjective f) (hpow : ∀ (a : A) (n : α), f (n a) = n f a) :

If f : A → B is a surjective homomorphism and A is α-divisible, then B is also α-divisible.

Equations
theorem RootableBy.surjective_pow (A : Type u_4) (α : Type u_5) [Monoid A] [Pow A α] [Zero α] [RootableBy A α] {n : α} (hn : n 0) :
Function.Surjective fun (a : A) => a ^ n
theorem DivisibleBy.surjective_smul (A : Type u_4) (α : Type u_5) [AddMonoid A] [SMul α A] [Zero α] [DivisibleBy A α] {n : α} (hn : n 0) :
Function.Surjective fun (a : A) => n a
noncomputable instance QuotientGroup.rootableBy {A : Type u_2} [CommGroup A] (B : Subgroup A) [RootableBy A ] :

Any quotient group of a rootable group is rootable.

Equations
noncomputable instance QuotientAddGroup.divisibleBy {A : Type u_2} [AddCommGroup A] (B : AddSubgroup A) [DivisibleBy A ] :

Any quotient group of a divisible group is divisible

Equations