module Cat.Abelian.Base where
Abelian categories🔗
This module defines the sequence of properties which “work up to” abelian categories: Ab-enriched categories, pre-additive categories, pre-abelian categories, and abelian categories. Each concept builds on the last by adding a new categorical property on top of a precategory.
Ab-enriched categories🔗
An category is one where each set carries the structure of an Abelian group, such that the composition map is bilinear, hence extending to an Abelian group homomorphism
where the term on the left is the tensor product of the corresponding As the name implies, every such category has a canonical (made monoidal using but we do not use the language of enriched category theory in our development of Abelian categories.
record Ab-category {o ℓ} (C : Precategory o ℓ) : Type (o ⊔ lsuc ℓ) where
open Cat C public
field
: ∀ A B → Abelian-group-on (Hom A B)
Abelian-group-on-hom
module Hom {A B} = Abelian-group-on (Abelian-group-on-hom A B) renaming (_⁻¹ to inverse)
open Hom
using (zero-diff)
renaming (_—_ to _-_ ; _*_ to _+_ ; 1g to 0m)
public
: ∀ A B → Abelian-group ℓ
Hom-grp = (el (Hom A B) (Hom-set A B)) , Abelian-group-on-hom A B
Hom-grp A B
field
-- Composition is multilinear:
∘-linear-l: ∀ {A B C} (f g : Hom B C) (h : Hom A B)
→ (f ∘ h) + (g ∘ h) ≡ (f + g) ∘ h
∘-linear-r: ∀ {A B C} (f : Hom B C) (g h : Hom A B)
→ (f ∘ g) + (f ∘ h) ≡ f ∘ (g + h)
: ∀ {A B C} → Ab.Hom (Hom-grp B C ⊗ Hom-grp A B) (Hom-grp A C)
∘map {A} {B} {C} =
∘map (Hom-grp B C) (Hom-grp A B) (Hom-grp A C)
from-bilinear-map (record { map = _∘_
; pres-*l = λ x y z → sym (∘-linear-l x y z)
; pres-*r = λ x y z → sym (∘-linear-r x y z)
})
Note that from multilinearity of composition, it follows that the addition of and composition1 operations satisfy familiar algebraic identities, e.g. etc.
: ∀ {A B C} {f : Hom B C} → f ∘ 0m {A} {B} ≡ 0m
∘-zero-r {f = f} =
∘-zero-r .intror Hom.inverser ⟩
f ∘ 0m ≡⟨ Hom(f ∘ 0m - f ∘ 0m) ≡⟨ Hom.associative ⟩
f ∘ 0m + (f ∘ 0m + f ∘ 0m) - f ∘ 0m ≡⟨ ap (_- f ∘ 0m) (∘-linear-r _ _ _) ⟩
(f ∘ (0m + 0m)) - f ∘ 0m ≡⟨ ap ((_- f ∘ 0m) ⊙ (f ∘_)) Hom.idl ⟩
(f ∘ 0m) - f ∘ 0m ≡⟨ Hom.inverser ⟩
0m ∎
: ∀ {A B C} {f : Hom A B} → 0m ∘ f ≡ 0m {A} {C}
∘-zero-l {f = f} =
∘-zero-l .introl Hom.inversel ⟩
0m ∘ f ≡⟨ Hom(Hom.inverse (0m ∘ f) + 0m ∘ f) + 0m ∘ f ≡⟨ sym Hom.associative ⟩
.inverse (0m ∘ f) + (0m ∘ f + 0m ∘ f) ≡⟨ ap (Hom.inverse (0m ∘ f) +_) (∘-linear-l _ _ _) ⟩
Hom.inverse (0m ∘ f) + ((0m + 0m) ∘ f) ≡⟨ ap ((Hom.inverse (0m ∘ f) +_) ⊙ (_∘ f)) Hom.idl ⟩
Hom.inverse (0m ∘ f) + (0m ∘ f) ≡⟨ Hom.inversel ⟩
Hom
0m ∎
∘-negatel: ∀ {A B C} {g : Hom B C} {h : Hom A B}
→ Hom.inverse (g ∘ h) ≡ Hom.inverse g ∘ h
{g = g} {h} = monoid-inverse-unique Hom.has-is-monoid (g ∘ h) _ _
∘-negatel .inversel
Hom(∘-linear-l _ _ _ ∙ ap (_∘ h) Hom.inverser ∙ ∘-zero-l)
∘-negater: ∀ {A B C} {g : Hom B C} {h : Hom A B}
→ Hom.inverse (g ∘ h) ≡ g ∘ Hom.inverse h
{g = g} {h} = monoid-inverse-unique Hom.has-is-monoid (g ∘ h) _ _
∘-negater .inversel
Hom(∘-linear-r _ _ _ ∙ ap (g ∘_) Hom.inverser ∙ ∘-zero-r)
∘-minus-l: ∀ {A B C} (f g : Hom B C) (h : Hom A B)
→ (f ∘ h) - (g ∘ h) ≡ (f - g) ∘ h
=
∘-minus-l f g h (f ∘ h +_) ∘-negatel ⟩
f ∘ h - g ∘ h ≡⟨ ap (Hom.inverse g ∘ h) ≡⟨ ∘-linear-l _ _ _ ⟩
f ∘ h + (f - g) ∘ h ∎
∘-minus-r: ∀ {A B C} (f : Hom B C) (g h : Hom A B)
→ (f ∘ g) - (f ∘ h) ≡ f ∘ (g - h)
=
∘-minus-r f g h (f ∘ g +_) ∘-negater ⟩
f ∘ g - f ∘ h ≡⟨ ap (f ∘ Hom.inverse h) ≡⟨ ∘-linear-r _ _ _ ⟩
f ∘ g + (g - h) ∎ f ∘
Before moving on, we note the following property of If is an object s.t. then is a zero object.
module _ {o ℓ} {C : Precategory o ℓ} (A : Ab-category C) where
private module A = Ab-category A
: ∀ {X} → A.id {X} ≡ A.0m → is-zero C X
id-zero→zero .is-zero.has-is-initial B = contr A.0m λ h → sym $
id-zero→zero idm .intror refl ⟩
h ≡⟨ A.∘ A.id ≡⟨ A.refl⟩∘⟨ idm ⟩
h A.∘ A.0m ≡⟨ A.∘-zero-r ⟩
h A.0m ∎
A.is-zero.has-is-terminal x = contr A.0m λ h → sym $
id-zero→zero idm .introl refl ⟩
h ≡⟨ A.id A.∘ h ≡⟨ idm A.⟩∘⟨refl ⟩
A.0m A.∘ h ≡⟨ A.∘-zero-l ⟩
A.0m ∎ A
Perhaps the simplest example of an is.. any ring! In the same way that a monoid is a category with one object, and a group is a groupoid with one object, a ring is a ringoid with one object; Ringoid being another word for rather than a horizontal categorification of the drummer for the Beatles. The next simplest example is itself:
module _ where
open Ab-category
: ∀ {ℓ} → Ab-category (Ab ℓ)
Ab-ab-category .Abelian-group-on-hom A B = Ab.Abelian-group-on-hom A B
Ab-ab-category .∘-linear-l f g h = trivial!
Ab-ab-category .∘-linear-r f g h = ext λ _ →
Ab-ab-category (f .preserves .is-group-hom.pres-⋆ _ _) sym
Additive categories🔗
An is additive when its underlying category has a terminal object and finite products; By the yoga above, this implies that the terminal object is also a zero object, and the finite products coincide with finite coproducts.
record is-additive {o ℓ} (C : Precategory o ℓ) : Type (o ⊔ lsuc ℓ) where
field has-ab : Ab-category C
open Ab-category has-ab public
field
: Terminal C
has-terminal : ∀ A B → Product C A B
has-prods
: Zero C
∅ .Zero.∅ = has-terminal .Terminal.top
∅ .Zero.has-is-zero = id-zero→zero has-ab $
∅ (has-terminal .Terminal.has⊤ _) _ _
is-contr→is-prop module ∅ = Zero ∅
: ∀ {A B} → ∅.zero→ {A} {B} ≡ 0m
0m-unique = ap₂ _∘_ (∅.has⊥ _ .paths _) refl ∙ ∘-zero-l 0m-unique
Coincidence of finite products and finite coproducts leads to an object commonly called a (finite) biproduct. The coproduct coprojections are given by the pair of maps
respectively, and the comultiplication of and is given by We can calculate, for the first coprojection followed by comultiplication,
and analogously for the second coprojection followed by comultiplication.
private module Prod = Binary-products C has-prods
open Prod
: ∀ A B → Coproduct C A B
has-coprods = coprod where
has-coprods A B open Coproduct
open is-coproduct
: Coproduct C A B
coprod .coapex = A ⊗₀ B
coprod .ι₁ = ⟨ id , 0m ⟩
coprod .ι₂ = ⟨ 0m , id ⟩
coprod .has-is-coproduct .[_,_] f g = f ∘ π₁ + g ∘ π₂
coprod .has-is-coproduct .[]∘ι₁ {inj0 = f} {g} =
coprod (f ∘ π₁ + g ∘ π₂) ∘ ⟨ id , 0m ⟩ ≡⟨ sym (∘-linear-l _ _ _) ⟩
(f ∘ π₁) ∘ ⟨ id , 0m ⟩ + _ ≡⟨ Hom.elimr (pullr π₂∘⟨⟩ ∙ ∘-zero-r) ⟩
(f ∘ π₁) ∘ ⟨ id , 0m ⟩ ≡⟨ cancelr π₁∘⟨⟩ ⟩
f ∎.has-is-coproduct .[]∘ι₂ {inj0 = f} {g} =
coprod (f ∘ π₁ + g ∘ π₂) ∘ ⟨ 0m , id ⟩ ≡⟨ sym (∘-linear-l _ _ _) ⟩
_ + (g ∘ π₂) ∘ ⟨ 0m , id ⟩ ≡⟨ Hom.eliml (pullr π₁∘⟨⟩ ∙ ∘-zero-r) ⟩
(g ∘ π₂) ∘ ⟨ 0m , id ⟩ ≡⟨ cancelr π₂∘⟨⟩ ⟩
g ∎
For uniqueness, we use distributivity of composition over addition of morphisms and the universal property of the product to establish the desired equation. Check it out:
.has-is-coproduct .unique {inj0 = f} {g} {other} p q = sym $
coprod _+_ (pushl (sym p)) (pushl (sym q)) ⟩
f ∘ π₁ + g ∘ π₂ ≡⟨ ap₂ (other ∘ ⟨ id , 0m ⟩ ∘ π₁) + (other ∘ ⟨ 0m , id ⟩ ∘ π₂) ≡⟨ ∘-linear-r _ _ _ ⟩
(⟨ id , 0m ⟩ ∘ π₁ + ⟨ 0m , id ⟩ ∘ π₂) ≡⟨ elimr lemma ⟩
other ∘
other ∎where
: ⟨ id , 0m ⟩ ∘ π₁ + ⟨ 0m , id ⟩ ∘ π₂
lemma
≡ id= Prod.unique₂ {pr1 = π₁} {pr2 = π₂}
lemma (sym (∘-linear-r _ _ _) ∙ ap₂ _+_ (cancell π₁∘⟨⟩) (pulll π₁∘⟨⟩ ∙ ∘-zero-l) ∙ Hom.elimr refl)
(sym (∘-linear-r _ _ _) ∙ ap₂ _+_ (pulll π₂∘⟨⟩ ∙ ∘-zero-l) (cancell π₂∘⟨⟩) ∙ Hom.eliml refl)
(elimr refl)
(elimr refl)
module Coprod = Binary-coproducts C has-coprods
open Coprod
Thus every additive category is semiadditive.
: is-semiadditive C
additive→semiadditive .is-semiadditive.has-zero = ∅
additive→semiadditive .is-semiadditive.has-biproducts {A} {B} = bp where
additive→semiadditive open is-biproduct
: Biproduct C A B
bp .Biproduct.biapex = A ⊗₀ B
bp .Biproduct.π₁ = π₁
bp .Biproduct.π₂ = π₂
bp .Biproduct.ι₁ = ι₁
bp .Biproduct.ι₂ = ι₂
bp .Biproduct.has-is-biproduct .has-is-product = Prod.has-is-product
bp .Biproduct.has-is-biproduct .has-is-coproduct = Coprod.has-is-coproduct
bp .Biproduct.has-is-biproduct .πι₁ = π₁∘⟨⟩
bp .Biproduct.has-is-biproduct .πι₂ = π₂∘⟨⟩
bp .Biproduct.has-is-biproduct .ιπ-comm =
bp
ι₁ ∘ π₁ ∘ ι₂ ∘ π₂ ≡⟨ refl⟩∘⟨ pulll π₁∘⟨⟩ ⟩
ι₁ ∘ 0m ∘ π₂ ≡⟨ pulll ∘-zero-r ∙ ∘-zero-l ⟩
0m ≡˘⟨ pulll ∘-zero-r ∙ ∘-zero-l ⟩
ι₂ ∘ 0m ∘ π₁ ≡˘⟨ refl⟩∘⟨ pulll π₂∘⟨⟩ ⟩
ι₂ ∘ π₂ ∘ ι₁ ∘ π₁ ∎
open is-semiadditive additive→semiadditive hiding (∘-linear-l; ∘-linear-r)
As described there, every semiadditive category has its
own enrichment in commutative monoids. Since we already know that the
zero morphisms agree (0m-unique
),
it would be natural to expect that the additions also agree; this is
straightforward to check by linearity.
: ∀ {A B} {f g : Hom A B} → f +→ g ≡ f + g
enrichments-agree {f = f} {g} =
enrichments-agree (id ∘ π₁ + id ∘ π₂) ∘ (f ⊗₁ g) ∘ δ ≡⟨ ap₂ _+_ (idl _) (idl _) ⟩∘⟨refl ⟩
(π₁ + π₂) ∘ (f ⊗₁ g) ∘ δ ≡˘⟨ ∘-linear-l _ _ _ ⟩
(π₁ ∘ (f ⊗₁ g) ∘ δ + π₂ ∘ (f ⊗₁ g) ∘ δ) ≡⟨ ap₂ _+_ (pulll π₁∘⟨⟩ ∙ cancelr π₁∘⟨⟩) (pulll π₂∘⟨⟩ ∙ cancelr π₂∘⟨⟩) ⟩
f + g ∎
Therefore, in order to get an additive category from a semiadditive category, it suffices to ask for inverses for every morphism, so that each becomes a group.
module _ {o ℓ} (C : Precategory o ℓ) (semiadditive : is-semiadditive C) where
open Cat C
open is-semiadditive semiadditive
semiadditive+group→additive: (inv : ∀ {A B} → Hom A B → Hom A B)
→ (invl : ∀ {A B} {f : Hom A B} → inv f +→ f ≡ zero→)
→ is-additive C
.is-additive.has-ab = ab where
semiadditive+group→additive inv invl : ∀ {A B} → make-abelian-group (Hom A B)
mk .make-abelian-group.ab-is-set = hlevel 2
mk .make-abelian-group.mul = _+→_
mk .make-abelian-group.inv = inv
mk .make-abelian-group.1g = zero→
mk .make-abelian-group.idl _ = +-idl
mk .make-abelian-group.assoc _ _ _ = +-assoc
mk .make-abelian-group.invl _ = invl
mk .make-abelian-group.comm _ _ = +-comm
mk
: Ab-category C
ab .Ab-category.Abelian-group-on-hom _ _ = to-abelian-group-on mk
ab .Ab-category.∘-linear-l _ _ _ = ∘-linear-l
ab .Ab-category.∘-linear-r _ _ _ = ∘-linear-r
ab
.is-additive.has-terminal = terminal
semiadditive+group→additive inv invl .is-additive.has-prods _ _ = Biprod.product semiadditive+group→additive inv invl
Pre-abelian & abelian categories🔗
An additive category is pre-abelian when it additionally has kernels and cokernels, hence binary equalisers and coequalisers where one of the maps is zero.
record is-pre-abelian {o ℓ} (C : Precategory o ℓ) : Type (o ⊔ lsuc ℓ) where
field has-additive : is-additive C
open is-additive has-additive public
field
: ∀ {A B} (f : Hom A B) → Kernel C ∅ f
kernel : ∀ {A B} (f : Hom A B) → Coequaliser C 0m f
cokernel
module Ker {A B} (f : Hom A B) = Kernel (kernel f)
module Coker {A B} (f : Hom A B) = Coequaliser (cokernel f)
Every morphism in a preabelian category admits a canonical decomposition as
where, as indicated, the map is an epimorphism (indeed a regular epimorphism, since it is a cokernel) and the map is a regular monomorphism.
decompose: ∀ {A B} (f : Hom A B)
→ Σ[ f' ∈ Hom (Coker.coapex (Ker.kernel f)) (Ker.ker (Coker.coeq f)) ]
(f ≡ Ker.kernel (Coker.coeq f) ∘ f' ∘ Coker.coeq (Ker.kernel f))
{A} {B} f = map , sym path
decompose where
: Hom (Coker.coapex (Ker.kernel f)) B
proj' = Coker.universal (Ker.kernel f) {e' = f} $ sym path proj'
where abstract
: f ∘ kernel f .Kernel.kernel ≡ f ∘ 0m
path = Ker.equal f
path .zero-∘r _
·· ∅_∘_ (∅.has⊥ _ .paths 0m) refl
·· ap₂ ·· ∘-zero-l ·· sym ∘-zero-r
: Hom (Coker.coapex (Ker.kernel f)) (Ker.ker (Coker.coeq f))
map = Ker.universal (Coker.coeq f) {e' = proj'} $ sym path map
The existence of the map and indeed of the maps and follow from the universal properties of kernels and cokernels. The map is the canonical quotient map and the map is the canonical subobject inclusion
where abstract
: ∅.zero→ ∘ proj' ≡ Coker.coeq f ∘ proj'
path = Coker.unique₂ (Ker.kernel f)
path {e' = 0m} (∘-zero-r ∙ sym ∘-zero-l)
(pushl (∅.zero-∘r _) ∙ pulll ( ap₂ _∘_ refl (∅.has⊤ _ .paths 0m)
)
∙ ∘-zero-r)
∙ ∘-zero-l(pullr (Coker.factors (Ker.kernel f)) ∙ sym (Coker.coequal _)
)
∙ ∘-zero-r
=
path .kernel (Coker.coeq f) ∘ map ∘ Coker.coeq (Ker.kernel f) ≡⟨ pulll (Ker.factors _) ⟩
Ker.coeq (Ker.kernel f) ≡⟨ Coker.factors _ ⟩
proj' ∘ Coker f ∎
A pre-abelian category is abelian when the map in the above decomposition is an isomorphism.
record is-abelian {o ℓ} (C : Precategory o ℓ) : Type (o ⊔ lsuc ℓ) where
field has-is-preab : is-pre-abelian C
open is-pre-abelian has-is-preab public
field
coker-ker≃ker-coker: ∀ {A B} (f : Hom A B) → is-invertible (decompose f .fst)
This implies in particular that any monomorphism is a kernel, and every epimorphism is a cokernel. Let’s investigate the case for “every mono is a kernel” first: Suppose that is some monomorphism; We’ll show that it’s isomorphic to in the slice category
module _ {A B} (f : Hom A B) (monic : is-monic f) where
private
module m = Cat (Slice C B)
The map is obtained as the composite
where the isomorphism is our canonical map from before.
: m.Hom (cut f) (cut (Ker.kernel (Coker.coeq f)))
f→kercoker ./-Hom.map = decompose f .fst ∘ Coker.coeq (Ker.kernel f)
f→kercoker ./-Hom.commutes = sym (decompose f .snd) f→kercoker
Conversely, map is the composite
where the second map arises from the universal property of the cokernel: We can map out of it with the map since (using that is mono), we have from
: m.Hom (cut (Ker.kernel (Coker.coeq f))) (cut f)
kercoker→f ./-Hom.map =
kercoker→f .universal (Ker.kernel f) {e' = id} (monic _ _ path) ∘
Coker.is-invertible.inv
coker-ker≃ker-coker f where abstract
: f ∘ id ∘ 0m ≡ f ∘ id ∘ Ker.kernel f
path =
path (f ∘_) (eliml refl) ∙ ∘-zero-r ⟩
f ∘ id ∘ 0m ≡⟨ ap .zero-∘r _ ∙ 0m-unique ⟩
0m ≡˘⟨ ∅(∅.zero→ ∘ Ker.kernel f) ≡˘⟨ Ker.equal f ⟩
.kernel f ≡⟨ ap (f ∘_) (introl refl) ⟩
f ∘ Ker.kernel f ∎ f ∘ id ∘ Ker
This is indeed a map in the slice using that both isomorphisms and coequalisers are epic to make progress.
./-Hom.commutes = path where
kercoker→f =
lemma (Coker.coeq _) (Coker.has-is-coeq _) _ _ $
is-coequaliser→is-epic (Coker.factors _)
pullr
·· elimr refl(decompose f .snd ∙ assoc _ _ _)
··
=
path (coker-ker≃ker-coker _) _ _ $
invertible→epic (f ∘ Coker.universal _ _ ∘ _) ∘ decompose f .fst ≡⟨ ap₂ _∘_ (assoc _ _ _) refl ⟩
((f ∘ Coker.universal _ _) ∘ _) ∘ decompose f .fst ≡⟨ cancelr (coker-ker≃ker-coker _ .is-invertible.invr) ⟩
.universal _ _ ≡⟨ lemma ⟩
f ∘ Coker.kernel _ ∘ decompose f .fst ∎ Ker
Using the universal property of the cokernel (both uniqueness and universality), we establish that the maps defined above are inverses in thus assemble into an isomorphism in the slice.
: cut f m.≅ cut (Ker.kernel (Coker.coeq f))
mono→kernel = m.make-iso f→kercoker kercoker→f f→kc→f kc→f→kc where
mono→kernel : f→kercoker m.∘ kercoker→f ≡ m.id
f→kc→f = ext $
f→kc→f (decompose f .fst ∘ Coker.coeq _) ∘ Coker.universal _ _ ∘ _ ≡⟨ cancel-inner lemma ⟩
.fst ∘ _ ≡⟨ coker-ker≃ker-coker f .is-invertible.invl ⟩
decompose f
id ∎where
= Coker.unique₂ _
lemma {e' = Coker.coeq (Ker.kernel f)}
(∘-zero-r ∙ sym (sym (Coker.coequal _) ∙ ∘-zero-r))
(pullr (Coker.factors (Ker.kernel f)) ∙ elimr refl)
(eliml refl)
: kercoker→f m.∘ f→kercoker ≡ m.id
kc→f→kc = ext $
kc→f→kc (Coker.universal _ _ ∘ _) ∘ decompose f .fst ∘ Coker.coeq _ ≡⟨ cancel-inner (coker-ker≃ker-coker f .is-invertible.invr) ⟩
.universal _ _ ∘ Coker.coeq _ ≡⟨ Coker.factors _ ⟩
Coker id ∎
“multiplication”↩︎