From bf86b57f07ed9c5765cf3b397d3ffa32a3098620 Mon Sep 17 00:00:00 2001 From: Reed Mullanix Date: Sun, 15 May 2022 10:13:30 -0700 Subject: [PATCH 01/14] defn: sections, retractions, split monos/epis --- src/Cat/Morphism.lagda.md | 99 +++++++++++++++++++++++++++++++++++++++ 1 file changed, 99 insertions(+) diff --git a/src/Cat/Morphism.lagda.md b/src/Cat/Morphism.lagda.md index abbbed862..dc236ce58 100644 --- a/src/Cat/Morphism.lagda.md +++ b/src/Cat/Morphism.lagda.md @@ -277,3 +277,102 @@ invertible→epic {f = f} invert g h p = open is-invertible invert ``` +## Sections and Retractions + +Given 2 maps $e : A \to B$ and $s : B \to A$, we say that $s$ is a section of +$e$ if $e \circ s = id$. The intuition behind the name is that $s$ we can think of +$e$ as taking some sort of quotient, and $s$ as picking out representatives +for each equivalence class formed by $e$. +In other words, $s$ picks out a cross-*section*. + +```agda +_is-section-of_ : (s : Hom b a) → (e : Hom a b) → Type _ +_is-section-of_ s e = e ∘ s ≡ id +``` + +Furthermore, we also that $e$ is a retraction of $s$ when $e \circ s = id$. +The intuition here is that the act of picking representatives of equivalence +classes induces an idempotent map $s \circ e : A \to A$, which *retracts* each +element of $A$ to the representative of the equivalence class formed by $e$. + +```agda +_is-retract-of_ : (e : Hom a b) → (s : Hom b a) → Type _ +_is-retract-of_ e s = e ∘ s ≡ id +``` + +It's worth showing that a retraction does indeed induce an idempotent morphism on +$A$. + +```agda +retract-idempotent : ∀ {e : Hom a b} {s : Hom b a} + → e is-retract-of s → (s ∘ e) ∘ (s ∘ e) ≡ s ∘ e +retract-idempotent {e = e} {s = s} retract = + (s ∘ e) ∘ s ∘ e ≡˘⟨ assoc s e (s ∘ e) ⟩ + s ∘ (e ∘ s ∘ e) ≡⟨ ap (s ∘_) (assoc e s e) ⟩ + s ∘ (e ∘ s) ∘ e ≡⟨ ap (λ ϕ → s ∘ ϕ ∘ e) retract ⟩ + s ∘ id ∘ e ≡⟨ ap (s ∘_) (idl e) ⟩ + s ∘ e ∎ +``` + +Earlier, I claimed that we ought to think of $e$ as taking some sort of quotient. +This is made precise by the fact that if $e$ has some section $s$, then $e$ must +be an epimorphism. + +```agda +section→is-epic : ∀ {e : Hom a b} {s : Hom b a} + → s is-section-of e → is-epic e +section→is-epic {e = e} {s = s} section g h p = + g ≡˘⟨ idr g ⟩ + g ∘ id ≡˘⟨ ap (g ∘_) section ⟩ + g ∘ e ∘ s ≡⟨ assoc g e s ⟩ + (g ∘ e) ∘ s ≡⟨ ap (_∘ s) p ⟩ + (h ∘ e) ∘ s ≡˘⟨ assoc h e s ⟩ + h ∘ e ∘ s ≡⟨ ap (h ∘_) section ⟩ + h ∘ id ≡⟨ idr h ⟩ + h ∎ +``` + +Dually, if $e$ is a retraction of $s$, then $s$ must be a monomorphism. +```agda +retract→is-monic : ∀ {e : Hom a b} {s : Hom b a} + → e is-retract-of s → is-monic s +retract→is-monic {e = e} {s = s} retract g h p = + g ≡˘⟨ idl g ⟩ + id ∘ g ≡˘⟨ ap (_∘ g) retract ⟩ + (e ∘ s) ∘ g ≡˘⟨ assoc e s g ⟩ + e ∘ s ∘ g ≡⟨ ap (e ∘_) p ⟩ + e ∘ s ∘ h ≡⟨ assoc e s h ⟩ + (e ∘ s) ∘ h ≡⟨ ap (_∘ h) retract ⟩ + id ∘ h ≡⟨ idl h ⟩ + h ∎ +``` + +## Split Epi and Monomorphisms + +The fact that the existence of a section of $f$ implies that $f$ is epic gives rise +to the notion of a split epimorphism. We can think of this +as some sort of quotient equipped with a choice of representatives. + +```agda +record is-split-epic (f : Hom a b) : Type (o ⊔ h) where + field + split : Hom b a + section : split is-section-of f + +is-split-epic→is-epic : is-split-epic f → is-epic f +is-split-epic→is-epic split-epic = section→is-epic section + where + open is-split-epic split-epic +``` + +Again by duality, we also have the notion of a split monomorphism. We can think +of these as some embedding along with some classifying map. Alternatively, we +can view a split monomorphisms as some means of selecting normal forms, along +with an evaluation map. + +```agda +record is-split-monic (f : Hom a b) : Type (o ⊔ h) where + field + split : Hom b a + retract : split is-retract-of f +``` From 9f8475de510e05a53ab58357859b85bc7efd3182 Mon Sep 17 00:00:00 2001 From: Reed Mullanix Date: Sun, 15 May 2022 10:45:42 -0700 Subject: [PATCH 02/14] =?UTF-8?q?defn:=20Add=20id=E2=82=82=20to=20Cat.Reas?= =?UTF-8?q?oning?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/Cat/Reasoning.lagda.md | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/Cat/Reasoning.lagda.md b/src/Cat/Reasoning.lagda.md index 425a0e21a..8d8cb46c5 100644 --- a/src/Cat/Reasoning.lagda.md +++ b/src/Cat/Reasoning.lagda.md @@ -37,6 +37,9 @@ id-comm {f = f} = idr f ∙ sym (idl f) id-comm-sym : ∀ {a b} {f : Hom a b} → id ∘ f ≡ f ∘ id id-comm-sym {f = f} = idl f ∙ sym (idr f) +id₂ : ∀ {a} → id ∘ id ≡ id {a} +id₂ = idl id + module _ (a≡id : a ≡ id) where eliml : a ∘ f ≡ f From 72ca5a3077bd33e3ff4575ddacb6e10c6d3f0e57 Mon Sep 17 00:00:00 2001 From: Reed Mullanix Date: Sun, 15 May 2022 10:46:48 -0700 Subject: [PATCH 03/14] defn: split coequalizers --- src/Cat/Diagram/Coequaliser/Split.lagda.md | 202 +++++++++++++++++++++ 1 file changed, 202 insertions(+) create mode 100644 src/Cat/Diagram/Coequaliser/Split.lagda.md diff --git a/src/Cat/Diagram/Coequaliser/Split.lagda.md b/src/Cat/Diagram/Coequaliser/Split.lagda.md new file mode 100644 index 000000000..43dc5cf87 --- /dev/null +++ b/src/Cat/Diagram/Coequaliser/Split.lagda.md @@ -0,0 +1,202 @@ +```agda +open import Cat.Prelude + +module Cat.Diagram.Coequaliser.Split {o ℓ} (C : Precategory o ℓ) where +open import Cat.Diagram.Coequaliser C +``` + + + +# Split Coequalizers + +Split Coequalizers are one of those definitions that seem utterly opaque when first +presented, but are actually quite natural when viewed through the correct lens. +With this in mind, we are going to provide the motivation _first_, and then +define the general construct. + +To start, let $R \subseteq B \times B$ be some equivalence relation. A natural thing +to consider is the quotient $B/R$, which gives us the following diagram: +~~~{.quiver} +\[\begin{tikzcd} + R & {B \times B} & B & {B/R} + \arrow[hook, from=1-1, to=1-2] + \arrow["{p_1}", shift left=2, from=1-2, to=1-3] + \arrow["{p_2}"', shift right=2, from=1-2, to=1-3] + \arrow["q", from=1-3, to=1-4] +\end{tikzcd}\] +~~~ + +Now, when one has a quotient, it's useful to have a means of picking representatives +for each equivalence class. This is essentially what normalization algorithms do, +which we can both agree are very useful indeed. This ends up defining a map +$s : B/R \to B$ that is a section of $q$ (IE: $q \circ s = id$). + +This gives us the following diagram (We've omited the injection of $R$ into +$B \times B$ for clarity). +~~~{.quiver} +\[\begin{tikzcd} + R & B & {B/R} + \arrow["q", shift left=2, from=1-2, to=1-3] + \arrow["s", shift left=2, from=1-3, to=1-2] + \arrow["{p_1}", shift left=2, from=1-1, to=1-2] + \arrow["{p_2}"', shift right=2, from=1-1, to=1-2] +\end{tikzcd}\] +~~~ + +This lets us define yet another map $r : B \to R$, which will witness the fact that +any $b : B$ is related to it's representative $s(q(b))$. We can define this map explicitly +as so: +$$ + r(b) = (b, s(q(b))) +$$ + +Now, how do we encode this diagramatically? To start, $p_1 \circ r = id$ by the +$\beta$ law for products. Furthermore, $p_2 \circ r = s \circ q$, also by the +$\beta$ law for products. This gives us a diagram that captures the essence of having +a quotient by an equivalence relation, along with a means of picking representatives. + +~~~{.quiver} +\[\begin{tikzcd} + R & B & {B/R} + \arrow["q", shift left=2, from=1-2, to=1-3] + \arrow["s", shift left=2, from=1-3, to=1-2] + \arrow["{p_1}", shift left=5, from=1-1, to=1-2] + \arrow["{p_2}"', shift right=5, from=1-1, to=1-2] + \arrow["r"', from=1-2, to=1-1] +\end{tikzcd}\] +~~~ + +Such a situation is called a *Split Coequaliser*. + +```agda +record is-split-coequaliser (f g : Hom A B) (e : Hom B E) + (s : Hom E B) (t : Hom B A) : Type (o ⊔ ℓ) where + field + coequal : e ∘ f ≡ e ∘ g + rep-section : e ∘ s ≡ id + witness-section : f ∘ t ≡ id + commute : s ∘ e ≡ g ∘ t +``` + +Now, let's show that this thing actually deserves the name Coequaliser. ```agda +is-split-coequaliser→is-coequalizer : + is-split-coequaliser f g e s t → is-coequaliser f g e +is-split-coequaliser→is-coequalizer + {f = f} {g = g} {e = e} {s = s} {t = t} split-coeq = + coequalizes + where + open is-split-coequaliser split-coeq +``` + +The proof here is mostly a diagram shuffle. We construct the universal +map by going back along $s$, and then following $e'$ to our destination, +like so: + +~~~{.quiver} +\[\begin{tikzcd} + A && B && E \\ + \\ + &&&& X + \arrow["q", shift left=2, from=1-3, to=1-5] + \arrow["s", shift left=2, from=1-5, to=1-3] + \arrow["{p_1}", shift left=5, from=1-1, to=1-3] + \arrow["{p_2}"', shift right=5, from=1-1, to=1-3] + \arrow["r"', from=1-3, to=1-1] + \arrow["{e'}", from=1-3, to=3-5] + \arrow["{e' \circ s}", color={rgb,255:red,214;green,92;blue,92}, from=1-5, to=3-5] +\end{tikzcd}\] +~~~ + +```agda + coequalizes : is-coequaliser f g e + coequalizes .is-coequaliser.coequal = coequal + coequalizes .is-coequaliser.coequalise {e′ = e′} _ = e′ ∘ s + coequalizes .is-coequaliser.universal {e′ = e′} {p = p} = + (e′ ∘ s) ∘ e ≡⟨ extendr commute ⟩ + (e′ ∘ g) ∘ t ≡˘⟨ p ⟩∘⟨refl ⟩ + (e′ ∘ f) ∘ t ≡⟨ cancelr witness-section ⟩ + e′ ∎ + coequalizes .is-coequaliser.unique {e′ = e′} {p} {colim′} q = + colim′ ≡⟨ intror rep-section ⟩ + colim′ ∘ e ∘ s ≡⟨ pulll (sym q) ⟩ + e′ ∘ s ∎ +``` + +Intuitively, we can think of this as constructing a map out of the quotient $B/R$ +from a map out of $B$ by picking a representative, and then applying the map out +of $B$. Universality follows by the fact that the representative is related to +the original element of $B$, and uniqueness by the fact that $s$ is a section. + +```agda +record Split-coequaliser (f g : Hom A B) : Type (o ⊔ ℓ) where + field + {coapex} : Ob + coeq : Hom B coapex + rep : Hom coapex B + witness : Hom B A + has-is-split-coeq : is-split-coequaliser f g coeq rep witness + + open is-split-coequaliser has-is-split-coeq public +``` + +## Split coequalizers and split epimorphisms + +Much like the situation with coequalizers, the coequalizing map of a +split coequalizer is a split epimorphism. This follows directly from the fact +that $s$ is a section of $e$. + +```agda +is-split-coequaliser→is-split-epic + : is-split-coequaliser f g e s t → is-split-epic e +is-split-coequaliser→is-split-epic {e = e} {s = s} split-coeq = + split-epic + where + open is-split-epic + open is-split-coequaliser split-coeq + + split-epic : is-split-epic e + split-epic .split = s + split-epic .section = rep-section +``` + +Also of note, if $e$ is a split epimorphism with splitting $s$, then +the following diagram is a split coequalizer. +~~~{.quiver} +\[\begin{tikzcd} + A & A & B + \arrow["id", shift left=3, from=1-1, to=1-2] + \arrow["{s \circ e}"', shift right=3, from=1-1, to=1-2] + \arrow["id"{description}, from=1-2, to=1-1] + \arrow["e", shift left=1, from=1-2, to=1-3] + \arrow["s", shift left=2, from=1-3, to=1-2] +\end{tikzcd}\] +~~~ + +Using the intuition that split coequalizers are quotients of equivalence relations +equipped with a choice of representatives, we can decode this diagram as constructing +an equivalence relation on $A$ out of a section of $e$, where $a ~ b$ if and only +if they get taken to the same cross section of $A$ via $s \circ e$. + +```agda +is-split-epic→is-split-coequalizer + : s is-section-of e → is-split-coequaliser id (s ∘ e) e s id +is-split-epic→is-split-coequalizer {s = s} {e = e} section = split-coeq + where + open is-split-coequaliser + + split-coeq : is-split-coequaliser id (s ∘ e) e s id + split-coeq .coequal = + e ∘ id ≡⟨ idr e ⟩ + e ≡⟨ insertl section ⟩ + e ∘ s ∘ e ∎ + split-coeq .rep-section = section + split-coeq .witness-section = id₂ + split-coeq .commute = sym (idr (s ∘ e)) +``` From b01a5368869e06e1c1d8df04c6ff144a3a965b1c Mon Sep 17 00:00:00 2001 From: Reed Mullanix Date: Sun, 15 May 2022 11:03:02 -0700 Subject: [PATCH 04/14] defn: absoluteness of split coequalisers --- .../Coequaliser/Split/Properties.lagda.md | 58 +++++++++++++++++++ 1 file changed, 58 insertions(+) create mode 100644 src/Cat/Diagram/Coequaliser/Split/Properties.lagda.md diff --git a/src/Cat/Diagram/Coequaliser/Split/Properties.lagda.md b/src/Cat/Diagram/Coequaliser/Split/Properties.lagda.md new file mode 100644 index 000000000..3495a8d03 --- /dev/null +++ b/src/Cat/Diagram/Coequaliser/Split/Properties.lagda.md @@ -0,0 +1,58 @@ +```agda +open import Cat.Prelude + +import Cat.Diagram.Coequaliser.Split as SplitCoeq +import Cat.Reasoning +import Cat.Functor.Reasoning + +module Cat.Diagram.Coequaliser.Split.Properties where +``` + +# Properties of split coequalizers + +This module proves some general properties of [split coequalisers]. + +[split coequalisers]: Cat.Diagram.Coequaliser.Split.html + +## Absoluteness + +The property of being a split coequaliser is a purely diagrammatic one, which has +the lovely property of being preserved by _all_ functors. We call such colimits +absolute. + +```agda +module _ {o o′ ℓ ℓ′} + {C : Precategory o ℓ} {D : Precategory o′ ℓ′} + (F : Functor C D) where +``` + + +The proof follows the fact that functors preserve diagrams, and reduces to a bit +of symbol shuffling. + +```agda + is-split-coequaliser-absolute + : is-split-coequaliser C f g e s t + → is-split-coequaliser D (F₁ f) (F₁ g) (F₁ e) (F₁ s) (F₁ t) + is-split-coequaliser-absolute + {f = f} {g = g} {e = e} {s = s} {t = t} split-coeq = F-split-coeq + where + open is-split-coequaliser split-coeq + + F-split-coeq : is-split-coequaliser D _ _ _ _ _ + F-split-coeq .coequal = weave coequal + F-split-coeq .rep-section = annihilate rep-section + F-split-coeq .witness-section = annihilate witness-section + F-split-coeq .commute = weave commute +``` From 4b774fb90b0014a6ae26c1ff47c3a454f3e843e4 Mon Sep 17 00:00:00 2001 From: Reed Mullanix Date: Sun, 15 May 2022 14:42:59 -0700 Subject: [PATCH 05/14] fix: tweak public exports of Finitely-complete --- src/Cat/Diagram/Limit/Finite.lagda.md | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Cat/Diagram/Limit/Finite.lagda.md b/src/Cat/Diagram/Limit/Finite.lagda.md index 9fc0d2d82..9a0fb2535 100644 --- a/src/Cat/Diagram/Limit/Finite.lagda.md +++ b/src/Cat/Diagram/Limit/Finite.lagda.md @@ -72,8 +72,8 @@ to give a terminal object and binary products. Pb : ∀ {A B C} (f : Hom A C) (g : Hom B C) → Ob Pb f g = pullbacks f g .Pullback.apex - module Cart = Cartesian C products - open Cart using (_⊗_) public + module Cart = Cartesian C products hiding (_⊗_) + open Cartesian C products using (_⊗_) public open Finitely-complete ``` From 8b8e7350dc172f49ba8ed919b41e7390cbf87d50 Mon Sep 17 00:00:00 2001 From: Reed Mullanix Date: Sun, 15 May 2022 14:43:20 -0700 Subject: [PATCH 06/14] defn: Split coequalizers induce quotients --- .../Coequaliser/Split/Properties.lagda.md | 95 +++++++++++++++++++ 1 file changed, 95 insertions(+) diff --git a/src/Cat/Diagram/Coequaliser/Split/Properties.lagda.md b/src/Cat/Diagram/Coequaliser/Split/Properties.lagda.md index 3495a8d03..09feae75a 100644 --- a/src/Cat/Diagram/Coequaliser/Split/Properties.lagda.md +++ b/src/Cat/Diagram/Coequaliser/Split/Properties.lagda.md @@ -1,7 +1,11 @@ ```agda open import Cat.Prelude +open import Cat.Diagram.Limit.Finite + import Cat.Diagram.Coequaliser.Split as SplitCoeq +import Cat.Diagram.Pullback +import Cat.Diagram.Congruence import Cat.Reasoning import Cat.Functor.Reasoning @@ -56,3 +60,94 @@ of symbol shuffling. F-split-coeq .witness-section = annihilate witness-section F-split-coeq .commute = weave commute ``` + +## Congruences and Quotients + +When motivating split coequalisers, we discussed how they arise naturally from +quotients equipped with a choice of representatives of equivalence classes. +Let's go the other direction, and show that split coequalizers induce [quotients]. +The rough idea of the construction is that we can construct an idempotent $A \to A$ +that takes every $a : A$ to it's canonical representative, and then to take +the kernel pair of that morphism to construct a congruence. + +[quotients]: Cat.Diagram.Congruence.html#quotient-objects +[kernel pair]: Cat.Diagram.Congruence.html#quotient-objects + +```agda +module _ {o ℓ} + {C : Precategory o ℓ} + (fc : Finitely-complete C) + where +``` + + + +```agda + split-coequaliser→is-quotient-of + : ∀ {R A A/R} {i r : Hom R A} {q : Hom A A/R} + {rep : Hom A/R A} {witness : Hom A R} + → is-split-coequaliser i r q rep witness + → is-quotient-of (Kernel-pair (r ∘ witness)) q + split-coequaliser→is-quotient-of + {i = i} {r = r} {q = q} {rep = rep} {witness = witness} split-coeq = + is-split-coequaliser→is-coequalizer quotient-split + where + open is-split-coequaliser split-coeq + module R′ = Pullback (pullbacks (r ∘ witness) (r ∘ witness)) +``` + +We will need to do a bit of symbol munging to get the right split coequaliser here, +as the morphisms don't precisely line up due to the fact that we want to be working +with the kernel pair, not `R`. + +However, we first prove a small helper lemma. If we use the intuition of `R` +as a set of pairs of elements and their canonical representatives, then this +lemma states that the representative of $a : A$ is the same as the representative _of_ +the representative of $a$. + +```agda + same-rep : (r ∘ witness) ∘ i ≡ (r ∘ witness) ∘ r + same-rep = + (r ∘ witness) ∘ i ≡˘⟨ commute ⟩∘⟨refl ⟩ + (rep ∘ q) ∘ i ≡⟨ extendr coequal ⟩ + (rep ∘ q) ∘ r ≡⟨ commute ⟩∘⟨refl ⟩ + (r ∘ witness) ∘ r ∎ +``` + +Now, onto the meat of the proof. This is mostly mindless algebraic manipulation +to get things to line up correctly. + +```agda + quotient-split : + is-split-coequaliser + (π₁ ∘ ⟨ R′.p₁ , R′.p₂ ⟩) (π₂ ∘ ⟨ R′.p₁ , R′.p₂ ⟩) + q rep (R′.limiting same-rep ∘ witness) + quotient-split .coequal = retract→is-monic rep-section _ _ $ + rep ∘ q ∘ π₁ ∘ ⟨ R′.p₁ , R′.p₂ ⟩ ≡⟨ refl⟩∘⟨ refl⟩∘⟨ π₁∘⟨⟩ ⟩ + rep ∘ q ∘ R′.p₁ ≡⟨ pulll commute ⟩ + (r ∘ witness) ∘ R′.p₁ ≡⟨ R′.square ⟩ + (r ∘ witness) ∘ R′.p₂ ≡⟨ pushl (sym commute) ⟩ + rep ∘ q ∘ R′.p₂ ≡˘⟨ refl⟩∘⟨ refl⟩∘⟨ π₂∘⟨⟩ ⟩ + rep ∘ q ∘ π₂ ∘ ⟨ R′.p₁ , R′.p₂ ⟩ ∎ + quotient-split .rep-section = rep-section + quotient-split .witness-section = + (π₁ ∘ ⟨ R′.p₁ , R′.p₂ ⟩) ∘ R′.limiting same-rep ∘ witness ≡⟨ π₁∘⟨⟩ ⟩∘⟨refl ⟩ + R′.p₁ ∘ R′.limiting same-rep ∘ witness ≡⟨ pulll R′.p₁∘limiting ⟩ + i ∘ witness ≡⟨ witness-section ⟩ + id ∎ + quotient-split .commute = + rep ∘ q ≡⟨ commute ⟩ + r ∘ witness ≡⟨ pushl (sym R′.p₂∘limiting) ⟩ + R′.p₂ ∘ R′.limiting same-rep ∘ witness ≡˘⟨ π₂∘⟨⟩ ⟩∘⟨refl ⟩ + (π₂ ∘ ⟨ R′.p₁ , R′.p₂ ⟩) ∘ R′.limiting same-rep ∘ witness ∎ +``` + From b010e02a286b65f6f33cbc89509f4b4a980cc3fa Mon Sep 17 00:00:00 2001 From: Reed Mullanix Date: Sun, 15 May 2022 14:45:30 -0700 Subject: [PATCH 07/14] format: formatting fixes for split coequalisers --- src/Cat/Diagram/Coequaliser/Split.lagda.md | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/src/Cat/Diagram/Coequaliser/Split.lagda.md b/src/Cat/Diagram/Coequaliser/Split.lagda.md index 43dc5cf87..ee0dc0506 100644 --- a/src/Cat/Diagram/Coequaliser/Split.lagda.md +++ b/src/Cat/Diagram/Coequaliser/Split.lagda.md @@ -36,7 +36,7 @@ to consider is the quotient $B/R$, which gives us the following diagram: Now, when one has a quotient, it's useful to have a means of picking representatives for each equivalence class. This is essentially what normalization algorithms do, which we can both agree are very useful indeed. This ends up defining a map -$s : B/R \to B$ that is a section of $q$ (IE: $q \circ s = id$). +$s : B/R \to B$ that is a section of $q$ (i.e: $q \circ s = id$). This gives us the following diagram (We've omited the injection of $R$ into $B \times B$ for clarity). @@ -51,7 +51,7 @@ $B \times B$ for clarity). ~~~ This lets us define yet another map $r : B \to R$, which will witness the fact that -any $b : B$ is related to it's representative $s(q(b))$. We can define this map explicitly +any $b : B$ is related to its representative $s(q(b))$. We can define this map explicitly as so: $$ r(b) = (b, s(q(b))) @@ -73,7 +73,7 @@ a quotient by an equivalence relation, along with a means of picking representat \end{tikzcd}\] ~~~ -Such a situation is called a *Split Coequaliser*. +Such a situation is called a **split coequaliser**. ```agda record is-split-coequaliser (f g : Hom A B) (e : Hom B E) @@ -85,7 +85,8 @@ record is-split-coequaliser (f g : Hom A B) (e : Hom B E) commute : s ∘ e ≡ g ∘ t ``` -Now, let's show that this thing actually deserves the name Coequaliser. ```agda +Now, let's show that this thing actually deserves the name Coequaliser. +```agda is-split-coequaliser→is-coequalizer : is-split-coequaliser f g e s t → is-coequaliser f g e is-split-coequaliser→is-coequalizer From bad2d4ca0c6d62a6a3e575e47e92ecef639a335f Mon Sep 17 00:00:00 2001 From: Reed Mullanix Date: Sun, 15 May 2022 14:46:52 -0700 Subject: [PATCH 08/14] format: formatting fixes for split coequaliser properties --- src/Cat/Diagram/Coequaliser/Split/Properties.lagda.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Cat/Diagram/Coequaliser/Split/Properties.lagda.md b/src/Cat/Diagram/Coequaliser/Split/Properties.lagda.md index 09feae75a..d0fda6505 100644 --- a/src/Cat/Diagram/Coequaliser/Split/Properties.lagda.md +++ b/src/Cat/Diagram/Coequaliser/Split/Properties.lagda.md @@ -22,7 +22,7 @@ This module proves some general properties of [split coequalisers]. The property of being a split coequaliser is a purely diagrammatic one, which has the lovely property of being preserved by _all_ functors. We call such colimits -absolute. +**absolute**. ```agda module _ {o o′ ℓ ℓ′} From df554543738254c92eb14a101722de4f272e0b3e Mon Sep 17 00:00:00 2001 From: Reed Mullanix Date: Sun, 15 May 2022 16:33:18 -0700 Subject: [PATCH 09/14] refactor: define product diagrams over the 2 object category --- src/Cat/Diagram/Limit/Product.lagda.md | 127 ++++++++------------ src/Cat/Functor/Adjoint/Continuous.lagda.md | 8 +- src/Cat/Instances/Shape/Pair.lagda.md | 71 +++++++++++ 3 files changed, 123 insertions(+), 83 deletions(-) create mode 100644 src/Cat/Instances/Shape/Pair.lagda.md diff --git a/src/Cat/Diagram/Limit/Product.lagda.md b/src/Cat/Diagram/Limit/Product.lagda.md index 20ce78923..80978f4bb 100644 --- a/src/Cat/Diagram/Limit/Product.lagda.md +++ b/src/Cat/Diagram/Limit/Product.lagda.md @@ -7,6 +7,7 @@ description: | ```agda open import Cat.Diagram.Limit.Base open import Cat.Instances.Discrete +open import Cat.Instances.Shape.Pair open import Cat.Instances.Functor open import Cat.Diagram.Terminal open import Cat.Diagram.Product @@ -14,13 +15,12 @@ open import Cat.Prelude open import Data.Bool -module Cat.Diagram.Limit.Product {o h} (C : Precategory o h) where +module Cat.Diagram.Limit.Product {o h} (Cat : Precategory o h) where ``` + +```agda +module _ {o ℓ} {C : Precategory o ℓ} where + open Precategory C + open Functor + + 2-objects→pair-diagram : ∀ (a b : Ob) → Functor ·· C + 2-objects→pair-diagram a b = funct where + funct : Functor _ _ + funct .F₀ false = a + funct .F₀ true = b + funct .F₁ {false} {false} _ = id + funct .F₁ {true} {true} _ = id + funct .F-id {false} = refl + funct .F-id {true} = refl + funct .F-∘ {false} {false} {false} _ _ = sym (idl _) + funct .F-∘ {true} {true} {true} _ _ = sym (idl _) +``` From 3508d94fde101f868f912946f43fe30d36be3dda Mon Sep 17 00:00:00 2001 From: Reed Mullanix Date: Mon, 16 May 2022 09:58:23 -0700 Subject: [PATCH 10/14] defn: split epis of kernel pairs are split coequalizers --- .../Coequaliser/Split/Properties.lagda.md | 110 ++++++++++++++++++ 1 file changed, 110 insertions(+) diff --git a/src/Cat/Diagram/Coequaliser/Split/Properties.lagda.md b/src/Cat/Diagram/Coequaliser/Split/Properties.lagda.md index d0fda6505..a7ec76349 100644 --- a/src/Cat/Diagram/Coequaliser/Split/Properties.lagda.md +++ b/src/Cat/Diagram/Coequaliser/Split/Properties.lagda.md @@ -61,6 +61,116 @@ of symbol shuffling. F-split-coeq .commute = weave commute ``` +## Split Epis and Kernel Pairs + +We've already seen one way to construct a split coequalizer from a split epimorphism, +via `is-split-epic→is-split-coequalizer`{.Agda}. However, this is not the only way! +A somewhat more intuitive construction can be used which involves taking the kernel +pair of some split epimorphism $e : B \to C$. + +```agda +module _ {o ℓ} {C : Precategory o ℓ} where + open Cat.Reasoning C + open Cat.Diagram.Pullback C + open SplitCoeq C +``` + +First, some preliminaries. We can construct a morphism from $B$ into the kernel +pair of $e$ by like so: + +~~~{.quiver} +\[\begin{tikzcd} + B \\ + & A & B \\ + & B & C + \arrow[from=2-2, to=3-2] + \arrow[from=2-2, to=2-3] + \arrow["e"', from=2-3, to=3-3] + \arrow["e", from=3-2, to=3-3] + \arrow["\lrcorner"{anchor=center, pos=0.125}, draw=none, from=2-2, to=3-3] + \arrow[color={rgb,255:red,214;green,92;blue,92}, dashed, from=1-1, to=2-2] + \arrow["id"', curve={height=12pt}, from=1-1, to=3-2] + \arrow["{s \circ e}", curve={height=-12pt}, from=1-1, to=2-3] +\end{tikzcd}\] +~~~ + +Using our intuition of split coequalisers as quotients with representatives, +this map will pair up an element of $b$ with it's representative. + +```agda + private + represent + : ∀ {A B C} {p₁ p₂ : Hom A B} {e : Hom B C} {s : Hom C B} + → (kp : is-pullback p₁ e p₂ e) + → (section : s is-section-of e) + → Hom B A + represent {e = e} {s = s} kp section = limiting commutes where + open is-pullback kp + abstract + commutes : e ∘ id ≡ e ∘ s ∘ e + commutes = + e ∘ id ≡⟨ id-comm ⟩ + id ∘ e ≡⟨ pushl (sym section) ⟩ + e ∘ s ∘ e ∎ +``` + +This map lets us construct the following split coequaliser: +~~~{.quiver} +\[\begin{tikzcd} + A & B & C + \arrow["{p_1}", shift left=3, from=1-1, to=1-2] + \arrow["{p_2}"', shift right=3, from=1-1, to=1-2] + \arrow["r"{description}, from=1-2, to=1-1] + \arrow["e", shift left=2, from=1-2, to=1-3] + \arrow["s", shift left=2, from=1-3, to=1-2] +\end{tikzcd}\] +~~~ + +All of the required commuting squares exist due to our clever choice +of map into the kernel pair. + +```agda + is-split-epic→is-split-coequalizer-of-kernel-pair + : ∀ {A B C} {p₁ p₂ : Hom A B} {e : Hom B C} {s : Hom C B} + → (kp : is-pullback p₁ e p₂ e) + → (section : s is-section-of e) + → is-split-coequaliser p₁ p₂ e s (represent kp section) + is-split-epic→is-split-coequalizer-of-kernel-pair + {p₁ = p₁} {p₂ = p₂} {e = e} {s = s} kp section = split-coeq where + open is-pullback kp + open is-split-coequaliser + + split-coeq : is-split-coequaliser _ _ _ _ _ + split-coeq .coequal = square + split-coeq .rep-section = section + split-coeq .witness-section = p₁∘limiting + split-coeq .commute = sym p₂∘limiting +``` + +We also provide a bundled form of this lemma. + +```agda + split-epic→split-coequalizer-of-kernel-pair + : ∀ {B C} {e : Hom B C} + → (kp : Pullback e e) + → is-split-epic e + → Split-coequaliser (Pullback.p₁ kp) (Pullback.p₂ kp) + split-epic→split-coequalizer-of-kernel-pair + {e = e} kp split-epic = split-coeq where + open Pullback kp + open is-split-epic split-epic + open Split-coequaliser + + split-coeq : Split-coequaliser _ _ + split-coeq .coapex = _ + split-coeq .coeq = e + split-coeq .rep = split + split-coeq .witness = represent has-is-pb section + split-coeq .has-is-split-coeq = + is-split-epic→is-split-coequalizer-of-kernel-pair has-is-pb section +``` + + ## Congruences and Quotients When motivating split coequalisers, we discussed how they arise naturally from From 1a79e3dc14d76cbb9a2d1a4c278aef8010db3f22 Mon Sep 17 00:00:00 2001 From: Reed Mullanix Date: Mon, 16 May 2022 10:23:27 -0700 Subject: [PATCH 11/14] defn: reflection of colimits, isomorphisms of colimit coapexes --- src/Cat/Diagram/Colimit/Base.lagda.md | 70 +++++++++++++++++++++++++-- 1 file changed, 67 insertions(+), 3 deletions(-) diff --git a/src/Cat/Diagram/Colimit/Base.lagda.md b/src/Cat/Diagram/Colimit/Base.lagda.md index cae078b9e..ccf28ada7 100644 --- a/src/Cat/Diagram/Colimit/Base.lagda.md +++ b/src/Cat/Diagram/Colimit/Base.lagda.md @@ -3,7 +3,7 @@ open import Cat.Diagram.Initial open import Cat.Prelude import Cat.Functor.Reasoning as Func -import Cat.Morphism +import Cat.Reasoning module Cat.Diagram.Colimit.Base where ``` @@ -64,8 +64,8 @@ We this pair of category and functor a _diagram_ in $C$. ```agda module _ {J : Precategory o ℓ} {C : Precategory o′ ℓ′} (F : Functor J C) where private - import Cat.Reasoning J as J - import Cat.Reasoning C as C + module J = Cat.Reasoning J + module C = Cat.Reasoning C module F = Functor F record Cocone : Type (o ⊔ ℓ ⊔ o′ ⊔ ℓ′) where @@ -200,6 +200,9 @@ cocones over that diagram. Colimit-apex : Colimit → C.Ob Colimit-apex x = coapex (Initial.bot x) + + Colimit-universal : (L : Colimit) → (K : Cocone) → C.Hom (Colimit-apex L) (coapex K) + Colimit-universal L K = hom (Initial.¡ L {K}) ``` @@ -247,6 +250,67 @@ say that it _preserves_ colimits. Preserves-colimit K = is-colimit Dia K → is-colimit (F F∘ Dia) (F-map-cocone K) ``` +## Reflection of colimits + +We say a functor __reflects__ colimits if the existence of a colimiting +cocone "downstairs" implies that we must have a limiting cocone "upstairs". + +More concretely, if the image of a cocone $F \circ K$ in $\ca{D}$ +is a colimiting cocone, then $K$ must have already been a +colimiting cocone in $\ca{C}$ + +```agda + Reflects-colimit : Cocone Dia → Type _ + Reflects-colimit K = is-colimit (F F∘ Dia) (F-map-cocone K) → is-colimit Dia K +``` + +# Uniqueness + + + +If the universal map $L \to K$ between coapexes of some colimit is +invertible, that means that $K$ is also a colimiting cocone. + +```agda + coapex-iso→is-colimit + : (K : Cocone F) + (L : Colimit F) + → C.is-invertible (Colimit-universal F L K) + → is-colimit F K + coapex-iso→is-colimit K L invert K′ = colimits where + module K = Cocone K + module K′ = Cocone K′ + module L = Cocone (Initial.bot L) + module universal K = Cocone-hom (Initial.¡ L {K}) + open C.is-invertible invert + + colimits : is-contr (Cocones.Hom K K′) + colimits .centre .Cocone-hom.hom = universal.hom K′ C.∘ inv + colimits .centre .Cocone-hom.commutes _ = + (universal.hom K′ C.∘ inv) C.∘ K.ψ _ ≡˘⟨ ap ((universal.hom K′ C.∘ inv) C.∘_) (universal.commutes K _) ⟩ + (universal.hom K′ C.∘ inv) C.∘ (universal.hom K C.∘ L.ψ _) ≡⟨ C.cancel-inner invr ⟩ + universal.hom K′ C.∘ L.ψ _ ≡⟨ universal.commutes K′ _ ⟩ + K′.ψ _ ∎ + colimits .paths f = + let module f = Cocone-hom f in + Cocone-hom-path F $ C.invertible→epic invert _ _ $ + (universal.hom K′ C.∘ inv) C.∘ universal.hom K ≡⟨ C.cancelr invr ⟩ + universal.hom K′ ≡⟨ ap Cocone-hom.hom (Initial.¡-unique L (f Cocones.∘ Initial.¡ L)) ⟩ + f.hom C.∘ universal.hom K ∎ +``` + + ## Cocompleteness A category is **cocomplete** if admits for limits of arbitrary shape. From 4f00251efe33b496145d319fce5f050b118437ca Mon Sep 17 00:00:00 2001 From: Reed Mullanix Date: Mon, 16 May 2022 11:58:40 -0700 Subject: [PATCH 12/14] defn: conservative functors preserve colimits --- src/Cat/Diagram/Colimit/Base.lagda.md | 29 ++++++++- src/Cat/Functor/Conservative.lagda.md | 85 ++++++++++++++++++++++++++- 2 files changed, 110 insertions(+), 4 deletions(-) diff --git a/src/Cat/Diagram/Colimit/Base.lagda.md b/src/Cat/Diagram/Colimit/Base.lagda.md index ccf28ada7..3d485ec46 100644 --- a/src/Cat/Diagram/Colimit/Base.lagda.md +++ b/src/Cat/Diagram/Colimit/Base.lagda.md @@ -198,8 +198,12 @@ cocones over that diagram. Colimit : Type _ Colimit = Initial Cocones + Colimit-cocone : Colimit → Cocone + Colimit-cocone = Initial.bot + Colimit-apex : Colimit → C.Ob - Colimit-apex x = coapex (Initial.bot x) + Colimit-apex x = coapex (Colimit-cocone x) + Colimit-universal : (L : Colimit) → (K : Cocone) → C.Hom (Colimit-apex L) (coapex K) Colimit-universal L K = hom (Initial.¡ L {K}) @@ -241,6 +245,22 @@ functors preserve commutative diagrams. F.₁ (Cocone.ψ x _) ∎ ``` +Note that this also lets us map morphisms between cocones into $\ca{D}$. + +```agda + F-map-cocone-hom + : {X Y : Cocone Dia} + → Cocone-hom Dia X Y + → Cocone-hom (F F∘ Dia) (F-map-cocone X) (F-map-cocone Y) + F-map-cocone-hom {X = X} {Y = Y} f = hom where + module f = Cocone-hom f + + hom : Cocone-hom (F F∘ Dia) (F-map-cocone X) (F-map-cocone Y) + hom .Cocone-hom.hom = F .F₁ f.hom + hom .Cocone-hom.commutes _ = F.collapse (f.commutes _) +``` + + Though functors must take cocones to cocones, they may not necessarily take colimiting cocones to colimiting cocones! When a functor does, we say that it _preserves_ colimits. @@ -250,6 +270,13 @@ say that it _preserves_ colimits. Preserves-colimit K = is-colimit Dia K → is-colimit (F F∘ Dia) (F-map-cocone K) ``` +```agda + F-map-colimit : (L : Colimit Dia) → Preserves-colimit (Colimit-cocone Dia L) → Colimit (F F∘ Dia) + F-map-colimit L preserves .Initial.bot = F-map-cocone (Initial.bot L) + F-map-colimit L preserves .Initial.has⊥ = preserves (Initial.has⊥ L) +``` + + ## Reflection of colimits We say a functor __reflects__ colimits if the existence of a colimiting diff --git a/src/Cat/Functor/Conservative.lagda.md b/src/Cat/Functor/Conservative.lagda.md index 7d8dfe969..1a5bef280 100644 --- a/src/Cat/Functor/Conservative.lagda.md +++ b/src/Cat/Functor/Conservative.lagda.md @@ -1,10 +1,14 @@ ```agda open import Cat.Diagram.Limit.Base +open import Cat.Diagram.Colimit.Base open import Cat.Diagram.Terminal +open import Cat.Diagram.Initial open import Cat.Functor.Base open import Cat.Morphism open import Cat.Prelude hiding (J) +import Cat.Reasoning + module Cat.Functor.Conservative where ``` @@ -44,9 +48,15 @@ in $C$ as well (see `apex-iso→is-limit`{.Agda}). ```agda module _ {F : Functor C D} (conservative : is-conservative F) where - conservative-reflects-limits : ∀ {Dia : Functor J C} → (L : Limit Dia) - → (∀ (K : Cone Dia) → Preserves-limit F K) - → (∀ (K : Cone Dia) → Reflects-limit F K) + private + module D = Cat.Reasoning D + module C = Cat.Reasoning C + module Cocones {o h o′ h′} {J : Precategory o h} {C : Precategory o′ h′} {Dia : Functor J C} = Cat.Reasoning (Cocones Dia) + + conservative-reflects-limits + : ∀ {Dia : Functor J C} → (L : Limit Dia) + → (∀ (K : Cone Dia) → Preserves-limit F K) + → (∀ (K : Cone Dia) → Reflects-limit F K) conservative-reflects-limits {Dia = Dia} L-lim preserves K limits = apex-iso→is-limit Dia K L-lim $ conservative @@ -73,3 +83,72 @@ module _ {F : Functor C D} (conservative : is-conservative F) where hom (F-map-cone-hom F (Terminal.! L-lim)) ≡⟨⟩ F .F₁ (hom L-lim.!) ∎ ``` + +We also have a dual theorem for colimits. + +```agda + conservative-reflects-colimits + : ∀ {Dia : Functor J C} → (L : Colimit Dia) + → (∀ (K : Cocone Dia) → Preserves-colimit F K) + → (∀ (K : Cocone Dia) → Reflects-colimit F K) + conservative-reflects-colimits + {Dia = Dia} L-colim preserves K F∘K-colimits = + coapex-iso→is-colimit Dia K L-colim + $ conservative + invert + where + + F∘K-colim : Colimit (F F∘ Dia) + F∘K-colim .Initial.bot = F-map-cocone F K + F∘K-colim .Initial.has⊥ = F∘K-colimits + + F∘L-colim : Colimit (F F∘ Dia) + F∘L-colim = F-map-colimit F L-colim (preserves (Colimit-cocone Dia L-colim)) + + module L-colim = Initial L-colim + module F∘L-colim = Initial F∘L-colim + module F∘K-colim = Initial F∘K-colim + open Cocone-hom + + L : Cocone Dia + L = L-colim.bot + + F∘L : Cocone (F F∘ Dia) + F∘L = F-map-cocone F L-colim.bot + + F∘K : Cocone (F F∘ Dia) + F∘K = F-map-cocone F K + + L-universal : (K′ : Cocone Dia) → Cocone-hom Dia L K′ + L-universal K′ = L-colim.¡ {K′} + + F∘L-universal : (K′ : Cocone (F F∘ Dia)) → Cocone-hom (F F∘ Dia) F∘L K′ + F∘L-universal K′ = F∘L-colim.¡ {K′} + + F∘K-universal : (K′ : Cocone (F F∘ Dia)) → Cocone-hom (F F∘ Dia) F∘K K′ + F∘K-universal K′ = F∘K-colim.¡ {K′} + + module F∘L-universal K′ = Cocone-hom (F∘L-universal K′) + module L-universal K′ = Cocone-hom (L-universal K′) + module F∘K-universal K′ = Cocone-hom (F∘K-universal K′) + + F-preserves-universal + : ∀ {K′} → F∘L-universal.hom (F-map-cocone F K′) ≡ F. F₁ (L-universal.hom K′) + F-preserves-universal {K′} = + F∘L-universal.hom (F-map-cocone F K′) ≡⟨ ap hom (F∘L-colim.¡-unique (F-map-cocone-hom F (L-universal K′))) ⟩ + hom (F-map-cocone-hom F (L-universal K′)) ≡⟨⟩ + F .F₁ (L-universal.hom K′) ∎ + + invert : is-invertible D (F .F₁ (Colimit-universal Dia L-colim K)) + invert .is-invertible.inv = F∘K-universal.hom F∘L + invert .is-invertible.inverses .Inverses.invl = + F .F₁ (L-universal.hom K) D.∘ F∘K-universal.hom F∘L ≡˘⟨ ap (D._∘ F∘K-universal.hom F∘L) F-preserves-universal ⟩ + F∘L-universal.hom F∘K D.∘ F∘K-universal.hom F∘L ≡⟨⟩ + hom (F∘L-universal F∘K Cocones.∘ F∘K-universal F∘L) ≡⟨ ap hom (F∘K-colim.¡-unique₂ (F∘L-universal F∘K Cocones.∘ F∘K-universal F∘L) Cocones.id) ⟩ + D.id ∎ + invert .is-invertible.inverses .Inverses.invr = + F∘K-universal.hom F∘L D.∘ F .F₁ (L-universal.hom K) ≡˘⟨ ap (F∘K-universal.hom F∘L D.∘_) F-preserves-universal ⟩ + F∘K-universal.hom F∘L D.∘ F∘L-universal.hom F∘K ≡⟨⟩ + hom (F∘K-universal F∘L Cocones.∘ F∘L-universal F∘K) ≡⟨ ap hom (F∘L-colim.¡-unique₂ (F∘K-universal F∘L Cocones.∘ F∘L-universal F∘K) Cocones.id) ⟩ + D.id ∎ +``` From 782abd7f0187a7f2be4bca774df7e2c24990b1d4 Mon Sep 17 00:00:00 2001 From: Reed Mullanix Date: Mon, 16 May 2022 12:24:37 -0700 Subject: [PATCH 13/14] defn: coequalisers are colimits --- src/Cat/Diagram/Colimit/Coequaliser.lagda.md | 94 ++++++++++++++++++++ 1 file changed, 94 insertions(+) create mode 100644 src/Cat/Diagram/Colimit/Coequaliser.lagda.md diff --git a/src/Cat/Diagram/Colimit/Coequaliser.lagda.md b/src/Cat/Diagram/Colimit/Coequaliser.lagda.md new file mode 100644 index 000000000..3cfa3f048 --- /dev/null +++ b/src/Cat/Diagram/Colimit/Coequaliser.lagda.md @@ -0,0 +1,94 @@ +```agda +open import Cat.Instances.Shape.Parallel + +open import Cat.Diagram.Colimit.Base +open import Cat.Diagram.Coequaliser +open import Cat.Instances.Functor +open import Cat.Diagram.Initial +open import Cat.Prelude + +open import Data.Bool + +module Cat.Diagram.Colimit.Coequaliser + {o ℓ} (Cat : Precategory o ℓ) where +``` + + + +We establish the correspondence between `Coequaliser`{.Agda} and the +`Colimit`{.Agda} of a [parallel arrows] diagram. + +[parallel arrows]: Cat.Instances.Shape.Parallel.html + +```agda +Fork→Cocone + : ∀ {E} (F : Functor ·⇉· Cat) + → (eq : Hom (F .F₀ true) E) + → eq ∘ F .F₁ {false} {true} false ≡ eq ∘ F .F₁ {false} {true} true + → Cocone F +Fork→Cocone F eq p .coapex = _ +Fork→Cocone F eq p .ψ false = eq ∘ F .F₁ false +Fork→Cocone F eq p .ψ true = eq +Fork→Cocone F eq p .commutes {false} {false} tt = elimr (F .F-id) +Fork→Cocone F eq p .commutes {false} {true} false = refl +Fork→Cocone F eq p .commutes {false} {true} true = sym p +Fork→Cocone F eq p .commutes {true} {true} tt = elimr (F .F-id) + +Coequaliser→Colimit + : ∀ {F : Functor ·⇉· Cat} + → Coequaliser Cat (F .F₁ false) (F .F₁ true) + → Colimit F +Coequaliser→Colimit {F = F} coeq = colim where + module coeq = Coequaliser coeq + colim : Colimit _ + colim .bot = Fork→Cocone F coeq.coeq coeq.coequal + colim .has⊥ K .centre .hom = + coeq.coequalise $ K .commutes {false} {true} false + ∙ sym (K .commutes true) + colim .has⊥ K .centre .commutes false = + pulll coeq.universal ∙ K .commutes {false} {true} false + colim .has⊥ K .centre .commutes true = coeq.universal + colim .has⊥ K .paths other = Cocone-hom-path _ (sym (coeq.unique p)) where + p : K .ψ true ≡ other .hom ∘ coeq.coeq + p = sym (other .commutes _) + +Colimit→Coequaliser + : ∀ {F : Functor ·⇉· Cat} + → Colimit F + → Coequaliser Cat (F .F₁ false) (F .F₁ true) +Colimit→Coequaliser {F = F} colim = coequ where + module colim = Initial colim + + coequ : Coequaliser Cat _ _ + coequ .coapex = _ + coequ .coeq = colim.bot .ψ true + coequ .has-is-coeq .coequal = + colim.bot .commutes {false} {true} false + ∙ sym (colim.bot .commutes true) + coequ .has-is-coeq .coequalise p = + colim.has⊥ (Fork→Cocone _ _ p) .centre .hom + coequ .has-is-coeq .universal {p = p} = + colim.has⊥ (Fork→Cocone _ _ p) .centre .commutes true + coequ .has-is-coeq .unique {e′ = e′} {p = p} {colim = colim'} x = + sym (ap hom (colim.has⊥ (Fork→Cocone _ _ p) .paths other)) + where + other : Cocone-hom _ _ _ + other .hom = colim' + other .commutes false = + colim' ∘ colim .bot .ψ false ≡⟨ pushr (sym $ colim.bot .commutes false) ⟩ + (colim' ∘ colim.bot .ψ true) ∘ F .F₁ false ≡˘⟨ ap (_∘ F .F₁ false) x ⟩ + e′ ∘ F .F₁ false ∎ + other .commutes true = sym x +``` From 1633a241622c1ed323bdc8f1e6be23cc0a9fc61a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Am=C3=A9lia=20Liao?= Date: Mon, 16 May 2022 17:07:05 -0300 Subject: [PATCH 14/14] fixup: avoid subst in conservative-reflects-limits Closes #77 --- src/Cat/Functor/Conservative.lagda.md | 29 ++++++++++++++++++--------- 1 file changed, 20 insertions(+), 9 deletions(-) diff --git a/src/Cat/Functor/Conservative.lagda.md b/src/Cat/Functor/Conservative.lagda.md index 1a5bef280..ee9c2fd9e 100644 --- a/src/Cat/Functor/Conservative.lagda.md +++ b/src/Cat/Functor/Conservative.lagda.md @@ -58,11 +58,7 @@ module _ {F : Functor C D} (conservative : is-conservative F) where → (∀ (K : Cone Dia) → Preserves-limit F K) → (∀ (K : Cone Dia) → Reflects-limit F K) conservative-reflects-limits {Dia = Dia} L-lim preserves K limits = - apex-iso→is-limit Dia K L-lim - $ conservative - $ subst (λ ϕ → is-invertible D ϕ) F-preserves-universal - $ Cone-invertible→apex-invertible (F F∘ Dia) - $ !-invertible (Cones (F F∘ Dia)) F∘L-lim K-lim + apex-iso→is-limit Dia K L-lim $ conservative invert where F∘L-lim : Limit (F F∘ Dia) F∘L-lim .Terminal.top = F-map-cone F (Terminal.top L-lim) @@ -77,11 +73,26 @@ module _ {F : Functor C D} (conservative : is-conservative F) where open Cone-hom F-preserves-universal - : hom F∘L-lim.! ≡ F .F₁ (hom {x = K} L-lim.!) + : F∘L-lim.! .hom ≡ F .F₁ (L-lim.! {x = K} .hom) F-preserves-universal = - hom F∘L-lim.! ≡⟨ ap hom (F∘L-lim.!-unique (F-map-cone-hom F L-lim.!)) ⟩ - hom (F-map-cone-hom F (Terminal.! L-lim)) ≡⟨⟩ - F .F₁ (hom L-lim.!) ∎ + F∘L-lim.! .hom ≡⟨ ap hom (F∘L-lim.!-unique (F-map-cone-hom F L-lim.!)) ⟩ + F-map-cone-hom F (Terminal.! L-lim) .hom ≡⟨⟩ + F .F₁ (L-lim.! .hom) ∎ + + open is-invertible + open Inverses + inverse = !-invertible (Cones (F F∘ Dia)) F∘L-lim K-lim + + invert : D.is-invertible (F .F₁ (L-lim.! .hom)) + invert .inv = inverse .inv .hom + invert .inverses .invl = + F .F₁ (L-lim.! .hom) D.∘ invert .inv ≡˘⟨ F-preserves-universal D.⟩∘⟨refl ⟩ + F∘L-lim.! .hom D.∘ invert .inv ≡⟨ ap hom (inverse .inverses .invl) ⟩ + D.id ∎ + invert .inverses .invr = + invert .inv D.∘ F .F₁ (L-lim.! .hom) ≡˘⟨ D.refl⟩∘⟨ F-preserves-universal ⟩ + invert .inv D.∘ F∘L-lim.! .hom ≡⟨ ap hom (inverse .inverses .invr) ⟩ + D.id ∎ ``` We also have a dual theorem for colimits.