module Algebra.Group.Instances.Integers where

The group of integersπŸ”—

The fundamental example of an abelian group is the group of integers, under addition.

β„€-ab : Abelian-group lzero
β„€-ab = to-ab mk-β„€ where
  open make-abelian-group
  mk-β„€ : make-abelian-group Int
  mk-β„€ .ab-is-set = hlevel 2
  mk-β„€ .mul x y = x +β„€ y
  mk-β„€ .inv x = negβ„€ x
  mk-β„€ .1g = 0
  mk-β„€ .idl x = +β„€-zerol x
  mk-β„€ .assoc x y z = +β„€-assoc x y z
  mk-β„€ .invl x = +β„€-invl x
  mk-β„€ .comm x y = +β„€-commutative x y

β„€ : Group lzero
℀ = Abelian→Group ℀-ab

We show that is the free group on one generator, i.e.Β the free object on relative to the forgetful functor from groups to sets, Grpβ†ͺSets.

We start by defining the group homomorphism that sends to (or, in additive notation, where is a group and is an element of using the universal property of the integers.

  module pow (x : ⌞ G ⌟) where
    pow : Int β†’ ⌞ G ⌟
    pow = β„€.map-out unit ((_⋆ x) , ⋆-equivr x)

    pow-sucr : βˆ€ a β†’ pow (sucβ„€ a) ≑ pow a ⋆ x
    pow-sucr = β„€.map-out-rotate _ _

    pow-+ : βˆ€ a b β†’ pow (a +β„€ b) ≑ pow a ⋆ pow b
    pow-+ a = β„€.induction
      (ap pow (+β„€-zeror a) βˆ™ sym idr)
      Ξ» b β†’
        pow (a +β„€ b)        ≑ pow a ⋆ pow b        β‰ƒβŸ¨ ap (_⋆ x) , equivβ†’cancellable (⋆-equivr x) βŸ©β‰ƒ
        pow (a +β„€ b) ⋆ x    ≑ (pow a ⋆ pow b) ⋆ x  β‰ƒβŸ¨ βˆ™-post-equiv (sym associative) βŸ©β‰ƒ
        pow (a +β„€ b) ⋆ x    ≑ pow a ⋆ pow b ⋆ x    β‰ƒβŸ¨ βˆ™-post-equiv (ap (pow a ⋆_) (sym (pow-sucr b))) βŸ©β‰ƒ
        pow (a +β„€ b) ⋆ x    ≑ pow a ⋆ pow (sucβ„€ b) β‰ƒβŸ¨ βˆ™-pre-equiv (pow-sucr (a +β„€ b)) βŸ©β‰ƒ
        pow (sucβ„€ (a +β„€ b)) ≑ pow a ⋆ pow (sucβ„€ b) β‰ƒβŸ¨ βˆ™-pre-equiv (ap pow (+β„€-sucr a b)) βŸ©β‰ƒ
        pow (a +β„€ sucβ„€ b)   ≑ pow a ⋆ pow (sucβ„€ b) β‰ƒβˆŽ

A type-theoretic interjection is necessary: the integers live on the zeroth universe, so to have an group of integers, we must lift it.

    pow-hom : Groups.Hom (LiftGroup β„“ β„€) G
    pow-hom .hom (lift i) = pow i
    pow-hom .preserves .pres-⋆ (lift a) (lift b) = pow-+ a b

This is the unique group homomorphism that sends to

    pow-unique : (g : Groups.Hom (LiftGroup β„“ β„€) G) β†’ g # 1 ≑ x β†’ g ≑ pow-hom
    pow-unique g g1≑x = ext $ β„€.map-out-unique (Ξ» i β†’ g # lift i)
      (pres-id (g .preserves))
      Ξ» y β†’
        g # lift ⌜ sucβ„€ y ⌝ β‰‘βŸ¨ ap! (sym (+β„€-oner y)) βŸ©β‰‘
        g # lift (y +β„€ 1)   β‰‘βŸ¨ g .preserves .pres-⋆ (lift y) 1 βŸ©β‰‘
        g # lift y ⋆ g # 1  β‰‘βŸ¨ ap (g # lift y ⋆_) g1≑x βŸ©β‰‘
        g # lift y ⋆ x      ∎

  open pow public

We prove some other useful basic properties of exponentiation. 1

  pow-sucl : βˆ€ x a β†’ pow x (sucβ„€ a) ≑ x ⋆ pow x a
  pow-0 : βˆ€ x β†’ pow x 0 ≑ unit
  pow-1 : βˆ€ x β†’ pow x 1 ≑ x
  pow-* : βˆ€ x a b β†’ pow x (a *β„€ b) ≑ pow (pow x a) b
  pow-unit : βˆ€ n β†’ pow unit n ≑ unit
  pow-inverse : βˆ€ x n β†’ pow (x ⁻¹) n ≑ pow x n ⁻¹
  pow-sucl x a =
    pow x (sucβ„€ a)    β‰‘Λ˜βŸ¨ ap (pow x) (+β„€-onel a) βŸ©β‰‘Λ˜
    pow x (1 +β„€ a)    β‰‘βŸ¨ pow-+ x 1 a βŸ©β‰‘
    pow x 1 ⋆ pow x a β‰‘βŸ¨ ap (_⋆ pow x a) (pow-1 x) βŸ©β‰‘
    x ⋆ pow x a       ∎

  pow-0 x = refl

  pow-1 x = idl

  pow-* x a = β„€.induction (ap (pow x) (*β„€-zeror a)) Ξ» b β†’
    pow x (a *β„€ b)           ≑ pow (pow x a) b           β‰ƒβŸ¨ _ , equivβ†’cancellable (⋆-equivr _) βŸ©β‰ƒ
    pow x (a *β„€ b) ⋆ pow x a ≑ pow (pow x a) b ⋆ pow x a β‰ƒβŸ¨ βˆ™-pre-equiv (pow-+ x (a *β„€ b) a) βŸ©β‰ƒ
    pow x (a *β„€ b +β„€ a)      ≑ pow (pow x a) b ⋆ pow x a β‰ƒβŸ¨ βˆ™-pre-equiv (ap (pow x) (*β„€-sucr a b)) βŸ©β‰ƒ
    pow x (a *β„€ sucβ„€ b)      ≑ pow (pow x a) b ⋆ pow x a β‰ƒβŸ¨ βˆ™-post-equiv (sym (pow-sucr (pow x a) b)) βŸ©β‰ƒ
    pow x (a *β„€ sucβ„€ b)      ≑ pow (pow x a) (sucβ„€ b)    β‰ƒβˆŽ

  pow-unit = β„€.induction refl (Ξ» x β†’ βˆ™-pre-equiv (pow-sucr unit x βˆ™ idr))

  pow-inverse x = β„€.induction (sym inv-unit) Ξ» n β†’
    pow (x ⁻¹) n        ≑ pow x n ⁻¹        β‰ƒβŸ¨ _ , equivβ†’cancellable (⋆-equivr (x ⁻¹)) βŸ©β‰ƒ
    pow (x ⁻¹) n ⋆ x ⁻¹ ≑ pow x n ⁻¹ ⋆ x ⁻¹ β‰ƒβŸ¨ βˆ™-pre-equiv (pow-sucr (x ⁻¹) n) βŸ©β‰ƒ
    pow (x ⁻¹) (sucβ„€ n) ≑ pow x n ⁻¹ ⋆ x ⁻¹ β‰ƒβŸ¨ βˆ™-post-equiv (sym inv-comm) βŸ©β‰ƒ
    pow (x ⁻¹) (sucβ„€ n) ≑ (x ⋆ pow x n) ⁻¹  β‰ƒβŸ¨ βˆ™-post-equiv (ap _⁻¹ (sym (pow-sucl x n))) βŸ©β‰ƒ
    pow (x ⁻¹) (sucβ„€ n) ≑ pow x (sucβ„€ n) ⁻¹ β‰ƒβˆŽ

Finally, the properties above imply that is the free group on one generator.

β„€-free : Free-object Grpβ†ͺSets (el! ⊀)
β„€-free .Free-object.free = LiftGroup lzero β„€
β„€-free .Free-object.unit _ = 1
β„€-free .Free-object.fold {G} x = pow-hom G (x _)
β„€-free .Free-object.commute {G} {x} = ext Ξ» _ β†’ pow-1 G (x _)
β„€-free .Free-object.unique {G} {x} g comm =
  pow-unique G (x _) g (unext comm _)
Note

While the notation for pow makes sense in multiplicative notation, we would instead write in additive notation. In fact, we can show that coincides with the multiplication in the group of integers itself.

pow-β„€ : βˆ€ a b β†’ pow β„€ a b ≑ a *β„€ b
pow-β„€ a = β„€.induction (sym (*β„€-zeror a)) Ξ» b β†’
  pow β„€ a b        ≑ a *β„€ b      β‰ƒβŸ¨ ap (_+β„€ a) , equivβ†’cancellable (Group-on.⋆-equivr (β„€ .snd) a) βŸ©β‰ƒ
  pow β„€ a b +β„€ a   ≑ a *β„€ b +β„€ a β‰ƒβŸ¨ βˆ™-pre-equiv (pow-sucr β„€ a b) βŸ©β‰ƒ
  pow β„€ a (sucβ„€ b) ≑ a *β„€ b +β„€ a β‰ƒβŸ¨ βˆ™-post-equiv (sym (*β„€-sucr a b)) βŸ©β‰ƒ
  pow β„€ a (sucβ„€ b) ≑ a *β„€ sucβ„€ b β‰ƒβˆŽ

Order of an elementπŸ”—

We define the order of an element of a group as the minimal positive2 integer such that if it exists.

In particular, if then we have that the order of divides We define this notion first in the code, then use it to define the order of

  order-divides : ⌞ G ⌟ β†’ Nat β†’ Type β„“
  order-divides x n = pow G x (pos n) ≑ unit

  has-finite-order : ⌞ G ⌟ β†’ Type β„“
  has-finite-order x = minimal-solution Ξ» n β†’
    Positive n Γ— order-divides x n

  order : (x : ⌞ G ⌟) β†’ has-finite-order x β†’ Nat
  order x (n , _) = n

An element with finite order is also called a torsion element. A group all of whose elements are torsion is called a torsion group, while a group whose only torsion element is the unit is called torsion-free.

  is-torsion-group : Type β„“
  is-torsion-group = βˆ€ g β†’ has-finite-order g

  is-torsion-free : Type β„“
  is-torsion-free = βˆ€ g β†’ has-finite-order g β†’ g ≑ unit
Note

Our definition of torsion groups requires being able to compute the order of every element of the group. There is a weaker definition that only requires every element to have some such that the two definitions are equivalent if the group is discrete, since is well-ordered.

We quickly note that is torsion-free, since multiplication by a nonzero integer is injective.

β„€-torsion-free : is-torsion-free β„€
β„€-torsion-free n (k , (k>0 , nk≑0) , _) = *β„€-injectiver (pos k) n 0
  (Ξ» p β†’ <-not-equal k>0 (pos-injective (sym p)))
  (sym (pow-β„€ n (pos k)) βˆ™ nk≑0)

  1. Pedantically, these properties say that pow is the unique near-ring homomorphism from the initial near-ring, to the near-ring of (pointed) endofunctions of a generalisation of endomorphism rings to non-abelian groups.β†©οΈŽ

  2. Without this requirement, every element would trivially have order β†©οΈŽ