module Cat.Diagram.Colimit.Finite where
module _ {ℓ ℓ'} (C : Precategory ℓ ℓ') where
open Cat C
Finitely cocomplete categories🔗
Finitely cocomplete categories are dual to finitely complete categories in that they admit colimits for all diagrams of finite shape.
: Typeω
is-finitely-cocomplete = ∀ {o ℓ} {D : Precategory o ℓ} → is-finite-precategory D
is-finitely-cocomplete → (F : Functor D C) → Colimit F
Equivalently, a finitely cocomplete category has:
- An initial object (colimit over the empty diagram)
- All binary coproducts (colimits over any diagrams of shape
- All binary coequalisers (colimits over any diagrams of shape
- All binary pushouts (colimits over any diagrams of shape
record Finitely-cocomplete : Type (ℓ ⊔ ℓ') where
field
: Initial C
initial : ∀ A B → Coproduct C A B
coproducts : ∀ {A B} (f g : Hom A B) → Coequaliser C f g
coequalisers : ∀ {A B X} (f : Hom X A) (g : Hom X B) → Pushout C f g
pushouts
: ∀ {A B} (f g : Hom A B) → Ob
Coequ = coequalisers f g .Coequaliser.coapex
Coequ f g
: ∀ {A B C} (f : Hom C A) (g : Hom C B) → Ob
Po = pushouts f g .Pushout.coapex
Po f g
open Finitely-cocomplete
As seen on the page for finitely complete categories, this definition equivalently states that a finitely cocomplete category has either of the following:
- An initial object, all binary coproducts, and all binary coequalisers
- An initial object and all binary pushouts
This follows from the fact that we can build coproducts and coequalisers from pushouts, and that conversely, we can build pushouts from coproducts and coequalisers. We begin by proving the latter.
With coequalisers🔗
We construct a pushout under a span as a quotient object of a coproduct i.e. the coequaliser of and where and are injections into the coproduct.
Where is some object which admits injections and from X and Y respectively. This coequaliser represents a pushout
where is the pushout’s coapex, or equivalently, the coequaliser of and
coproduct-coequaliser→pushout: ∀ {E P X Y Z} {in1 : Hom X P} {in2 : Hom Y P} {f : Hom Z X}
{g : Hom Z Y} {e : Hom P E}
→ is-coproduct C in1 in2
→ is-coequaliser C (in1 ∘ f) (in2 ∘ g) e
→ is-pushout C f (e ∘ in1) g (e ∘ in2)
{in1 = in1} {in2} {f} {g} {e} cp cq = po where
coproduct-coequaliser→pushout open is-pushout
module cq = is-coequaliser cq
module cp = is-coproduct cp
: is-pushout C _ _ _ _
po .square = sym (assoc _ _ _) ∙ cq.coequal ∙ assoc _ _ _
po .universal {i₁' = i₁'} {i₂'} p =
po .universal {e' = cp.[ i₁' , i₂' ]} (
cq.[ i₁' , i₂' ] ∘ (in1 ∘ f) ≡⟨ pulll cp.[]∘ι₁ ⟩
cp
i₁' ∘ f ≡⟨ p ⟩.[]∘ι₂ ⟩
i₂' ∘ g ≡˘⟨ pulll cp.[ i₁' , i₂' ] ∘ (in2 ∘ g) ∎
cp)
.universal∘i₁ = pulll cq.factors ∙ cp.[]∘ι₁
po .universal∘i₂ = pulll cq.factors ∙ cp.[]∘ι₂
po .unique p q =
po .unique ((cp.unique (sym (assoc _ _ _) ∙ p) (sym (assoc _ _ _) ∙ q))) cq
Thus, if a category has an initial object, binary coproducts, and binary coequalisers, it is finitely cocomplete.
with-coequalisers: Initial C
→ (∀ A B → Coproduct C A B)
→ (∀ {A B} (f g : Hom A B) → Coequaliser C f g)
→ Finitely-cocomplete
.initial = init
with-coequalisers init copr coeq .coproducts = copr
with-coequalisers init copr coeq .coequalisers = coeq
with-coequalisers init copr coeq .pushouts {A} {B} {X} f g =
with-coequalisers init copr coeq record { has-is-po = coproduct-coequaliser→pushout Copr.has-is-coproduct Coequ.has-is-coeq }
where
module Copr = Coproduct (copr A B)
module Coequ = Coequaliser (coeq (Copr.ι₁ ∘ f) (Copr.ι₂ ∘ g))
With pushouts🔗
A coproduct is a pushout under a span whose vertex is the initial object.
initial-pushout→coproduct: ∀ {P X Y I} {in1 : Hom X P} {in2 : Hom Y P} {f : Hom I X} {g : Hom I Y}
→ is-initial C I → is-pushout C f in1 g in2 → is-coproduct C in1 in2
{in1 = in1} {in2} {f} {g} init po = coprod where
initial-pushout→coproduct module Po = is-pushout po
: is-coproduct C in1 in2
coprod .is-coproduct.[_,_] in1' in2' =
coprod .universal {i₁' = in1'} {i₂' = in2'} (is-contr→is-prop (init _) _ _)
Po.is-coproduct.[]∘ι₁ = Po.universal∘i₁
coprod .is-coproduct.[]∘ι₂ = Po.universal∘i₂
coprod .is-coproduct.unique p q = Po.unique p q
coprod
with-pushouts: Initial C
→ (∀ {A B X} (f : Hom X A) (g : Hom X B) → Pushout C f g)
→ Finitely-cocomplete
= fcc where
with-pushouts bot po module bot = Initial bot
: ∀ A B → Coproduct C A B
mkcoprod = record { has-is-coproduct = initial-pushout→coproduct bot.has⊥ po' }
mkcoprod A B where po' = po (bot.has⊥ A .centre) (bot.has⊥ B .centre) .Pushout.has-is-po
: ∀ {A B} (f g : Hom A B) → Coequaliser C f g
mkcoeq {A = A} {B} f g = coequ where mkcoeq
The construction of coequalisers from pushouts follows its dual.
module A+A = Coproduct (mkcoprod A A)
: Hom A+A.coapex A
[id,id] = A+A.[ id , id ]
[id,id]
: Hom A+A.coapex B
[f,g] = A+A.[ f , g ]
[f,g]
module Po = Pushout (po [f,g] [id,id])
open is-coequaliser
open Coequaliser
: Coequaliser C f g
coequ .coapex = Po.coapex
coequ .coeq = Po.i₁
coequ .has-is-coeq .coequal =
coequ .i₁ ∘ f ≡˘⟨ ap (Po.i₁ ∘_) A+A.[]∘ι₁ ⟩
Po.i₁ ∘ [f,g] ∘ A+A.ι₁ ≡⟨ assoc _ _ _ ∙ ap (_∘ A+A.ι₁) Po.square ⟩
Po(Po.i₂ ∘ [id,id]) ∘ A+A.ι₁ ≡⟨ sym (assoc _ _ _) ∙ pushr (A+A.[]∘ι₁ ∙ sym A+A.[]∘ι₂) ⟩
(Po.i₂ ∘ [id,id]) ∘ A+A.ι₂ ≡⟨ ap (_∘ A+A.ι₂) (sym Po.square) ⟩
(Po.i₁ ∘ [f,g]) ∘ A+A.ι₂ ≡⟨ sym (assoc _ _ _) ∙ ap (Po.i₁ ∘_) A+A.[]∘ι₂ ⟩
.i₁ ∘ g ∎
Po.has-is-coeq .universal {e' = e'} p =
coequ .universal (A+A.unique₂ refl refl (in1) (in2))
Powhere
: ((e' ∘ f) ∘ [id,id]) ∘ A+A.ι₁ ≡ (e' ∘ [f,g]) ∘ A+A.ι₁
in1 =
in1 ((e' ∘ f) ∘ [id,id]) ∘ A+A.ι₁ ≡⟨ cancelr A+A.[]∘ι₁ ⟩ -- ≡⟨ cancell A+A.[]∘ι₁ ⟩
.[]∘ι₁ ⟩ -- ≡˘⟨ pulll A+A.[]∘ι₁ ⟩
e' ∘ f ≡˘⟨ pullr A+A(e' ∘ [f,g]) ∘ A+A.ι₁ ∎
: ((e' ∘ f) ∘ [id,id]) ∘ A+A.ι₂ ≡ (e' ∘ [f,g]) ∘ A+A.ι₂
in2 =
in2 ((e' ∘ f) ∘ [id,id]) ∘ A+A.ι₂ ≡⟨ cancelr A+A.[]∘ι₂ ⟩
e' ∘ f ≡⟨ p ⟩.[]∘ι₂ ⟩
e' ∘ g ≡˘⟨ pullr A+A(e' ∘ [f,g]) ∘ A+A.ι₂ ∎
.has-is-coeq .factors = Po.universal∘i₁
coequ .has-is-coeq .unique {F} {e' = e'} {colim = colim} e'=col∘i₁ =
coequ .unique e'=col∘i₁ path
Powhere
: colim ∘ Po.i₂ ≡ e' ∘ f
path =
path .i₂ ≡⟨ insertr A+A.[]∘ι₁ ⟩
colim ∘ Po((colim ∘ Po.i₂) ∘ [id,id]) ∘ A+A.ι₁ ≡⟨ ap (_∘ A+A.ι₁) (extendr (sym Po.square)) ⟩
((colim ∘ Po.i₁) ∘ [f,g]) ∘ A+A.ι₁ ≡⟨ ap (_∘ A+A.ι₁) (ap (_∘ [f,g]) e'=col∘i₁) ⟩
(e' ∘ [f,g]) ∘ A+A.ι₁ ≡⟨ pullr A+A.[]∘ι₁ ⟩
e' ∘ f ∎
: Finitely-cocomplete
fcc .initial = bot
fcc .coproducts = mkcoprod
fcc .coequalisers = mkcoeq
fcc .pushouts = po fcc
coproduct→initial-pushout: ∀ {P X Y I} {in1 : Hom X P} {in2 : Hom Y P} {f : Hom I X} {g : Hom I Y}
→ is-initial C I → is-coproduct C in1 in2 → is-pushout C f in1 g in2
= po where
coproduct→initial-pushout i r open is-pushout
: is-pushout C _ _ _ _
po .square = is-contr→is-prop (i _) _ _
po .universal _ = r .is-coproduct.[_,_] _ _
po .universal∘i₁ = r .is-coproduct.[]∘ι₁
po .universal∘i₂ = r .is-coproduct.[]∘ι₂
po .unique p q = r .is-coproduct.unique p q po
Rex functors🔗
A right exact, or rex, functor is a functor that preserves finite colimits.
module _ {o ℓ o' ℓ'} {C : Precategory o ℓ} {D : Precategory o' ℓ'} where
private module C = Cat C
private module D = Cat D
record is-rex (F : Functor C D) : Type (o ⊔ ℓ ⊔ o' ⊔ ℓ') where
private module F = Functor F
field
: ∀ {I} → is-initial C I → is-initial D (F.₀ I)
pres-⊥
pres-pushout: ∀ {P X Y Z} {in1 : C.Hom X P} {in2 : C.Hom Y P}
{f : C.Hom Z X} {g : C.Hom Z Y}
→ is-pushout C f in1 g in2
→ is-pushout D (F.₁ f) (F.₁ in1) (F.₁ g) (F.₁ in2)
pres-coproduct: ∀ {P A B I} {in1 : C.Hom A P} {in2 : C.Hom B P}
→ is-initial C I
→ is-coproduct C in1 in2
→ is-coproduct D (F.₁ in1) (F.₁ in2)
= initial-pushout→coproduct D (pres-⊥ init)
pres-coproduct init copr (pres-pushout {f = init _ .centre} {g = init _ .centre}
(coproduct→initial-pushout C init copr))
: ∀ {A B} {f : C.Hom A B} → C.is-epic f → D.is-epic (F.₁ f)
pres-epis {f = f} epi = is-pushout→is-epic
pres-epis (subst (λ x → is-pushout D _ x _ x) F.F-id
(pres-pushout
(is-epic→is-pushout epi)))