module Algebra.Group.Ab where
Abelian groups🔗
A very important class of groups (which includes most familiar examples of groups: the integers, all finite cyclic groups, etc) are those with a commutative group operation, that is, those for which Accordingly, these have a name reflecting their importance and ubiquity: They are called commutative groups. Just kidding! They’re named abelian groups, named after a person, because nothing can have self-explicative names in mathematics. It’s the law.
private variable
: Level
ℓ : Type ℓ
G
: Group-on G → Type _
Group-on-is-abelian = ∀ x y → Group-on._⋆_ G x y ≡ Group-on._⋆_ G y x
Group-on-is-abelian G
: (g : Group-on G) → is-prop (Group-on-is-abelian g)
Group-on-is-abelian-is-prop = Π-is-hlevel² 1 λ _ _ → g .Group-on.has-is-set _ _ Group-on-is-abelian-is-prop g
This module does the usual algebraic structure dance to set up the category of Abelian groups.
record is-abelian-group (_*_ : G → G → G) : Type (level-of G) where
no-eta-equality
field
: is-group _*_
has-is-group : ∀ {x y} → x * y ≡ y * x
commutes open is-group has-is-group renaming (unit to 1g) public
: ∀ a b c d → a * b ≡ c * d → a — c ≡ d — b
equal-sum→equal-diff = commutes ∙ swizzle p inverser inversel equal-sum→equal-diff a b c d p
private unquoteDecl eqv = declare-record-iso eqv (quote is-abelian-group)
instance
H-Level-is-abelian-group: ∀ {n} {* : G → G → G} → H-Level (is-abelian-group *) (suc n)
= prop-instance $ Iso→is-hlevel 1 eqv $
H-Level-is-abelian-group 1 (hlevel 1) λ x → Π-is-hlevel²' 1 λ _ _ →
Σ-is-hlevel .has-is-set x _ _ is-group
record Abelian-group-on (T : Type ℓ) : Type ℓ where
no-eta-equality
field
_*_ : T → T → T
: is-abelian-group _*_
has-is-ab open is-abelian-group has-is-ab renaming (inverse to infixl 30 _⁻¹) public
: Group-on T
Abelian→Group-on .Group-on._⋆_ = _*_
Abelian→Group-on .Group-on.has-is-group = has-is-group
Abelian→Group-on
: Group-on-is-abelian Abelian→Group-on
Abelian→Group-on-abelian _ _ = commutes
Abelian→Group-on-abelian
infixr 20 _*_
open Abelian-group-on using (Abelian→Group-on; Abelian→Group-on-abelian) public
: ∀ ℓ → Thin-structure ℓ Abelian-group-on
Abelian-group-structure .is-hom f G₁ G₂ ∣ =
∣ Abelian-group-structure ℓ (Abelian→Group-on G₁) (Abelian→Group-on G₂) f
is-group-hom .is-hom f G₁ G₂ .is-tr = hlevel 1
Abelian-group-structure ℓ .id-is-hom .is-group-hom.pres-⋆ x y = refl
Abelian-group-structure ℓ .∘-is-hom f g α β .is-group-hom.pres-⋆ x y =
Abelian-group-structure ℓ (β .is-group-hom.pres-⋆ x y) ∙ α .is-group-hom.pres-⋆ (g x) (g y)
ap f .id-hom-unique {s = s} {t} α _ = p where
Abelian-group-structure ℓ open Abelian-group-on
: s ≡ t
p ._*_ x y = α .is-group-hom.pres-⋆ x y i
p i .has-is-ab = is-prop→pathp
p i (λ i → hlevel {T = is-abelian-group (λ x y → p i ._*_ x y)} 1)
(s .has-is-ab) (t .has-is-ab) i
: ∀ ℓ → Precategory (lsuc ℓ) ℓ
Ab = Structured-objects (Abelian-group-structure ℓ)
Ab ℓ
module Ab {ℓ} = Cat.Reasoning (Ab ℓ)
instance
: ∀ {ℓ} → is-equational (Abelian-group-structure ℓ)
Ab-equational .is-equational.invert-id-hom = Groups-equational .is-equational.invert-id-hom Ab-equational
: (ℓ : Level) → Type (lsuc ℓ)
Abelian-group _ = Ab.Ob
Abelian-group
: ∀ {ℓ} → Abelian-group ℓ → Group ℓ
Abelian→Group = G .fst , Abelian→Group-on (G .snd)
Abelian→Group G
record make-abelian-group (T : Type ℓ) : Type ℓ where
no-eta-equality
field
: is-set T
ab-is-set : T → T → T
mul : T → T
inv : T
1g : ∀ x → mul 1g x ≡ x
idl : ∀ x y z → mul x (mul y z) ≡ mul (mul x y) z
assoc : ∀ x → mul (inv x) x ≡ 1g
invl : ∀ x y → mul x y ≡ mul y x
comm
: make-group T
make-abelian-group→make-group = mg where
make-abelian-group→make-group : make-group T
mg .make-group.group-is-set = ab-is-set
mg .make-group.unit = 1g
mg .make-group.mul = mul
mg .make-group.inv = inv
mg .make-group.assoc = assoc
mg .make-group.invl = invl
mg .make-group.idl = idl
mg
: Group-on T
to-group-on-ab = to-group-on make-abelian-group→make-group
to-group-on-ab
: Abelian-group-on T
to-abelian-group-on .Abelian-group-on._*_ = mul
to-abelian-group-on .Abelian-group-on.has-is-ab .is-abelian-group.has-is-group =
to-abelian-group-on .has-is-group to-group-on-ab
Group-on.Abelian-group-on.has-is-ab .is-abelian-group.commutes =
to-abelian-group-on _ _
comm
: Abelian-group ℓ
to-ab .fst ∣ = T
∣ to-ab .fst .is-tr = ab-is-set
to-ab .snd = to-abelian-group-on
to-ab
: ∀ {ℓ} → Group ℓ → Type ℓ
is-commutative-group = Group-on-is-abelian (G .snd)
is-commutative-group G
from-commutative-group: ∀ {ℓ} (G : Group ℓ)
→ is-commutative-group G
→ Abelian-group ℓ
.fst = G .fst
from-commutative-group G comm .snd .Abelian-group-on._*_ =
from-commutative-group G comm ._⋆_ (G .snd)
Group-on.snd .Abelian-group-on.has-is-ab .is-abelian-group.has-is-group =
from-commutative-group G comm .has-is-group (G .snd)
Group-on.snd .Abelian-group-on.has-is-ab .is-abelian-group.commutes =
from-commutative-group G comm _ _
comm
Grp→Ab→Grp: ∀ {ℓ} (G : Group ℓ) (c : is-commutative-group G)
→ Abelian→Group (from-commutative-group G c) ≡ G
= Σ-pathp refl go where
Grp→Ab→Grp G c : Abelian→Group-on (from-commutative-group G c .snd) ≡ G .snd
go .Group-on._⋆_ = G .snd .Group-on._⋆_
go i .Group-on.has-is-group = G .snd .Group-on.has-is-group
go i
open make-abelian-group using (make-abelian-group→make-group ; to-group-on-ab ; to-abelian-group-on ; to-ab) public
open Functor
: ∀ {ℓ} → Functor (Ab ℓ) (Groups ℓ)
Ab↪Grp .F₀ = Abelian→Group
Ab↪Grp .F₁ f .hom = f .hom
Ab↪Grp .F₁ f .preserves = f .preserves
Ab↪Grp .F-id = trivial!
Ab↪Grp .F-∘ f g = trivial!
Ab↪Grp
: ∀ {ℓ} → is-fully-faithful (Ab↪Grp {ℓ})
Ab↪Grp-is-ff {x = A} {B} = is-iso→is-equiv $ iso
Ab↪Grp-is-ff (λ _ → trivial!) (λ _ → trivial!)
promote where
: Groups.Hom (Abelian→Group A) (Abelian→Group B) → Ab.Hom A B
promote .hom = f .hom
promote f .preserves = f .preserves
promote f
: ∀ {ℓ} → Functor (Ab ℓ) (Sets ℓ)
Ab↪Sets = Grp↪Sets F∘ Ab↪Grp Ab↪Sets
The fundamental example of an abelian group is the group of integers.
Given an abelian group we can define the negation automorphism which inverts every element: since the group operation is commutative, we have so this is a homomorphism.
module _ {ℓ} (G : Abelian-group ℓ) where
open Abelian-group-on (G .snd)
: G Ab.≅ G
negation = total-iso
negation (_⁻¹ , is-involutive→is-equiv (λ _ → inv-inv))
(record { pres-⋆ = λ x y → inv-comm ∙ commutes })