module Algebra.Group.Cat.FinitelyComplete {ℓ} where
Finite limits of groups🔗
We present explicit computations of finite limits in the category of groups, though do note that the terminal group is also initial (i.e. it is a zero object). Knowing that the category of groups admits a right adjoint functor into the category of sets (the underlying set functor) drives us in computing limits of groups as limits of sets, and equipping those with a group structure: we are forced to do this since right adjoints preserve limits.
The zero group🔗
The zero object in the category of groups is given by the unit type, equipped with its unique group structure. Correspondingly, we may refer to this group in prose either as or as
: ∀ {ℓ} → Group ℓ
Zero-group = to-group zg where
Zero-group : make-group (Lift _ ⊤)
zg .make-group.group-is-set x y p q i j = lift tt
zg .make-group.unit = lift tt
zg .make-group.mul = λ x y → lift tt
zg .make-group.inv x = lift tt
zg .make-group.assoc x y z = refl
zg .make-group.invl x = refl
zg .make-group.idl x = refl
zg
: is-initial (Groups ℓ) Zero-group
Zero-group-is-initial (_ , G) .centre = total-hom (λ x → G.unit) gh where
Zero-group-is-initial module G = Group-on G
: is-group-hom _ _ (λ x → G.unit)
gh .pres-⋆ x y =
gh .unit ≡˘⟨ G.idl ⟩
G.unit G.⋆ G.unit ∎
G(_ , G) .paths x =
Zero-group-is-initial λ _ → sym (is-group-hom.pres-id (x .preserves))
ext
: is-terminal (Groups ℓ) Zero-group
Zero-group-is-terminal _ .centre =
Zero-group-is-terminal (λ _ → lift tt) record { pres-⋆ = λ _ _ _ → lift tt }
total-hom _ .paths x = trivial!
Zero-group-is-terminal
: is-zero (Groups ℓ) Zero-group
Zero-group-is-zero = record
Zero-group-is-zero { has-is-initial = Zero-group-is-initial
; has-is-terminal = Zero-group-is-terminal
}
: Zero (Groups ℓ)
∅ᴳ .Zero.∅ = Zero-group
∅ᴳ .Zero.has-is-zero = Zero-group-is-zero ∅ᴳ
Direct products🔗
We compute the product of two groups as the product of their underlying sets, equipped with the operation of “pointwise addition”.
: Group ℓ → Group ℓ → Group ℓ
Direct-product (G , Gg) (H , Hg) = to-group G×Hg where
Direct-product module G = Group-on Gg
module H = Group-on Hg
: make-group (∣ G ∣ × ∣ H ∣)
G×Hg .make-group.group-is-set = hlevel 2
G×Hg .make-group.unit = G.unit , H.unit
G×Hg .make-group.mul (a , x) (b , y) = a G.⋆ b , x H.⋆ y
G×Hg .make-group.inv (a , x) = a G.⁻¹ , x H.⁻¹
G×Hg .make-group.assoc x y z = ap₂ _,_ G.associative H.associative
G×Hg .make-group.invl x = ap₂ _,_ G.inversel H.inversel
G×Hg .make-group.idl x = ap₂ _,_ G.idl H.idl G×Hg
The projection maps and universal factoring are all given exactly as for the category of sets.
: Groups.Hom (Direct-product G H) G
proj₁ .hom = fst
proj₁ .preserves .pres-⋆ x y = refl
proj₁
: Groups.Hom (Direct-product G H) H
proj₂ .hom = snd
proj₂ .preserves .pres-⋆ x y = refl
proj₂
: Groups.Hom G H → Groups.Hom G K → Groups.Hom G (Direct-product H K)
factor .hom x = f # x , g # x
factor f g .preserves .pres-⋆ x y = ap₂ _,_ (f .preserves .pres-⋆ _ _) (g .preserves .pres-⋆ _ _)
factor f g
: is-product (Groups ℓ) {G} {H} proj₁ proj₂
Direct-product-is-product {G} {H} = p where
Direct-product-is-product open is-product
: is-product _ _ _
p .⟨_,_⟩ = factor
p .π₁∘factor = Grp↪Sets-is-faithful refl
p .π₂∘factor = Grp↪Sets-is-faithful refl
p .unique other p q = Grp↪Sets-is-faithful (funext λ x →
p _,_ (happly (ap hom p) x) (happly (ap hom q) x)) ap₂
What sets the direct product of groups apart from (e.g.) the cartesian product of sets is that both “factors” embed into the direct product, by taking the identity as the other coordinate: Indeed, in the category of abelian groups, the direct product is also a coproduct.
: G Groups.↪ Direct-product G H
inj₁ {G} {H} .mor .hom x = x , H .snd .unit
inj₁ {G} {H} .mor .preserves .pres-⋆ x y = ap (_ ,_) (sym (H .snd .idl))
inj₁ {G} {H} .monic g h x = Grp↪Sets-is-faithful (funext λ e i → (x i # e) .fst)
inj₁
: H Groups.↪ Direct-product G H
inj₂ {H} {G} .mor .hom x = G .snd .unit , x
inj₂ {H} {G} .mor .preserves .pres-⋆ x y = ap (_, _) (sym (G .snd .idl))
inj₂ {H} {G} .monic g h x = Grp↪Sets-is-faithful (funext λ e i → (x i # e) .snd) inj₂
Equalisers🔗
open import Cat.Diagram.Equaliser
The equaliser of two group homomorphisms
is given by their equaliser as Set-morphisms, equipped with the evident
group structure. Indeed, we go ahead and compute the actual Equaliser
in sets, and re-use all of its
infrastructure to make an equaliser in Groups
.
module _ {G H : Group ℓ} (f g : Groups.Hom G H) where
private
module G = Group-on (G .snd)
module H = Group-on (H .snd)
module f = is-group-hom (f .preserves)
module g = is-group-hom (g .preserves)
module seq = Equaliser
(SL.Sets-equalisers
{A = G.underlying-set}
{B = H.underlying-set}
(f .hom) (g .hom))
Recall that points there are elements of the domain (here, a point together with a proof that To “lift” the group structure from to we must prove that, if and then But this follows from and being group homomorphisms:
: Group ℓ
Equaliser-group = to-group equ-group where
Equaliser-group : ∣ seq.apex ∣ → ∣ seq.apex ∣ → ∣ seq.apex ∣
equ-⋆ (a , p) (b , q) = (a G.⋆ b) , r where abstract
equ-⋆ : f # (G .snd ._⋆_ a b) ≡ g # (G .snd ._⋆_ a b)
r = f.pres-⋆ a b ·· ap₂ H._⋆_ p q ·· sym (g.pres-⋆ _ _)
r
: ∣ seq.apex ∣ → ∣ seq.apex ∣
equ-inv (x , p) = x G.⁻¹ , q where abstract
equ-inv : f # (G.inverse x) ≡ g # (G.inverse x)
q = f.pres-inv ·· ap H._⁻¹ p ·· sym g.pres-inv
q
abstract
: f # G.unit ≡ g # G.unit
invs = f.pres-id ∙ sym g.pres-id invs
Similar yoga must be done for the inverse maps and the group unit.
: make-group ∣ seq.apex ∣
equ-group .make-group.group-is-set = seq.apex .is-tr
equ-group .make-group.unit = G.unit , invs
equ-group .make-group.mul = equ-⋆
equ-group .make-group.inv = equ-inv
equ-group .make-group.assoc x y z = Σ-prop-path! G.associative
equ-group .make-group.invl x = Σ-prop-path! G.inversel
equ-group .make-group.idl x = Σ-prop-path! G.idl
equ-group
open is-equaliser
open Equaliser
We can then, pretty effortlessly, prove that the Equaliser-group
, together with the
canonical injection
equalise the group homomorphisms
and
: Equaliser (Groups ℓ) f g
Groups-equalisers .apex = Equaliser-group
Groups-equalisers .equ = total-hom fst record { pres-⋆ = λ x y → refl }
Groups-equalisers .has-is-eq .equal = Grp↪Sets-is-faithful seq.equal
Groups-equalisers .has-is-eq .universal {F = F} {e'} p = total-hom go lim-gh where
Groups-equalisers = seq.universal {F = underlying-set (F .snd)} (ap hom p)
go
: is-group-hom _ _ go
lim-gh .pres-⋆ x y = Σ-prop-path! (e' .preserves .pres-⋆ _ _)
lim-gh
.has-is-eq .factors {F = F} {p = p} = Grp↪Sets-is-faithful
Groups-equalisers (seq.factors {F = underlying-set (F .snd)} {p = ap hom p})
.has-is-eq .unique {F = F} {p = p} q = Grp↪Sets-is-faithful
Groups-equalisers (seq.unique {F = underlying-set (F .snd)} {p = ap hom p} (ap hom q))
Putting all of these constructions together, we get the proof that
Groups
is finitely complete, since we can compute pullbacks
as certain equalisers.
open import Cat.Diagram.Limit.Finite
: Finitely-complete (Groups ℓ)
Groups-finitely-complete = with-equalisers (Groups ℓ) top prod Groups-equalisers
Groups-finitely-complete where
: Terminal (Groups ℓ)
top .Terminal.top = Zero-group
top .Terminal.has⊤ = Zero-group-is-terminal
top
: ∀ A B → Product (Groups ℓ) A B
prod .Product.apex = Direct-product A B
prod A B .Product.π₁ = proj₁
prod A B .Product.π₂ = proj₂
prod A B .Product.has-is-product = Direct-product-is-product prod A B