module Cat.Diagram.Colimit.Base where
Colimits🔗
Idea🔗
Colimits are dual to limits; much like their duals, they generalize constructions in several settings to arbitrary categories. A colimit (if it exists), is the “best solution” to an “identification problem”. This is in contrast to the limit, which acts as a solution to an “equational problem”.
Therefore, we define colimits in a similar way to limits. the only difference being that we define the colimit of a diagram as a left Kan extension instead of a right Kan extension. This gives us the expected “mapping out” universal property, as opposed to the “mapping in” property associated to limits.
Note that approach to colimits is not what normally presented in introductory material. Instead, most books opt to define colimits via cocones, as they are less abstract, though harder to work with in the long run.
module _ {J : Precategory o₁ h₁} {C : Precategory o₂ h₂} (Diagram : Functor J C) where
private
module C = Precategory C
: ∀ {x : C.Ob} → (Diagram => Const x) → Diagram => const! x F∘ !F
cocone→unit unquoteDef cocone→unit = define-coherence cocone→unit
: (x : C.Ob) → Diagram => Const x → Type _
is-colimit =
is-colimit x cocone (const! x) (cocone→unit cocone)
is-lan !F Diagram
: Type _
Colimit = Lan !F Diagram Colimit
Concretely🔗
As mentioned, our definition is very abstract, meaning we can directly re-use definitions and theorems about Kan extensions in the setting of colimits. The trade-off is that while working with colimits in general is easier, working with specific colimits becomes more difficult, as the data we actually care about has been obfuscated.
One particularly egregious failure is… actually constructing
colimits. The definition in terms of Lan
hides the concrete data behind a few
abstractions, which would be very tedious to write out each time. To
work around this, we provide an auxiliary record type, make-is-colimit
, as an intermediate step
in constructing left extensions.
First, we require morphisms from the every value of the diagram to the coapex; taken as a family, we call it . Moreover, if is a morphism in the “shape” category , we require , which encodes the relevant naturality.
field
: (j : J.Ob) → C.Hom (F₀ j) coapex
ψ : ∀ {x y} (f : J.Hom x y) → ψ y C.∘ F₁ f ≡ ψ x commutes
The rest of the data ensures that is the universal family of maps with this property; if is another natural family, then each factors through the coapex by a unique universal morphism:
universal: ∀ {x : C.Ob}
→ (eps : ∀ j → C.Hom (F₀ j) x)
→ (∀ {x y} (f : J.Hom x y) → eps y C.∘ F₁ f ≡ eps x)
→ C.Hom coapex x
factors: ∀ {j : J.Ob} {x : C.Ob}
→ (eps : ∀ j → C.Hom (F₀ j) x)
→ (p : ∀ {x y} (f : J.Hom x y) → eps y C.∘ F₁ f ≡ eps x)
→ universal eps p C.∘ ψ j ≡ eps j
unique: ∀ {x : C.Ob}
→ (eps : ∀ j → C.Hom (F₀ j) x)
→ (p : ∀ {x y} (f : J.Hom x y) → eps y C.∘ F₁ f ≡ eps x)
→ (other : C.Hom coapex x)
→ (∀ j → other C.∘ ψ j ≡ eps j)
→ other ≡ universal eps p
Once we have this data, we can use it to construct a value of type
is-colimit
. The naturality
condition we required above may seem too weak, but the full naturality
condition can be derived from it and the rest of the data.
to-is-colimit: ∀ {Diagram : Functor J C} {coapex}
→ (mc : make-is-colimit Diagram coapex)
→ is-colimit Diagram coapex (to-cocone mc)
{Diagram} {coapex} mkcolim = colim where
to-is-colimit open make-is-colimit mkcolim
open is-lan
open Functor
: is-colimit Diagram coapex (to-cocone mkcolim)
colim .σ {M = M} α .η _ =
colim (α .η) (λ f → α .is-natural _ _ f ∙ C.eliml (M .F-id))
universal .σ {M = M} α .is-natural _ _ _ =
colim .idr _ ∙ C.introl (M .F-id)
C.σ-comm {α = α} = Nat-path λ j →
colim (α .η) _
factors .σ-uniq {α = α} {σ' = σ'} p = Nat-path λ _ →
colim (α .η) _ (σ' .η _) (λ j → sym (p ηₚ j)) sym $ unique
The concrete interface of make-is-colimit
is also handy for
consuming specific colimits. To enable this use case, we
provide a function which unmakes a colimit.
unmake-colimit: ∀ {D : Functor J C} {F : Functor ⊤Cat C} {eta}
→ is-lan !F D F eta
→ make-is-colimit D (Functor.F₀ F tt)
{D} {F} {eta} colim = mc module unmake-colimit where
unmake-colimit = Functor.F₀ F tt
coapex module eta = _=>_ eta
open is-lan colim
open Functor D
open make-is-colimit
open _=>_
module _ {x} (eps : ∀ j → C.Hom (F₀ j) x)
(p : ∀ {x y} (f : J.Hom x y) → eps y C.∘ F₁ f ≡ eps x)
where
: D => const! x F∘ !F
eps-nt .η = eps
eps-nt .is-natural _ _ f = p f ∙ sym (C.idl _)
eps-nt
: C.Hom coapex x
hom = σ {M = const! x} eps-nt .η tt
hom
: make-is-colimit D coapex
mc .ψ = eta.η
mc .commutes f = eta.is-natural _ _ f ∙ C.eliml (Functor.F-id F)
mc .universal = hom
mc .factors e p = σ-comm {α = eps-nt e p} ηₚ _
mc .unique {x = x} eta p other q =
mc {σ' = other-nt} (Nat-path λ j → sym (q j)) ηₚ tt
sym $ σ-uniq where
: F => const! x
other-nt .η _ = other
other-nt .is-natural _ _ _ = C.elimr (Functor.F-id F) ∙ sym (C.idl _) other-nt
We also provide a similar interface for the bundled form of colimits.
module Colimit
{J : Precategory o₁ h₁} {C : Precategory o₂ h₂} {D : Functor J C} (L : Colimit D)
where
The coapex of the colimit can be obtained by applying the extension
functor to the single object of ⊤Cat
.
: C.Ob
coapex = Ext .F₀ tt coapex
Furthermore, we can show that the apex is the colimit, in the sense
of is-colimit
, of the diagram.
You’d think this is immediate, but unfortunately, proof assistants:
is-colimit
asks for the
constant functor functor
with value coapex
to be a Kan extension, but Colimit
, being an instance of Lan
, packages an arbitrary
functor
.
Since Agda does not compare functors for -equality, we have to shuffle our data around manually. Fortunately, this isn’t a very long computation.
: D => Const coapex
cocone .η = eta .η
cocone .is-natural x y f =
cocone .is-natural x y f ∙ ap (C._∘ _) (Ext .F-id)
eta
: is-colimit D coapex cocone
has-colimit .is-lan.σ α .η = σ α .η
has-colimit .is-lan.σ α .is-natural x y f =
has-colimit (_ C.∘_) (sym (Ext .F-id)) ∙ σ α .is-natural tt tt tt
ap .is-lan.σ-comm =
has-colimit (λ _ → σ-comm ηₚ _)
Nat-path .is-lan.σ-uniq {M = M} {σ' = σ'} p =
has-colimit (λ _ → σ-uniq {σ' = nt} (Nat-path (λ j → p ηₚ j)) ηₚ _)
Nat-path where
: Ext => M
nt .η = σ' .η
nt .is-natural x y f = ap (_ C.∘_) (Ext .F-id) ∙ σ' .is-natural x y f
nt
open is-colimit has-colimit public
Uniqueness🔗
Much like limits, colimits are unique up to isomorphism. This all follows from general properties of Kan extensions, combined with the fact that natural isomorphisms between functors correspond with isomorphisms in .
colimits→inversesp: ∀ {f : C.Hom x y} {g : C.Hom y x}
→ (∀ {j : J.Ob} → f C.∘ Cx.ψ j ≡ Cy.ψ j)
→ (∀ {j : J.Ob} → g C.∘ Cy.ψ j ≡ Cx.ψ j)
→ C.Inverses f g
colimits→invertiblep: ∀ {f : C.Hom x y}
→ (∀ {j : J.Ob} → f C.∘ Cx.ψ j ≡ Cy.ψ j)
→ C.is-invertible f
: x C.≅ y
colimits-unique : C.is-invertible (Cx.universal Cy.ψ Cy.commutes)
colimits→invertible
colimits→inverses: C.Inverses (Cx.universal Cy.ψ Cy.commutes) (Cy.universal Cx.ψ Cx.commutes)
Furthermore, if the universal map is invertible, then that means its domain is also a colimit of the diagram. This also follows from a general theorem of Kan extensions, though some golfin is required to obtain the correct inverse definitionally.
is-invertible→is-colimitp: ∀ {K' : Functor ⊤Cat C} {eta : D => K' F∘ !F}
→ (eps : ∀ j → C.Hom (D.₀ j) (K' .F₀ tt))
→ (p : ∀ {x y} (f : J.Hom x y) → eps y C.∘ D.₁ f ≡ eps x)
→ (∀ {j} → eps j ≡ eta .η j)
→ C.is-invertible (Cy.universal eps p)
→ is-lan !F D K' eta
{K' = K'} {eta = eta} eps p q invert =
is-invertible→is-colimitp
generalize-colimitp(is-invertible→is-lan Cy $ invertible→invertibleⁿ _ λ _ → invert)
q
Another useful fact is that if is a colimit of some diagram , and is naturally isomorphic to some other diagram , then the coapex of is also a colimit of .
natural-iso-diagram→is-colimitp: ∀ {D' : Functor J C} {eta : D' => K F∘ !F}
→ (isos : D ≅ⁿ D')
→ (∀ {j} → Cy.ψ j C.∘ Isoⁿ.from isos .η j ≡ eta .η j)
→ is-lan !F D' K eta
{D' = D'} isos q = generalize-colimitp
natural-iso-diagram→is-colimitp (natural-iso-of→is-lan Cy isos)
q
Since is-colimit
is a
proposition, and the colimiting cocones are all unique (“up to
isomorphism”), if we’re talking about univalent categories,
then Colimit
itself is a
proposition. This is also an instance of the more general uniqueness of Kan extensions.
: is-category C → is-prop (Colimit Diagram)
Colimit-is-prop = Lan-is-prop cat Colimit-is-prop cat
Preservation of colimits🔗
The definitions here are the same idea as preservation of limits, just dualized.
: Type _
preserves-colimit =
preserves-colimit ∀ {K : Functor ⊤Cat C} {eta : Diagram => K F∘ !F}
→ (colim : is-lan !F Diagram K eta)
→ preserves-lan F colim
: Type _
reflects-colimit =
reflects-colimit ∀ {K : Functor ⊤Cat C} {eps : Diagram => K F∘ !F}
→ (lan : is-lan !F (F F∘ Diagram) (F F∘ K) (nat-assoc-to (F ▸ eps)))
→ reflects-lan F lan
Cocontinuity🔗
is-cocontinuous: ∀ (oshape hshape : Level)
{C : Precategory o₁ h₁}
{D : Precategory o₂ h₂}
→ Functor C D → Type _
A cocontinuous functor is one that, for every shape of diagram
J
, and every diagram diagram
of shape J
in
C
, preserves the colimit for that diagram.
{C = C} F =
is-cocontinuous oshape hshape ∀ {J : Precategory oshape hshape} {Diagram : Functor J C}
→ preserves-colimit F Diagram
Cocompleteness🔗
A category is cocomplete if it admits colimits for diagrams of arbitrary shape. However, in the presence of excluded middle, if a category admits coproducts indexed by its class of morphisms, then it is automatically thin. Since excluded middle is independent of type theory, we can not prove that any non-thin categories have arbitrary colimits.
Instead, categories are cocomplete with respect to a pair of universes: A category is -cocomplete if it has colimits for any diagram indexed by a precategory with objects in and morphisms in .
: ∀ {oc ℓc} o ℓ → Precategory oc ℓc → Type _
is-cocomplete = ∀ {J : Precategory oj ℓj} (F : Functor J C) → Colimit F is-cocomplete oj ℓj C
While this condition might sound very strong, and thus that it would be hard to come by, it turns out we can get away with only two fundamental types of colimits: coproducts and coequalisers. In order to construct the colimit for a diagram of shape , we will require coproducts indexed by ’s type of objects and by its type of morphisms.
colimit-as-coequaliser-of-coproduct: ∀ {oj ℓj} {J : Precategory oj ℓj}
→ has-coproducts-indexed-by C (Precategory.Ob J)
→ has-coproducts-indexed-by C (Arrows J)
→ has-coequalisers C
→ (F : Functor J C) → Colimit F
{oj} {ℓj} {J} has-Ob-cop has-Arrows-cop has-coeq F =
colimit-as-coequaliser-of-coproduct (to-is-colimit colim) where to-colimit
Given a diagram , we start by building the coproduct of all the objects appearing in the diagram.
: Indexed-coproduct C λ o → F₀ o
Obs = has-Ob-cop _ Obs
Our colimit will arise as a quotient object of this coproduct-of-objects, namely the coequaliser of two carefully chosen morphisms.
As a guiding example, the pushout of and should be the quotient of by the equivalence relation generated by . In full generality, for each arrow in our diagram, we should have that injecting into the th component of our coproduct should give the same result as precomposing with and injecting into the th component.
This suggests to build another indexed coproduct of all the domains of arrows in the diagram, taking the first morphism to be the injection into the domain component and the second morphism to be the injection into the codomain component precomposed with :
: Indexed-coproduct C {Idx = Arrows J} λ (a , b , f) → F₀ a
Dom = has-Arrows-cop _
Dom
: C.Hom (Dom .ΣF) (Obs .ΣF)
s t = Dom .match λ (a , b , f) → Obs .ι b C.∘ F₁ f
s = Dom .match λ (a , b , f) → Obs .ι a
t
: Coequaliser C s t
coequ = has-coeq _ _
coequ
: make-is-colimit F (coequ .coapex) colim
The rest of the proof amounts to repackaging the data of the coequaliser and coproducts as the data for a colimit.
.ψ c = coequ .coeq C.∘ Obs .ι c
colim .commutes {a} {b} f =
colim (coequ .coeq C.∘ Obs .ι b) C.∘ F₁ f ≡˘⟨ C.extendr (Dom .commute) ⟩
.coeq C.∘ s ⌝ C.∘ Dom .ι (a , b , f) ≡⟨ ap! (coequ .coequal) ⟩
⌜ coequ (coequ .coeq C.∘ t) C.∘ Dom .ι (a , b , f) ≡⟨ C.pullr (Dom .commute) ⟩
.coeq C.∘ Obs .ι a ∎
coequ .universal {x} e comm = coequ .universal comm' where
colim : C.Hom (Obs .ΣF) x
e' = Obs .match e
e' : e' C.∘ s ≡ e' C.∘ t
comm' = Indexed-coproduct.unique₂ Dom λ i@(a , b , f) →
comm' (e' C.∘ s) C.∘ Dom .ι i ≡⟨ C.extendr (Dom .commute) ⟩
.∘ Obs .ι b ⌝ C.∘ F₁ f ≡⟨ ap! (Obs .commute) ⟩
⌜ e' C.∘ F₁ f ≡⟨ comm f ⟩
e b C.commute ⟩
e a ≡˘⟨ Obs .∘ Obs .ι a ≡˘⟨ C.pullr (Dom .commute) ⟩
e' C(e' C.∘ t) C.∘ Dom .ι i ∎
.factors {j} e comm =
colim .universal e comm C.∘ (coequ .coeq C.∘ Obs .ι j) ≡⟨ C.pulll (coequ .factors) ⟩
colim .match e C.∘ Obs .ι j ≡⟨ Obs .commute ⟩
Obs
e j ∎.unique e comm u' fac = coequ .unique $ Obs .unique _
colim λ i → sym (C.assoc _ _ _) ∙ fac i
This implies that a category with coequalisers and large enough indexed coproducts has all colimits.
coproducts+coequalisers→cocomplete: ∀ {oj ℓj}
→ has-indexed-coproducts C (oj ⊔ ℓj)
→ has-coequalisers C
→ is-cocomplete oj ℓj C
{oj} {ℓj} has-cop has-coeq =
coproducts+coequalisers→cocomplete
colimit-as-coequaliser-of-coproduct(λ _ → Lift-Indexed-coproduct C ℓj (has-cop _))
(λ _ → has-cop _)
has-coeq