From f6b6f0d4f4eeaa390d052465f28ff4fe278b481d Mon Sep 17 00:00:00 2001 From: Reed Mullanix Date: Tue, 21 Mar 2023 10:02:37 -0700 Subject: [PATCH] Redefine (co)limits in terms of kan extensions (#186) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit # Description This PR refactors (co)limits to be defined in terms of kan extensions. The main motivation for doing so is to make preservation of limits less annoying to deal with; as it currently stands we need to define special functions for mapping cones, which makes life much harder. ## Notes As noted on discord, representability does not work due to (annoying!) universe level issues. This PR is also based off #184; I'll rebase when the time comes. --------- Co-authored-by: Amélia Liao --- src/1Lab/Path.lagda.md | 12 + .../Group/Cat/FinitelyComplete.lagda.md | 8 +- src/Algebra/Group/Subgroup.lagda.md | 6 +- src/Algebra/Ring/Module.lagda.md | 2 +- src/Cat/Abelian/Base.lagda.md | 40 +- src/Cat/Abelian/Images.lagda.md | 22 +- src/Cat/Abelian/Limits.lagda.md | 4 +- src/Cat/Base.lagda.md | 13 + src/Cat/Bi/Base.lagda.md | 2 +- src/Cat/Bi/Instances/Spans.lagda.md | 142 +-- src/Cat/CartesianClosed/Base.lagda.md | 8 +- src/Cat/CartesianClosed/Instances/PSh.agda | 20 +- src/Cat/CartesianClosed/Locally.lagda.md | 8 +- src/Cat/Diagram/Coend.lagda.md | 2 +- src/Cat/Diagram/Coend/Formula.lagda.md | 67 +- src/Cat/Diagram/Coequaliser.lagda.md | 39 +- .../Diagram/Coequaliser/RegularEpi.lagda.md | 10 +- src/Cat/Diagram/Colimit/Base.lagda.md | 697 +++++++--- src/Cat/Diagram/Colimit/Cocone.lagda.md | 196 +++ src/Cat/Diagram/Colimit/Finite.lagda.md | 32 +- .../Diagram/Colimit/Representable.lagda.md | 77 ++ src/Cat/Diagram/Congruence.lagda.md | 26 +- src/Cat/Diagram/Coproduct/Copower.lagda.md | 2 +- src/Cat/Diagram/Coproduct/Indexed.lagda.md | 115 +- src/Cat/Diagram/Duals.lagda.md | 92 +- src/Cat/Diagram/Equaliser.lagda.md | 41 +- src/Cat/Diagram/Equaliser/Kernel.lagda.md | 7 +- .../Diagram/Equaliser/RegularMono.lagda.md | 10 +- src/Cat/Diagram/Limit/Base.lagda.md | 1118 ++++++++++------- src/Cat/Diagram/Limit/Cone.lagda.md | 245 ++++ src/Cat/Diagram/Limit/Equaliser.lagda.md | 135 +- src/Cat/Diagram/Limit/Finite.lagda.md | 61 +- src/Cat/Diagram/Limit/Product.lagda.md | 114 +- src/Cat/Diagram/Limit/Pullback.lagda.md | 46 +- src/Cat/Diagram/Monad/Codensity.lagda.md | 5 +- src/Cat/Diagram/Monad/Limits.lagda.md | 263 ++-- src/Cat/Diagram/Product.lagda.md | 210 ++-- src/Cat/Diagram/Product/Finite.lagda.md | 6 +- src/Cat/Diagram/Product/Indexed.lagda.md | 108 +- src/Cat/Diagram/Product/Solver.lagda.md | 10 +- src/Cat/Diagram/Pullback.lagda.md | 23 +- src/Cat/Diagram/Pullback/Properties.lagda.md | 74 +- src/Cat/Diagram/Pushout.lagda.md | 8 +- src/Cat/Diagram/Pushout/Properties.lagda.md | 8 +- src/Cat/Diagram/Zero.lagda.md | 6 + src/Cat/Displayed/Bifibration.lagda.md | 28 +- src/Cat/Displayed/Cartesian.lagda.md | 4 +- src/Cat/Displayed/Cartesian/Discrete.lagda.md | 4 +- src/Cat/Displayed/Cartesian/Identity.lagda.md | 2 +- src/Cat/Displayed/Cartesian/Right.lagda.md | 6 +- src/Cat/Displayed/Cartesian/Weak.lagda.md | 43 +- src/Cat/Displayed/Cocartesian.lagda.md | 10 +- .../Displayed/Cocartesian/Indexing.lagda.md | 4 +- src/Cat/Displayed/Cocartesian/Weak.lagda.md | 17 +- src/Cat/Displayed/Composition.lagda.md | 8 +- .../Displayed/Instances/CT-Structure.lagda.md | 8 +- src/Cat/Displayed/Instances/Simple.lagda.md | 8 +- src/Cat/Displayed/Instances/Slice.lagda.md | 39 +- src/Cat/Displayed/Instances/Trivial.lagda.md | 13 +- src/Cat/Displayed/Morphism.lagda.md | 4 +- src/Cat/Displayed/Morphism/Duality.lagda.md | 11 +- src/Cat/Displayed/Total.lagda.md | 2 +- src/Cat/Displayed/Total/Op.lagda.md | 17 +- src/Cat/Displayed/Univalence/Thin.lagda.md | 2 +- src/Cat/Functor/Adjoint.lagda.md | 21 +- src/Cat/Functor/Adjoint/Continuous.lagda.md | 263 ++-- src/Cat/Functor/Adjoint/Hom.lagda.md | 10 +- src/Cat/Functor/Adjoint/Kan.lagda.md | 221 ++++ src/Cat/Functor/Coherence.agda | 96 ++ src/Cat/Functor/Conservative.lagda.md | 62 +- src/Cat/Functor/Dense.lagda.md | 33 +- src/Cat/Functor/Equivalence/Complete.lagda.md | 2 +- src/Cat/Functor/Everything.agda | 5 +- src/Cat/Functor/Final.lagda.md | 119 +- src/Cat/Functor/Hom.lagda.md | 246 +--- src/Cat/Functor/Hom/Cocompletion.lagda.md | 29 +- src/Cat/Functor/Hom/Coyoneda.lagda.md | 247 ++++ src/Cat/Functor/Hom/Displayed.lagda.md | 6 +- src/Cat/Functor/Hom/Representable.lagda.md | 209 ++- src/Cat/Functor/Kan.lagda.md | 413 ------ src/Cat/Functor/Kan/Adjoint.lagda.md | 113 ++ src/Cat/Functor/Kan/Base.lagda.md | 348 +++++ src/Cat/Functor/Kan/Duality.lagda.md | 186 +++ src/Cat/Functor/Kan/Global.lagda.md | 15 +- src/Cat/Functor/Kan/Nerve.lagda.md | 343 +++-- src/Cat/Functor/Kan/Pointwise.lagda.md | 621 +++++++++ src/Cat/Functor/Kan/Representable.lagda.md | 159 +++ src/Cat/Functor/Kan/Right.lagda.md | 133 -- src/Cat/Functor/Kan/Unique.lagda.md | 632 ++++++++-- src/Cat/Functor/Monadic/Beck.lagda.md | 46 +- src/Cat/Functor/Monadic/Crude.lagda.md | 32 +- src/Cat/Functor/Pullback.lagda.md | 28 +- src/Cat/Instances/Comma.lagda.md | 2 + src/Cat/Instances/Discrete.lagda.md | 37 +- src/Cat/Instances/Functor.lagda.md | 91 +- src/Cat/Instances/Functor/Compose.lagda.md | 61 +- src/Cat/Instances/Functor/Limits.lagda.md | 211 +--- src/Cat/Instances/Product.lagda.md | 2 +- src/Cat/Instances/Sets/Cocomplete.lagda.md | 47 +- src/Cat/Instances/Sets/Complete.lagda.md | 53 +- src/Cat/Instances/Sets/Congruences.lagda.md | 12 +- src/Cat/Instances/Shape/Initial.lagda.md | 42 + src/Cat/Instances/Shape/Parallel.lagda.md | 28 +- src/Cat/Instances/Shape/Terminal.lagda.md | 62 + src/Cat/Instances/Slice.lagda.md | 119 +- src/Cat/Instances/Twisted.lagda.md | 1 + src/Cat/Monoidal/Instances/Cartesian.lagda.md | 2 +- src/Cat/Morphism.lagda.md | 8 + src/Cat/Morphism/Duality.lagda.md | 1 + src/Cat/Morphism/StrongEpi.lagda.md | 10 +- src/Cat/Reasoning.lagda.md | 6 + src/Cat/Regular.lagda.md | 6 +- src/Data/Set/Surjection.lagda.md | 6 +- src/Order/Frame/Free.lagda.md | 4 +- .../Instances/Lower/Cocompletion.lagda.md | 2 +- src/Order/Instances/Subobjects.lagda.md | 2 +- src/Topoi/Base.lagda.md | 31 +- src/Topoi/Classifying/Diaconescu.lagda.md | 10 +- src/Topoi/Reasoning.lagda.md | 14 +- src/index.lagda.md | 9 +- 120 files changed, 6620 insertions(+), 3267 deletions(-) create mode 100644 src/Cat/Diagram/Colimit/Cocone.lagda.md create mode 100644 src/Cat/Diagram/Colimit/Representable.lagda.md create mode 100644 src/Cat/Diagram/Limit/Cone.lagda.md create mode 100644 src/Cat/Functor/Adjoint/Kan.lagda.md create mode 100644 src/Cat/Functor/Coherence.agda create mode 100644 src/Cat/Functor/Hom/Coyoneda.lagda.md delete mode 100644 src/Cat/Functor/Kan.lagda.md create mode 100644 src/Cat/Functor/Kan/Adjoint.lagda.md create mode 100644 src/Cat/Functor/Kan/Base.lagda.md create mode 100644 src/Cat/Functor/Kan/Duality.lagda.md create mode 100644 src/Cat/Functor/Kan/Pointwise.lagda.md create mode 100644 src/Cat/Functor/Kan/Representable.lagda.md delete mode 100644 src/Cat/Functor/Kan/Right.lagda.md create mode 100644 src/Cat/Instances/Shape/Initial.lagda.md diff --git a/src/1Lab/Path.lagda.md b/src/1Lab/Path.lagda.md index 62118e787..c2fecf9c5 100644 --- a/src/1Lab/Path.lagda.md +++ b/src/1Lab/Path.lagda.md @@ -1200,6 +1200,18 @@ ap₂ : ∀ {a b c} {A : Type a} {B : A → Type b} {C : (x : A) → B x → Typ ap₂ f p q i = f (p i) (q i) ``` + + This operation satisfies many identities definitionally that are only propositional when `ap`{.Agda} is defined in terms of `J`{.Agda}. For instance: diff --git a/src/Algebra/Group/Cat/FinitelyComplete.lagda.md b/src/Algebra/Group/Cat/FinitelyComplete.lagda.md index f62d260d8..fa8fce8b5 100644 --- a/src/Algebra/Group/Cat/FinitelyComplete.lagda.md +++ b/src/Algebra/Group/Cat/FinitelyComplete.lagda.md @@ -234,14 +234,14 @@ $g$. Groups-equalisers .apex = Equaliser-group Groups-equalisers .equ = total-hom fst record { pres-⋆ = λ x y → refl } Groups-equalisers .has-is-eq .equal = Forget-is-faithful seq.equal - Groups-equalisers .has-is-eq .limiting {F = F} {e′} p = total-hom map lim-gh where - map = seq.limiting {F = underlying-set (F .snd)} (ap hom p) + Groups-equalisers .has-is-eq .universal {F = F} {e′} p = total-hom map lim-gh where + map = seq.universal {F = underlying-set (F .snd)} (ap hom p) lim-gh : is-group-hom _ _ map lim-gh .pres-⋆ x y = Σ-prop-path (λ _ → H.has-is-set _ _) (e′ .preserves .pres-⋆ _ _) - Groups-equalisers .has-is-eq .universal {F = F} {p = p} = Forget-is-faithful - (seq.universal {F = underlying-set (F .snd)} {p = ap hom p}) + Groups-equalisers .has-is-eq .factors {F = F} {p = p} = Forget-is-faithful + (seq.factors {F = underlying-set (F .snd)} {p = ap hom p}) Groups-equalisers .has-is-eq .unique {F = F} {p = p} q = Forget-is-faithful (seq.unique {F = underlying-set (F .snd)} {p = ap hom p} (ap hom q)) diff --git a/src/Algebra/Group/Subgroup.lagda.md b/src/Algebra/Group/Subgroup.lagda.md index d5812e39b..4b815cc12 100644 --- a/src/Algebra/Group/Subgroup.lagda.md +++ b/src/Algebra/Group/Subgroup.lagda.md @@ -319,7 +319,7 @@ will compute. path : (x : ⌞ Kerf.ker ⌟) → A→im # A.unit ≡ A→im # (x .fst) path (x* , p) = Tpath (f.pres-id ∙ sym p) - coeq .coequalise {F = F} {e′ = e'} p = gh where + coeq .universal {F = F} {e′ = e'} p = gh where module F = Group-on (F .snd) module e' = is-group-hom (e' .preserves) @@ -330,7 +330,7 @@ will compute. {P = λ q r → elim p (((x , q) Ak.⋆ (y , r)) .snd) ≡ elim p q F.⋆ elim p r} (λ _ _ → F.has-is-set _ _) (λ x y → e'.pres-⋆ _ _) q r - coeq .universal = Forget-is-faithful refl + coeq .factors = Forget-is-faithful refl coeq .unique {F} {p = p} {colim = colim} prf = Forget-is-faithful (funext path) where abstract @@ -341,7 +341,7 @@ will compute. {P = λ q → colim # (x , q) ≡ elim p q} (λ _ → F.has-is-set _ _) (λ { (f , fp) → ap (colim #_) (Σ-prop-path (λ _ → squash) (sym fp)) - ∙ sym (happly (ap hom prf) f) }) + ∙ (happly (ap hom prf) f) }) t ``` diff --git a/src/Algebra/Ring/Module.lagda.md b/src/Algebra/Ring/Module.lagda.md index a76a2dd34..b9ee26ed3 100644 --- a/src/Algebra/Ring/Module.lagda.md +++ b/src/Algebra/Ring/Module.lagda.md @@ -12,8 +12,8 @@ open import Cat.Displayed.Fibre open import Cat.Displayed.Total open import Cat.Functor.Adjoint open import Cat.Displayed.Base -open import Cat.Abelian.Endo open import Cat.Abelian.Base +open import Cat.Abelian.Endo open import Cat.Prelude import Cat.Reasoning diff --git a/src/Cat/Abelian/Base.lagda.md b/src/Cat/Abelian/Base.lagda.md index 70f6027ac..7a65d3c50 100644 --- a/src/Cat/Abelian/Base.lagda.md +++ b/src/Cat/Abelian/Base.lagda.md @@ -316,7 +316,7 @@ monomorphism]. decompose {A} {B} f = map , sym path where proj′ : Hom (Coker.coapex (Ker.kernel f)) B - proj′ = Coker.coequalise (Ker.kernel f) {e′ = f} $ sym path + proj′ = Coker.universal (Ker.kernel f) {e′ = f} $ sym path ``` @@ -410,7 +410,7 @@ $f$ is mono), we have $0 = \ker f$ from $f0 = f\ker f$. ```agda kercoker→f : m.Hom (cut (Ker.kernel (Coker.coeq f))) (cut f) kercoker→f ./-Hom.map = - Coker.coequalise (Ker.kernel f) {e′ = id} (monic _ _ path) ∘ + Coker.universal (Ker.kernel f) {e′ = id} (monic _ _ path) ∘ coker-ker≃ker-coker f .is-invertible.inv where abstract path : f ∘ id ∘ 0m ≡ f ∘ id ∘ Ker.kernel f @@ -429,15 +429,15 @@ coequalisers are epic to make progress. kercoker→f ./-Hom.commutes = path where lemma = is-coequaliser→is-epic (Coker.coeq _) (Coker.has-is-coeq _) _ _ $ - pullr (Coker.universal _) + pullr (Coker.factors _) ·· elimr refl ·· (decompose f .snd ∙ assoc _ _ _) path = invertible→epic (coker-ker≃ker-coker _) _ _ $ - (f ∘ Coker.coequalise _ _ ∘ _) ∘ decompose f .fst ≡⟨ ap₂ _∘_ (assoc _ _ _) refl ⟩ - ((f ∘ Coker.coequalise _ _) ∘ _) ∘ decompose f .fst ≡⟨ cancelr (coker-ker≃ker-coker _ .is-invertible.invr) ⟩ - f ∘ Coker.coequalise _ _ ≡⟨ lemma ⟩ + (f ∘ Coker.universal _ _ ∘ _) ∘ decompose f .fst ≡⟨ ap₂ _∘_ (assoc _ _ _) refl ⟩ + ((f ∘ Coker.universal _ _) ∘ _) ∘ decompose f .fst ≡⟨ cancelr (coker-ker≃ker-coker _ .is-invertible.invr) ⟩ + f ∘ Coker.universal _ _ ≡⟨ lemma ⟩ Ker.kernel _ ∘ decompose f .fst ∎ ``` @@ -450,19 +450,19 @@ $\cA$, thus assemble into an isomorphism in the slice. mono→kernel = m.make-iso f→kercoker kercoker→f f→kc→f kc→f→kc where f→kc→f : f→kercoker m.∘ kercoker→f ≡ m.id f→kc→f = /-Hom-path $ - (decompose f .fst ∘ Coker.coeq _) ∘ Coker.coequalise _ _ ∘ _ ≡⟨ cancel-inner lemma ⟩ + (decompose f .fst ∘ Coker.coeq _) ∘ Coker.universal _ _ ∘ _ ≡⟨ cancel-inner lemma ⟩ decompose f .fst ∘ _ ≡⟨ coker-ker≃ker-coker f .is-invertible.invl ⟩ id ∎ where lemma = Coker.unique₂ _ {e′ = Coker.coeq (Ker.kernel f)} - {p = ∘-zero-r ∙ sym (sym (Coker.coequal _) ∙ ∘-zero-r)} - (sym (pullr (Coker.universal (Ker.kernel f)) ∙ elimr refl)) - (introl refl) + (∘-zero-r ∙ sym (sym (Coker.coequal _) ∙ ∘-zero-r)) + (pullr (Coker.factors (Ker.kernel f)) ∙ elimr refl) + (eliml refl) kc→f→kc : kercoker→f m.∘ f→kercoker ≡ m.id kc→f→kc = /-Hom-path $ - (Coker.coequalise _ _ ∘ _) ∘ decompose f .fst ∘ Coker.coeq _ ≡⟨ cancel-inner (coker-ker≃ker-coker f .is-invertible.invr) ⟩ - Coker.coequalise _ _ ∘ Coker.coeq _ ≡⟨ Coker.universal _ ⟩ + (Coker.universal _ _ ∘ _) ∘ decompose f .fst ∘ Coker.coeq _ ≡⟨ cancel-inner (coker-ker≃ker-coker f .is-invertible.invr) ⟩ + Coker.universal _ _ ∘ Coker.coeq _ ≡⟨ Coker.factors _ ⟩ id ∎ ``` diff --git a/src/Cat/Abelian/Images.lagda.md b/src/Cat/Abelian/Images.lagda.md index de14e57fc..7c5066640 100644 --- a/src/Cat/Abelian/Images.lagda.md +++ b/src/Cat/Abelian/Images.lagda.md @@ -81,7 +81,7 @@ $$ ```agda the-img .map ./-Hom.map = decompose f .fst ∘ Coker.coeq _ - the-img .map ./-Hom.commutes = pulll (Ker.universal _) ∙ Coker.universal _ + the-img .map ./-Hom.commutes = pulll (Ker.factors _) ∙ Coker.factors _ ``` ## Universality @@ -115,7 +115,7 @@ commutes. factor : ↓Hom (const! (cut f)) Forget-full-subcat the-img other factor .α = tt factor .β ./-Hom.map = - Coker.coequalise (Ker.kernel f) {e′ = other .map .map} path + Coker.universal (Ker.kernel f) {e′ = other .map .map} path ∘ coker-ker≃ker-coker f .is-invertible.inv ``` @@ -144,19 +144,21 @@ is the image of $f$.
Here's the tedious isomorphism algebra. + ```agda factor .β ./-Hom.commutes = invertible→epic (coker-ker≃ker-coker f) _ _ $ - Coker.unique₂ (Ker.kernel f) {e′ = f} - {p = sym (Ker.equal f ∙ ∅.zero-∘r _ ∙ 0m-unique ∙ sym ∘-zero-r)} - (sym ( ap₂ _∘_ ( sym (assoc _ _ _) + Coker.unique₂ (Ker.kernel f) + (sym (Ker.equal f ∙ ∅.zero-∘r _ ∙ 0m-unique ∙ sym ∘-zero-r)) + (ap₂ _∘_ ( sym (assoc _ _ _) ∙ ap₂ _∘_ refl (cancelr (coker-ker≃ker-coker f .is-invertible.invr))) refl - ∙ pullr (Coker.universal _) ∙ other .map .commutes)) - (decompose f .snd ∙ assoc _ _ _) + ∙ pullr (Coker.factors _) ∙ other .map .commutes) + (sym (decompose f .snd ∙ assoc _ _ _)) factor .sq = /-Hom-path $ sym $ other .y .witness _ _ $ - pulll (factor .β .commutes) - ·· the-img .map .commutes - ·· (sym (other .map .commutes) ∙ ap (other .y .object .map ∘_) (intror refl)) + pulll (factor .β .commutes) + ∙ the-img .map .commutes + ·· sym (other .map .commutes) + ·· ap (other .y .object .map ∘_) (intror refl) unique : ∀ x → factor ≡ x unique x = ↓Hom-path _ _ refl $ /-Hom-path $ other .y .witness _ _ $ diff --git a/src/Cat/Abelian/Limits.lagda.md b/src/Cat/Abelian/Limits.lagda.md index e8ff46720..6e173d376 100644 --- a/src/Cat/Abelian/Limits.lagda.md +++ b/src/Cat/Abelian/Limits.lagda.md @@ -61,13 +61,13 @@ module _ (A : is-pre-abelian C) where (f - g) ∘ Ker.kernel (f - g) ≡⟨ Ker.equal (f - g) ⟩ ∅.zero→ ∘ Ker.kernel (f - g) ≡⟨ ∅.zero-∘r _ ∙ 0m-unique ⟩ 0m ∎ - equ .limiting {e′ = e′} p = Ker.limiting (f - g) {e′ = e′} $ + equ .universal {e′ = e′} p = Ker.universal (f - g) {e′ = e′} $ (f - g) ∘ e′ ≡˘⟨ ∘-minus-l _ _ _ ⟩ f ∘ e′ - g ∘ e′ ≡⟨ ap (f ∘ e′ -_) (sym p) ⟩ f ∘ e′ - f ∘ e′ ≡⟨ Hom.inverser ⟩ 0m ≡˘⟨ ∅.zero-∘r _ ∙ 0m-unique ⟩ Zero.zero→ ∅ ∘ e′ ∎ - equ .universal = Ker.universal _ + equ .factors = Ker.factors _ equ .unique = Ker.unique (f - g) ``` diff --git a/src/Cat/Base.lagda.md b/src/Cat/Base.lagda.md index 32c414257..67035a3bc 100644 --- a/src/Cat/Base.lagda.md +++ b/src/Cat/Base.lagda.md @@ -473,6 +473,12 @@ module _ where Const {D = D} x .F-id = refl Const {D = D} x .F-∘ _ _ = sym (idr D _) + const-nt : ∀ {o ℓ o′ ℓ′} {C : Precategory o ℓ} {D : Precategory o′ ℓ′} + → {x y : Ob D} → Hom D x y + → Const {C = C} {D = D} x => Const {C = C} {D = D} y + const-nt f ._=>_.η _ = f + const-nt {D = D} f ._=>_.is-natural _ _ _ = idr D _ ∙ sym (idl D _) + infixr 30 _F∘_ infix 20 _=>_ @@ -486,6 +492,7 @@ module _ {o₁ h₁ o₂ h₂} module D = Precategory D module C = Precategory C + open Functor open _=>_ ``` --> @@ -531,5 +538,11 @@ is a proposition: _ηₚ_ : ∀ {a b : F => G} → a ≡ b → ∀ x → a .η x ≡ b .η x p ηₚ x = ap (λ e → e .η x) p + _ηᵈ_ : ∀ {F' G' : Functor C D} {p : F ≡ F'} {q : G ≡ G'} + → {a : F => G} {b : F' => G'} + → PathP (λ i → p i => q i) a b + → ∀ x → PathP (λ i → D.Hom (p i .F₀ x) (q i .F₀ x)) (a .η x) (b .η x) + p ηᵈ x = apd (λ i e → e .η x) p + infixl 45 _ηₚ_ ``` diff --git a/src/Cat/Bi/Base.lagda.md b/src/Cat/Bi/Base.lagda.md index a856128f5..714c675a3 100644 --- a/src/Cat/Bi/Base.lagda.md +++ b/src/Cat/Bi/Base.lagda.md @@ -1,5 +1,5 @@ ```agda -open import Cat.Instances.Functor.Compose +open import Cat.Instances.Functor.Compose renaming (_◆_ to _◇_) open import Cat.Instances.Functor open import Cat.Instances.Product open import Cat.Functor.Hom diff --git a/src/Cat/Bi/Instances/Spans.lagda.md b/src/Cat/Bi/Instances/Spans.lagda.md index 8b8656c5a..b27bad115 100644 --- a/src/Cat/Bi/Instances/Spans.lagda.md +++ b/src/Cat/Bi/Instances/Spans.lagda.md @@ -142,7 +142,7 @@ module _ (pb : ∀ {a b c} (f : Hom a b) (g : Hom c b) → Pullback C f g) where module y = Pullback (pb (y1 .left) (y2 .right)) x→y : Hom x.apex y.apex - x→y = y.limiting {p₁' = f .map ∘ x.p₁} {p₂' = g .map ∘ x.p₂} comm + x→y = y.universal {p₁' = f .map ∘ x.p₁} {p₂' = g .map ∘ x.p₂} comm where abstract open Pullback comm : y1 .left ∘ f .map ∘ x.p₁ ≡ y2 .right ∘ g .map ∘ x.p₂ @@ -150,16 +150,16 @@ module _ (pb : ∀ {a b c} (f : Hom a b) (g : Hom c b) → Pullback C f g) where res : Span-hom _ _ res .map = x→y - res .left = sym (pullr y.p₂∘limiting ∙ pulll (sym (g .left))) - res .right = sym (pullr y.p₁∘limiting ∙ pulll (sym (f .right))) + res .left = sym (pullr y.p₂∘universal ∙ pulll (sym (g .left))) + res .right = sym (pullr y.p₁∘universal ∙ pulll (sym (f .right))) Span-∘ .F-id {x1 , x2} = Span-hom-path $ sym $ x.unique id-comm id-comm where module x = Pullback (pb (x1 .left) (x2 .right)) Span-∘ .F-∘ {x1 , x2} {y1 , y2} {z1 , z2} f g = Span-hom-path $ sym $ z.unique - (pulll z.p₁∘limiting ∙ pullr y.p₁∘limiting ∙ assoc _ _ _) - (pulll z.p₂∘limiting ∙ pullr y.p₂∘limiting ∙ assoc _ _ _) + (pulll z.p₁∘universal ∙ pullr y.p₁∘universal ∙ assoc _ _ _) + (pulll z.p₂∘universal ∙ pullr y.p₂∘universal ∙ assoc _ _ _) where module x = Pullback (pb (x1 .left) (x2 .right)) module y = Pullback (pb (y1 .left) (y2 .right)) @@ -212,14 +212,14 @@ S \To S$. The right unitor is analogous. f ¤ g = Span-∘ .F₀ (f , g) sλ← : ∀ {A B} (x : Span A B) → Span-hom x (span _ C.id C.id ¤ x) - sλ← x .map = pb _ _ .limiting id-comm-sym - sλ← x .left = sym $ pullr (pb _ _ .p₂∘limiting) ∙ idr _ - sλ← x .right = sym $ pullr (pb _ _ .p₁∘limiting) ∙ idl _ + sλ← x .map = pb _ _ .universal id-comm-sym + sλ← x .left = sym $ pullr (pb _ _ .p₂∘universal) ∙ idr _ + sλ← x .right = sym $ pullr (pb _ _ .p₁∘universal) ∙ idl _ sρ← : ∀ {A B} (x : Span A B) → Span-hom x (x ¤ span _ C.id C.id) - sρ← x .map = pb _ _ .limiting id-comm - sρ← x .left = sym $ pullr (pb _ _ .p₂∘limiting) ∙ idl _ - sρ← x .right = sym $ pullr (pb _ _ .p₁∘limiting) ∙ idr _ + sρ← x .map = pb _ _ .universal id-comm + sρ← x .left = sym $ pullr (pb _ _ .p₂∘universal) ∙ idl _ + sρ← x .right = sym $ pullr (pb _ _ .p₁∘universal) ∙ idr _ ``` For the associator, while doing the construction in elementary terms is @@ -258,43 +258,43 @@ variables and) satisfy the triangle and pentagon identities. ```agda sα← : ∀ {A B C D} ((f , g , h) : Span C D × Span B C × Span A B) → Span-hom ((f ¤ g) ¤ h) (f ¤ (g ¤ h)) - sα← (f , g , h) .map = pb _ _ .limiting resp′ where + sα← (f , g , h) .map = pb _ _ .universal resp′ where abstract resp : g .left C.∘ pb (f .left) (g .right) .p₂ C.∘ pb ((f ¤ g) .left) (h .right) .p₁ ≡ h .right C.∘ pb ((f ¤ g) .left) (h .right) .p₂ resp = assoc _ _ _ ∙ pb _ _ .square - shuffle = pb _ _ .limiting {p₁' = pb _ _ .p₂ C.∘ pb _ _ .p₁} {p₂' = pb _ _ .p₂} resp + shuffle = pb _ _ .universal {p₁' = pb _ _ .p₂ C.∘ pb _ _ .p₁} {p₂' = pb _ _ .p₂} resp abstract resp′ : f .left C.∘ pb (f .left) (g .right) .p₁ C.∘ pb ((f ¤ g) .left) (h .right) .p₁ ≡ (g ¤ h) .right C.∘ shuffle - resp′ = sym $ pullr (pb _ _ .p₁∘limiting) ∙ extendl (sym (pb _ _ .square)) + resp′ = sym $ pullr (pb _ _ .p₁∘universal) ∙ extendl (sym (pb _ _ .square)) - sα← (f , g , h) .left = sym $ pullr (pb _ _ .p₂∘limiting) ∙ pullr (pb _ _ .p₂∘limiting) - sα← (f , g , h) .right = sym $ pullr (pb _ _ .p₁∘limiting) ∙ assoc _ _ _ + sα← (f , g , h) .left = sym $ pullr (pb _ _ .p₂∘universal) ∙ pullr (pb _ _ .p₂∘universal) + sα← (f , g , h) .right = sym $ pullr (pb _ _ .p₁∘universal) ∙ assoc _ _ _ sα→ : ∀ {A B C D} ((f , g , h) : Span C D × Span B C × Span A B) → Span-hom (f ¤ (g ¤ h)) ((f ¤ g) ¤ h) - sα→ (f , g , h) .map = pb _ _ .limiting resp′ where + sα→ (f , g , h) .map = pb _ _ .universal resp′ where abstract resp : f .left C.∘ pb (f .left) ((g ¤ h) .right) .p₁ ≡ g .right C.∘ pb (g .left) (h .right) .p₁ C.∘ pb (f .left) ((g ¤ h) .right) .p₂ resp = pb _ _ .square ∙ sym (assoc _ _ _) - shuffle = pb _ _ .limiting {p₁' = pb _ _ .p₁} {p₂' = pb _ _ .p₁ C.∘ pb _ _ .p₂} resp + shuffle = pb _ _ .universal {p₁' = pb _ _ .p₁} {p₂' = pb _ _ .p₁ C.∘ pb _ _ .p₂} resp abstract resp′ : (f ¤ g) .left C.∘ shuffle ≡ h .right C.∘ pb (g .left) (h .right) .p₂ C.∘ pb (f .left) ((g ¤ h) .right) .p₂ - resp′ = pullr (pb _ _ .p₂∘limiting) ∙ extendl (pb _ _ .square) + resp′ = pullr (pb _ _ .p₂∘universal) ∙ extendl (pb _ _ .square) - sα→ (f , g , h) .left = sym $ pullr (pb _ _ .p₂∘limiting) ∙ assoc _ _ _ - sα→ (f , g , h) .right = sym $ pullr (pb _ _ .p₁∘limiting) ∙ pullr (pb _ _ .p₁∘limiting) + sα→ (f , g , h) .left = sym $ pullr (pb _ _ .p₂∘universal) ∙ assoc _ _ _ + sα→ (f , g , h) .right = sym $ pullr (pb _ _ .p₁∘universal) ∙ pullr (pb _ _ .p₁∘universal) open make-natural-iso {-# TERMINATING #-} @@ -310,17 +310,17 @@ variables and) satisfy the triangle and pentagon identities. ni .inv x .left = refl ni .inv x .right = pb _ _ .square ni .eta∘inv x = Span-hom-path (Pullback.unique₂ (pb _ _) {p = idl _ ∙ ap₂ C._∘_ refl (introl refl)} - (pulll (pb _ _ .p₁∘limiting)) - (pulll (pb _ _ .p₂∘limiting)) + (pulll (pb _ _ .p₁∘universal)) + (pulll (pb _ _ .p₂∘universal)) (id-comm ∙ pb _ _ .square) id-comm) - ni .inv∘eta x = Span-hom-path (pb _ _ .p₂∘limiting) + ni .inv∘eta x = Span-hom-path (pb _ _ .p₂∘universal) ni .natural x y f = Span-hom-path $ Pullback.unique₂ (pb _ _) {p = idl _ ∙ f .right} - (pulll (pb _ _ .p₁∘limiting) ∙ pullr (pb _ _ .p₁∘limiting) ∙ idl _) - (pulll (pb _ _ .p₂∘limiting) ∙ pullr (pb _ _ .p₂∘limiting) ∙ idr _) - (pulll (pb _ _ .p₁∘limiting) ∙ sym (f .right)) - (pulll (pb _ _ .p₂∘limiting) ∙ idl _) + (pulll (pb _ _ .p₁∘universal) ∙ pullr (pb _ _ .p₁∘universal) ∙ idl _) + (pulll (pb _ _ .p₂∘universal) ∙ pullr (pb _ _ .p₂∘universal) ∙ idr _) + (pulll (pb _ _ .p₁∘universal) ∙ sym (f .right)) + (pulll (pb _ _ .p₂∘universal) ∙ idl _) Spanᵇ .unitor-r = to-natural-iso ni where ni : make-natural-iso (Id {C = Spans _ _}) _ ni .eta = sρ← @@ -328,75 +328,75 @@ variables and) satisfy the triangle and pentagon identities. ni .inv _ .left = sym (pb _ _ .square) ni .inv _ .right = refl ni .eta∘inv x = Span-hom-path (Pullback.unique₂ (pb _ _) {p = introl refl} - (pulll (pb _ _ .p₁∘limiting) ∙ idl _) - (pulll (pb _ _ .p₂∘limiting)) + (pulll (pb _ _ .p₁∘universal) ∙ idl _) + (pulll (pb _ _ .p₂∘universal)) (idr _) (id-comm ∙ sym (pb _ _ .square))) - ni .inv∘eta x = Span-hom-path (pb _ _ .p₁∘limiting) + ni .inv∘eta x = Span-hom-path (pb _ _ .p₁∘universal) ni .natural x y f = Span-hom-path $ Pullback.unique₂ (pb _ _) {p = sym (f .left) ∙ introl refl} - (pulll (pb _ _ .p₁∘limiting) ∙ pullr (pb _ _ .p₁∘limiting) ∙ idr _) - (pulll (pb _ _ .p₂∘limiting) ∙ pullr (pb _ _ .p₂∘limiting) ∙ idl _) - (pulll (pb _ _ .p₁∘limiting) ∙ idl _) - (pulll (pb _ _ .p₂∘limiting) ∙ sym (f .left)) + (pulll (pb _ _ .p₁∘universal) ∙ pullr (pb _ _ .p₁∘universal) ∙ idr _) + (pulll (pb _ _ .p₂∘universal) ∙ pullr (pb _ _ .p₂∘universal) ∙ idl _) + (pulll (pb _ _ .p₁∘universal) ∙ idl _) + (pulll (pb _ _ .p₂∘universal) ∙ sym (f .left)) Spanᵇ .associator = to-natural-iso ni where ni : make-natural-iso _ _ ni .eta = sα← ni .inv = sα→ ni .eta∘inv x = Span-hom-path $ Pullback.unique₂ (pb _ _) {p = pb _ _ .square} - (pulll (pb _ _ .p₁∘limiting) ∙ pullr (pb _ _ .p₁∘limiting) ∙ pb _ _ .p₁∘limiting) - (pulll (pb _ _ .p₂∘limiting) ∙ unique₂ (pb _ _) {p = extendl (pb _ _ .square)} - (pulll (pb _ _ .p₁∘limiting) ∙ pullr (pb _ _ .p₁∘limiting) ∙ pb _ _ .p₂∘limiting) - (pulll (pb _ _ .p₂∘limiting) ∙ pb _ _ .p₂∘limiting) + (pulll (pb _ _ .p₁∘universal) ∙ pullr (pb _ _ .p₁∘universal) ∙ pb _ _ .p₁∘universal) + (pulll (pb _ _ .p₂∘universal) ∙ unique₂ (pb _ _) {p = extendl (pb _ _ .square)} + (pulll (pb _ _ .p₁∘universal) ∙ pullr (pb _ _ .p₁∘universal) ∙ pb _ _ .p₂∘universal) + (pulll (pb _ _ .p₂∘universal) ∙ pb _ _ .p₂∘universal) refl refl) (idr _) (idr _) ni .inv∘eta x = Span-hom-path $ Pullback.unique₂ (pb _ _) {p = pb _ _ .square} - (pulll (pb _ _ .p₁∘limiting) ∙ unique₂ (pb _ _) {p = extendl (pb _ _ .square)} - (pulll (pb _ _ .p₁∘limiting) ∙ pb _ _ .p₁∘limiting) - (pulll (pb _ _ .p₂∘limiting) ∙ pullr (pb _ _ .p₂∘limiting) ∙ pb _ _ .p₁∘limiting) + (pulll (pb _ _ .p₁∘universal) ∙ unique₂ (pb _ _) {p = extendl (pb _ _ .square)} + (pulll (pb _ _ .p₁∘universal) ∙ pb _ _ .p₁∘universal) + (pulll (pb _ _ .p₂∘universal) ∙ pullr (pb _ _ .p₂∘universal) ∙ pb _ _ .p₁∘universal) refl refl) - (pulll (pb _ _ .p₂∘limiting) ∙ pullr (pb _ _ .p₂∘limiting) ∙ pb _ _ .p₂∘limiting) + (pulll (pb _ _ .p₂∘universal) ∙ pullr (pb _ _ .p₂∘universal) ∙ pb _ _ .p₂∘universal) (idr _) (idr _) ni .natural x y f = Span-hom-path $ Pullback.unique₂ (pb _ _) {p₁' = f .fst .map C.∘ pb _ _ .p₁ C.∘ pb _ _ .p₁} - {p₂' = pb _ _ .limiting + {p₂' = pb _ _ .universal {p₁' = f .snd .fst .map C.∘ pb _ _ .p₂ C.∘ pb _ _ .p₁} {p₂' = f .snd .snd .map C.∘ pb _ _ .p₂} (pulll (sym (f .snd .fst .left)) ∙ assoc _ _ _ ∙ pb _ _ .square ∙ pushl (f .snd .snd .right))} - {p = sym $ pullr (pb _ _ .p₁∘limiting) ∙ pulll (sym (f .snd .fst .right)) ∙ extendl (sym (pb _ _ .square)) ∙ pushl (f .fst .left)} - (pulll (pb _ _ .p₁∘limiting) ∙ pullr (pb _ _ .p₁∘limiting)) - (pulll (pb _ _ .p₂∘limiting) ∙ pb _ _ .unique - (pulll (extendl (pb _ _ .p₁∘limiting)) ∙ pullr (pullr (pb _ _ .p₂∘limiting)) ∙ ap₂ C._∘_ refl (pb _ _ .p₁∘limiting)) - (pulll (extendl (pb _ _ .p₂∘limiting)) ∙ pullr (pullr (pb _ _ .p₂∘limiting)) ∙ ap₂ C._∘_ refl (pb _ _ .p₂∘limiting))) - (pulll (pb _ _ .p₁∘limiting) ∙ pullr (pb _ _ .p₁∘limiting) ∙ pulll (pb _ _ .p₁∘limiting) ∙ sym (assoc _ _ _)) - (pulll (pb _ _ .p₂∘limiting) ∙ pb _ _ .unique - (pulll (pb _ _ .p₁∘limiting) ∙ pullr (pb _ _ .p₁∘limiting) ∙ extendl (pb _ _ .p₂∘limiting)) - (pulll (pb _ _ .p₂∘limiting) ∙ pb _ _ .p₂∘limiting)) + {p = sym $ pullr (pb _ _ .p₁∘universal) ∙ pulll (sym (f .snd .fst .right)) ∙ extendl (sym (pb _ _ .square)) ∙ pushl (f .fst .left)} + (pulll (pb _ _ .p₁∘universal) ∙ pullr (pb _ _ .p₁∘universal)) + (pulll (pb _ _ .p₂∘universal) ∙ pb _ _ .unique + (pulll (extendl (pb _ _ .p₁∘universal)) ∙ pullr (pullr (pb _ _ .p₂∘universal)) ∙ ap₂ C._∘_ refl (pb _ _ .p₁∘universal)) + (pulll (extendl (pb _ _ .p₂∘universal)) ∙ pullr (pullr (pb _ _ .p₂∘universal)) ∙ ap₂ C._∘_ refl (pb _ _ .p₂∘universal))) + (pulll (pb _ _ .p₁∘universal) ∙ pullr (pb _ _ .p₁∘universal) ∙ pulll (pb _ _ .p₁∘universal) ∙ sym (assoc _ _ _)) + (pulll (pb _ _ .p₂∘universal) ∙ pb _ _ .unique + (pulll (pb _ _ .p₁∘universal) ∙ pullr (pb _ _ .p₁∘universal) ∙ extendl (pb _ _ .p₂∘universal)) + (pulll (pb _ _ .p₂∘universal) ∙ pb _ _ .p₂∘universal)) Spanᵇ .triangle f g = Span-hom-path $ pb _ _ .unique - (pulll (pb _ _ .p₁∘limiting) ∙ pullr (pb _ _ .p₁∘limiting) ∙ pb _ _ .p₁∘limiting ∙ introl refl) - (pulll (pb _ _ .p₂∘limiting) ∙ pullr (pb _ _ .p₂∘limiting) ∙ eliml refl) + (pulll (pb _ _ .p₁∘universal) ∙ pullr (pb _ _ .p₁∘universal) ∙ pb _ _ .p₁∘universal ∙ introl refl) + (pulll (pb _ _ .p₂∘universal) ∙ pullr (pb _ _ .p₂∘universal) ∙ eliml refl) Spanᵇ .pentagon f g h i = Span-hom-path $ Pullback.unique₂ (pb _ _) - {p = pullr (pulll (pb _ _ .p₂∘limiting) ∙ pullr (pulll (pb _ _ .p₂∘limiting) ∙ pullr (pb _ _ .p₂∘limiting)) ∙ ap₂ C._∘_ refl (pulll (pb _ _ .p₁∘limiting))) - ∙ ap₂ C._∘_ refl (extendl (pb _ _ .p₂∘limiting)) ∙ sym (ap₂ C._∘_ refl (idl _ ∙ extendl (pb _ _ .p₂∘limiting)) ∙ extendl (sym (pb _ _ .square)))} - (pulll (pb _ _ .p₁∘limiting) ∙ pullr (pulll (pb _ _ .p₁∘limiting))) - (pulll (pb _ _ .p₂∘limiting) ∙ pullr (pulll (pb _ _ .p₂∘limiting) ∙ pullr (pb _ _ .p₂∘limiting))) - (pulll (pb _ _ .p₁∘limiting) - ∙ Pullback.unique₂ (pb _ _) {p = pullr (pb _ _ .p₂∘limiting) ∙ extendl (pb _ _ .square) ∙ sym (assoc _ _ _)} - (pulll (pb _ _ .p₁∘limiting) ∙ pb _ _ .p₁∘limiting) - (pulll (pb _ _ .p₂∘limiting) ∙ pullr (pb _ _ .p₂∘limiting)) - (pulll (pb _ _ .p₁∘limiting) ∙ pb _ _ .unique - (pulll (pb _ _ .p₁∘limiting) ∙ pulll (pb _ _ .p₁∘limiting) ∙ pb _ _ .p₁∘limiting ∙ idl _) - (pulll (pb _ _ .p₂∘limiting) ∙ pulll (pullr (pb _ _ .p₂∘limiting)) ∙ pullr (pullr (pb _ _ .p₂∘limiting) ∙ pulll (pb _ _ .p₁∘limiting)) ∙ pulll (pb _ _ .p₁∘limiting))) - (pulll (pb _ _ .p₂∘limiting) ∙ pullr (pulll (pb _ _ .p₂∘limiting) ∙ pullr (pb _ _ .p₂∘limiting)) - ∙ ap₂ C._∘_ refl (pulll (pb _ _ .p₁∘limiting)) ∙ pulll (pb _ _ .p₂∘limiting) ∙ sym (assoc _ _ _))) - ( pulll (pb _ _ .p₂∘limiting) - ·· pullr (pb _ _ .p₂∘limiting) - ·· sym (idl _ ·· pulll (pb _ _ .p₂∘limiting) ·· sym (assoc _ _ _))) + {p = pullr (pulll (pb _ _ .p₂∘universal) ∙ pullr (pulll (pb _ _ .p₂∘universal) ∙ pullr (pb _ _ .p₂∘universal)) ∙ ap₂ C._∘_ refl (pulll (pb _ _ .p₁∘universal))) + ∙ ap₂ C._∘_ refl (extendl (pb _ _ .p₂∘universal)) ∙ sym (ap₂ C._∘_ refl (idl _ ∙ extendl (pb _ _ .p₂∘universal)) ∙ extendl (sym (pb _ _ .square)))} + (pulll (pb _ _ .p₁∘universal) ∙ pullr (pulll (pb _ _ .p₁∘universal))) + (pulll (pb _ _ .p₂∘universal) ∙ pullr (pulll (pb _ _ .p₂∘universal) ∙ pullr (pb _ _ .p₂∘universal))) + (pulll (pb _ _ .p₁∘universal) + ∙ Pullback.unique₂ (pb _ _) {p = pullr (pb _ _ .p₂∘universal) ∙ extendl (pb _ _ .square) ∙ sym (assoc _ _ _)} + (pulll (pb _ _ .p₁∘universal) ∙ pb _ _ .p₁∘universal) + (pulll (pb _ _ .p₂∘universal) ∙ pullr (pb _ _ .p₂∘universal)) + (pulll (pb _ _ .p₁∘universal) ∙ pb _ _ .unique + (pulll (pb _ _ .p₁∘universal) ∙ pulll (pb _ _ .p₁∘universal) ∙ pb _ _ .p₁∘universal ∙ idl _) + (pulll (pb _ _ .p₂∘universal) ∙ pulll (pullr (pb _ _ .p₂∘universal)) ∙ pullr (pullr (pb _ _ .p₂∘universal) ∙ pulll (pb _ _ .p₁∘universal)) ∙ pulll (pb _ _ .p₁∘universal))) + (pulll (pb _ _ .p₂∘universal) ∙ pullr (pulll (pb _ _ .p₂∘universal) ∙ pullr (pb _ _ .p₂∘universal)) + ∙ ap₂ C._∘_ refl (pulll (pb _ _ .p₁∘universal)) ∙ pulll (pb _ _ .p₂∘universal) ∙ sym (assoc _ _ _))) + ( pulll (pb _ _ .p₂∘universal) + ·· pullr (pb _ _ .p₂∘universal) + ·· sym (idl _ ·· pulll (pb _ _ .p₂∘universal) ·· sym (assoc _ _ _))) ```
diff --git a/src/Cat/CartesianClosed/Base.lagda.md b/src/Cat/CartesianClosed/Base.lagda.md index 0eccb80bb..631e14744 100644 --- a/src/Cat/CartesianClosed/Base.lagda.md +++ b/src/Cat/CartesianClosed/Base.lagda.md @@ -51,7 +51,7 @@ record is-cc {o ℓ} (C : Precategory o ℓ) (cartesian : ∀ A B → Product C terminal : Terminal C open Cat.Reasoning C - open Cartesian C cartesian public + open Binary-products C cartesian public private module ×-Bifunctor = Bifunctor {C = C} {C} {C} ×-functor @@ -86,13 +86,13 @@ uncurrying transformations: [adjuncts]: Cat.Functor.Adjoint.html#adjuncts ```agda - ev : Hom ([ X , Y ] ⊗ X) Y + ev : Hom ([ X , Y ] ⊗₀ X) Y ev = T⊣H.counit.ε _ - curry : Hom (X ⊗ Y) Z → Hom X [ Y , Z ] + curry : Hom (X ⊗₀ Y) Z → Hom X [ Y , Z ] curry = L-adjunct (tensor⊣hom _) - uncurry : Hom X [ Y , Z ] → Hom (X ⊗ Y) Z + uncurry : Hom X [ Y , Z ] → Hom (X ⊗₀ Y) Z uncurry = R-adjunct (tensor⊣hom _) ``` diff --git a/src/Cat/CartesianClosed/Instances/PSh.agda b/src/Cat/CartesianClosed/Instances/PSh.agda index 21f6e992d..fdb6c2a41 100644 --- a/src/Cat/CartesianClosed/Instances/PSh.agda +++ b/src/Cat/CartesianClosed/Instances/PSh.agda @@ -77,11 +77,11 @@ module _ {o ℓ κ} {C : Precategory o ℓ} where pb .p₂ .η x (a , b , _) = b pb .p₂ .is-natural _ _ _ = refl pb .has-is-pb .square = Nat-path λ _ → funext λ (_ , _ , p) → p - pb .has-is-pb .limiting path .η idx arg = _ , _ , (path ηₚ idx $ₚ arg) - pb .has-is-pb .limiting {p₁' = p₁'} {p₂'} path .is-natural x y f = + pb .has-is-pb .universal path .η idx arg = _ , _ , (path ηₚ idx $ₚ arg) + pb .has-is-pb .universal {p₁' = p₁'} {p₂'} path .is-natural x y f = funext λ x → pb-path (happly (p₁' .is-natural _ _ _) _) (happly (p₂' .is-natural _ _ _) _) - pb .has-is-pb .p₁∘limiting = Nat-path λ _ → refl - pb .has-is-pb .p₂∘limiting = Nat-path λ _ → refl + pb .has-is-pb .p₁∘universal = Nat-path λ _ → refl + pb .has-is-pb .p₂∘universal = Nat-path λ _ → refl pb .has-is-pb .unique p q = Nat-path λ _ → funext λ _ → pb-path (p ηₚ _ $ₚ _) (q ηₚ _ $ₚ _) @@ -165,13 +165,13 @@ module _ {o ℓ κ} {C : Precategory o ℓ} where coequ .coeq .η i = incq coequ .coeq .is-natural x y f = refl coequ .has-is-coeq .coequal = Nat-path λ _ → funext λ x → glue x - coequ .has-is-coeq .coequalise {F = F} {e′ = e′} p .η x = + coequ .has-is-coeq .universal {F = F} {e′ = e′} p .η x = Coeq-rec (F .F₀ x .is-tr) (e′ .η x) (p ηₚ x $ₚ_) - coequ .has-is-coeq .coequalise {F = F} {e′ = e′} p .is-natural x y f = funext $ + coequ .has-is-coeq .universal {F = F} {e′ = e′} p .is-natural x y f = funext $ Coeq-elim-prop (λ _ → F .F₀ _ .is-tr _ _) λ _ → happly (e′ .is-natural _ _ _) _ - coequ .has-is-coeq .universal = Nat-path λ _ → refl + coequ .has-is-coeq .factors = Nat-path λ _ → refl coequ .has-is-coeq .unique {F = F} p = Nat-path λ i → funext $ - Coeq-elim-prop (λ _ → F .F₀ _ .is-tr _ _) λ x → sym p ηₚ i $ₚ x + Coeq-elim-prop (λ _ → F .F₀ _ .is-tr _ _) λ x → p ηₚ i $ₚ x module _ {κ} {C : Precategory κ κ} where private @@ -182,7 +182,7 @@ module _ {κ} {C : Precategory κ κ} where PSh-closed = cc where cat = PSh κ C - open Cartesian cat (PSh-products {C = C}) public + open Binary-products cat (PSh-products {C = C}) public module _ (A : PSh.Ob) where module A = Functor A @@ -191,7 +191,7 @@ module _ {κ} {C : Precategory κ κ} where hom₀ B = F where module B = Functor B F : PSh.Ob - F .F₀ c = el ((よ₀ C c ⊗ A) => B) Nat-is-set + F .F₀ c = el ((よ₀ C c ⊗₀ A) => B) Nat-is-set F .F₁ f nt .η i (g , x) = nt .η i (f C.∘ g , x) F .F₁ f nt .is-natural x y g = funext λ o → ap (nt .η y) (Σ-pathp (C.assoc _ _ _) refl) ∙ happly (nt .is-natural _ _ _) _ diff --git a/src/Cat/CartesianClosed/Locally.lagda.md b/src/Cat/CartesianClosed/Locally.lagda.md index c1fbcb1f3..1b28a0006 100644 --- a/src/Cat/CartesianClosed/Locally.lagda.md +++ b/src/Cat/CartesianClosed/Locally.lagda.md @@ -69,7 +69,7 @@ automatically Cartesian categories, since products in $\cC/a$ are computed as pullbacks in $\cC$. ```agda - module /-Prods (a : Ob) = Cartesian (Slice C a) + module /-Prods (a : Ob) = Binary-products (Slice C a) (λ A B → Pullback→Fibre-product (pullbacks (A .map) (B .map))) ``` @@ -135,7 +135,7 @@ adjunction $(\sum_f f^*) \dashv (\prod_f f^*)$. tensor⊣hom′ = LF⊣GR (f*⊣Πf _) (Σf⊣f* pullbacks f.map) -- The product functor we have to give an adjoint to... - product = Cartesian.×-functor (Slice C a) (λ A B → Pullback→Fibre-product (pullbacks (A .map) (B .map))) + product = Binary-products.×-functor (Slice C a) (λ A B → Pullback→Fibre-product (pullbacks (A .map) (B .map))) a×- = Left product f -- ... is the same that we already proved is left adjoint to hom! @@ -148,8 +148,8 @@ adjunction $(\sum_f f^*) \dashv (\prod_f f^*)$. → PathP (λ i → /-Hom (p1 x i) (p1 y i)) (F₁ tensor g) (F₁ a×- g) p2 {x} {y} g = /-Hom-pathp _ _ (Pullback.unique₂ (pullbacks (y .map) (f .map)) {p = path} - (Pullback.p₁∘limiting (pullbacks _ _)) (Pullback.p₂∘limiting (pullbacks _ _)) - (Pullback.p₁∘limiting (pullbacks _ _)) (Pullback.p₂∘limiting (pullbacks _ _) ∙ idl _)) + (Pullback.p₁∘universal (pullbacks _ _)) (Pullback.p₂∘universal (pullbacks _ _)) + (Pullback.p₁∘universal (pullbacks _ _)) (Pullback.p₂∘universal (pullbacks _ _) ∙ idl _)) where path = y .map ∘ g .map ∘ Pullback.p₁ (pullbacks _ _) ≡⟨ pulll (g .commutes) ⟩ diff --git a/src/Cat/Diagram/Coend.lagda.md b/src/Cat/Diagram/Coend.lagda.md index c9f28d1a4..6b136fe57 100644 --- a/src/Cat/Diagram/Coend.lagda.md +++ b/src/Cat/Diagram/Coend.lagda.md @@ -66,7 +66,7 @@ if our codomain category $\cD$ does not have a well-behaved calculus of tensor products. As a motivating example, the computation of [left Kan extensions][lan] as certain colimits has this form! -[lan]: Cat.Functor.Kan.html#a-formula +[lan]: Cat.Functor.Kan.Pointwise.html#computing-pointwise-extensions We call such an object a **coend** of the functor $F$, and denote it by $\int^c F(c,c)$ (or $\int F$ for short). Being a type of [colimit], diff --git a/src/Cat/Diagram/Coend/Formula.lagda.md b/src/Cat/Diagram/Coend/Formula.lagda.md index 251b0ef10..b89c8faf7 100644 --- a/src/Cat/Diagram/Coend/Formula.lagda.md +++ b/src/Cat/Diagram/Coend/Formula.lagda.md @@ -17,9 +17,7 @@ module Cat.Diagram.Coend.Formula @@ -52,24 +50,26 @@ module _ (F : Functor (C ^op ×ᶜ C) D) where module C = Cat C module D = Cat D module F = F-r F - - cocone→cowedge : Cocone (twistᵒᵖ F) → Cowedge F - cocone→cowedge x .nadir = x .coapex - cocone→cowedge x .ψ c = x .ψ ((c , c) , C.id) - cocone→cowedge x .extranatural f = - x .ψ (_ , C.id) D.∘ Bifunctor.second F f ≡⟨ x .commutes (record { commutes = C.eliml (C.idl C.id) }) ⟩ - x .ψ (_ , f) ≡˘⟨ x .commutes (record { commutes = C.cancelr (C.idl C.id) }) ⟩ - x .ψ (_ , C.id) D.∘ Bifunctor.first F f ∎ - - cowedge→cocone : Cowedge F → Cocone (twistᵒᵖ F) - cowedge→cocone W .coapex = W .nadir - cowedge→cocone W .ψ ((c , c′) , f) = W .ψ c D.∘ Bifunctor.second F f - cowedge→cocone W .commutes {(a , b) , f} {(x , y) , g} h = - (W .ψ x D.∘ Bifunctor.second F g) D.∘ F.₁ (_ , _) ≡⟨ W .extranatural g D.⟩∘⟨refl ⟩ - (W .ψ y D.∘ Bifunctor.first F g) D.∘ F.₁ (_ , _) ≡⟨ D.pullr (F.weave (Σ-pathp (C.introl refl) refl)) ⟩ - W .ψ y D.∘ Bifunctor.first F (Twist.before h C.∘ g) D.∘ Bifunctor.second F (Twist.after h) ≡⟨ D.extendl (sym (W .extranatural _)) ⟩ - W .ψ a D.∘ Bifunctor.second F (Twist.before h C.∘ g) D.∘ Bifunctor.second F (Twist.after h) ≡⟨ D.refl⟩∘⟨ sym (Bifunctor.second∘second F) ∙ ap (Bifunctor.second F) (h .Twist.commutes) ⟩ - W .ψ a D.∘ Bifunctor.second F f ∎ + open _=>_ + open Twist + open Bifunctor + + cocone→cowedge : ∀ {x} → twistᵒᵖ F => Const x → Cowedge F + cocone→cowedge eta .nadir = _ + cocone→cowedge eta .ψ c = eta .η ((c , c) , C.id) + cocone→cowedge eta .extranatural f = + eta .is-natural _ _ (twist _ _ (C.eliml (C.idl _))) + ∙ (sym $ eta .is-natural _ _ (twist _ _ (C.cancelr (C.idl _)))) + + cowedge→cocone : (W : Cowedge F) → twistᵒᵖ F => Const (W .nadir) + cowedge→cocone W .η ((c , c') , f) = W .ψ c D.∘ second F f + cowedge→cocone W .is-natural ((a , b) , f) ((x , y) , g) h = + (W .ψ x D.∘ F.F₁ (C.id , g)) D.∘ F.F₁ (_ , _) ≡⟨ W .extranatural g D.⟩∘⟨refl ⟩ + (W .ψ y D.∘ F.F₁ (g , C.id)) D.∘ F.F₁ (h .before , h .after) ≡⟨ D.pullr (F.weave (C.introl refl ,ₚ refl)) ⟩ + W .ψ y D.∘ ((F.F₁ (h .before C.∘ g , C.id)) D.∘ F.F₁ (C.id , h .after)) ≡⟨ D.extendl (sym (W .extranatural _)) ⟩ + (W .ψ a D.∘ (F.F₁ (C.id , h .before C.∘ g) D.∘ F.F₁ (C.id , h .after))) ≡⟨ D.refl⟩∘⟨ sym (Bifunctor.second∘second F) ∙ ap (Bifunctor.second F) (h .commutes) ⟩ + W .ψ a D.∘ F.F₁ (C.id , f) ≡⟨ sym (D.idl _) ⟩ + D.id D.∘ W .ψ a D.∘ F.F₁ (C.id , f) ∎ ``` We can now extend that correspondence to calculating coends as certain @@ -79,20 +79,21 @@ colimits: $\cD$ has a coend for $F$ if it has a colimit for $F\pi_t$. colimit→coend : Colimit (twistᵒᵖ F) → Coend F colimit→coend colim = coend where open Coend - module W = Initial colim + module W = Colimit colim coend : Coend F - - coend .Coend.cowedge = cocone→cowedge (colim .Initial.bot) - coend .Coend.factor W′ = W.¡ {x = cowedge→cocone W′} .hom - coend .Coend.commutes = W.¡ .commutes _ ∙ D.elimr F.F-id - coend .Coend.unique {W = W} {g = g} comm = - ap hom ⊙ sym $ W.¡-unique $ cocone-hom _ λ o → sym $ square (o .snd) where - square : ∀ {a b} (f : C.Hom b a) → _ - square {a} {b} f = - W .ψ a D.∘ Bifunctor.second F f ≡⟨ W .extranatural f ⟩ - W .ψ b D.∘ Bifunctor.first F f ≡⟨ D.pushl (sym comm) ⟩ - g D.∘ W.bot .ψ (_ , C.id) D.∘ Bifunctor.first F f ≡⟨ D.refl⟩∘⟨ W.bot .commutes (record { before = f ; after = C.id ; commutes = C.cancelr (C.idl _) }) ⟩ - g D.∘ W.bot .ψ (_ , f) ∎ + coend .cowedge = cocone→cowedge W.cocone + coend .factor W′ = + W.universal + (cowedge→cocone W′ .η) + (λ f → cowedge→cocone W′ .is-natural _ _ f ∙ D.idl _) + coend .commutes {W = W′} = + W.factors _ _ ∙ D.elimr (Bifunctor.second-id F) + coend .unique {W = W′} comm = + W.unique _ _ _ $ λ j → + sym $ + W′ .extranatural _ + ·· D.pushl (sym comm) + ·· (D.refl⟩∘⟨ (W.commutes (twist _ _ (C.cancelr (C.idl _))))) cocomplete→coend : is-cocomplete (o ⊔ ℓ) ℓ D → Coend F cocomplete→coend colim = colimit→coend (colim _) diff --git a/src/Cat/Diagram/Coequaliser.lagda.md b/src/Cat/Diagram/Coequaliser.lagda.md index 4167618ac..4449f96a7 100644 --- a/src/Cat/Diagram/Coequaliser.lagda.md +++ b/src/Cat/Diagram/Coequaliser.lagda.md @@ -36,23 +36,24 @@ and $g$. record is-coequaliser {E} (f g : Hom A B) (coeq : Hom B E) : Type (o ⊔ ℓ) where field coequal : coeq ∘ f ≡ coeq ∘ g - coequalise : ∀ {F} {e′ : Hom B F} (p : e′ ∘ f ≡ e′ ∘ g) → Hom E F - universal : ∀ {F} {e′ : Hom B F} {p : e′ ∘ f ≡ e′ ∘ g} - → coequalise p ∘ coeq ≡ e′ + universal : ∀ {F} {e′ : Hom B F} (p : e′ ∘ f ≡ e′ ∘ g) → Hom E F + factors : ∀ {F} {e′ : Hom B F} {p : e′ ∘ f ≡ e′ ∘ g} + → universal p ∘ coeq ≡ e′ unique : ∀ {F} {e′ : Hom B F} {p : e′ ∘ f ≡ e′ ∘ g} {colim : Hom E F} - → e′ ≡ colim ∘ coeq - → colim ≡ coequalise p + → colim ∘ coeq ≡ e′ + → colim ≡ universal p unique₂ - : ∀ {F} {e′ : Hom B F} {p : e′ ∘ f ≡ e′ ∘ g} {colim' colim'' : Hom E F} - → e′ ≡ colim' ∘ coeq - → e′ ≡ colim'' ∘ coeq - → colim' ≡ colim'' - unique₂ {p = p} q r = unique {p = p} q ∙ sym (unique r) - - id-coequalise : id ≡ coequalise coequal - id-coequalise = unique (sym (idl _)) + : ∀ {F} {e′ : Hom B F} {o1 o2 : Hom E F} + → (e′ ∘ f ≡ e′ ∘ g) + → o1 ∘ coeq ≡ e′ + → o2 ∘ coeq ≡ e′ + → o1 ≡ o2 + unique₂ p q r = unique {p = p} q ∙ sym (unique r) + + id-coequalise : id ≡ universal coequal + id-coequalise = unique (idl _) ``` There is also a convenient bundling of an coequalising arrow together with @@ -81,8 +82,8 @@ is-coequaliser→is-epic → is-coequaliser f g coequ → is-epic coequ is-coequaliser→is-epic {f = f} {g = g} equ equalises h i p = - h ≡⟨ unique (sym p) ⟩ - coequalise (extendr coequal) ≡˘⟨ unique refl ⟩ + h ≡⟨ unique p ⟩ + universal (extendr coequal) ≡˘⟨ unique refl ⟩ i ∎ where open is-coequaliser equalises @@ -93,9 +94,9 @@ coequaliser-unique → E ≅ E′ coequaliser-unique {c1 = c1} {c2} co1 co2 = make-iso - (co1 .coequalise {e′ = c2} (co2 .coequal)) - (co2 .coequalise {e′ = c1} (co1 .coequal)) - (unique₂ co2 {p = co2 .coequal} (sym (pullr (co2 .universal) ∙ co1 .universal)) (introl refl)) - (unique₂ co1 {p = co1 .coequal} (sym (pullr (co1 .universal) ∙ co2 .universal)) (introl refl)) + (co1 .universal {e′ = c2} (co2 .coequal)) + (co2 .universal {e′ = c1} (co1 .coequal)) + (unique₂ co2 (co2 .coequal) (pullr (co2 .factors) ∙ co1 .factors) (idl _)) + (unique₂ co1 (co1 .coequal) (pullr (co1 .factors) ∙ co2 .factors) (idl _)) where open is-coequaliser ``` diff --git a/src/Cat/Diagram/Coequaliser/RegularEpi.lagda.md b/src/Cat/Diagram/Coequaliser/RegularEpi.lagda.md index 6a037fdcd..f065f9517 100644 --- a/src/Cat/Diagram/Coequaliser/RegularEpi.lagda.md +++ b/src/Cat/Diagram/Coequaliser/RegularEpi.lagda.md @@ -106,13 +106,13 @@ is-regular-epi→is-effective-epi {f = f} kp reg = epi where epi .p₂ = kp.p₂ epi .is-kernel-pair = kp.has-is-pb epi .has-is-coeq .coequal = kp.square - epi .has-is-coeq .coequalise {F = F} {e′} p = reg.coequalise q where + epi .has-is-coeq .universal {F = F} {e′} p = reg.universal q where q : e′ ∘ reg.arr₁ ≡ e′ ∘ reg.arr₂ q = - e′ ∘ reg.arr₁ ≡⟨ ap (e′ ∘_) (sym kp.p₂∘limiting) ⟩ - e′ ∘ kp.p₂ ∘ kp.limiting (sym reg.coequal) ≡⟨ pulll (sym p) ⟩ - (e′ ∘ kp.p₁) ∘ kp.limiting _ ≡⟨ pullr kp.p₁∘limiting ⟩ + e′ ∘ reg.arr₁ ≡⟨ ap (e′ ∘_) (sym kp.p₂∘universal) ⟩ + e′ ∘ kp.p₂ ∘ kp.universal (sym reg.coequal) ≡⟨ pulll (sym p) ⟩ + (e′ ∘ kp.p₁) ∘ kp.universal _ ≡⟨ pullr kp.p₁∘universal ⟩ e′ ∘ reg.arr₂ ∎ - epi .has-is-coeq .universal = reg.universal + epi .has-is-coeq .factors = reg.factors epi .has-is-coeq .unique = reg.unique ``` diff --git a/src/Cat/Diagram/Colimit/Base.lagda.md b/src/Cat/Diagram/Colimit/Base.lagda.md index dd35adf0a..b2dee6794 100644 --- a/src/Cat/Diagram/Colimit/Base.lagda.md +++ b/src/Cat/Diagram/Colimit/Base.lagda.md @@ -1,9 +1,13 @@ ```agda -open import Cat.Diagram.Initial +open import Cat.Instances.Shape.Terminal +open import Cat.Functor.Kan.Unique +open import Cat.Functor.Coherence +open import Cat.Instances.Functor +open import Cat.Functor.Kan.Base open import Cat.Prelude import Cat.Functor.Reasoning as Func -import Cat.Morphism +import Cat.Reasoning module Cat.Diagram.Colimit.Base where ``` @@ -18,232 +22,633 @@ private variable # Idea Colimits are dual to limits [limit]; much like their cousins, 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". +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". [limit]: Cat.Diagram.Limit.Base.html -# Construction +We define colimits in a similar way to limits; the only difference being +that we define a colimits of a diagram $F$ as left [Kan extensions] +instead of right [Kan extensions]. This gives us the expected "mapping +out" universal property, as opposed to the "mapping in" property +associated to limits. -Every concrete colimit ([coproducts], [coequalisers], [initial objects]) -we have seen so far consists of roughly the same data. We begin with -some collection of objects and morphisms, and then state that the -colimit of that collection is some object with a universal property -relating all of those objects and morphisms. +[Kan extensions]: Cat.Functor.Kan.Base.html -[coproducts]: Cat.Diagram.Coproduct.html -[coequalisers]: Cat.Diagram.Coequaliser.html -[initial objects]: Cat.Diagram.Initial.html +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. -It would be convenient to be able to talk about _all_ of these -situations at once, as opposed to on a case-by-case basis. To do this, -we need to introduce a bit of categorical machinery: the Cocone. +[cocones]: Cat.Diagram.Colimit.Cocone.html -The first step is to generalize the "collection of objects and -morphisms" involved. Luckily, this step involves no new ideas, just -a change in perspective. We already have a way of describing a -collection of objects and morphisms: it's a category! As an example, -the starting data of a coproduct is a pair of objects, which can -be thought of as a very simple category, with only identity morphisms. + + +```agda +module _ {J : Precategory o₁ h₁} {C : Precategory o₂ h₂} (Diagram : Functor J C) where + private + module C = Precategory C + + cocone→unit : ∀ {x : C.Ob} → (Diagram => Const x) → Diagram => const! x F∘ !F + unquoteDef cocone→unit = define-coherence cocone→unit + + is-colimit : (x : C.Ob) → Diagram => Const x → Type _ + is-colimit x cocone = + is-lan !F Diagram (const! x) (cocone→unit cocone) + + Colimit : Type _ + Colimit = Lan !F Diagram +``` -~~~{.quiver .short-2} -\[\begin{tikzcd} -A & B -\end{tikzcd}\] -~~~ +## Concretely -The next step also involves nothing more than a change in perspective. -Let's denote our "diagram" category from earlier as $J$. Then, a means -of picking out a $J$ shaped figure in $C$ is... a functor $F : J \to C$! -Going back to the coproduct example, a functor from our 2 object -category into $C$ selects 2 (not necessarily distinct!) objects in $C$. -We this pair of category and functor a _diagram_ in $C$. +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`{.Agda} 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`{.Agda}, as an intermediate step in constructing left +extensions. + + -Now, for the actual machinery! If we want to capture the essence of -all of our concrete examples of colimits, we a notion of an object -equipped with maps _from_ every object in our diagram. We call this -designated object the "coapex" of the cocone. - +First, we require morphisms from the every value of the diagram to the +coapex; taken as a family, we call it $\phi$. Moreover, if $f : x \to y$ +is a morphism in the "shape" category $\cJ$, we require $\psi y \circ Ff += \psi x$, which encodes the relevant naturality. ```agda field - coapex : C.Ob - ψ : (x : J.Ob) → C.Hom (F.₀ x) coapex + ψ : (j : J.Ob) → C.Hom (F₀ j) coapex + commutes : ∀ {x y} (f : J.Hom x y) → ψ y C.∘ F₁ f ≡ ψ x +``` + +The rest of the data ensures that $\psi$ is the universal family of maps +with this property; if $\eps_j : Fj \to x$ is another natural family, +then each $\eps_j$ factors through the coapex by a _unique_ universal +morphism: + +```agda + 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`{.Agda}. 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. + + -If our diagram consisted of only objects, we would be done! However, -some diagrams contain non-identity morphisms, so we need to take those -into account as well. This bit is best understood through the lens of -the coequaliser: in order to describe the commuting condition there, -we want every injection map from the codomain of a morphism to -factor through that morphism. +```agda + to-is-colimit + : ∀ {Diagram : Functor J C} {coapex} + → (mc : make-is-colimit Diagram coapex) + → is-colimit Diagram coapex (to-cocone mc) + to-is-colimit {Diagram} {coapex} mkcolim = colim where + open make-is-colimit mkcolim + open is-lan + open Functor + + colim : is-colimit Diagram coapex (to-cocone mkcolim) + colim .σ {M = M} α .η _ = + universal (α .η) (λ f → α .is-natural _ _ f ∙ C.eliml (M .F-id)) + colim .σ {M = M} α .is-natural _ _ _ = + C.idr _ ∙ C.introl (M .F-id) + colim .σ-comm {α = α} = Nat-path λ j → + factors (α .η) _ + colim .σ-uniq {α = α} {σ′ = σ′} p = Nat-path λ _ → + sym $ unique (α .η) _ (σ′ .η _) (λ j → sym (p ηₚ j)) +``` + -As per usual, we define a helper lemma charaterizing the path space -of cones: +The concrete interface of `make-is-colimit`{.Agda} is also handy for +_consuming_ specific colimits. To enable this use case, we provide a +function which **un**makes a colimit. ```agda - open Cocone + 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) + unmake-colimit {D} {F} {eta} colim = mc module unmake-colimit where + coapex = Functor.F₀ F tt + 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 + + eps-nt : D => const! x F∘ !F + eps-nt .η = eps + eps-nt .is-natural _ _ f = p f ∙ sym (C.idl _) + + hom : C.Hom coapex x + hom = σ {M = const! x} eps-nt .η tt + + mc : 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 = + sym $ σ-uniq {σ′ = other-nt} (Nat-path λ j → sym (q j)) ηₚ tt + where + other-nt : F => const! x + other-nt .η _ = other + other-nt .is-natural _ _ _ = C.elimr (Functor.F-id F) ∙ sym (C.idl _) +``` - Cocone-path : {x y : Cocone} - → (p : coapex x ≡ coapex y) - → (∀ o → PathP (λ i → C.Hom (F.₀ o) (p i)) (ψ x o) (ψ y o)) - → x ≡ y - Cocone-path p q i .coapex = p i - Cocone-path p q i .ψ o = q o i - Cocone-path {x = x} {y = y} p q i .commutes {x = a} {y = b} f = - is-prop→pathp (λ i → C.Hom-set _ _ (q b i C.∘ F.₁ f) (q a i)) - (x .commutes f) (y .commutes f) i + + + -Now that we've defined cocones, we need a way to figure out how to -express universal properties. Like most things categorical, we begin -by considering a "cocone morphism", which will give us a category -that we can work within. The idea here is that a morphism of cocones -is a morphism in $C$ between the coapicies, such that all of the -injection maps commute. +We also provide a similar interface for the bundled form of colimits. ```agda - record Cocone-hom (x y : Cocone) : Type (o ⊔ ℓ′) where - no-eta-equality - constructor cocone-hom - field - hom : C.Hom (x .coapex) (y .coapex) - commutes : ∀ o → hom C.∘ x .ψ o ≡ y .ψ o +module Colimit + {J : Precategory o₁ h₁} {C : Precategory o₂ h₂} {D : Functor J C} (L : Colimit D) + where ``` -We define yet another helper lemma that describes the path space -of cocone morphisms. +The coapex of the colimit can be obtained by applying the extension functor +to the single object of `⊤Cat`{.Agda}. + +```agda + coapex : C.Ob + coapex = Ext .F₀ tt +``` + +Furthermore, we can show that the apex is the colimit, in the sense of +`is-colimit`{.Agda}, of the diagram. You'd think this is immediate, but +unfortunately, proof assistants: `is-colimit`{.Agda} asks for _the_ +constant functor functor $\{*\} \to \cC$ with value `coapex` to be a Kan +extension, but `Colimit`{.Agda}, being an instance of `Lan`{.Agda}, +packages an _arbitrary_ functor $\{*\} \to \cC$. + +Since Agda does not compare functors for $\eta$-equality, we have to +shuffle our data around manually. Fortunately, this isn't a very long +computation. + +```agda + cocone : D => Const coapex + cocone .η = eta .η + cocone .is-natural x y f = + eta .is-natural x y f ∙ ap (C._∘ _) (Ext .F-id) + + has-colimit : is-colimit D coapex cocone + has-colimit .is-lan.σ α .η = σ α .η + has-colimit .is-lan.σ α .is-natural x y f = + ap (_ C.∘_) (sym (Ext .F-id)) ∙ σ α .is-natural tt tt tt + has-colimit .is-lan.σ-comm = + Nat-path (λ _ → σ-comm ηₚ _) + has-colimit .is-lan.σ-uniq {M = M} {σ′ = σ′} p = + Nat-path (λ _ → σ-uniq {σ′ = nt} (Nat-path (λ j → p ηₚ j)) ηₚ _) + where + nt : Ext => M + nt .η = σ′ .η + nt .is-natural x y f = ap (_ C.∘_) (Ext .F-id) ∙ σ′ .is-natural x y f + + 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 $\top \to \cC$ +correspond with isomorphisms in $\cC$. + +[Much like limits]: Cat.Diagram.Limit.Base.html#uniqueness + + ```agda - open Cocone-hom + 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 + + colimits-unique : x C.≅ y + colimits→invertible : C.is-invertible (Cx.universal Cy.ψ Cy.commutes) + colimits→inverses + : C.Inverses (Cx.universal Cy.ψ Cy.commutes) (Cy.universal Cx.ψ Cx.commutes) +``` - Cocone-hom-path : ∀ {x y} {f g : Cocone-hom x y} → f .hom ≡ g .hom → f ≡ g - Cocone-hom-path p i .hom = p i - Cocone-hom-path {x = x} {y = y} {f = f} {g = g} p i .commutes o j = - is-set→squarep (λ i j → C.Hom-set _ _) - (λ j → p j C.∘ x .ψ o) (f .commutes o) (g .commutes o) refl i j + + +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. -Now, we can define the category of cocones over a given diagram: +[general theorem of Kan extensions]: Cat.Functor.Kan.Unique.html#is-invertible→is-lan + - compose : ∀ {x y z} → Cocone-hom y z → Cocone-hom x y → Cocone-hom x z - compose K L .hom = K .hom C.∘ L .hom - compose {x = x} {y = y} {z = z} K L .commutes o = - (K .hom C.∘ L .hom) C.∘ x .ψ o ≡⟨ C.pullr (L .commutes o) ⟩ - K .hom C.∘ y .ψ o ≡⟨ K .commutes o ⟩ - z .ψ o ∎ +```agda + 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 + is-invertible→is-colimitp {K' = K'} {eta = eta} eps p q invert = + generalize-colimitp + (is-invertible→is-lan Cy $ componentwise-invertible→invertible _ λ _ → invert) + q +``` - cat : Precategory _ _ - cat .Ob = Cocone - cat .Hom = Cocone-hom - cat .id = cocone-hom C.id (λ _ → C.idl _) - cat ._∘_ = compose - cat .idr f = Cocone-hom-path (C.idr (f .hom)) - cat .idl f = Cocone-hom-path (C.idl (f .hom)) - cat .assoc f g h = Cocone-hom-path (C.assoc (f .hom) (g .hom) (h .hom)) +Another useful fact is that if $C$ is a colimit of some diagram $Dia$, +and $Dia$ is naturally isomorphic to some other diagram $Dia'$, then the +coapex of $C$ is also a colimit of $Dia'$. +```agda + natural-iso-diagram→is-colimitp + : ∀ {D′ : Functor J C} {eta : D′ => K F∘ !F} + → (isos : natural-iso D D′) + → (∀ {j} → Cy.ψ j C.∘ natural-iso.from isos .η j ≡ eta .η j) + → is-lan !F D′ K eta + natural-iso-diagram→is-colimitp {D′ = D′} isos q = generalize-colimitp + (natural-iso-of→is-lan Cy isos) + q ``` -## Colimits + - Colimit : Type _ - Colimit = Initial Cocones +Since `is-colimit`{.Agda} is a proposition, and the colimiting cocones +are all unique (“up to isomorphism”), if we're talking about univalent +categories, then `Colimit`{.Agda} _itself_ is a proposition. This is +also an instance of the more general [uniqueness of Kan extensions]. - Colimit-apex : Colimit → C.Ob - Colimit-apex x = coapex (Initial.bot x) +[uniqueness of Kan extensions]: Cat.Functor.Kan.Unique.html + + + +```agda + Colimit-is-prop : is-category C → is-prop (Colimit Diagram) + Colimit-is-prop cat = Lan-is-prop cat ``` +# Preservation of Colimits + +The definitions here are the same idea as [preservation of limits], just +dualized. + +[preservation of limits]: Cat.Diagram.Limit.Base.html#preservation-of-limits + + +```agda + preserves-colimit : Type _ + preserves-colimit = + ∀ {K : Functor ⊤Cat C} {eta : Diagram => K F∘ !F} + → (colim : is-lan !F Diagram K eta) + → preserves-lan F colim + + reflects-colimit : Type _ + 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 +``` + + -# Preservation of Colimits - -Because a cocone is a commutative diagram, any given functor $F : \cC -\to \cD$ takes cocones in $\cC$ to cocones in $\cD$, as -functors preserve commutative diagrams. +## Cocontinuity ```agda - F-map-cocone : Cocone Dia → Cocone (F F∘ Dia) - F-map-cocone x .Cocone.coapex = F.₀ (Cocone.coapex x) - F-map-cocone x .Cocone.ψ f = F.₁ (Cocone.ψ x f) - F-map-cocone x .Cocone.commutes {y = y} f = - F.₁ (Cocone.ψ x y) D.∘ F .F₁ _ ≡⟨ F.collapse (Cocone.commutes x _) ⟩ - F.₁ (Cocone.ψ x _) ∎ +is-cocontinuous + : ∀ (oshape hshape : Level) + {C : Precategory o₁ h₁} + {D : Precategory o₂ h₂} + → Functor C D → Type _ ``` -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. +A cocontinuous functor is one that, for every shape of diagram `J`, +and every diagram `diagram`{.Agda} of shape `J` in `C`, preserves the +colimit for that diagram. ```agda - Preserves-colimit : Cocone Dia → Type _ - Preserves-colimit K = is-colimit Dia K → is-colimit (F F∘ Dia) (F-map-cocone K) +is-cocontinuous oshape hshape {C = C} F = + ∀ {J : Precategory oshape hshape} {Diagram : Functor J C} + → preserves-colimit F Diagram ``` ## Cocompleteness diff --git a/src/Cat/Diagram/Colimit/Cocone.lagda.md b/src/Cat/Diagram/Colimit/Cocone.lagda.md new file mode 100644 index 000000000..6864bfb7c --- /dev/null +++ b/src/Cat/Diagram/Colimit/Cocone.lagda.md @@ -0,0 +1,196 @@ +```agda +open import Cat.Diagram.Colimit.Base +open import Cat.Instances.Functor +open import Cat.Diagram.Initial +open import Cat.Prelude + +import Cat.Functor.Reasoning as Func +import Cat.Reasoning + +module Cat.Diagram.Colimit.Cocone where +``` + + + +# Colimits via Cocones + +As noted in the main page on [colimits], most introductory texts opt +to define colimits via categorical gadgets called **cocones**. + +[colimits]: Cat.Diagram.Colimit.Base.html + +A `Cocone`{.Agda} over $F$ is given by an object (the `coapex`{.agda}) +together with a family of maps `ψ`{.Agda} --- one for each object in the +indexing category `J`{.Agda} --- such that "everything in sight +commutes". + + + +```agda + record Cocone : Type (o ⊔ ℓ ⊔ o′ ⊔ ℓ′) where + no-eta-equality + constructor cocone + field + coapex : C.Ob + ψ : (x : J.Ob) → C.Hom (F.₀ x) coapex + commutes : ∀ {x y} (f : J.Hom x y) → ψ y C.∘ F.₁ f ≡ ψ x +``` + + + +## Cocone Maps + +To express the universal property of a colimit in terms of cocones, we +now have to define the notion of **cocone homomorphism**. We define a +cocone homomorphism to be a map between the coapices which commutes with +the family $\psi$. + +```agda + record Cocone-hom (x y : Cocone) : Type (o ⊔ ℓ′) where + no-eta-equality + constructor cocone-hom + field + hom : C.Hom (x .coapex) (y .coapex) + commutes : ∀ o → hom C.∘ x .ψ o ≡ y .ψ o +``` + + + +Since cocone homomorphisms are closed under composition in the base +category, it's immediate that they form a category. + +```agda + Cocones : Precategory _ _ + Cocones = cat where + open Precategory + + compose : ∀ {x y z} → Cocone-hom y z → Cocone-hom x y → Cocone-hom x z + compose K L .hom = K .hom C.∘ L .hom + compose {x = x} {y = y} {z = z} K L .commutes o = + (K .hom C.∘ L .hom) C.∘ x .ψ o ≡⟨ C.pullr (L .commutes o) ⟩ + K .hom C.∘ y .ψ o ≡⟨ K .commutes o ⟩ + z .ψ o ∎ +``` + + + + +## Initial Cocones as Colimits + +A cocone over some diagram $F$ contains the same data as natural +transformation from $F$ to a constant functor. Since we have defined a +colimit to consist of (a functor equipped with) a natural transformation +into a constant functor, there is an equivalence between the cocones +defined here and those considered in the definition of colimit. + +```agda + Cocone→cocone : (K : Cocone) → F => Const (Cocone.coapex K) + Cocone→cocone K .η = K .Cocone.ψ + Cocone→cocone K .is-natural x y f = K .Cocone.commutes f ∙ sym (C.idl _) +``` + +We can then rephrase the universality from the definition of left Kan +extension by asking that a particular cocone be [initial] in the +category we have just constructed. + +[initial object]: Cat.Diagram.Initial.html + +```agda + is-initial-cocone→is-colimit + : ∀ {K : Cocone} + → is-initial Cocones K + → is-colimit F (Cocone.coapex K) (Cocone→cocone K) + is-initial-cocone→is-colimit {K = K} init = to-is-colimitp colim refl where + open make-is-colimit + open Cocone + open Cocone-hom + + colim : make-is-colimit F (Cocone.coapex K) + colim .ψ = K .ψ + colim .commutes = K .commutes + colim .universal eta p = init (cocone _ eta p) .centre .hom + colim .factors eta p = init (cocone _ eta p) .centre .commutes _ + colim .unique eta p other q = + ap hom (sym (init (cocone _ eta p) .paths (cocone-hom other q))) +``` + +To finish concretising the correspondence, note that this process is +invertible: From a colimit, we can extract an initial cocone. + +```agda + is-colimit→is-initial-cocone + : ∀ {x} {eta : F => Const x} + → (L : is-colimit F x eta) + → is-initial Cocones (cocone x (is-colimit.ψ L) (is-colimit.commutes L)) +``` + +
+The proof consists of more data shuffling, so we omit it. + + +```agda + is-colimit→is-initial-cocone {x = x} L K = init where + module L = is-colimit L + module K = Cocone K + open Cocone-hom + + init : is-contr (Cocone-hom (cocone x L.ψ L.commutes) K) + init .centre .hom = L.universal K.ψ K.commutes + init .centre .commutes _ = L.factors K.ψ K.commutes + init .paths f = + Cocone-hom-path (sym (L.unique K.ψ K.commutes (f .hom) (f .commutes))) +``` +
diff --git a/src/Cat/Diagram/Colimit/Finite.lagda.md b/src/Cat/Diagram/Colimit/Finite.lagda.md index 0cc7f99b1..ea18f5b74 100644 --- a/src/Cat/Diagram/Colimit/Finite.lagda.md +++ b/src/Cat/Diagram/Colimit/Finite.lagda.md @@ -117,17 +117,17 @@ of $in_0f$ and $in_1g$. po : is-pushout C _ _ _ _ po .square = sym (assoc _ _ _) ∙ cq.coequal ∙ assoc _ _ _ - po .colimiting {i₁′ = i₁′} {i₂′} p = - cq.coequalise {e′ = cp.[ i₁′ , i₂′ ]} ( + po .universal {i₁′ = i₁′} {i₂′} p = + cq.universal {e′ = cp.[ i₁′ , i₂′ ]} ( cp.[ i₁′ , i₂′ ] ∘ (in1 ∘ f) ≡⟨ pulll cp.in₀∘factor ⟩ i₁′ ∘ f ≡⟨ p ⟩ i₂′ ∘ g ≡˘⟨ pulll cp.in₁∘factor ⟩ cp.[ i₁′ , i₂′ ] ∘ (in2 ∘ g) ∎ ) - po .i₁∘colimiting = pulll cq.universal ∙ cp.in₀∘factor - po .i₂∘colimiting = pulll cq.universal ∙ cp.in₁∘factor + po .i₁∘universal = pulll cq.factors ∙ cp.in₀∘factor + po .i₂∘universal = pulll cq.factors ∙ cp.in₁∘factor po .unique p q = - cq.unique (sym (cp.unique _ (sym (assoc _ _ _) ∙ p) (sym (assoc _ _ _) ∙ q))) + cq.unique ((cp.unique _ (sym (assoc _ _ _) ∙ p) (sym (assoc _ _ _) ∙ q))) ``` Thus, if a category has an initial object, binary coproducts, and @@ -162,9 +162,9 @@ A coproduct is a pushout under a span whose vertex is the initial object. coprod : is-coproduct C in1 in2 coprod .is-coproduct.[_,_] in1′ in2′ = - Po.colimiting {i₁′ = in1′} {i₂′ = in2′} (is-contr→is-prop (init _) _ _) - coprod .is-coproduct.in₀∘factor = Po.i₁∘colimiting - coprod .is-coproduct.in₁∘factor = Po.i₂∘colimiting + Po.universal {i₁′ = in1′} {i₂′ = in2′} (is-contr→is-prop (init _) _ _) + coprod .is-coproduct.in₀∘factor = Po.i₁∘universal + coprod .is-coproduct.in₁∘factor = Po.i₂∘universal coprod .is-coproduct.unique other p q = Po.unique p q with-pushouts @@ -226,8 +226,8 @@ The construction of coequalisers from pushouts follows its (Po.i₂ ∘ [id,id]) ∘ A+A.in₁ ≡⟨ ap (_∘ A+A.in₁) (sym Po.square) ⟩ (Po.i₁ ∘ [f,g]) ∘ A+A.in₁ ≡⟨ sym (assoc _ _ _) ∙ ap (Po.i₁ ∘_) A+A.in₁∘factor ⟩ Po.i₁ ∘ g ∎ - coequ .has-is-coeq .coequalise {e′ = e′} p = - Po.colimiting (A+A.unique₂ _ refl refl _ (in1) (in2)) + coequ .has-is-coeq .universal {e′ = e′} p = + Po.universal (A+A.unique₂ _ refl refl _ (in1) (in2)) where in1 : ((e′ ∘ f) ∘ [id,id]) ∘ A+A.in₀ ≡ (e′ ∘ [f,g]) ∘ A+A.in₀ in1 = @@ -242,15 +242,15 @@ The construction of coequalisers from pushouts follows its e′ ∘ g ≡˘⟨ pullr A+A.in₁∘factor ⟩ (e′ ∘ [f,g]) ∘ A+A.in₁ ∎ - coequ .has-is-coeq .universal = Po.i₁∘colimiting + coequ .has-is-coeq .factors = Po.i₁∘universal coequ .has-is-coeq .unique {F} {e′ = e′} {colim = colim} e′=col∘i₁ = - Po.unique (sym e′=col∘i₁) path + Po.unique e′=col∘i₁ path where path : colim ∘ Po.i₂ ≡ e′ ∘ f path = colim ∘ Po.i₂ ≡⟨ insertr A+A.in₀∘factor ⟩ ((colim ∘ Po.i₂) ∘ [id,id]) ∘ A+A.in₀ ≡⟨ ap (_∘ A+A.in₀) (extendr (sym Po.square)) ⟩ - ((colim ∘ Po.i₁) ∘ [f,g]) ∘ A+A.in₀ ≡⟨ ap (_∘ A+A.in₀) (ap (_∘ [f,g]) (sym e′=col∘i₁)) ⟩ + ((colim ∘ Po.i₁) ∘ [f,g]) ∘ A+A.in₀ ≡⟨ ap (_∘ A+A.in₀) (ap (_∘ [f,g]) e′=col∘i₁) ⟩ (e′ ∘ [f,g]) ∘ A+A.in₀ ≡⟨ pullr A+A.in₀∘factor ⟩ e′ ∘ f ∎ @@ -270,9 +270,9 @@ The construction of coequalisers from pushouts follows its open is-pushout po : is-pushout C _ _ _ _ po .square = is-contr→is-prop (i _) _ _ - po .colimiting _ = r .is-coproduct.[_,_] _ _ - po .i₁∘colimiting = r .is-coproduct.in₀∘factor - po .i₂∘colimiting = r .is-coproduct.in₁∘factor + po .universal _ = r .is-coproduct.[_,_] _ _ + po .i₁∘universal = r .is-coproduct.in₀∘factor + po .i₂∘universal = r .is-coproduct.in₁∘factor po .unique p q = r .is-coproduct.unique _ p q ``` --> diff --git a/src/Cat/Diagram/Colimit/Representable.lagda.md b/src/Cat/Diagram/Colimit/Representable.lagda.md new file mode 100644 index 000000000..be61917e6 --- /dev/null +++ b/src/Cat/Diagram/Colimit/Representable.lagda.md @@ -0,0 +1,77 @@ +```agda +open import Cat.Functor.Hom.Representable +open import Cat.Instances.Functor.Compose +open import Cat.Instances.Shape.Terminal +open import Cat.Instances.Sets.Complete +open import Cat.Diagram.Colimit.Base +open import Cat.Diagram.Limit.Base +open import Cat.Instances.Functor +open import Cat.Functor.Kan.Base +open import Cat.Functor.Hom +open import Cat.Prelude + +import Cat.Reasoning + +module Cat.Diagram.Colimit.Representable where +``` + +## Representability of Colimits + +Since [colimits] are defined by universal property, we can also phrase +the definition in terms of an equivalence between $\hom$-functors. + +[colimits]: Cat.Diagram.Colimit.Base.html + + + +Let $\mathrm{Dia} : \cJ \to \cC$ be some diagram in $\cC$. If +$\mathrm{Dia}$ has a colimit $c$, then that means that maps **out** of +$c$ are in bijection with a product of maps $\pi_i$, subject to some +conditions. + +```agda + Lim[C[F-,=]] : Functor C (Sets ℓ) + Lim[C[F-,=]] .F₀ c = el (Dia => Const c) Nat-is-set + Lim[C[F-,=]] .F₁ f α = const-nt f ∘nt α + Lim[C[F-,=]] .F-id = funext λ _ → Nat-path λ _ → C.idl _ + Lim[C[F-,=]] .F-∘ _ _ = funext λ _ → Nat-path λ _ → sym $ C.assoc _ _ _ + + Hom-into-inj + : ∀ {c : C.Ob} (eta : Dia => Const c) + → Hom-from C c => Lim[C[F-,=]] + Hom-into-inj eta .η x f = const-nt f ∘nt eta + Hom-into-inj eta .is-natural x y f = funext λ g → Nat-path λ _ → + sym $ C.assoc _ _ _ + + represents→is-colimit + : ∀ {c : C.Ob} {eta : Dia => Const c} + → is-natural-invertible (Hom-into-inj eta) + → is-colimit Dia c eta + represents→is-colimit {c} {eta} nat-inv = colim where + module nat-inv = is-natural-invertible nat-inv + + colim : is-colimit Dia c eta + colim .σ {M} α = + hom→⊤-natural-trans $ nat-inv.inv .η _ (idnat-constr ∘nt α) + colim .σ-comm {M} {α} = Nat-path λ j → + nat-inv.invl ηₚ _ $ₚ _ ηₚ j ∙ C.idl _ + colim .σ-uniq {M} {α} {σ′} q = Nat-path λ j → + nat-inv.inv .η _ (idnat-constr ∘nt ⌜ α ⌝) ≡⟨ ap! q ⟩ + nat-inv.inv .η _ ⌜ idnat-constr ∘nt (σ′ ◂ !F) ∘nt cocone→unit Dia eta ⌝ ≡⟨ ap! (Nat-path λ _ → C.idl _) ⟩ + nat-inv.inv .η (M .F₀ tt) (const-nt (σ′ .η j) ∘nt eta) ≡⟨ nat-inv.invr ηₚ _ $ₚ _ ⟩ + σ′ .η tt ∎ +``` diff --git a/src/Cat/Diagram/Congruence.lagda.md b/src/Cat/Diagram/Congruence.lagda.md index 94a9bf732..04347e66d 100644 --- a/src/Cat/Diagram/Congruence.lagda.md +++ b/src/Cat/Diagram/Congruence.lagda.md @@ -298,20 +298,20 @@ which is characterised by $p_1 f = p_2 f = \mathrm{id}$; This expresses, diagramatically, that $f(x) = f(x)$. ```agda - cg .has-refl = Kp.limiting {p₁' = id} {p₂' = id} refl - cg .refl-p₁ = ap (_∘ Kp.limiting refl) a×a.π₁∘factor ∙ Kp.p₁∘limiting - cg .refl-p₂ = ap (_∘ Kp.limiting refl) a×a.π₂∘factor ∙ Kp.p₂∘limiting + cg .has-refl = Kp.universal {p₁' = id} {p₂' = id} refl + cg .refl-p₁ = ap (_∘ Kp.universal refl) a×a.π₁∘factor ∙ Kp.p₁∘universal + cg .refl-p₂ = ap (_∘ Kp.universal refl) a×a.π₂∘factor ∙ Kp.p₂∘universal ``` Symmetry is witnessed by the map $(A \times_B A) \to (A \times_B A)$ which swaps the components. This one's pretty simple. ```agda - cg .has-sym = Kp.limiting {p₁' = Kp.p₂} {p₂' = Kp.p₁} (sym Kp.square) - cg .sym-p₁ = ap (_∘ Kp.limiting (sym Kp.square)) a×a.π₁∘factor - ∙ sym (a×a.π₂∘factor ∙ sym Kp.p₁∘limiting) - cg .sym-p₂ = ap (_∘ Kp.limiting (sym Kp.square)) a×a.π₂∘factor - ∙ sym (a×a.π₁∘factor ∙ sym Kp.p₂∘limiting) + cg .has-sym = Kp.universal {p₁' = Kp.p₂} {p₂' = Kp.p₁} (sym Kp.square) + cg .sym-p₁ = ap (_∘ Kp.universal (sym Kp.square)) a×a.π₁∘factor + ∙ sym (a×a.π₂∘factor ∙ sym Kp.p₁∘universal) + cg .sym-p₂ = ap (_∘ Kp.universal (sym Kp.square)) a×a.π₂∘factor + ∙ sym (a×a.π₁∘factor ∙ sym Kp.p₂∘universal) ```
@@ -321,7 +321,7 @@ Understanding the transitivity map is left as an exercise to the reader. ```agda cg .has-trans = - Kp.limiting {p₁' = Kp.p₁ ∘ rel.p₂} {p₂' = Kp.p₂ ∘ rel.p₁} path + Kp.universal {p₁' = Kp.p₁ ∘ rel.p₂} {p₂' = Kp.p₂ ∘ rel.p₁} path where abstract path : f ∘ Kp.p₁ ∘ rel.p₂ ≡ f ∘ Kp.p₂ ∘ rel.p₁ path = @@ -332,11 +332,11 @@ Understanding the transitivity map is left as an exercise to the reader. cg .trans-factors = sym ( - kernel-pair ∘ Kp.limiting _ + kernel-pair ∘ Kp.universal _ ≡⟨ a×a.⟨⟩∘ _ ⟩ - a×a.⟨ Kp.p₁ ∘ Kp.limiting _ , Kp.p₂ ∘ Kp.limiting _ ⟩ - ≡⟨ ap₂ a×a.⟨_,_⟩ (Kp.p₁∘limiting ∙ ap₂ _∘_ (sym a×a.π₁∘factor) refl) - (Kp.p₂∘limiting ∙ ap₂ _∘_ (sym a×a.π₂∘factor) refl) ⟩ + a×a.⟨ Kp.p₁ ∘ Kp.universal _ , Kp.p₂ ∘ Kp.universal _ ⟩ + ≡⟨ ap₂ a×a.⟨_,_⟩ (Kp.p₁∘universal ∙ ap₂ _∘_ (sym a×a.π₁∘factor) refl) + (Kp.p₂∘universal ∙ ap₂ _∘_ (sym a×a.π₂∘factor) refl) ⟩ a×a.⟨ (a×a.π₁ ∘ kernel-pair) ∘ rel.p₂ , (a×a.π₂ ∘ kernel-pair) ∘ rel.p₁ ⟩ ∎) diff --git a/src/Cat/Diagram/Coproduct/Copower.lagda.md b/src/Cat/Diagram/Coproduct/Copower.lagda.md index c074ecab5..53217e2f6 100644 --- a/src/Cat/Diagram/Coproduct/Copower.lagda.md +++ b/src/Cat/Diagram/Coproduct/Copower.lagda.md @@ -78,5 +78,5 @@ cocomplete→copowering : ∀ {o ℓ} {C : Precategory o ℓ} → is-cocomplete ℓ ℓ C → Functor (Sets ℓ ×ᶜ C) C cocomplete→copowering colim = Copowering λ S F → - Colimit→IC _ (is-hlevel-suc 2 (S .is-tr)) (Disc-adjunct F) (colim _) + Colimit→IC _ (is-hlevel-suc 2 (S .is-tr)) F (colim _) ``` diff --git a/src/Cat/Diagram/Coproduct/Indexed.lagda.md b/src/Cat/Diagram/Coproduct/Indexed.lagda.md index b8707c982..ff2e8ac0e 100644 --- a/src/Cat/Diagram/Coproduct/Indexed.lagda.md +++ b/src/Cat/Diagram/Coproduct/Indexed.lagda.md @@ -1,7 +1,9 @@ ```agda +open import Cat.Instances.Shape.Terminal open import Cat.Diagram.Colimit.Base open import Cat.Instances.Discrete open import Cat.Diagram.Pullback +open import Cat.Functor.Kan.Base open import Cat.Diagram.Initial open import Cat.Prelude @@ -67,59 +69,72 @@ has-indexed-coproducts ℓ = ∀ {I : Type ℓ} (F : I → C.Ob) → Indexed-cop ## As colimits Similarly to the product case, when $I$ is a groupoid, indexed -coproducts correspond to discrete diagrams of shape $I$. +coproducts correspond to colimits of discrete diagrams of shape $I$. ```agda module _ {I : Type ℓ'} (i-is-grpd : is-groupoid I) (F : I → C.Ob) where - open Cocone-hom - open Initial - open Cocone - - IC→Colimit : Indexed-coproduct F → Colimit {C = C} (Disc-adjunct {iss = i-is-grpd} F) - IC→Colimit IC = colim where - module IC = Indexed-coproduct IC - - thecolim : Cocone _ - thecolim .coapex = IC.ΣF - thecolim .ψ = IC.ι - thecolim .commutes {x} = - J (λ y p → IC.ι y C.∘ subst (C.Hom (F x) ⊙ F) p C.id ≡ IC.ι x) - (C.elimr (transport-refl _)) - - colim : Colimit _ - colim .bot = thecolim - colim .has⊥ x .centre .hom = IC.match (x .ψ) - colim .has⊥ x .centre .commutes o = IC.commute - colim .has⊥ x .paths h = Cocone-hom-path _ (sym (IC.unique _ λ i → h .commutes _)) - -module _ {I : Type ℓ'} (isg : is-groupoid I) (F : Functor (Disc I isg) C) where - private module F = Functor F - open is-indexed-coproduct - open Indexed-coproduct - open Cocone-hom - open Initial - open Cocone - - Inj→Cocone : ∀ {Y} → (∀ i → C.Hom (F.₀ i) Y) → Cocone F - Inj→Cocone f .coapex = _ - Inj→Cocone f .ψ = f - Inj→Cocone f .commutes {x} = J (λ y p → f y C.∘ F.₁ p ≡ f x) (C.elimr F.F-id) - - Colimit→IC : Colimit {C = C} F → Indexed-coproduct F.₀ - Colimit→IC colim = the-ic where - module colim = Cocone (colim .bot) - - the-ic : Indexed-coproduct _ - the-ic .ΣF = colim.coapex - the-ic .ι = colim.ψ - the-ic .has-is-ic .match f = colim .has⊥ (Inj→Cocone f) .centre .hom - the-ic .has-is-ic .commute = colim .has⊥ _ .centre .commutes _ - the-ic .has-is-ic .unique {h = h} f p i = - colim .has⊥ (Inj→Cocone f) .paths h′ (~ i) .hom - where - h′ : Cocone-hom _ _ _ - h′ .hom = h - h′ .commutes o = p _ + open _=>_ + + Inj→Cocone : ∀ {x} → (∀ i → C.Hom (F i) x) + → Disc-adjunct {C = C} {iss = i-is-grpd} F => Const x + Inj→Cocone inj .η i = inj i + Inj→Cocone inj .is-natural i j p = + J (λ j p → inj j C.∘ subst (C.Hom (F i) ⊙ F) p C.id ≡ C.id C.∘ inj i) + (C.elimr (transport-refl C.id) ∙ sym (C.idl _)) p + + is-indexed-coproduct→is-colimit + : ∀ {x} {inj : ∀ i → C.Hom (F i) x} + → is-indexed-coproduct F inj + → is-colimit (Disc-adjunct F) x (Inj→Cocone inj) + is-indexed-coproduct→is-colimit {x = x} {inj} ic = + to-is-colimitp mc refl + where + module ic = is-indexed-coproduct ic + open make-is-colimit + + mc : make-is-colimit (Disc-adjunct F) x + mc .ψ i = inj i + mc .commutes {i} {j} p = + J (λ j p → inj j C.∘ subst (C.Hom (F i) ⊙ F) p C.id ≡ inj i) + (C.elimr (transport-refl C.id)) + p + mc .universal eta p = ic.match eta + mc .factors eta p = ic.commute + mc .unique eta p other q = ic.unique eta q + + is-colimit→is-indexed-coprduct + : ∀ {K : Functor ⊤Cat C} + → {eps : Disc-adjunct {iss = i-is-grpd} F => K F∘ !F} + → is-lan !F (Disc-adjunct F) K eps + → is-indexed-coproduct F (eps .η) + is-colimit→is-indexed-coprduct {K = K} {eps} colim = ic where + module colim = is-colimit colim + open is-indexed-coproduct + + ic : is-indexed-coproduct F (eps .η) + ic .match k = + colim.universal k + (J (λ j p → k j C.∘ subst (C.Hom (F _) ⊙ F) p C.id ≡ k _) + (C.elimr (transport-refl _))) + ic .commute = + colim.factors _ _ + ic .unique k comm = + colim.unique _ _ _ comm + + IC→Colimit + : Indexed-coproduct F + → Colimit {C = C} (Disc-adjunct {iss = i-is-grpd} F) + IC→Colimit ic = + to-colimit (is-indexed-coproduct→is-colimit has-is-ic) + where open Indexed-coproduct ic + + Colimit→IC + : Colimit {C = C} (Disc-adjunct {iss = i-is-grpd} F) + → Indexed-coproduct F + Colimit→IC colim .Indexed-coproduct.ΣF = _ + Colimit→IC colim .Indexed-coproduct.ι = _ + Colimit→IC colim .Indexed-coproduct.has-is-ic = + is-colimit→is-indexed-coprduct (Colimit.has-colimit colim) ``` # Disjoint coproducts diff --git a/src/Cat/Diagram/Duals.lagda.md b/src/Cat/Diagram/Duals.lagda.md index 462860f7c..9654964e0 100644 --- a/src/Cat/Diagram/Duals.lagda.md +++ b/src/Cat/Diagram/Duals.lagda.md @@ -73,8 +73,8 @@ is-co-equaliser→is-coequaliser is-co-equaliser→is-coequaliser eq = record { coequal = eq.equal - ; coequalise = eq.limiting ; universal = eq.universal + ; factors = eq.factors ; unique = eq.unique } where module eq = Co-equ.is-equaliser eq @@ -85,8 +85,8 @@ is-coequaliser→is-co-equaliser is-coequaliser→is-co-equaliser coeq = record { equal = coeq.coequal - ; limiting = coeq.coequalise ; universal = coeq.universal + ; factors = coeq.factors ; unique = coeq.unique } where module coeq = Coequ.is-coequaliser coeq @@ -119,9 +119,9 @@ is-co-pullback→is-pushout is-co-pullback→is-pushout pb = record { square = pb.square - ; colimiting = pb.limiting - ; i₁∘colimiting = pb.p₁∘limiting - ; i₂∘colimiting = pb.p₂∘limiting + ; universal = pb.universal + ; i₁∘universal = pb.p₁∘universal + ; i₂∘universal = pb.p₂∘universal ; unique = pb.unique } where module pb = Co-pull.is-pullback pb @@ -132,9 +132,9 @@ is-pushout→is-co-pullback is-pushout→is-co-pullback po = record { square = po.square - ; limiting = po.colimiting - ; p₁∘limiting = po.i₁∘colimiting - ; p₂∘limiting = po.i₂∘colimiting + ; universal = po.universal + ; p₁∘universal = po.i₁∘universal + ; p₂∘universal = po.i₂∘universal ; unique = po.unique } where module po = Push.is-pushout po @@ -144,7 +144,9 @@ is-pushout→is-co-pullback po = ```agda open import Cat.Diagram.Colimit.Base +open import Cat.Diagram.Colimit.Cocone open import Cat.Diagram.Limit.Base +open import Cat.Diagram.Limit.Cone open import Cat.Diagram.Terminal open import Cat.Diagram.Initial @@ -195,55 +197,39 @@ module _ {o ℓ} {J : Precategory o ℓ} {F : Functor J C} where ## Co/limits +Limits and colimits are defined via [Kan extensions], so it's reasonable +to expect that [duality of Kan extensions] would apply to (co)limits. +Unfortunately, proof assistants: (co)limits are extensions of +`!F`{.Agda}, but duality of Kan extensions inserts an extra `Functor.op`. +We could work around this, but it's easier to just do the proofs by hand. + +[Kan extensions]: Cat.Functor.Kan.Base.html +[duality of Kan extensions]: Cat.Functor.Kan.Duality.html + ```agda Colimit→Co-limit : Colimit F → Limit F^op - Colimit→Co-limit colim = lim where - lim : Limit F^op - lim .top = Cocone→Co-cone (colim .bot) - lim .has⊤ co-cone = - retract→is-contr f g fg - (colim .has⊥ (Co-cone→Cocone co-cone)) - where - f : _ → _ - f x = subst (λ e → Cone-hom F^op e _) - (Co-cone→Cocone→Co-cone _) - (Cocone-hom→Co-cone-hom x) - - g : _ → _ - g x = subst (λ e → Cocone-hom F e _) - (Cocone→Co-cone→Cocone _) - (Co-cone-hom→Cocone-hom x) - - fg : is-left-inverse f g - fg x = Cone-hom-path _ (transport-refl _ ∙ transport-refl _) + Colimit→Co-limit colim = to-limit (to-is-limit ml) where + module colim = Colimit colim + open make-is-limit + + ml : make-is-limit F^op colim.coapex + ml .ψ = colim.ψ + ml .commutes = colim.commutes + ml .universal eta p = colim.universal eta p + ml .factors eta p = colim.factors eta p + ml .unique eta p other q = colim.unique eta p other q Co-limit→Colimit : Limit F^op → Colimit F - Co-limit→Colimit lim = colim where - colim : Colimit F - colim .bot = Co-cone→Cocone (lim .top) - colim .has⊥ cocon = - retract→is-contr f g fg - (lim .has⊤ (Cocone→Co-cone cocon)) - where - f : _ → _ - f x = subst (λ e → Cocone-hom F _ e) - (Cocone→Co-cone→Cocone _) - (Co-cone-hom→Cocone-hom x) - - g : _ → _ - g x = subst (λ e → Cone-hom F^op _ e) - (Co-cone→Cocone→Co-cone _) - (Cocone-hom→Co-cone-hom x) - - fg : is-left-inverse f g - fg x = Cocone-hom-path _ (transport-refl _ ∙ transport-refl _) - -module _ {o ℓ} {J : Precategory o ℓ} {F F′ : Functor J C} where - private module JC = Cat.Reasoning Cat[ J , C ] - - Colimit-ap-iso : F JC.≅ F′ → Colimit F → Colimit F′ - Colimit-ap-iso f cl = - Co-limit→Colimit (Limit-ap-iso (op-natural-iso f) (Colimit→Co-limit cl)) + Co-limit→Colimit lim = to-colimit (to-is-colimit mc) where + module lim = Limit lim + open make-is-colimit + + mc : make-is-colimit F lim.apex + mc .ψ = lim.ψ + mc .commutes = lim.commutes + mc .universal eta p = lim.universal eta p + mc .factors eta p = lim.factors eta p + mc .unique eta p other q = lim.unique eta p other q ``` diff --git a/src/Cat/Diagram/Equaliser.lagda.md b/src/Cat/Diagram/Equaliser.lagda.md index 187a67478..6aed323ff 100644 --- a/src/Cat/Diagram/Equaliser.lagda.md +++ b/src/Cat/Diagram/Equaliser.lagda.md @@ -26,12 +26,12 @@ right-hand-sides agree. record is-equaliser {E} (f g : Hom A B) (equ : Hom E A) : Type (ℓ ⊔ ℓ′) where field equal : f ∘ equ ≡ g ∘ equ - limiting : ∀ {F} {e′ : Hom F A} (p : f ∘ e′ ≡ g ∘ e′) → Hom F E - universal : ∀ {F} {e′ : Hom F A} {p : f ∘ e′ ≡ g ∘ e′} → equ ∘ limiting p ≡ e′ + universal : ∀ {F} {e′ : Hom F A} (p : f ∘ e′ ≡ g ∘ e′) → Hom F E + factors : ∀ {F} {e′ : Hom F A} {p : f ∘ e′ ≡ g ∘ e′} → equ ∘ universal p ≡ e′ unique - : ∀ {F} {e′ : Hom F A} {p : f ∘ e′ ≡ g ∘ e′} {lim' : Hom F E} - → e′ ≡ equ ∘ lim' - → lim' ≡ limiting p + : ∀ {F} {e′ : Hom F A} {p : f ∘ e′ ≡ g ∘ e′} {other : Hom F E} + → equ ∘ other ≡ e′ + → other ≡ universal p equal-∘ : f ∘ equ ∘ h ≡ g ∘ equ ∘ h equal-∘ {h = h} = @@ -39,11 +39,12 @@ record is-equaliser {E} (f g : Hom A B) (equ : Hom E A) : Type (ℓ ⊔ ℓ′) g ∘ equ ∘ h ∎ unique₂ - : ∀ {F} {e′ : Hom F A} {p : f ∘ e′ ≡ g ∘ e′} {lim' lim'' : Hom F E} - → e′ ≡ equ ∘ lim' - → e′ ≡ equ ∘ lim'' - → lim' ≡ lim'' - unique₂ {p = p} q r = unique {p = p} q ∙ sym (unique r) + : ∀ {F} {e′ : Hom F A} {o1 o2 : Hom F E} + → f ∘ e′ ≡ g ∘ e′ + → equ ∘ o1 ≡ e′ + → equ ∘ o2 ≡ e′ + → o1 ≡ o2 + unique₂ p q r = unique {p = p} q ∙ sym (unique r) ``` We can visualise the situation using the commutative diagram below: @@ -86,8 +87,22 @@ is-equaliser→is-monic → is-equaliser f g equ → is-monic equ is-equaliser→is-monic equ equalises g h p = - g ≡⟨ unique (sym p) ⟩ - limiting equal-∘ ≡˘⟨ unique refl ⟩ - h ∎ + unique₂ (extendl equal) p refl where open is-equaliser equalises ``` + +# Categories with all equalisers + +We also define a helper module for working with categories that have +equalisers of all parallel pairs of morphisms. + +```agda +has-equalisers : Type _ +has-equalisers = ∀ {a b} (f g : Hom a b) → Equaliser f g + +module Equalisers (all-equalisers : has-equalisers) where + module equaliser {a b} (f g : Hom a b) = Equaliser (all-equalisers f g) + + Equ : ∀ {a b} (f g : Hom a b) → Ob + Equ = equaliser.apex +``` diff --git a/src/Cat/Diagram/Equaliser/Kernel.lagda.md b/src/Cat/Diagram/Equaliser/Kernel.lagda.md index cbbb561c3..32a016791 100644 --- a/src/Cat/Diagram/Equaliser/Kernel.lagda.md +++ b/src/Cat/Diagram/Equaliser/Kernel.lagda.md @@ -95,7 +95,8 @@ $$ ```agda p : ¡ ∘ ! ≡ id - p = KKf.unique₂ {p = zero-∘l _ ∙ sym (zero-∘r _)} (sym (zero-∘l _)) - ( Kf.unique₂ {p = zero-∘l _ ∙ sym (zero-∘r _)} (sym (zero-∘l _)) - (sym (KKf.equal ∙ zero-∘r _)) ∙ intror refl) + p = KKf.unique₂ (zero-comm _ _) (zero-∘l _) + (Kf.unique₂ (extendl (zero-comm _ _)) + (pulll KKf.equal ∙ idr _) + (zero-comm _ _)) ``` diff --git a/src/Cat/Diagram/Equaliser/RegularMono.lagda.md b/src/Cat/Diagram/Equaliser/RegularMono.lagda.md index 4ff40c66f..90f5e0e74 100644 --- a/src/Cat/Diagram/Equaliser/RegularMono.lagda.md +++ b/src/Cat/Diagram/Equaliser/RegularMono.lagda.md @@ -117,7 +117,7 @@ $\phi \circ i_2 = \rm{arr_2}$. ```agda phi : Hom f⊔f.coapex reg.c - phi = f⊔f.colimiting reg.equal + phi = f⊔f.universal reg.equal open is-effective-mono mon : is-effective-mono f @@ -143,14 +143,14 @@ universal map $E \to a$ which commutes with "everything in sight": open is-equaliser eq : is-equaliser _ _ _ _ eq .equal = f⊔f.square - eq .limiting {F = F} {e′ = e′} p = reg.limiting p′ where + eq .universal {F = F} {e′ = e′} p = reg.universal p′ where p′ : reg.arr₁ ∘ e′ ≡ reg.arr₂ ∘ e′ p′ = - reg.arr₁ ∘ e′ ≡˘⟨ ap (_∘ e′) f⊔f.i₁∘colimiting ⟩ + reg.arr₁ ∘ e′ ≡˘⟨ ap (_∘ e′) f⊔f.i₁∘universal ⟩ (phi ∘ f⊔f.i₁) ∘ e′ ≡⟨ extendr p ⟩ - (phi ∘ f⊔f.i₂) ∘ e′ ≡⟨ ap (_∘ e′) f⊔f.i₂∘colimiting ⟩ + (phi ∘ f⊔f.i₂) ∘ e′ ≡⟨ ap (_∘ e′) f⊔f.i₂∘universal ⟩ reg.arr₂ ∘ e′ ∎ - eq .universal = reg.universal + eq .factors = reg.factors eq .unique = reg.unique ``` diff --git a/src/Cat/Diagram/Limit/Base.lagda.md b/src/Cat/Diagram/Limit/Base.lagda.md index a2b64e6a1..d4cfbb003 100644 --- a/src/Cat/Diagram/Limit/Base.lagda.md +++ b/src/Cat/Diagram/Limit/Base.lagda.md @@ -1,6 +1,9 @@ ```agda +open import Cat.Instances.Shape.Terminal +open import Cat.Functor.Kan.Unique +open import Cat.Functor.Coherence open import Cat.Instances.Functor -open import Cat.Diagram.Terminal +open import Cat.Functor.Kan.Base open import Cat.Prelude import Cat.Functor.Reasoning as Func @@ -11,555 +14,788 @@ module Cat.Diagram.Limit.Base where # Idea -The idea of limits generalises many concrete constructions in -mathematics - from several settings, such as set theory, topology and -algebra - to arbitrary categories. A limit, _if it exists_, is "the best -solution" to an "equational problem". For a first intuition, we can -build them graphically, working directly "on top" of a diagram. - -## Products - -**Note**: Products are described explicitly in -[`Cat.Diagram.Product`]. The description there might be easier to parse. - -[`Cat.Diagram.Product`]: Cat.Diagram.Product.html - -Consider a _discrete_ diagram - one that has no interesting morphisms, -only identities, such as the collection of objects below. - -~~~{.quiver .short-2} -\[\begin{tikzcd} -A & B -\end{tikzcd}\] -~~~ - -To consider a limit for this diagram, we first have to go through an -intermediate step - a `Cone`{.Agda}. Draw a new object, the -`apex`{.Agda} - let's call it $P$ - and a family of maps from $P$ "down -onto" all of the objects of the diagram. The new maps have to make -everything in sight commute, but in the case of a discrete diagram, we -can pick any maps. There are no obligations. - -~~~{.quiver .short-1} -\[\begin{tikzcd} - & P \\ - A && B - \arrow["\pi_1"', from=1-2, to=2-1] - \arrow["\pi_2", from=1-2, to=2-3] -\end{tikzcd}\] -~~~ - -It'll be helpful to think of the maps as projections - which is why -they're labelled with the greek letter $\pi$, for **p**rojection. -However, for an arbitrary cone, the maps are.. well, arbitrary. To -consider a concrete example, we can pretend our diagram was in -$\rm{Set}$ all along, and that $A$ was the set $\bb{Q}$ and $B$ -was the set $\bb{R}$. Then the following is a cone over it: - -~~~{.quiver .short-1} -\[\begin{tikzcd} - & {\mathbb{N}} \\ - {\mathbb{Q}} && {\mathbb{R}} - \arrow["{x \mapsto x}"', hook, from=1-2, to=2-1] - \arrow["{x \mapsto \sqrt{x}}", from=1-2, to=2-3] -\end{tikzcd}\] +**Note**: This page describes the general definition of limits, and +assumes some familiarity with some concrete examples, in particular +[terminal objects], [products], [equalisers], and [pullbacks]. It might +be a good idea to check out those pages before continuing! + +[terminal objects]: Cat.Diagram.Terminal.html +[products]: Cat.Diagram.Product.html +[equalisers]: Cat.Diagram.Equaliser.html +[pullbacks]: Cat.Diagram.Pullback.html + +To motivate limits, note how all the above examples have roughly the +same structure. They all consist of some object, a bunch of maps out +of said object, some commutativity conditions, and a universal property +that states that we can construct unique maps into the object under +certain conditions. + +Our first step towards making this vague intuition precise is to construct +a mathematical widget that picks out a collection of objects, arrows, +and commutativity conditions in a category. This is required to describe +the collection of maps out of our special objects, and the equations +they satisfy. Luckily, we already have such a widget: functors! + +To see how this works, let's take a very simple example: a functor out +of the [single object category with one morphism] into some category +$\cC$. If we look at the image of such a functor, we can see that it +picks out a single object; the single morphism must be taken to the +identity morphism due to functoriality. We can scale this example up +to functors from [discrete category] with $n$ elements; if we look at +the image of such a functor, we shall see that it selects out $n$ objects +of $\cC$. + +[single object category with one morphism]: Cat.Instances.Shape.Terminal.html +[discrete category]: Cat.Instances.Discrete.html + +We can perform the same trick with non-discrete categories; for instance, +a functor out of the [parallel arrows category] selects a pair of +parallel arrows in $\cC$, a functor out of the [isomorphism category] +selects an isomorphism in $\cC$, and so on. We call such functors +**diagrams** to suggest that we should think of them as picking out some +shape in $\cC$. + +[parallel arrows category]: Cat.Instances.Shape.Parallel.html +[isomorphism category]: Cat.Instances.Shape.Isomorphism.html + +We can use this newfound insight to describe the shape that each of our +examples is defined over. Products are defined over a diagram that +consists of a pair of objects; these diagrams correspond to functors +out of the category with 2 objects and only identity morphisms. +Equalisers are defined over a diagram that consists of a pair of parallel +morphisms; such diagrams are given by functors out of the aforementioned +parallel arrows category. Pullbacks defined over a diagram of the shape, +$\bullet \to \bullet \leftarrow \bullet$; again, these diagrams are +given by functors out of the category with that exact shape. Terminal +objects may seem to break this trend, but we can think of them as +being defined over the empty diagram, the unique functor from +the category with no objects. + +We now move our attention to the maps out of our special object into +the objects of the diagram. Note that these maps need to commute with +any morphisms in the diagram, as is the case with pullbacks and +equalisers. This is where our definition diverges from many other +introductory sources. The typical approach is to define a bespoke +widget called a [cone] that encodes the required morphisms and commuting +conditions, and then proceeding from there. + +[cone]: Cat.Diagram.Limit.Cone.html + +However, this approach is somewhat unsatisfying. Why did we have to +invent a new object just to bundle up the data we needed? Furthermore, +cones are somewhat disconnected from the rest of category theory, which +makes it more difficult to integrate results about limits into the larger +body of work. It would be great if we could encode the data we needed +using existing objects! + +Luckily, we can! If we take a step back, we can notice that we are trying +to construct a map into a functor. What are maps into functors? Natural +transformations! Concretely, let $D : \cJ \to cC$ be some diagram. +We can encode the same data as a cone in a natural transformation +$\eta : {!x} \circ \mathord{!} \to D$, where $!x : \top \to \cC$ denotes +the constant functor that maps object to $x$ and every morphism to $id$, +and $! : \cJ \to \top$ denotes the unique functor into the terminal +category. The components of such a natural transformation yield maps from +$x \to D(j)$ for every $j : \cJ$, and naturality ensures that these +maps must commute with the rest of the diagram. We can describe this +situation diagrammatically like so: + +~~~{.quiver} +\begin{tikzcd} + && \{*\} \\ + \\ + {\cJ} &&& {} & {\cC} + \arrow[from=3-1, to=1-3] + \arrow["{!X}"', from=1-3, to=3-5] + \arrow[""{name=0, anchor=center, inner sep=0}, from=3-1, to=3-5] + \arrow["\eta"{description}, shorten <=4pt, shorten >=4pt, Rightarrow, from=1-3, to=0] +\end{tikzcd} ~~~ -Abstracting again, there is a canonical definition of _cone homomorphism_ - -A map between the apices that makes everything in sight commute. If $P'$ -and $P$ were both apices for our original discrete diagram, we would -draw a cone homomorphism $f : P' \to P$ as the dashed arrow in following -commutative Starfleet comms badge. - -~~~{.quiver .tall-2} -\[\begin{tikzcd} - & {P'} \\ +All that remains is the universal property. If we translate this into +our existing machinery, that means that $!x$ must be the universal +functor equipped with a natural transformation $\eta$; that is, for any +other $K : \{*\} \to \cC$ equipped with $\tau : K \circ \mathord{!} \to +D$, we have a unique natural transformation $\sigma : K \to {!x}$ that +factors $\tau$. This is a bit of a mouthful, so let's look at a diagram +instead. + +~~~{.quiver} +\begin{tikzcd} + && {\{*\}} \\ \\ - & P \\ - A && B - \arrow["{f}"{description}, dashed, from=1-2, to=3-2] - \arrow["{\pi_1}", from=3-2, to=4-1] - \arrow["{\pi_2}"', from=3-2, to=4-3] - \arrow[curve={height=6pt}, from=1-2, to=4-1] - \arrow[curve={height=-6pt}, from=1-2, to=4-3] -\end{tikzcd}\] + {\cJ} &&& {} & {\cC} + \arrow[from=3-1, to=1-3] + \arrow[""{name=0, anchor=center, inner sep=0}, "{!x}"', from=1-3, to=3-5] + \arrow[""{name=1, anchor=center, inner sep=0}, from=3-1, to=3-5] + \arrow[""{name=2, anchor=center, inner sep=0}, "K", curve={height=-18pt}, from=1-3, to=3-5] + \arrow["\eta"{description}, shorten <=4pt, shorten >=4pt, Rightarrow, from=1-3, to=1] + \arrow["\sigma", shorten <=3pt, shorten >=3pt, Rightarrow, from=2, to=0] +\end{tikzcd} ~~~ -These assemble into a category, `Cones`{.agda}, with composition and -identity given by the composition and identity of ambient category. A -`Limit`{.agda} is, then, a [terminal object] in this category. For $P$ -to be a limit in our diagram above, we require that, for any other cone -$P'$ there exists a _unique_ arrow $P' \to P$ that makes everything -commute. - -[terminal object]: Cat.Diagram.Terminal.html - -The limit over a discrete diagram is called a **product**, and it's -important to note that the diagram need not be finite. Here are concrete -examples of products in categories: - -- In $\rm{Sets}$, the limit is the _Cartesian product_ of the objects -of the diagram, and the arrows are the projections onto the factors. - -- In $\rm{Top}$, the limit is the _product space_ of the objects, -and the arrows are projections, considered as continuous maps. The -product topology can be defined as the coarsest topology that makes the -projections continuous. - -- In a partially ordered set, considered as a category, the limit is the -_greatest lower bound_ of the objects - In a poset, we consider that -there exists a map $a \to b$ whenever $a \le b$. +We might be tempted to stop here and call it a day, but we can go one +step further. It turns out that these universal functors have a name: +they are [right Kan extensions]. This allows for an extremely concise +definition of limits: $x : \cC$ is the limit of a diagram +$D : \cJ \to \cC$ when the constant functor $!x : \{*\} \to \cC$ is +a right Kan extension of $! : \cJ \to \{*\}$ along $D$. - Normalising the definition of limit slightly, it's works out to be an - object $p$ such that $p \le a$ and $p \le b$, and if there are any - other $p'$s satisfying that, then $p' \le p$. - -This last example also demonstrates that, while we can always _describe_ -the limit over a diagram, it does not necessarily exist. Consider the -poset $(\bb{R} \setminus {0}, \le)$ of real numbers except zero, -with the usual ordering. Then the product indexed by $\{ x \in -\bb{R} : x > 0 \}$ - which is normally $0$ - does not exist. Not -every category has every limit. Some categories have no limits at all! -If a category has every limit, it's called _complete_. - -## Terminal objects - -**Note**: Terminal objects are described explicitly in -[`Cat.Diagram.Terminal`]. The description there might be easier to -parse. - -[`Cat.Diagram.Terminal`]: Cat.Diagram.Terminal.html - -Perhaps a simpler example of a limit is the one over the empty diagram. -Since the empty diagram is discrete, what we'll end up with is a kind of -product - an empty product. A cone over the empty diagram is an object - -the family of maps, indexed over the objects of the diagram, is empty. - -In this case, a cone homomorphism works out to be a regular ol' -morphism, so the terminal object in the cone category is a terminal -object in the ambient category: an object $\top$ such that, for every -other $A$, there's a unique arrow $A \to \top$. - -# Construction - -Cones are always considered over a _diagram_. Diagram just means -"functor", but it's a concept with an attitude: Whereas, in general, -functors can be a lot more involved than the name "diagram" would -suggest, _every_ functor can be considered a diagram! However, for the -purpose of constructing limits, we generally work with functors out of -"shapes", tiny categories like $\bullet \to \bullet \leftarrow \bullet$. +[right Kan extensions]: Cat.Functor.Kan.Base.html ```agda -module _ {J : Precategory o₁ h₁} {C : Precategory o₂ h₂} (F : Functor J C) where +module _ {J : Precategory o₁ h₁} {C : Precategory o₂ h₂} (Diagram : Functor J C) where private - import Cat.Reasoning J as J - import Cat.Reasoning C as C - module F = Functor F + module C = Precategory C + open _=>_ + open Functor - record Cone : Type (o₁ ⊔ o₂ ⊔ h₁ ⊔ h₂) where - no-eta-equality + cone→counit : ∀ {x : C.Ob} → (Const x => Diagram) → const! x F∘ !F => Diagram + unquoteDef cone→counit = define-coherence cone→counit + + counit→cone : ∀ {K : Functor ⊤Cat C} → K F∘ !F => Diagram → (Const (K .F₀ tt) => Diagram) + counit→cone {K = K} eta .η = eta .η + counit→cone {K = K} eta .is-natural x y f = + ap (_ C.∘_) (sym (K .F-id)) ∙ eta .is-natural x y f + + is-limit : (x : C.Ob) → Const x => Diagram → Type _ + is-limit x cone = is-ran !F Diagram (const! x) (cone→counit cone) ``` -A `Cone`{.Agda} over $F$ is given by an object (the `apex`{.agda}) -together with a family of maps `ψ`{.Agda} --- one for each object in the -indexing category `J`{.Agda} --- such that "everything in sight -commutes". +In a "bundled" form, we may define the _type of limits_ for a diagram +$D$ as the type of right extensions of $D$ along the terminating functor +$\mathord{!} : \cJ \to \{*\}$. ```agda - field - apex : C.Ob - ψ : (x : J.Ob) → C.Hom apex (F.₀ x) + Limit : Type _ + Limit = Ran !F Diagram ``` -For every map $f : x \to y$ in the indexing category, we require that -the diagram below commutes. This encompasses "everything" since the only -non-trivial composites that can be formed with the data at hand are of -the form $F(f) \circ \psi_x$. +## Concretely + +The definition above is very concise, and it has the benefit of being +abstract: We can re-use definitions and theorems originally stated for +Kan extensions to limits. However, it has the downside of being +abstract: it's good for working with _limits in general_, but working +with a _specific_ limit is penalised, as the data we want to get at is +"buried". +The definition above is also hard to _instantiate_, since you have to.. +bury the data, and some of it is really quite deep! What we do is +provide an auxiliary record, `make-is-limit`{.Agda}, which computes +right extensions to the point. + + -These non-trivial composites consist of following the left leg of the -diagram below, followed by the bottom leg. For it to commute, that path -has to be the same as following the right leg. +We solve this by defining a _concretised_ version of `is-limit`{.Agda}, +called `make-is-limit`{.Agda}, which exposes the following data. First, +we have morphisms from the apex to every value in the diagram, a family +called $\psi$. Moreover, if $f : x \to y$ is a morphism in the "shape" +category $\cJ$, then $F(f)\psi_x = \psi_y$, i.e., the $\psi$ maps fit into +triangles -~~~{.quiver .short-05} +~~~{.quiver .tall-15} \[\begin{tikzcd} - & {\operatorname{apex}} \\ - {F(x)} && {F(y)} - \arrow["{F(f)}"', from=2-1, to=2-3] - \arrow["{\psi_x}"', from=1-2, to=2-1] - \arrow["{\psi_y}", from=1-2, to=2-3] + & {\mathrm{apex}} \\ + \\ + Fx && {Fy\text{.}} + \arrow["{\psi_x}"', curve={height=6pt}, from=1-2, to=3-1] + \arrow["{\psi_y}", curve={height=-6pt}, from=1-2, to=3-3] + \arrow["Ff"', from=3-1, to=3-3] \end{tikzcd}\] ~~~ -Since paths in Hom-spaces `are propositions`{.Agda ident=Hom-set}, to -test equality of cones, it suffices for the apices to be equal, and for -their $\psi$s to assign equal morphisms for every object in the shape -category. - ```agda - Cone-path : {x y : Cone} - → (p : Cone.apex x ≡ Cone.apex y) - → (∀ o → PathP (λ i → C.Hom (p i) (F.₀ o)) (Cone.ψ x o) (Cone.ψ y o)) - → x ≡ y - Cone-path p q i .Cone.apex = p i - Cone-path p q i .Cone.ψ o = q o i - Cone-path {x = x} {y} p q i .Cone.commutes {x = a} {y = b} f = - is-prop→pathp (λ i → C.Hom-set _ _ (F.₁ f C.∘ q a i) (q b i)) - (x .commutes f) (y .commutes f) i - where open Cone + field + ψ : (j : J.Ob) → C.Hom apex (F₀ j) + commutes : ∀ {x y} (f : J.Hom x y) → F₁ f C.∘ ψ x ≡ ψ y ``` -## Cone maps +The rest of the data says that $\psi$ is the universal family of maps +with this property: If $\eta_j : x \to Fj$ is another family of maps +with the same commutativty property, then each $\eta_j$ factors through +the apex by a single, _unique_ universal morphism: ```agda - record Cone-hom (x y : Cone) : Type (o₁ ⊔ h₂) where - no-eta-equality + universal + : ∀ {x : C.Ob} + → (eta : ∀ j → C.Hom x (F₀ j)) + → (∀ {x y} (f : J.Hom x y) → F₁ f C.∘ eta x ≡ eta y) + → C.Hom x apex + + factors + : ∀ {j : J.Ob} {x : C.Ob} + → (eta : ∀ j → C.Hom x (F₀ j)) + → (p : ∀ {x y} (f : J.Hom x y) → F₁ f C.∘ eta x ≡ eta y) + → ψ j C.∘ universal eta p ≡ eta j + + unique + : ∀ {x : C.Ob} + → (eta : ∀ j → C.Hom x (F₀ j)) + → (p : ∀ {x y} (f : J.Hom x y) → F₁ f C.∘ eta x ≡ eta y) + → (other : C.Hom x apex) + → (∀ j → ψ j C.∘ other ≡ eta j) + → other ≡ universal eta p ``` -A `Cone homomorphism`{.Agda ident="Cone-hom"} is -- like the introduction -says -- a map `hom`{.Agda} in the ambient category between the apices, -such that "everything in sight `commutes`{.Agda ident="Cone-hom.commutes"}". -Specifically, for any choice of object $o$ in the index category, the -composition of `hom`{.Agda} with the domain cone's `ψ`{.Agda} (at that -object) must be equal to the codomain's `ψ`{.Agda}. - - + + +If we have this data, then we can make a value of `is-limit`{.Agda}. It +might seem like naturality, required for a Kan extension, is missing +from `make-is-limit`{.Agda}, but it can be derived from the other data +we have been given: -Since cone homomorphisms are morphisms in the underlying category with -extra properties, we can lift the laws from the underlying category to -the category of `Cones`{.Agda}. The definition of `compose`{.Agda} is the -enlightening part, since we have to prove that two cone homomorphisms -again preserve _all_ the commutativities. - -```agda - Cones : Precategory _ _ - Cones = cat where - open Precategory - - compose : {x y z : _} → Cone-hom y z → Cone-hom x y → Cone-hom x z - compose {x} {y} {z} F G = r where - open Cone-hom - r : Cone-hom x z - r .hom = hom F C.∘ hom G - r .commutes o = - Cone.ψ z o C.∘ hom F C.∘ hom G ≡⟨ C.pulll (commutes F o) ⟩ - Cone.ψ y o C.∘ hom G ≡⟨ commutes G o ⟩ - Cone.ψ x o ∎ - - cat : Precategory _ _ - cat .Ob = Cone - cat .Hom = Cone-hom - cat .id = record { hom = C.id ; commutes = λ _ → C.idr _ } - cat ._∘_ = compose - cat .idr f = Cone-hom-path (C.idr _) - cat .idl f = Cone-hom-path (C.idl _) - cat .assoc f g h = Cone-hom-path (C.assoc _ _ _) +```agda + to-is-limit : ∀ {D : Functor J C} {apex} + → (mk : make-is-limit D apex) + → is-limit D apex (to-cone mk) + to-is-limit {Diagram} {apex} mklim = lim where + open make-is-limit mklim + open is-ran + open Functor + open _=>_ + + lim : is-limit Diagram apex (to-cone mklim) + lim .σ {M = M} α .η _ = + universal (α .η) (λ f → sym (α .is-natural _ _ f) ∙ C.elimr (M .F-id)) + lim .σ {M = M} α .is-natural _ _ _ = + lim .σ α .η _ C.∘ M .F₁ tt ≡⟨ C.elimr (M .F-id) ⟩ + lim .σ α .η _ ≡˘⟨ C.idl _ ⟩ + C.id C.∘ lim .σ α .η _ ∎ + lim .σ-comm {β = β} = Nat-path λ j → + factors (β .η) _ + lim .σ-uniq {β = β} {σ′ = σ′} p = Nat-path λ _ → + sym $ unique (β .η) _ (σ′ .η tt) (λ j → sym (p ηₚ j)) ``` -## Limits +To _use_ the data of `is-limit`, we provide a function for *un*making a +limit: -At risk of repeating myself: *A `Limit`{.agda} is, then, a [terminal -object] in this category.* +```agda + unmake-limit + : ∀ {D : Functor J C} {F : Functor ⊤Cat C} {eps} + → is-ran !F D F eps + → make-is-limit D (Functor.F₀ F tt) +``` -[terminal object]: Cat.Diagram.Terminal.html + + - Limit-universal : (L : Limit) → (K : Cone) → C.Hom (Cone.apex K) (Limit-apex L) - Limit-universal L K = Cone-hom.hom (Terminal.! L {K}) +We also provide a similar interface for the bundled form of limits. - is-limit : Cone → Type _ - is-limit K = is-terminal Cones K +```agda +module Limit + {J : Precategory o₁ h₁} {C : Precategory o₂ h₂} {D : Functor J C} (L : Limit D) + where ``` -# Preservation of limits - -Since a cone is, in particular, a commutative diagram, and every functor -preserves commutativity of diagrams, a functor $F : \cC \to \cD$ -acts on a cone over $\rm{Dia}$ (in $\cC$), sending it to a cone over -$F \circ \rm{Dia}$ (in $\cD$). +The "apex" object of the limit is the single value in the image of the +extension functor. ```agda - F-map-cone : Cone Dia → Cone (F F∘ Dia) - Cone.apex (F-map-cone x) = F₀ F (Cone.apex x) - Cone.ψ (F-map-cone x) x₁ = F₁ F (Cone.ψ x x₁) - Cone.commutes (F-map-cone x) {y = y} f = - F.₁ (F₁ Dia f) D.∘ F.₁ (Cone.ψ x _) ≡⟨ F.collapse (Cone.commutes x _) ⟩ - F.₁ (Cone.ψ x y) ∎ + apex : C.Ob + apex = Ext.₀ tt +``` + +Furthermore, we can show that the apex is the limit, in the sense of +`is-limit`{.Agda}, of the diagram. You'd think this is immediate, but +unfortunately, proof assistants: `is-limit`{.Agda} asks for _the_ +constant functor functor $\{*\} \to \cC$ with value `apex` to be a Kan +extension, but `Limit`{.Agda}, being an instance of `Ran`{.Agda}, +packages an _arbitrary_ functor $\{*\} \to \cC$. +Since Agda does not compare functors for $\eta$-equality, we have to +shuffle our data around manually. Fortunately, this isn't a very long +computation. + +```agda + cone : Const apex => D + cone .η x = eps .η x + cone .is-natural x y f = ap (_ C.∘_) (sym $ Ext .F-id) ∙ eps .is-natural x y f + + has-limit : is-limit D apex cone + has-limit .is-ran.σ α .η = σ α .η + has-limit .is-ran.σ α .is-natural x y f = + σ α .is-natural tt tt tt ∙ ap (C._∘ _) (Ext .F-id) + has-limit .is-ran.σ-comm = + Nat-path (λ _ → σ-comm ηₚ _) + has-limit .is-ran.σ-uniq {M = M} {σ′ = σ′} p = + Nat-path (λ _ → σ-uniq {σ′ = nt} (Nat-path (λ j → p ηₚ j)) ηₚ _) where + nt : M => Ext + nt .η = σ′ .η + nt .is-natural x y f = σ′ .is-natural x y f ∙ ap (C._∘ _) (sym $ Ext .F-id) + + open is-limit has-limit public ``` -Note that this also lets us map morphisms between cones into $\cD$. +# Uniqueness + + +Above, there has been mention of _the_ limit. The limit of a diagram, if +it exists, is unique up to isomorphism. This follows directly from +[uniqueness of Kan extensions]. -Suppose you have a limit $L$ of $\rm{Dia}$ --- which is, to reiterate, a -terminal object in the category of cones over $\rm{Dia}$. We say that -$F$ *preserves $L$* if $F(L)$, as defined right above, is a terminal -object in the category of cones over $F \circ \rm{Dia}$. +[uniqueness of Kan extensions]: Cat.Functor.Kan.Unique.html + +We show a slightly more general result first: if there exist a pair of +maps $f$, $g$ between the apexes of the 2 limits, and these maps commute +with the 2 limits, then $f$ and $g$ are inverses. ```agda - Preserves-limit : Cone Dia → Type _ - Preserves-limit K = is-limit Dia K → is-limit (F F∘ Dia) (F-map-cone K) + limits→inversesp + : ∀ {f : C.Hom x y} {g : C.Hom y x} + → (∀ {j : J.Ob} → Ly.ψ j C.∘ f ≡ Lx.ψ j) + → (∀ {j : J.Ob} → Lx.ψ j C.∘ g ≡ Ly.ψ j) + → C.Inverses f g + limits→inversesp {f = f} {g = g} f-factor g-factor = + natural-inverses→inverses + {α = hom→⊤-natural-trans f} + {β = hom→⊤-natural-trans g} + (Ran-unique.σ-inversesp Ly Lx + (Nat-path λ j → f-factor {j}) + (Nat-path λ j → g-factor {j})) + tt ``` -This definition is necessary because $\cD$ will not, in general, -possess an operation assigning a limit to every diagram --- therefore, -there might not be a "canonical limit" of $F\circ\rm{Dia}$ we could -compare $F(L)$ to. However, since limits are described by a universal -property (in particular, being terminal), we don't _need_ such an -object! Any limit is as good as any other. +Furthermore, any morphism between apexes that commutes with the limit +must be invertible. -In more concise terms, we say a functor preserves limits if it takes -limiting cones "upstairs" to limiting cones "downstairs". +```agda + limits→invertiblep + : ∀ {f : C.Hom x y} + → (∀ {j : J.Ob} → Ly.ψ j C.∘ f ≡ Lx.ψ j) + → C.is-invertible f + limits→invertiblep {f = f} f-factor = + is-natural-invertible→invertible + {α = hom→⊤-natural-trans f} + (Ran-unique.σ-is-invertiblep + Ly + Lx + (Nat-path λ j → f-factor {j})) + tt +``` -## Reflection of limits +This implies that the universal maps must also be inverses. -Using our analogy from before, we say a functor _reflects_ limits -if it takes limiting cones "downstairs" to limiting cones "upstairs". +```agda + limits→inverses + : C.Inverses (Ly.universal Lx.ψ Lx.commutes) (Lx.universal Ly.ψ Ly.commutes) + limits→inverses = + limits→inversesp (Ly.factors Lx.ψ Lx.commutes) (Lx.factors Ly.ψ Ly.commutes) + + limits→invertible + : C.is-invertible (Ly.universal Lx.ψ Lx.commutes) + limits→invertible = limits→invertiblep (Ly.factors Lx.ψ Lx.commutes) +``` -More concretely, if we have some limiting cone in $\cD$ of $F \circ -\rm{Dia}$ with apex $F(a)$, then $a$ was _already the limit_ of -$\rm{Dia}$! +Finally, we can bundle this data up to show that the apexes are isomorphic. ```agda - Reflects-limit : Cone Dia → Type _ - Reflects-limit K = is-limit (F F∘ Dia) (F-map-cone K) → is-limit Dia K + limits-unique : x C.≅ y + limits-unique = + Nat-iso→Iso (Ran-unique.unique Lx Ly) tt ``` -## Creation of limits -Finally, we say a functor _creates_ limits of shape $\rm{Dia}$ if it -both preserves _and_ reflects those limits. Intuitively, this means that -the limits of shape $\rm{Dia}$ in $\cC$ are in a 1-1 correspondence -with the limits $F \circ \rm{Dia}$ in $\cD$. +Furthermore, if the universal map is invertible, then that means that +its domain is _also_ a limit of the diagram. + ```agda -is-continuous - : ∀ {oshape hshape} - {C : Precategory o₁ h₁} - {D : Precategory o₂ h₂} - → Functor C D → Type _ + is-invertible→is-limitp + : ∀ {K' : Functor ⊤Cat C} {eps : K' F∘ !F => D} + → (eta : ∀ j → C.Hom (K' .F₀ tt) (D.₀ j)) + → (p : ∀ {x y} (f : J.Hom x y) → D.₁ f C.∘ eta x ≡ eta y) + → (∀ {j} → eta j ≡ eps .η j) + → C.is-invertible (Ly.universal eta p) + → is-ran !F D K' eps + is-invertible→is-limitp {K' = K'} eta p q invert = + generalize-limitp + (is-invertible→is-ran Ly $ componentwise-invertible→invertible _ (λ _ → invert)) + q ``` -A continuous functor is one that --- for every shape of diagram `J`, and -every diagram `diagram`{.Agda} of shape `J` in `C` --- preserves the -limit for that diagram. +Another useful fact is that if $L$ is a limit of some diagram $Dia$, and +$Dia$ is naturally isomorphic to some other diagram $Dia'$, then the +apex of $L$ is also a limit of $Dia'$. ```agda -is-continuous {oshape = oshape} {hshape} {C = C} F = - ∀ {J : Precategory oshape hshape} {diagram : Functor J C} - → (K : Cone diagram) → Preserves-limit F K + natural-iso-diagram→is-limitp + : ∀ {D′ : Functor J C} {eps : K F∘ !F => D′} + → (isos : natural-iso D D′) + → (∀ {j} → natural-iso.to isos .η j C.∘ Ly.ψ j ≡ eps .η j) + → is-ran !F D′ K eps + natural-iso-diagram→is-limitp {D′ = D′} isos p = + generalize-limitp + (natural-iso-of→is-ran Ly isos) + p ``` -# Uniqueness -Above, there has been mention of _the_ limit. The limit of a diagram, if -it exists, is unique up to isomorphism. We prove that here. The argument -is as follows: Fixing a diagram $F$, suppose that $X$ and $Y$ are both -limiting cones for for $F$. +Since limits are unique “up to isomorphism”, if $C$ is a univalent +category, then `Limit`{.Agda} itself is a proposition! This is an +instance of the more general [uniqueness of Kan extensions]. + +[uniqueness of Kan extensions]: Cat.Functor.Kan.Unique.html + -It follows from $Y$ (resp. $X$) being a terminal object that there is a -unique cone homomorphism $f : X \to Y$ (resp $g : Y \to X$). - +```agda + Limit-is-prop : is-category C → is-prop (Limit Diagram) + Limit-is-prop cat = Ran-is-prop cat ``` - f : Cones.Hom X-cone Y-cone - f = Terminal.! Y {X-cone} - g : Cones.Hom Y-cone X-cone - g = Terminal.! X {Y-cone} -``` -To show that $g$ is an inverse to $f$, consider the composition $g \circ -f$ (the other case is symmetric): It is a map $g \circ f : X \to X$. -Since $X$ is a terminal object, we have that the space of cone -homomorphisms $X \to X$ is contractible - and thus any two such maps are -equal. Thus, $g \circ f = \id{X} : X \to X$. +# Preservation of Limits +Suppose you have a limit $L$ of a diagram $\rm{Dia}$. We say that $F$ +**preserves $L$** if $F(L)$ is also a limit of $F \circ \rm{Dia}$. + + + +Suppose you have a limit $L$ of a diagram $\rm{Dia}$. We say that $F$ +**preserves $L$** if $F(L)$ is also a limit of $F \circ \rm{Dia}$. + +This definition is necessary because $\cD$ will not, in general, +possess an operation assigning a limit to every diagram --- therefore, +there might not be a "canonical limit" of $F\circ\rm{Dia}$ we could +compare $F(L)$ to. However, since limits are described by a universal +property (in particular, being terminal), we don't _need_ such an +object! Any limit is as good as any other. - g∘f≡id : (g Cones.∘ f) ≡ Cones.id - g∘f≡id = Terminal.!-unique₂ X (g Cones.∘ f) Cones.id +In more concise terms, we say a functor preserves limits if it takes +limiting cones "upstairs" to limiting cones "downstairs". + +```agda + preserves-limit : Type _ + preserves-limit = + ∀ {K : Functor ⊤Cat C} {eps : K F∘ !F => Diagram} + → (lim : is-ran !F Diagram K eps) + → preserves-ran F lim ``` -There is an evident functor from `Cones`{.Agda} to `C`, which sends -cones to their apices and cone homomorphisms to their underlying maps. -Being a functor, it preserves isomorphisms, so we get an isomorphism of -the limit _objects_ from the isomorphism of limit _cones_. +## Reflection of limits + +Using the terminology from before, we say a functor **reflects limits** +if it takes limiting cones "downstairs" to limiting cones "upstairs". +More concretely, if we have a limit in $\cD$ of $F \circ \rm{Dia}$ with +apex $F(a)$, then $F$ reflects this limit means that $a$ _was already_ +the limit of $\rm{Dia}$! -Contrary to the paragraph above, though, rather than defining a functor, -there is a direct proof that an isomorphism of limits results in an -isomorphism of apices. Fortunately, the direct proof -`Cone≅→apex≅`{.Agda} is exactly what one would get had they defined the -functor and used the proof that it preserves isomorphisms. +```agda + reflects-limit : Type _ + reflects-limit = + ∀ {K : Functor ⊤Cat C} {eps : K F∘ !F => Diagram} + → (ran : is-ran !F (F F∘ Diagram) (F F∘ K) (nat-assoc-from (F ▸ eps))) + → reflects-ran F ran +``` + - Cone-invertible→apex-invertible : {X Y : Cone F} {f : Cones.Hom X Y} - → Cones.is-invertible f - → C.is-invertible (Cone-hom.hom f) - Cone-invertible→apex-invertible {f = f} f-invert = - C.make-invertible (Cone-hom.hom inv) (ap Cone-hom.hom invl) (ap Cone-hom.hom invr) - where open Cones.is-invertible f-invert +## Continuity - Limit-unique - : {X Y : Limit F} - → Cone.apex (Terminal.top X) C.≅ Cone.apex (Terminal.top Y) - Limit-unique {X} {Y} = Cone≅→apex≅ (Limiting-cone-unique X Y) +```agda +is-continuous + : ∀ (oshape hshape : Level) + {C : Precategory o₁ h₁} + {D : Precategory o₂ h₂} + → Functor C D → Type _ ``` -If the universal map $K \to L$ between apexes of some limit -is invertible, then that means that $K$ is also a limiting cone. +A continuous functor is one that --- for every shape of diagram `J`, and +every diagram `diagram`{.Agda} of shape `J` in `C` --- preserves the +limit for that diagram. ```agda - apex-iso→is-limit - : (K : Cone F) - (L : Limit F) - → C.is-invertible (Limit-universal F L K) - → is-limit F K - apex-iso→is-limit K L invert K′ = limits - where - module K = Cone K - module K′ = Cone K′ - module L = Cone (Terminal.top L) - module universal {K} = Cone-hom (Terminal.! L {K}) - open C.is-invertible invert - - limits : is-contr (Cones.Hom K′ K) - limits .centre .Cone-hom.hom = inv C.∘ Limit-universal F L K′ - limits .centre .Cone-hom.commutes _ = - (K.ψ _) C.∘ (inv C.∘ universal.hom) ≡˘⟨ ap ( C._∘ (inv C.∘ universal.hom)) (universal.commutes _) ⟩ - (L.ψ _ C.∘ universal.hom) C.∘ (inv C.∘ universal.hom) ≡⟨ C.cancel-inner invl ⟩ - L.ψ _ C.∘ universal.hom ≡⟨ universal.commutes _ ⟩ - K′.ψ _ ∎ - limits .paths f = Cone-hom-path F $ C.invertible→monic invert _ _ $ - universal.hom C.∘ (inv C.∘ universal.hom) ≡⟨ C.cancell invl ⟩ - universal.hom ≡⟨ ap Cone-hom.hom (Terminal.!-unique L (Terminal.! L Cones.∘ f)) ⟩ - universal.hom C.∘ Cone-hom.hom f ∎ -``` - -## Completeness +is-continuous oshape hshape {C = C} F = + ∀ {J : Precategory oshape hshape} {Diagram : Functor J C} + → preserves-limit F Diagram +``` + +# Completeness A category is **complete** if admits for limits of arbitrary shape. However, in the presence of excluded middle, if a category admits @@ -576,65 +812,3 @@ indexed by a precategory with objects in $\ty\ o$ and morphisms in $\ty\ is-complete : ∀ {oc ℓc} o ℓ → Precategory oc ℓc → Type _ is-complete o ℓ C = ∀ {D : Precategory o ℓ} (F : Functor D C) → Limit F ``` - - diff --git a/src/Cat/Diagram/Limit/Cone.lagda.md b/src/Cat/Diagram/Limit/Cone.lagda.md new file mode 100644 index 000000000..51b1f41b6 --- /dev/null +++ b/src/Cat/Diagram/Limit/Cone.lagda.md @@ -0,0 +1,245 @@ +```agda +open import Cat.Instances.Shape.Terminal +open import Cat.Diagram.Limit.Base +open import Cat.Instances.Functor +open import Cat.Diagram.Terminal +open import Cat.Functor.Kan.Base +open import Cat.Prelude + +import Cat.Functor.Reasoning as Func +import Cat.Reasoning + +module Cat.Diagram.Limit.Cone where +``` + +# Limits via Cones + +As noted in the main page on [limits], most introductory material defines +limits via a mathematical widget called a **cone**. + +[limits]: Cat.Diagram.Limit.Base.html + + + +```agda +module _ {J : Precategory o₁ h₁} {C : Precategory o₂ h₂} (F : Functor J C) where + private + import Cat.Reasoning J as J + import Cat.Reasoning C as C + module F = Functor F + + record Cone : Type (o₁ ⊔ o₂ ⊔ h₁ ⊔ h₂) where + no-eta-equality +``` + +A `Cone`{.Agda} over $F$ is given by an object (the `apex`{.agda}) +together with a family of maps `ψ`{.Agda} --- one for each object in the +indexing category `J`{.Agda} --- such that "everything in sight +commutes". + +```agda + field + apex : C.Ob + ψ : (x : J.Ob) → C.Hom apex (F.₀ x) +``` + +For every map $f : x \to y$ in the indexing category, we require that +the diagram below commutes. This encompasses "everything" since the only +non-trivial composites that can be formed with the data at hand are of +the form $F(f) \circ \psi_x$. + +```agda + commutes : ∀ {x y} (f : J.Hom x y) → F.₁ f C.∘ ψ x ≡ ψ y +``` + +These non-trivial composites consist of following the left leg of the +diagram below, followed by the bottom leg. For it to commute, that path +has to be the same as following the right leg. + +~~~{.quiver .short-05} +\[\begin{tikzcd} + & {\operatorname{apex}} \\ + {F(x)} && {F(y)} + \arrow["{F(f)}"', from=2-1, to=2-3] + \arrow["{\psi_x}"', from=1-2, to=2-1] + \arrow["{\psi_y}", from=1-2, to=2-3] +\end{tikzcd}\] +~~~ + +Since paths in Hom-spaces `are propositions`{.Agda ident=Hom-set}, to +test equality of cones, it suffices for the apices to be equal, and for +their $\psi$s to assign equal morphisms for every object in the shape +category. + +```agda + Cone-path : {x y : Cone} + → (p : Cone.apex x ≡ Cone.apex y) + → (∀ o → PathP (λ i → C.Hom (p i) (F.₀ o)) (Cone.ψ x o) (Cone.ψ y o)) + → x ≡ y + Cone-path p q i .Cone.apex = p i + Cone-path p q i .Cone.ψ o = q o i + Cone-path {x = x} {y} p q i .Cone.commutes {x = a} {y = b} f = + is-prop→pathp (λ i → C.Hom-set _ _ (F.₁ f C.∘ q a i) (q b i)) + (x .commutes f) (y .commutes f) i + where open Cone +``` + +## Cone maps + +```agda + record Cone-hom (x y : Cone) : Type (o₁ ⊔ h₂) where + no-eta-equality + constructor cone-hom +``` + +A `Cone homomorphism`{.Agda ident="Cone-hom"} is -- like the introduction +says -- a map `hom`{.Agda} in the ambient category between the apices, +such that "everything in sight `commutes`{.Agda ident="Cone-hom.commutes"}". +Specifically, for any choice of object $o$ in the index category, the +composition of `hom`{.Agda} with the domain cone's `ψ`{.Agda} (at that +object) must be equal to the codomain's `ψ`{.Agda}. + + +```agda + field + hom : C.Hom (Cone.apex x) (Cone.apex y) + commutes : ∀ o → Cone.ψ y o C.∘ hom ≡ Cone.ψ x o +``` + + + +Since cone homomorphisms are morphisms in the underlying category with +extra properties, we can lift the laws from the underlying category to +the category of `Cones`{.Agda}. The definition of `compose`{.Agda} is the +enlightening part, since we have to prove that two cone homomorphisms +again preserve _all_ the commutativities. + +```agda + Cones : Precategory _ _ + Cones = cat where + open Precategory + + compose : {x y z : _} → Cone-hom y z → Cone-hom x y → Cone-hom x z + compose {x} {y} {z} F G = r where + open Cone-hom + r : Cone-hom x z + r .hom = hom F C.∘ hom G + r .commutes o = + Cone.ψ z o C.∘ hom F C.∘ hom G ≡⟨ C.pulll (commutes F o) ⟩ + Cone.ψ y o C.∘ hom G ≡⟨ commutes G o ⟩ + Cone.ψ x o ∎ + + cat : Precategory _ _ + cat .Ob = Cone + cat .Hom = Cone-hom + cat .id = record { hom = C.id ; commutes = λ _ → C.idr _ } + cat ._∘_ = compose + cat .idr f = Cone-hom-path (C.idr _) + cat .idl f = Cone-hom-path (C.idl _) + cat .assoc f g h = Cone-hom-path (C.assoc _ _ _) +``` + + + +## Terminal Cones as Limits + +Note that cones over some diagram $F$ contain the exact same data as +natural transformations from a constant functor to $F$. +To obtain a limit, all we need is a way of stating that a given cone is +universal. In particular, the [terminal object] in the category of cones +over a diagram $F$ (if it exists!) is the limit of $F$. + +[terminal object]: Cat.Diagram.Terminal.html + +The proof here is just shuffling data around: this is not totally +surprising, as both constructions contain the same data, just organized +differently. + +```agda + Cone→cone : (K : Cone) → Const (Cone.apex K) => F + Cone→cone K .η = K .Cone.ψ + Cone→cone K .is-natural x y f = C.idr _ ∙ sym (K .Cone.commutes f) + + is-terminal-cone→is-limit + : ∀ {K : Cone} + → is-terminal Cones K + → is-limit F (Cone.apex K) (Cone→cone K) + is-terminal-cone→is-limit {K = K} term = isl where + open Cone-hom + open is-ran + open Cone + + isl : is-ran _ F _ (cone→counit F (Cone→cone K)) + isl .σ {M = M} α = nt where + α′ : Cone + α′ .apex = M .Functor.F₀ tt + α′ .ψ x = α .η x + α′ .commutes f = sym (α .is-natural _ _ f) ∙ C.elimr (M .Functor.F-id) + + nt : M => const! (K .apex) + nt .η x = term α′ .centre .hom + nt .is-natural tt tt tt = C.elimr (M .Functor.F-id) ∙ C.introl refl + isl .σ-comm = Nat-path λ x → term _ .centre .commutes _ + isl .σ-uniq {σ′ = σ′} x = Nat-path λ _ → ap hom $ term _ .paths λ where + .hom → σ′ .η _ + .commutes _ → sym (x ηₚ _) +``` + +The inverse direction of this equivalence also consists of packing and +unpacking data. + +```agda + is-limit→is-terminal-cone + : ∀ {x} {eps : Const x => F} + → (L : is-limit F x eps) + → is-terminal Cones (record { commutes = is-limit.commutes L }) + is-limit→is-terminal-cone {x = x} L K = term where + module L = is-limit L + module K = Cone K + open Cone-hom + + term : is-contr (Cone-hom K _) + term .centre .hom = + L.universal K.ψ K.commutes + term .centre .commutes _ = + L.factors K.ψ K.commutes + term .paths f = + Cone-hom-path (sym (L.unique K.ψ K.commutes (f .hom) (f .commutes))) +``` + + diff --git a/src/Cat/Diagram/Limit/Equaliser.lagda.md b/src/Cat/Diagram/Limit/Equaliser.lagda.md index bdc56afb4..dd0d3c115 100644 --- a/src/Cat/Diagram/Limit/Equaliser.lagda.md +++ b/src/Cat/Diagram/Limit/Equaliser.lagda.md @@ -1,26 +1,25 @@ ```agda open import Cat.Instances.Shape.Parallel +open import Cat.Instances.Shape.Terminal open import Cat.Diagram.Limit.Base open import Cat.Diagram.Equaliser open import Cat.Instances.Functor -open import Cat.Diagram.Terminal +open import Cat.Functor.Kan.Base open import Cat.Prelude open import Data.Bool -module Cat.Diagram.Limit.Equaliser {o h} (Cat : Precategory o h) where +module Cat.Diagram.Limit.Equaliser {o h} (C : Precategory o h) where ``` @@ -30,58 +29,78 @@ We establish the correspondence between `Equaliser`{.Agda} and the [parallel arrows]: Cat.Instances.Shape.Parallel.html ```agda -Fork→Cone - : ∀ {E} (F : Functor ·⇉· Cat) - → (eq : Hom E (F .F₀ false)) - → F .F₁ {false} {true} false ∘ eq ≡ F .F₁ {false} {true} true ∘ eq - → Cone F -Fork→Cone F eq p .apex = _ -Fork→Cone F eq p .ψ false = eq -Fork→Cone F eq p .ψ true = F .F₁ false ∘ eq -Fork→Cone F eq p .commutes {false} {false} tt = eliml (F .F-id) -Fork→Cone F eq p .commutes {false} {true} false = refl -Fork→Cone F eq p .commutes {false} {true} true = sym p -Fork→Cone F eq p .commutes {true} {true} tt = eliml (F .F-id) +is-equaliser→is-limit + : ∀ {e} (F : Functor ·⇉· C) {equ : Hom e (F .F₀ false)} + → (eq : is-equaliser C (forkl F) (forkr F) equ) + → is-limit {C = C} F e (Fork→Cone F (is-equaliser.equal eq)) +is-equaliser→is-limit {e} F {equ} is-eq = + to-is-limitp ml λ where + {true} → refl + {false} → refl + where + module is-eq = is-equaliser is-eq + open make-is-limit -Equaliser→Limit - : ∀ {F : Functor ·⇉· Cat} - → Equaliser Cat (F .F₁ false) (F .F₁ true) - → Limit F -Equaliser→Limit {F = F} eq = lim where - module eq = Equaliser eq - lim : Limit _ - lim .top = Fork→Cone F eq.equ eq.equal - lim .has⊤ cone .centre .hom = - eq.limiting (cone .commutes {false} {true} false ∙ sym (cone .commutes true)) - lim .has⊤ cone .centre .commutes false = eq.universal - lim .has⊤ cone .centre .commutes true = pullr eq.universal ∙ cone .commutes {false} {true} false - lim .has⊤ cone .paths other = Cone-hom-path _ (sym (eq.unique p)) where - p : cone .ψ false ≡ eq .equ ∘ other .hom - p = sym (other .commutes _) + ml : make-is-limit F e + ml .ψ true = forkl F ∘ equ + ml .ψ false = equ + ml .commutes {true} {true} tt = eliml (F .F-id) + ml .commutes {false} {true} true = sym is-eq.equal + ml .commutes {false} {true} false = refl + ml .commutes {false} {false} tt = eliml (F .F-id) + ml .universal eta p = + is-eq.universal (p {false} {true} false ∙ sym (p {false} {true} true)) + ml .factors {true} eta p = + pullr is-eq.factors ∙ p {false} {true} false + ml .factors {false} eta p = + is-eq.factors + ml .unique eta p other q = + is-eq.unique (q false) -Limit→Equaliser - : ∀ {F : Functor ·⇉· Cat} - → Limit F - → Equaliser Cat (F .F₁ {false} {true} false) (F .F₁ true) -Limit→Equaliser {F} lim = eq where - module lim = Terminal lim - eq : Equaliser Cat _ _ - eq .apex = _ - eq .equ = lim.top .ψ false - eq .has-is-eq .equal = - lim.top .commutes {false} {true} false ∙ sym (lim.top .commutes true) - eq .has-is-eq .limiting p = lim.has⊤ (Fork→Cone _ _ p) .centre .hom - eq .has-is-eq .universal {p = p} = - lim.has⊤ (Fork→Cone _ _ p) .centre .commutes false - eq .has-is-eq .unique {e′ = e'} {p = p} {lim' = lim'} x = - sym (ap hom (lim.has⊤ (Fork→Cone _ _ p) .paths other)) - where - other : Cone-hom _ _ _ - other .hom = _ - other .commutes false = sym x - other .commutes true = sym ( - F .F₁ false ∘ e' ≡⟨ ap (_ ∘_) x ⟩ - F .F₁ false ∘ lim.top .ψ false ∘ lim' ≡⟨ pulll (lim.top .commutes false) ⟩ - lim.top .ψ true ∘ lim' ∎ - ) +is-limit→is-equaliser + : ∀ (F : Functor ·⇉· C) {K : Functor ⊤Cat C} + → {eta : K F∘ !F => F} + → is-ran !F F K eta + → is-equaliser C (forkl F) (forkr F) (eta .η false) +is-limit→is-equaliser F {K} {eta} lim = eq where + module lim = is-limit lim + + parallel + : ∀ {x} → Hom x (F .F₀ false) + → (j : Bool) → Hom x (F .F₀ j) + parallel e′ true = forkl F ∘ e′ + parallel e′ false = e′ + + parallel-commutes + : ∀ {x} {e′ : Hom x (F .F₀ false)} + → forkl F ∘ e′ ≡ forkr F ∘ e′ + → ∀ i j → (h : Precategory.Hom ·⇉· i j) + → F .F₁ {i} {j} h ∘ parallel e′ i ≡ parallel e′ j + parallel-commutes p true true tt = eliml (F .F-id) + parallel-commutes p false true true = sym p + parallel-commutes p false true false = refl + parallel-commutes p false false tt = eliml (F .F-id) + + eq : is-equaliser C (forkl F) (forkr F) (eta .η false) + eq .equal = + sym (eta .is-natural false true false) ∙ eta .is-natural false true true + eq .universal {e′ = e′} p = + lim.universal (parallel e′) (λ {i} {j} h → parallel-commutes p i j h) + eq .factors = lim.factors {j = false} _ _ + eq .unique {p = p} {other = other} q = + lim.unique _ _ _ λ where + true → + ap (_∘ other) (intror (K .F-id) ∙ eta .is-natural false true true) + ·· pullr q + ·· sym p + false → q + +Equaliser→Limit : ∀ {F : Functor ·⇉· C} → Equaliser C (forkl F) (forkr F) → Limit F +Equaliser→Limit {F = F} eq = to-limit (is-equaliser→is-limit F (has-is-eq eq)) + +Limit→Equaliser : ∀ {F : Functor ·⇉· C} → Limit F → Equaliser C (forkl F) (forkr F) +Limit→Equaliser lim .apex = _ +Limit→Equaliser lim .equ = _ +Limit→Equaliser {F = F} lim .has-is-eq = + is-limit→is-equaliser F (Limit.has-limit lim) ``` diff --git a/src/Cat/Diagram/Limit/Finite.lagda.md b/src/Cat/Diagram/Limit/Finite.lagda.md index a117b6ebc..4a6c6f1b5 100644 --- a/src/Cat/Diagram/Limit/Finite.lagda.md +++ b/src/Cat/Diagram/Limit/Finite.lagda.md @@ -5,6 +5,7 @@ open import Cat.Diagram.Limit.Pullback open import Cat.Instances.Shape.Cospan open import Cat.Diagram.Limit.Product open import Cat.Diagram.Limit.Base +open import Cat.Diagram.Limit.Cone open import Cat.Instances.Discrete open import Cat.Diagram.Equaliser open import Cat.Diagram.Pullback @@ -62,8 +63,8 @@ to give a terminal object and binary products. record Finitely-complete : Type (ℓ ⊔ ℓ') where field terminal : Terminal C - products : ∀ A B → Product C A B - equalisers : ∀ {A B} (f g : Hom A B) → Equaliser C f g + products : has-products C + equalisers : has-equalisers C pullbacks : has-pullbacks C Eq : ∀ {A B} (f g : Hom A B) → Ob @@ -72,8 +73,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 Products = Binary-products C products + open Products using (_⊗₀_) public open Finitely-complete ``` @@ -153,17 +154,17 @@ equalisers to factor _that_ as a unique arrow $P' \to X \times_Z Y$. ~~~ ```agda - pb .limiting {p₁' = p₁'} {p₂' = p₂'} p = - eq.limiting {e′ = pr.⟨ p₁' , p₂' ⟩} ( + pb .universal {p₁' = p₁'} {p₂' = p₂'} p = + eq.universal {e′ = pr.⟨ p₁' , p₂' ⟩} ( (f ∘ p1) ∘ pr.⟨ p₁' , p₂' ⟩ ≡⟨ pullr pr.π₁∘factor ⟩ f ∘ p₁' ≡⟨ p ⟩ g ∘ p₂' ≡˘⟨ pullr pr.π₂∘factor ⟩ (g ∘ p2) ∘ pr.⟨ p₁' , p₂' ⟩ ∎ ) - pb .p₁∘limiting = pullr eq.universal ∙ pr.π₁∘factor - pb .p₂∘limiting = pullr eq.universal ∙ pr.π₂∘factor + pb .p₁∘universal = pullr eq.factors ∙ pr.π₁∘factor + pb .p₂∘universal = pullr eq.factors ∙ pr.π₂∘factor pb .unique p q = - eq.unique (sym (pr.unique _ (assoc _ _ _ ∙ p) (assoc _ _ _ ∙ q))) + eq.unique ((pr.unique _ (assoc _ _ _ ∙ p) (assoc _ _ _ ∙ q))) ``` Hence, assuming that a category has a terminal object, binary products @@ -241,9 +242,9 @@ object $*$. prod : is-product C p1 p2 prod .is-product.⟨_,_⟩ p1' p2' = - Pb.limiting {p₁' = p1'} {p₂' = p2'} (is-contr→is-prop (term _) _ _) - prod .is-product.π₁∘factor = Pb.p₁∘limiting - prod .is-product.π₂∘factor = Pb.p₂∘limiting + Pb.universal {p₁' = p1'} {p₂' = p2'} (is-contr→is-prop (term _) _ _) + prod .is-product.π₁∘factor = Pb.p₁∘universal + prod .is-product.π₂∘factor = Pb.p₂∘universal prod .is-product.unique other p q = Pb.unique p q with-pullbacks @@ -348,8 +349,8 @@ top face $\rm{equ} : \rm{eq}(f,g) \to A$ in our pullback diagram is indeed the equaliser of $f$ and $g$. ```agda - eq .has-is-eq .limiting {e′ = e′} p = - Pb.limiting (Bb.unique₂ refl refl (sym p1) (sym p2)) + eq .has-is-eq .universal {e′ = e′} p = + Pb.universal (Bb.unique₂ refl refl (sym p1) (sym p2)) where p1 : Bb.π₁ ∘ ⟨id,id⟩ ∘ f ∘ e′ ≡ Bb.π₁ ∘ ⟨f,g⟩ ∘ e′ p1 = @@ -364,17 +365,17 @@ is indeed the equaliser of $f$ and $g$. g ∘ e′ ≡˘⟨ pulll Bb.π₂∘factor ⟩ Bb.π₂ ∘ ⟨f,g⟩ ∘ e′ ∎ - eq .has-is-eq .universal = Pb.p₂∘limiting - eq .has-is-eq .unique {F} {e′ = e′} {lim' = lim'} e′=p₂∘l = - Pb.unique path (sym e′=p₂∘l) + eq .has-is-eq .factors = Pb.p₂∘universal + eq .has-is-eq .unique {F} {e′ = e′} {other = other} p₂∘l=e′ = + Pb.unique path p₂∘l=e′ where - path : Pb.p₁ ∘ lim' ≡ f ∘ e′ + path : Pb.p₁ ∘ other ≡ f ∘ e′ path = - Pb.p₁ ∘ lim' ≡⟨ insertl Bb.π₁∘factor ⟩ - Bb.π₁ ∘ ⟨id,id⟩ ∘ Pb.p₁ ∘ lim' ≡⟨ ap (Bb.π₁ ∘_) (extendl Pb.square) ⟩ - Bb.π₁ ∘ ⟨f,g⟩ ∘ Pb.p₂ ∘ lim' ≡⟨ ap (Bb.π₁ ∘_) (ap (⟨f,g⟩ ∘_) (sym e′=p₂∘l)) ⟩ - Bb.π₁ ∘ ⟨f,g⟩ ∘ e′ ≡⟨ pulll Bb.π₁∘factor ⟩ - f ∘ e′ ∎ + Pb.p₁ ∘ other ≡⟨ insertl Bb.π₁∘factor ⟩ + Bb.π₁ ∘ ⟨id,id⟩ ∘ Pb.p₁ ∘ other ≡⟨ ap (Bb.π₁ ∘_) (extendl Pb.square) ⟩ + Bb.π₁ ∘ ⟨f,g⟩ ∘ Pb.p₂ ∘ other ≡⟨ ap (Bb.π₁ ∘_) (ap (⟨f,g⟩ ∘_) p₂∘l=e′) ⟩ + Bb.π₁ ∘ ⟨f,g⟩ ∘ e′ ≡⟨ pulll Bb.π₁∘factor ⟩ + f ∘ e′ ∎ ``` Putting it all together into a record we get our proof of finite completeness: @@ -396,9 +397,9 @@ Putting it all together into a record we get our proof of finite completeness: open is-pullback pb : is-pullback C _ _ _ _ pb .square = is-contr→is-prop (t _) _ _ - pb .limiting _ = r .is-product.⟨_,_⟩ _ _ - pb .p₁∘limiting = r .is-product.π₁∘factor - pb .p₂∘limiting = r .is-product.π₂∘factor + pb .universal _ = r .is-product.⟨_,_⟩ _ _ + pb .p₁∘universal = r .is-product.π₁∘factor + pb .p₂∘universal = r .is-product.π₂∘factor pb .unique p q = r .is-product.unique _ p q is-complete→finitely @@ -424,13 +425,11 @@ Putting it all together into a record we get our proof of finite completeness: open Cone term′ : Terminal C - term′ = record { top = limF .top .apex ; has⊤ = limiting } where + term′ = record { top = Limit.apex limF ; has⊤ = limiting } where limiting : ∀ x → is-contr _ limiting x = - contr (limF .has⊤ (h′ x) .centre .hom) λ h → - ap hom (limF .has⊤ (h′ x) .paths (record { hom = h ; commutes = λ { () } })) - where h′ : ∀ x → Cone _ - h′ x = record { apex = x ; ψ = λ { () } ; commutes = λ { {()} } } + contr (Limit.universal limF (λ { () }) (λ { {()} })) λ h → + sym (Limit.unique limF _ _ h λ { () }) ``` --> diff --git a/src/Cat/Diagram/Limit/Product.lagda.md b/src/Cat/Diagram/Limit/Product.lagda.md index 5287e5013..266da232f 100644 --- a/src/Cat/Diagram/Limit/Product.lagda.md +++ b/src/Cat/Diagram/Limit/Product.lagda.md @@ -5,10 +5,11 @@ description: | --- ```agda +open import Cat.Instances.Shape.Terminal open import Cat.Diagram.Limit.Base open import Cat.Instances.Discrete open import Cat.Instances.Functor -open import Cat.Diagram.Terminal +open import Cat.Functor.Kan.Base open import Cat.Diagram.Product open import Cat.Prelude @@ -21,13 +22,10 @@ module Cat.Diagram.Limit.Product {o h} (C : Precategory o h) where ```agda open import Cat.Reasoning C --- Yikes: open is-product -open Terminal -open Cone-hom open Product open Functor -open Cone +open _=>_ ``` --> @@ -40,9 +38,9 @@ existing infrastructure): ```agda 2-object-diagram : ∀ {iss} → Ob → Ob → Functor (Disc′ (el Bool iss)) C -2-object-diagram A B = Disc-diagram Discrete-Bool λ where - false → A - true → B +2-object-diagram a b = Disc-diagram Discrete-Bool λ where + true → a + false → b ``` With that out of the way, we can establish the claimed equivalence. We @@ -53,43 +51,81 @@ but they're arranged in very different ways: ```agda Pair→Cone - : ∀ {iss} {Q A B} → Hom Q A → Hom Q B - → Cone (2-object-diagram {iss = iss} A B) -Pair→Cone {Q = Q} _ _ .apex = Q -Pair→Cone p1 p2 .ψ false = p1 -Pair→Cone p1 p2 .ψ true = p2 -Pair→Cone _ _ .commutes {false} {false} f = ap (_∘ _) (transport-refl _) ∙ idl _ -Pair→Cone _ _ .commutes {false} {true} f = absurd (true≠false (sym f)) -Pair→Cone _ _ .commutes {true} {false} f = absurd (true≠false f) -Pair→Cone _ _ .commutes {true} {true} f = ap (_∘ _) (transport-refl _) ∙ idl _ + : ∀ {iss} {q a b} → Hom q a → Hom q b + → Const q => (2-object-diagram {iss} a b) +Pair→Cone f g = Disc-natural λ where + true → f + false → g ``` The two-way correspondence between products and terminal cones is then simple enough to establish by appropriately shuffling the data. ```agda -Prod→Lim : ∀ {iss} {A B} → Product C A B → Limit (2-object-diagram {iss = iss} A B) -Prod→Lim prod .top = Pair→Cone (prod .π₁) (prod .π₂) -Prod→Lim prod .has⊤ x .centre .hom = prod .⟨_,_⟩ (x .ψ false) (x .ψ true) -Prod→Lim prod .has⊤ x .centre .commutes false = prod .π₁∘factor -Prod→Lim prod .has⊤ x .centre .commutes true = prod .π₂∘factor -Prod→Lim prod .has⊤ x .paths y = Cone-hom-path (2-object-diagram _ _) - (sym (prod .unique (y .hom) (y .commutes _) (y .commutes _))) - -Lim→Prod : ∀ {iss} {A B} → Limit (2-object-diagram {iss = iss} A B) → Product C A B -Lim→Prod x .apex = x .top .apex -Lim→Prod x .π₁ = x .top .ψ false -Lim→Prod x .π₂ = x .top .ψ true -Lim→Prod x .has-is-product .⟨_,_⟩ p1 p2 = x .has⊤ (Pair→Cone p1 p2) .centre .hom -Lim→Prod x .has-is-product .π₁∘factor = x .has⊤ (Pair→Cone _ _) .centre .commutes _ -Lim→Prod x .has-is-product .π₂∘factor = x .has⊤ (Pair→Cone _ _) .centre .commutes _ -Lim→Prod x .has-is-product .unique f p q = - sym (ap hom (x .has⊤ (Pair→Cone _ _) .paths other)) +is-product→is-limit + : ∀ {iss} {x a b} {p1 : Hom x a} {p2 : Hom x b} + → is-product C p1 p2 + → is-limit {C = C} (2-object-diagram {iss} a b) x (Pair→Cone p1 p2) +is-product→is-limit {x = x} {a} {b} {p1} {p2} is-prod = + to-is-limitp ml λ where + {true} → refl + {false} → refl where - other : Cone-hom (2-object-diagram _ _) _ _ - other .hom = f - other .commutes false = p - other .commutes true = q + module is-prod = is-product is-prod + open make-is-limit + + ml : make-is-limit (2-object-diagram a b) x + ml .ψ true = p1 + ml .ψ false = p2 + ml .commutes {true} {true} f = ap (_∘ p1) (transport-refl id) ∙ idl _ + ml .commutes {true} {false} f = absurd (true≠false f) + ml .commutes {false} {true} f = absurd (true≠false $ sym f) + ml .commutes {false} {false} f = ap (_∘ p2) (transport-refl id) ∙ idl _ + ml .universal eta _ = is-prod.⟨ eta true , eta false ⟩ + ml .factors {true} eta _ = is-prod.π₁∘factor + ml .factors {false} eta _ = is-prod.π₂∘factor + ml .unique eta p other q = is-prod.unique other (q true) (q false) + +is-limit→is-product + : ∀ {iss} {a b} {K : Functor ⊤Cat C} + → {eta : K F∘ !F => 2-object-diagram {iss} a b} + → is-ran !F (2-object-diagram {iss} a b) K eta + → is-product C (eta .η true) (eta .η false) +is-limit→is-product {iss = iss} {a} {b} {K} {eta} lim = prod where + module lim = is-limit lim + + pair + : ∀ {y} → Hom y a → Hom y b + → ∀ (j : Bool) → Hom y (2-object-diagram {iss} a b .F₀ j) + pair p1 p2 true = p1 + pair p1 p2 false = p2 + + pair-commutes + : ∀ {y} {p1 : Hom y a} {p2 : Hom y b} + → {i j : Bool} + → (p : i ≡ j) + → 2-object-diagram {iss} a b .F₁ p ∘ pair p1 p2 i ≡ pair p1 p2 j + pair-commutes {p1 = p1} {p2 = p2} = + J (λ _ p → 2-object-diagram {iss} a b .F₁ p ∘ pair p1 p2 _ ≡ pair p1 p2 _) + (eliml (2-object-diagram {iss} a b .F-id)) + + prod : is-product C (eta .η true) (eta .η false) + prod .⟨_,_⟩ f g = lim.universal (pair f g) pair-commutes + prod .π₁∘factor {_} {p1'} {p2'} = lim.factors (pair p1' p2') pair-commutes + prod .π₂∘factor {_} {p1'} {p2'} = lim.factors (pair p1' p2') pair-commutes + prod .unique other p q = lim.unique _ _ other λ where + true → p + false → q + +Product→Limit : ∀ {iss} {a b} → Product C a b → Limit (2-object-diagram {iss} a b) +Product→Limit prod = to-limit (is-product→is-limit (has-is-product prod)) + +Limit→Product : ∀ {iss} {a b} → Limit (2-object-diagram {iss} a b) → Product C a b +Limit→Product lim .apex = Limit.apex lim +Limit→Product lim .π₁ = Limit.cone lim .η true +Limit→Product lim .π₂ = Limit.cone lim .η false +Limit→Product lim .has-is-product = + is-limit→is-product (Limit.has-limit lim) ``` We note that _any_ functor $F : \rm{Disc}(\{0,1\}) \to \cC$ is @@ -99,7 +135,7 @@ defined above. ```agda canonical-functors : ∀ {iss} (F : Functor (Disc′ (el Bool iss)) C) - → F ≡ 2-object-diagram (F₀ F false) (F₀ F true) + → F ≡ 2-object-diagram (F₀ F true) (F₀ F false) canonical-functors {iss = iss} F = Functor-path p q where p : ∀ x → _ p false = refl diff --git a/src/Cat/Diagram/Limit/Pullback.lagda.md b/src/Cat/Diagram/Limit/Pullback.lagda.md index 345f018ac..84cc6caf2 100644 --- a/src/Cat/Diagram/Limit/Pullback.lagda.md +++ b/src/Cat/Diagram/Limit/Pullback.lagda.md @@ -1,6 +1,7 @@ ```agda open import Cat.Instances.Shape.Cospan open import Cat.Diagram.Limit.Base +open import Cat.Diagram.Limit.Cone open import Cat.Instances.Functor open import Cat.Diagram.Pullback open import Cat.Diagram.Terminal @@ -44,35 +45,35 @@ Square→Cone {F = F} p1 p2 square .commutes {cs-b} {cs-b} _ = eliml (F .F-id) Square→Cone {F = F} p1 p2 square .commutes {cs-b} {cs-c} _ = sym square Square→Cone {F = F} p1 p2 square .commutes {cs-c} {cs-c} _ = eliml (F .F-id) -Pullback→Limit +Pullback→Terminal-cone : ∀ {x y} {A B C} {f : Hom A C} {g : Hom B C} → Pullback Cat f g - → Limit (cospan→cospan-diagram x y {C = Cat} f g) -Pullback→Limit {f = f} {g} pb = lim where + → Terminal (Cones (cospan→cospan-diagram x y {C = Cat} f g)) +Pullback→Terminal-cone {f = f} {g} pb = lim where module pb = Pullback pb - lim : Limit _ + lim : Terminal (Cones _) lim .top = Square→Cone _ _ pb.square - lim .has⊤ cone .centre .hom = pb.limiting (cone .commutes (lift tt) ∙ sym (cone .commutes {cs-b} {cs-c} (lift tt))) - lim .has⊤ cone .centre .commutes cs-a = pb.p₁∘limiting - lim .has⊤ cone .centre .commutes cs-b = pb.p₂∘limiting - lim .has⊤ cone .centre .commutes cs-c = pullr pb.p₁∘limiting ∙ cone .commutes (lift tt) + lim .has⊤ cone .centre .hom = pb.universal (cone .commutes (lift tt) ∙ sym (cone .commutes {cs-b} {cs-c} (lift tt))) + lim .has⊤ cone .centre .commutes cs-a = pb.p₁∘universal + lim .has⊤ cone .centre .commutes cs-b = pb.p₂∘universal + lim .has⊤ cone .centre .commutes cs-c = pullr pb.p₁∘universal ∙ cone .commutes (lift tt) lim .has⊤ cone .paths otherhom = Cone-hom-path _ (sym (pb.unique (otherhom .commutes _) (otherhom .commutes _))) -Limit→Pullback +Terminal-cone→Pullback : ∀ {x y} → {F : Functor (·→·←· {x} {y}) Cat} - → Limit F + → Terminal (Cones F) → Pullback Cat (F .F₁ {cs-a} {cs-c} _) (F .F₁ {cs-b} {cs-c} _) -Limit→Pullback {F = F} lim = pb where +Terminal-cone→Pullback {F = F} lim = pb where module lim = Terminal lim pb : Pullback Cat _ _ pb .apex = lim.top .apex pb .p₁ = lim.top .ψ cs-a pb .p₂ = lim.top .ψ cs-b pb .has-is-pb .square = lim.top .commutes _ ∙ sym (lim.top .commutes {cs-b} {cs-c} _) - pb .has-is-pb .limiting x = lim.has⊤ (Square→Cone _ _ x) .centre .hom - pb .has-is-pb .p₁∘limiting {p = p} = lim.has⊤ (Square→Cone _ _ p) .centre .commutes cs-a - pb .has-is-pb .p₂∘limiting {p = p} = lim.has⊤ (Square→Cone _ _ p) .centre .commutes cs-b + pb .has-is-pb .universal x = lim.has⊤ (Square→Cone _ _ x) .centre .hom + pb .has-is-pb .p₁∘universal {p = p} = lim.has⊤ (Square→Cone _ _ p) .centre .commutes cs-a + pb .has-is-pb .p₂∘universal {p = p} = lim.has⊤ (Square→Cone _ _ p) .centre .commutes cs-b pb .has-is-pb .unique {p₁' = p₁'} {p₂'} {p} {lim'} a b = sym (ap hom (lim.has⊤ (Square→Cone _ _ p) .paths other)) where @@ -85,3 +86,20 @@ Limit→Pullback {F = F} lim = pb where F .F₁ {cs-a} {cs-c} _ ∘ lim.top .ψ cs-a ∘ lim' ≡⟨ ap (_ ∘_) a ⟩ F .F₁ {cs-a} {cs-c} _ ∘ p₁' ∎ ``` + + diff --git a/src/Cat/Diagram/Monad/Codensity.lagda.md b/src/Cat/Diagram/Monad/Codensity.lagda.md index e617927dd..7abe43f65 100644 --- a/src/Cat/Diagram/Monad/Codensity.lagda.md +++ b/src/Cat/Diagram/Monad/Codensity.lagda.md @@ -1,6 +1,6 @@ ```agda -open import Cat.Functor.Kan.Right open import Cat.Instances.Functor +open import Cat.Functor.Kan.Base open import Cat.Diagram.Monad open import Cat.Functor.Base open import Cat.Prelude @@ -47,7 +47,7 @@ general categorical machinery: [right Kan extensions]. [limits]: Cat.Diagram.Limit.Base.html [size]: 1Lab.intro.html#universes-and-size-issues -[right Kan extensions]: Cat.Functor.Kan.Right.html +[right Kan extensions]: Cat.Functor.Kan.Base.html The really, really short of it is that the codensity monad of $F$ is the right Kan extension of $F$ along itself, $\Ran_F F$. @@ -59,7 +59,6 @@ module _ (F : Functor A B) (R : Ran F F) where private module A = Cat A module B = Cat B - module Ext = Func Ext module F = Func F ``` --> diff --git a/src/Cat/Diagram/Monad/Limits.lagda.md b/src/Cat/Diagram/Monad/Limits.lagda.md index 5433ebcec..a829e5c88 100644 --- a/src/Cat/Diagram/Monad/Limits.lagda.md +++ b/src/Cat/Diagram/Monad/Limits.lagda.md @@ -1,9 +1,11 @@ ```agda open import Cat.Functor.Equivalence.Complete +open import Cat.Instances.Shape.Terminal open import Cat.Functor.Conservative open import Cat.Functor.Equivalence open import Cat.Diagram.Limit.Base open import Cat.Diagram.Terminal +open import Cat.Functor.Kan.Base open import Cat.Diagram.Monad open import Cat.Prelude @@ -21,8 +23,6 @@ private open Algebra-hom open Algebra-on -open Cone-hom -open Cone ``` --> @@ -55,184 +55,157 @@ Suppose we have a diagram as in the setup --- we'll show that the functor $U : \cC^M \to \cC$ both preserves and _reflects_ limits, in that $K$ is a limiting cone if, and only if, $U(K)$ is. + -Let $L$ be a cone over $F$: Since $U(K)$ is a limiting cone, then we -have a unique map of $U(L) \to U(K)$, which we must show extends to a -map of _algebras_ $L \to K$, which by definition means $! \nu = \nu -M_1(!)$. But those are maps $M_0(L) \to U(K)$ --- so if $M_0(L)$ was a -cone over $U \circ F$, and those two were maps of cones, then they would -be equal! +We begin with the following key lemma: Write $K : \cC$ for the limit of +a diagram $\cJ \xto{F} \cC^M \xto{U} \cC$. If $K$ carries an $M$-algebra +structure $\nu$, and the limit projections $\psi : K \to F(j)$ are +$M$-algebra morphisms, then $(K, \nu)$ is the limit of $F$ in $\cC^M$. ```agda - ! : Cone-hom _ other K - ! .hom .morphism = !′ .hom - ! .hom .commutes = - ap hom $ is-contr→is-prop (uniq cone′) - (record { hom = !′ .hom C.∘ apex other .snd .ν - ; commutes = λ o → C.pulll (!′ .commutes o) - }) - (record { hom = apex K .snd .ν C.∘ M.M₁ (!′ .hom) - ; commutes = λ o → C.pulll (K .ψ o .commutes) - ·· C.pullr (sym (M.M-∘ _ _) ∙ ap M.M₁ (!′ .commutes o)) - ·· sym (other .ψ o .commutes) - }) + make-algebra-limit + : ∀ {K : Functor ⊤Cat C} {eps : K F∘ !F => Forget C M F∘ F} + → (lim : is-ran !F (Forget C M F∘ F) K eps) + → (nu : Algebra-on C M (K .F₀ tt)) + → (∀ j → is-limit.ψ lim j C.∘ nu .ν ≡ FAlg.ν j C.∘ M.M₁ (is-limit.ψ lim j)) + → make-is-limit F (K .F₀ tt , nu) + make-algebra-limit lim apex-alg comm = em-lim where + module lim = is-limit lim + open make-is-limit + module apex = Algebra-on apex-alg + open _=>_ ``` -The cone structure on $M_0(L)$ is given by composites $\psi_x \nu$, -which commute because $\psi$ is also a cone structure. More explicitly, -what we must show is $F_1(o) \psi_x \nu = \psi_y \nu$, which follows -immediately. +The assumptions to this lemma are actually rather hefty: they pretty +much directly say that $K$ was already the limit of $F$. However, with +this more "elementary" rephrasing, we gain a bit of extra control which +will be necessary for _constructing_ limits in $\cC^M$ out of limits in +$\cC$ later. ```agda - where - cone′ : Cone (Forget C M F∘ F) - cone′ .apex = M.M₀ (apex other .fst) - cone′ .ψ x = morphism (ψ other x) C.∘ apex other .snd .ν - cone′ .commutes o = C.pulll (ap morphism (commutes other o)) - - ! .commutes o = Algebra-hom-path _ - (uniq (F-map-cone (Forget C M) other) .centre .commutes o) + em-lim : make-is-limit F _ + em-lim .ψ j .morphism = lim.ψ j + em-lim .ψ j .commutes = comm j + em-lim .commutes f = Algebra-hom-path C (lim.commutes f) + em-lim .universal eta p .morphism = + lim.universal (λ j → eta j .morphism) (λ f i → p f i .morphism) + em-lim .factors eta p = + Algebra-hom-path C (lim.factors _ _) + em-lim .unique eta p other q = + Algebra-hom-path C (lim.unique _ _ _ λ j i → q j i .morphism) + em-lim .universal eta p .commutes = lim.unique₂ _ + (λ f → C.pulll (F.F₁ f .commutes) + ∙ C.pullr (sym (M.M-∘ _ _) ∙ ap M.M₁ (ap morphism (p f)))) + (λ j → C.pulll (lim.factors _ _) + ∙ eta j .commutes) + (λ j → C.pulll (comm j) + ∙ C.pullr (sym (M.M-∘ _ _) ∙ ap M.M₁ (lim.factors _ _))) ``` -For uniqueness, we use that the map $U(L) \to U(K)$ is unique, and that -the functor $U$ is faithful. +This key lemma also doubles as a proof that the underlying object +functor $U$ reflects limits: We already had an algebra structure +"upstairs"! ```agda - unique : ∀ x → ! ≡ x - unique x = Cone-hom-path _ $ Algebra-hom-path _ $ - ap hom (uniq (F-map-cone (Forget _ M) other) .paths hom′) - where - hom′ : Cone-hom _ _ _ - hom′ .hom = hom x .morphism - hom′ .commutes o = ap morphism (x .commutes o) + Forget-reflects-limits : reflects-limit (Forget C M) F + Forget-reflects-limits {K} {eps} lim = to-is-limitp + (make-algebra-limit lim (K .F₀ tt .snd) (λ j → eps .η j .commutes)) + (Algebra-hom-path C refl) ``` -I hope you like appealing to uniqueness of maps into limits, by the way. -We now relax the conditions on the theorem above, which relies on the -pre-existence of a cone $K$. In fact, what we have shown is that -`Forget` reflects the property of _being a limit_ --- what we now show -is that it reflects limit _objects_, too: if $U \circ F$ has a limit, -then so does $F$. +Having shown that $U$ reflects the property of _being a limit_, we now +turn to showing it reflects limits in general. Using our key lemma, it +will suffice to construct an algebra structure on $\lim_j UF(j)$, then +show that the projection maps $\psi_j : (\lim_j UF(j)) \to UF(j)$ extend +to algebra homomorphisms. ```agda - Forget-lift-limit : Limit (Forget _ M F∘ F) → Limit F - Forget-lift-limit lim-over = - record { top = cone′ - ; has⊤ = Forget-reflects-limits cone′ $ - subst (is-limit _) (sym U$L≡L) (lim-over .has⊤) - } + Forget-lift-limit : Limit (Forget C M F∘ F) → Limit F + Forget-lift-limit lim-over = to-limit $ to-is-limit $ make-algebra-limit + (Limit.has-limit lim-over) apex-algebra (λ j → lim-over.factors _ _) where - open Terminal - module cone = Cone (lim-over .top) + module lim-over = Limit lim-over + module lim = Ran lim-over ``` -What we must do, essentially, is prove that $\lim (U \circ F)$ admits an -algebra structure, much like we did for products of groups. In this, -we'll use two auxiliary cones over $U \circ F$, one with underlying -object given by $M(\lim (U \circ F))$ and one by $M^2(\lim (U \circ -F))$. We construct the one with a single $M$ first, and re-use its maps -in the construction of the one with $M^2$. - -The maps out of $M_0(\lim (U \circ F))$ are given by the composite -below, which assembles into a cone since $F_1(f)$ is a map of algebras -and $\psi$ is a cone. +An algebra structure consists, centrally, of a map $M(\lim_j UF(j)) \to +\lim_j UF(j)$: a map _into_ the limit. Maps into limits are the happy +case, since $\lim_j UF(j)$ is essentially defined by a (natural) +isomorphism between the sets $\hom(X, \lim_j UF(j))$ and $\lim_j \hom(X, +UF(j))$, the latter limit being computed in $\Sets$. We understand +limits of sets very well: they're (subsets of) sets of tuples! -$$ -M_0 (\lim (U \circ F)) \xto{M_1(\psi_x)} M_0 (F_0(x)) \xto{\nu} F_0(x) -$$ +Our algebra map is thus defined _componentwise_: Since we have a bunch +of maps $\nu_j : M(UF(j)) \to UF(j)$, coming from the algebra structures +on each $F(j)$, we can "tuple" them into a big map $\nu = \langle \nu_j +\psi_j \rangle _j$. Abusing notation slightly, we're defining $\nu(a, b, +...)$ to be $(\nu_a(a), \nu_b(b), ...)$. ```agda - cone₂ : Cone (Forget _ M F∘ F) - cone₂ .apex = M.M₀ cone.apex - cone₂ .ψ x = F.₀ x .snd .ν C.∘ M.M₁ (cone.ψ x) - cone₂ .commutes {x} {y} f = - F.₁ f .morphism C.∘ F.₀ x .snd .ν C.∘ M.M₁ (cone.ψ x) ≡⟨ C.pulll (F.₁ f .commutes) ⟩ - (F.₀ y .snd .ν C.∘ M.M₁ (F.₁ f .morphism)) C.∘ M.M₁ (cone.ψ x) ≡⟨ C.pullr (sym (M.M-∘ _ _) ∙ ap M.M₁ (cone.commutes f)) ⟩ - F.₀ y .snd .ν C.∘ M.M₁ (cone.ψ y) ∎ + apex-algebra : Algebra-on C M lim-over.apex + apex-algebra .ν = + lim-over.universal (λ j → FAlg.ν j C.∘ M.M₁ (lim-over.ψ j)) comm where abstract + comm : ∀ {x y} (f : J.Hom x y) + → F.₁ f .morphism C.∘ FAlg.ν x C.∘ M.M₁ (lim-over.ψ x) + ≡ FAlg.ν y C.∘ M.M₁ (lim-over.ψ y) + comm {x} {y} f = + F.₁ f .morphism C.∘ FAlg.ν x C.∘ M.M₁ (lim-over.ψ x) ≡⟨ C.extendl (F.₁ f .commutes) ⟩ + FAlg.ν y C.∘ M.M₁ (F.₁ f .morphism) C.∘ M.M₁ (lim-over.ψ x) ≡˘⟨ C.refl⟩∘⟨ M.M-∘ _ _ ⟩ + FAlg.ν y C.∘ M.M₁ (F.₁ f .morphism C.∘ lim-over.ψ x) ≡⟨ C.refl⟩∘⟨ ap M.M₁ (lim-over.commutes f) ⟩ + FAlg.ν y C.∘ M.M₁ (lim-over.ψ y) ∎ ``` -Below, we can reuse the work we did above by precomposing with $M$'s -multiplication $\mu$. +To show that $\nu$ really is an algebra structure, we'll reason +componentwise, too: we must show that $\nu(\eta_a, \eta_b, ...)$ is +the identity map: but we can compute -```agda - cone² : Cone (Forget _ M F∘ F) - cone² .apex = M.M₀ (M.M₀ cone.apex) - cone² .ψ x = cone₂ .ψ x C.∘ M.mult.η _ - cone² .commutes f = C.pulll (cone₂ .commutes f) -``` - -We now define the algebra structure on $\lim (U \circ F)$. It's very -tedious, but the multiplication is uniquely defined since it's a map -$M(\lim (U \circ F)) \to \lim (U \circ F)$ into a limit, and the -algebraic identities follow from again from limits being terminal. - -```agda - cone′ : Cone F - cone′ .apex = cone.apex , alg where -``` +$$ +\nu(\eta_a, \eta_b, ...) = (\nu_a\eta_a, \nu_b\eta_b, ...) = (\id, \id, ...) = \id\text{!} +$$ - +
+ +The other condition, compatibility with $M$'s multiplication, is +analogous. Formally, the computation is a bit less pretty, but it's no +more complicated. + ```agda - alg : Algebra-on _ M cone.apex - alg .ν = lim-over .has⊤ cone₂ .centre .hom - alg .ν-unit = ap hom $ - is-contr→is-prop (lim-over .has⊤ (lim-over .top)) - (record { hom = alg .ν C.∘ M.unit.η cone.apex ; commutes = comm1 }) - (record { hom = C.id ; commutes = λ _ → C.idr _ }) - - alg .ν-mult = ap hom $ is-contr→is-prop (lim-over .has⊤ cone²) - (record { hom = alg .ν C.∘ M.M₁ (alg .ν) ; commutes = comm2 }) - (record { hom = alg .ν C.∘ M.mult.η cone.apex - ; commutes = λ o → C.pulll (lim-over .has⊤ cone₂ .centre .commutes o) - }) + apex-algebra .ν-unit = lim-over.unique₂ _ lim-over.commutes + (λ j → C.pulll (lim-over.factors _ _) + ·· C.pullr (sym $ M.unit.is-natural _ _ _) + ·· C.cancell (FAlg.ν-unit j)) + (λ j → C.idr _) + apex-algebra .ν-mult = lim-over.unique₂ _ + (λ f → C.pulll $ C.pulll (F.₁ f .commutes) + ∙ C.pullr (sym (M.M-∘ _ _) ∙ ap M.M₁ (lim-over.commutes f))) + (λ j → C.pulll (lim-over.factors _ _) + ·· C.pullr (sym (M.M-∘ _ _) ∙ ap M.M₁ (lim-over.factors _ _) ∙ M.M-∘ _ _) + ·· C.extendl (FAlg.ν-mult j) + ·· ap (FAlg.ν j C.∘_) (M.mult.is-natural _ _ _) + ·· C.assoc _ _ _) + (λ j → C.pulll (lim-over.factors _ _)) ``` -The cone maps in $\cC^M$ are given by the cone maps we started with ---- specialising again to groups, we're essentially showing that the -projection map $\pi_1 : G \times H \to G$ _between sets_ is actually a -group homomorphism. - -```agda - cone′ .ψ x .morphism = cone.ψ x - cone′ .ψ x .commutes = lim-over .has⊤ cone₂ .centre .commutes x - cone′ .commutes f = Algebra-hom-path _ (cone.commutes f) - - U$L≡L : F-map-cone (Forget _ M) cone′ ≡ lim-over .top - U$L≡L = Cone-path _ refl λ o → refl -``` +
-We conclude by saying that, if $\cC$ is a complete category, then so -is $\cC^M$, with no assumptions on $M$. +Putting our previous results together, we conclude by observing that, if +$\cC$ is a complete category, then so is $\cC^M$, regardless of how +ill-behaved the monad $M$ might be. ```agda Eilenberg-Moore-is-complete : ∀ {a b} → is-complete a b C → is-complete a b (Eilenberg-Moore _ M) -Eilenberg-Moore-is-complete complete F = Forget-lift-limit F (complete _) +Eilenberg-Moore-is-complete complete F = + Forget-lift-limit F (complete _) ``` diff --git a/src/Cat/Diagram/Product.lagda.md b/src/Cat/Diagram/Product.lagda.md index d1519703c..98ef5dfc9 100644 --- a/src/Cat/Diagram/Product.lagda.md +++ b/src/Cat/Diagram/Product.lagda.md @@ -75,6 +75,9 @@ record is-product {A B P} (π₁ : Hom P A) (π₂ : Hom P B) : Type (o ⊔ h) w ⟨⟩∘ : ∀ {Q R} {p1 : Hom Q A} {p2 : Hom Q B} (f : Hom R Q) → ⟨ p1 , p2 ⟩ ∘ f ≡ ⟨ p1 ∘ f , p2 ∘ f ⟩ ⟨⟩∘ f = unique _ (pulll π₁∘factor) (pulll π₂∘factor) + + ⟨⟩-η : ⟨ π₁ , π₂ ⟩ ≡ id + ⟨⟩-η = sym $ unique id (idr _) (idr _) ``` A product of $A$ and $B$ is an explicit choice of product diagram: @@ -90,11 +93,17 @@ record Product (A B : Ob) : Type (o ⊔ h) where open is-product has-is-product public -open Product hiding (⟨_,_⟩ ; π₁ ; π₂ ; ⟨⟩∘) ``` ## Uniqueness + + Products, when they exist, are unique. It's easiest to see this with a diagrammatic argument: If we have product diagrams $A \ot P \to B$ and $A \ot P' \to B$, we can fit them into a "commutative diamond" like the @@ -134,133 +143,130 @@ We construct the map $P \to P'$ as the pairing of the projections from $P$, and symmetrically for $P' \to P$. ```agda -×-Unique : (p1 p2 : Product A B) → apex p1 ≅ apex p2 -×-Unique p1 p2 = make-iso p1→p2 p2→p1 p1→p2→p1 p2→p1→p2 - where - module p1 = Product p1 - module p2 = Product p2 + ×-Unique : (p1 p2 : Product A B) → apex p1 ≅ apex p2 + ×-Unique p1 p2 = make-iso p1→p2 p2→p1 p1→p2→p1 p2→p1→p2 + where + module p1 = Product p1 + module p2 = Product p2 - p1→p2 : Hom (apex p1) (apex p2) - p1→p2 = p2.⟨ p1.π₁ , p1.π₂ ⟩ + p1→p2 : Hom (apex p1) (apex p2) + p1→p2 = p2.⟨ p1.π₁ , p1.π₂ ⟩ - p2→p1 : Hom (apex p2) (apex p1) - p2→p1 = p1.⟨ p2.π₁ , p2.π₂ ⟩ + p2→p1 : Hom (apex p2) (apex p1) + p2→p1 = p1.⟨ p2.π₁ , p2.π₂ ⟩ ``` These are unique because they are maps into products which commute with the projections. ```agda - p1→p2→p1 : p1→p2 ∘ p2→p1 ≡ id - p1→p2→p1 = - p2.unique₂ - (assoc _ _ _ ·· ap (_∘ _) p2.π₁∘factor ·· p1.π₁∘factor) - (assoc _ _ _ ·· ap (_∘ _) p2.π₂∘factor ·· p1.π₂∘factor) - (idr _) (idr _) - - p2→p1→p2 : p2→p1 ∘ p1→p2 ≡ id - p2→p1→p2 = - p1.unique₂ - (assoc _ _ _ ·· ap (_∘ _) p1.π₁∘factor ·· p2.π₁∘factor) - (assoc _ _ _ ·· ap (_∘ _) p1.π₂∘factor ·· p2.π₂∘factor) - (idr _) (idr _) - -is-product-iso - : ∀ {A A′ B B′ P} {π₁ : Hom P A} {π₂ : Hom P B} - {f : Hom A A′} {g : Hom B B′} - → is-invertible f - → is-invertible g - → is-product π₁ π₂ - → is-product (f ∘ π₁) (g ∘ π₂) -is-product-iso f-iso g-iso prod = prod′ where - module fi = is-invertible f-iso - module gi = is-invertible g-iso - - open is-product - prod′ : is-product _ _ - prod′ .⟨_,_⟩ qa qb = prod .⟨_,_⟩ (fi.inv ∘ qa) (gi.inv ∘ qb) - prod′ .π₁∘factor = pullr (prod .π₁∘factor) ∙ cancell fi.invl - prod′ .π₂∘factor = pullr (prod .π₂∘factor) ∙ cancell gi.invl - prod′ .unique other p q = prod .unique other - (sym (ap (_ ∘_) (sym p) ∙ pulll (cancell fi.invr))) - (sym (ap (_ ∘_) (sym q) ∙ pulll (cancell gi.invr))) + p1→p2→p1 : p1→p2 ∘ p2→p1 ≡ id + p1→p2→p1 = + p2.unique₂ + (assoc _ _ _ ·· ap (_∘ _) p2.π₁∘factor ·· p1.π₁∘factor) + (assoc _ _ _ ·· ap (_∘ _) p2.π₂∘factor ·· p1.π₂∘factor) + (idr _) (idr _) + + p2→p1→p2 : p2→p1 ∘ p1→p2 ≡ id + p2→p1→p2 = + p1.unique₂ + (assoc _ _ _ ·· ap (_∘ _) p1.π₁∘factor ·· p2.π₁∘factor) + (assoc _ _ _ ·· ap (_∘ _) p1.π₂∘factor ·· p2.π₂∘factor) + (idr _) (idr _) + + is-product-iso + : ∀ {A A′ B B′ P} {π₁ : Hom P A} {π₂ : Hom P B} + {f : Hom A A′} {g : Hom B B′} + → is-invertible f + → is-invertible g + → is-product π₁ π₂ + → is-product (f ∘ π₁) (g ∘ π₂) + is-product-iso f-iso g-iso prod = prod′ where + module fi = is-invertible f-iso + module gi = is-invertible g-iso + + open is-product + prod′ : is-product _ _ + prod′ .⟨_,_⟩ qa qb = prod .⟨_,_⟩ (fi.inv ∘ qa) (gi.inv ∘ qb) + prod′ .π₁∘factor = pullr (prod .π₁∘factor) ∙ cancell fi.invl + prod′ .π₂∘factor = pullr (prod .π₂∘factor) ∙ cancell gi.invl + prod′ .unique other p q = prod .unique other + (sym (ap (_ ∘_) (sym p) ∙ pulll (cancell fi.invr))) + (sym (ap (_ ∘_) (sym q) ∙ pulll (cancell gi.invr))) ``` -# The product functor - -If $\cC$ admits products of all pairs of objects, then the assignment -$(A, B) \mapsto (A \times B)$ extends to a [bifunctor] $(\cC \times -\cC) \to \cC$. +# Categories with all binary products -[bifunctor]: Cat.Functor.Bifunctor.html +Categories with all binary products are quite common, so we define +a module for working with them. ```agda -module Cartesian (hasprods : ∀ A B → Product A B) where +has-products : Type _ +has-products = ∀ a b → Product a b + +module Binary-products (all-products : has-products) where + + module product {a} {b} = Product (all-products a b) + + open product renaming + (unique to ⟨⟩-unique; π₁∘factor to π₁∘⟨⟩; π₂∘factor to π₂∘⟨⟩) public open Functor - ×-functor : Functor (C ×ᶜ C) C - ×-functor .F₀ (A , B) = hasprods A B .apex - ×-functor .F₁ {a , x} {b , y} (f , g) = - hasprods b y .has-is-product .is-product.⟨_,_⟩ - (f ∘ hasprods a x .Product.π₁) (g ∘ hasprods a x .Product.π₂) - - ×-functor .F-id {a , b} = - unique₂ (hasprods a b) - (hasprods a b .π₁∘factor) - (hasprods a b .π₂∘factor) - id-comm id-comm - - ×-functor .F-∘ {a , b} {c , d} {e , f} x y = - unique₂ (hasprods e f) - (hasprods e f .π₁∘factor) - (hasprods e f .π₂∘factor) - ( pulll (hasprods e f .π₁∘factor) - ·· pullr (hasprods c d .π₁∘factor) - ·· assoc _ _ _) - ( pulll (hasprods e f .π₂∘factor) - ·· pullr (hasprods c d .π₂∘factor) - ·· assoc _ _ _) + infixr 7 _⊗₀_ + infix 50 _⊗₁_ ``` -We refer to a category admitting all binary products as **cartesian**. -When working with products, a Cartesian category is the place to be, -since we can work with the "canonical" product operations --- rather -than requiring different product data for any pair of objects we need a -product for. +We start by defining a "global" product-assigning operation. -Here we extract the data of the "global" product-assigning operation to -separate top-level definitions: +```agda + _⊗₀_ : Ob → Ob → Ob + a ⊗₀ b = product.apex {a} {b} +``` + +This operation extends to a bifunctor $\cC \times \cC \to \cC$. ```agda - _⊗_ : Ob → Ob → Ob - A ⊗ B = F₀ ×-functor (A , B) + _⊗₁_ : ∀ {a b x y} → Hom a x → Hom b y → Hom (a ⊗₀ b) (x ⊗₀ y) + f ⊗₁ g = ⟨ f ∘ π₁ , g ∘ π₂ ⟩ - ⟨_,_⟩ : Hom a b → Hom a c → Hom a (b ⊗ c) - ⟨ f , g ⟩ = hasprods _ _ .has-is-product .is-product.⟨_,_⟩ f g - π₁ : Hom (a ⊗ b) a - π₁ = hasprods _ _ .Product.π₁ + ×-functor : Functor (C ×ᶜ C) C + ×-functor .F₀ (a , b) = a ⊗₀ b + ×-functor .F₁ (f , g) = f ⊗₁ g + ×-functor .F-id = sym $ ⟨⟩-unique id id-comm id-comm + ×-functor .F-∘ (f , g) (h , i) = + sym $ ⟨⟩-unique (f ⊗₁ g ∘ h ⊗₁ i) + (pulll π₁∘⟨⟩ ∙ extendr π₁∘⟨⟩) + (pulll π₂∘⟨⟩ ∙ extendr π₂∘⟨⟩) +``` - π₂ : Hom (a ⊗ b) b - π₂ = hasprods _ _ .Product.π₂ +We also define a handful of common morphisms. - π₁∘⟨⟩ : {f : Hom a b} {g : Hom a c} → π₁ ∘ ⟨ f , g ⟩ ≡ f - π₁∘⟨⟩ = hasprods _ _ .has-is-product .is-product.π₁∘factor +```agda + δ : Hom a (a ⊗₀ a) + δ = ⟨ id , id ⟩ - π₂∘⟨⟩ : {f : Hom a b} {g : Hom a c} → π₂ ∘ ⟨ f , g ⟩ ≡ g - π₂∘⟨⟩ = hasprods _ _ .has-is-product .is-product.π₂∘factor + swap : Hom (a ⊗₀ b) (b ⊗₀ a) + swap = ⟨ π₂ , π₁ ⟩ +``` - ⟨⟩∘ : ∀ {Q R} {p1 : Hom Q a} {p2 : Hom Q b} (f : Hom R Q) - → ⟨ p1 , p2 ⟩ ∘ f ≡ ⟨ p1 ∘ f , p2 ∘ f ⟩ - ⟨⟩∘ f = is-product.⟨⟩∘ (hasprods _ _ .has-is-product) f +## Representability of products - ⟨⟩-unique : ∀ {Q} {p1 : Hom Q A} {p2 : Hom Q B} - → (other : Hom Q (A ⊗ B)) - → π₁ ∘ other ≡ p1 - → π₂ ∘ other ≡ p2 - → other ≡ ⟨ p1 , p2 ⟩ - ⟨⟩-unique = hasprods _ _ .has-is-product .is-product.unique +The collection of maps into a product $a \times b$ is equivalent to +the collection of pairs of maps into $a$ and $b$. The forward direction +of the equivalence is given by postcomposition of the projections, and +the reverse direction by the universal property of products. - ⟨⟩-η : ∀ {A B} → ⟨ π₁ , π₂ ⟩ ≡ id {A ⊗ B} - ⟨⟩-η = sym $ ⟨⟩-unique id (idr π₁) (idr π₂) +```agda +product-repr + : ∀ {a b} + → (prod : Product a b) + → (x : Ob) + → Hom x (Product.apex prod) ≃ (Hom x a × Hom x b) +product-repr prod x = Iso→Equiv λ where + .fst f → π₁ ∘ f , π₂ ∘ f + .snd .is-iso.inv (f , g) → ⟨ f , g ⟩ + .snd .is-iso.rinv (f , g) → π₁∘factor ,ₚ π₂∘factor + .snd .is-iso.linv f → sym (⟨⟩∘ f) ∙ eliml ⟨⟩-η + where open Product prod ``` diff --git a/src/Cat/Diagram/Product/Finite.lagda.md b/src/Cat/Diagram/Product/Finite.lagda.md index 0ed4673a8..26e9aa75b 100644 --- a/src/Cat/Diagram/Product/Finite.lagda.md +++ b/src/Cat/Diagram/Product/Finite.lagda.md @@ -26,8 +26,8 @@ open is-product open Terminal open Product open Cr C -private module Cart = Cartesian C products -open Cart using (_⊗_) +private module Cart = Binary-products C products +open Cart using (_⊗₀_) ``` --> @@ -56,7 +56,7 @@ Cartesian→standard-finite-products F = prod where F-apex : ∀ {n} → (F : Fin n → Ob) → Ob F-apex {zero} F = terminal .top F-apex {suc zero} F = F fzero - F-apex {suc (suc n)} F = F fzero ⊗ F-apex (λ e → F (fsuc e)) + F-apex {suc (suc n)} F = F fzero ⊗₀ F-apex (λ e → F (fsuc e)) F-pi : ∀ {n} (F : Fin n → Ob) (i : Fin n) → Hom (F-apex F) (F i) F-pi {suc zero} F fzero = id diff --git a/src/Cat/Diagram/Product/Indexed.lagda.md b/src/Cat/Diagram/Product/Indexed.lagda.md index ac37e8542..c48618963 100644 --- a/src/Cat/Diagram/Product/Indexed.lagda.md +++ b/src/Cat/Diagram/Product/Indexed.lagda.md @@ -1,7 +1,8 @@ ```agda +open import Cat.Instances.Shape.Terminal open import Cat.Diagram.Limit.Base open import Cat.Instances.Discrete -open import Cat.Diagram.Terminal +open import Cat.Functor.Kan.Base open import Cat.Prelude module Cat.Diagram.Product.Indexed {o ℓ} (C : Precategory o ℓ) where @@ -96,56 +97,65 @@ sets, and so the `Disc`{.Agda} construction does not apply to it. ```agda module _ {I : Type ℓ'} (i-is-grpd : is-groupoid I) (F : I → C.Ob) where - open Cone-hom - open Terminal - open Cone + open _=>_ + + Proj→Cone : ∀ {x} → (∀ i → C.Hom x (F i)) + → Const x => Disc-adjunct {C = C} {iss = i-is-grpd} F + Proj→Cone π .η i = π i + Proj→Cone π .is-natural i j p = + J (λ j p → π j C.∘ C.id ≡ subst (C.Hom (F i) ⊙ F) p C.id C.∘ π i) + (C.idr _ ∙ C.introl (transport-refl C.id)) + p + + is-indexed-product→is-limit + : ∀ {x} {π : ∀ i → C.Hom x (F i)} + → is-indexed-product F π + → is-limit (Disc-adjunct F) x (Proj→Cone π) + is-indexed-product→is-limit {x = x} {π} ip = + to-is-limitp ml refl + where + module ip = is-indexed-product ip + open make-is-limit + + ml : make-is-limit (Disc-adjunct F) x + ml .ψ j = π j + ml .commutes {i} {j} p = + J (λ j p → subst (C.Hom (F i) ⊙ F) p C.id C.∘ π i ≡ π j) + (C.eliml (transport-refl _)) + p + ml .universal eta p = ip.tuple eta + ml .factors eta p = ip.commute + ml .unique eta p other q = ip.unique eta q + + is-limit→is-indexed-product + : ∀ {K : Functor ⊤Cat C} + → {eta : K F∘ !F => Disc-adjunct {iss = i-is-grpd} F} + → is-ran !F (Disc-adjunct F) K eta + → is-indexed-product F (eta .η) + is-limit→is-indexed-product {K = K} {eta} lim = ip where + module lim = is-limit lim + open is-indexed-product hiding (eta) + + ip : is-indexed-product F (eta .η) + ip .tuple k = + lim.universal k + (J (λ j p → subst (C.Hom (F _) ⊙ F) p C.id C.∘ k _ ≡ k j) + (C.eliml (transport-refl _))) + ip .commute = + lim.factors _ _ + ip .unique k comm = + lim.unique _ _ _ comm IP→Limit : Indexed-product F → Limit {C = C} (Disc-adjunct {iss = i-is-grpd} F) - IP→Limit IP = lim where - module IP = Indexed-product IP - - thelim : Cone _ - thelim .apex = IP.ΠF - thelim .ψ = IP.π - thelim .commutes {x} = - J (λ y p → subst (C.Hom (F x) ⊙ F) p C.id C.∘ (IP.π x) ≡ IP.π y) - (C.eliml (transport-refl _)) - - lim : Limit _ - lim .top = thelim - lim .has⊤ x .centre .hom = IP.tuple (x .ψ) - lim .has⊤ x .centre .commutes o = IP.commute - lim .has⊤ x .paths h = Cone-hom-path _ (sym (IP.unique _ λ i → h .commutes _)) - -module _ {I : Type ℓ'} (isg : is-groupoid I) (F : Functor (Disc I isg) C) where - private module F = Functor F - open is-indexed-product - open Indexed-product - open Cone-hom - open Terminal - open Cone - - Proj→Cone : ∀ {Y} → (∀ i → C.Hom Y (F.₀ i)) → Cone F - Proj→Cone f .apex = _ - Proj→Cone f .ψ = f - Proj→Cone f .commutes {x} = - J (λ y p → F.₁ p C.∘ f x ≡ f y) (C.eliml F.F-id) - - Limit→IP : Limit {C = C} F → Indexed-product F.₀ - Limit→IP lim = the-ip where - module lim = Cone (lim .top) - - the-ip : Indexed-product _ - the-ip .ΠF = lim.apex - the-ip .π = lim.ψ - the-ip .has-is-ip .tuple f = lim .has⊤ (Proj→Cone f) .centre .hom - the-ip .has-is-ip .commute = lim .has⊤ (Proj→Cone _) .centre .commutes _ - the-ip .has-is-ip .unique {h = h} f p i = - lim .has⊤ (Proj→Cone f) .paths other (~ i) .hom - where - other : Cone-hom _ _ _ - other .hom = h - other .commutes i = p i + IP→Limit ip = + to-limit (is-indexed-product→is-limit has-is-ip) + where open Indexed-product ip + + Limit→IP : Limit {C = C} (Disc-adjunct {iss = i-is-grpd} F) → Indexed-product F + Limit→IP lim .Indexed-product.ΠF = _ + Limit→IP lim .Indexed-product.π = _ + Limit→IP lim .Indexed-product.has-is-ip = + is-limit→is-indexed-product (Limit.has-limit lim) ``` ## Uniqueness diff --git a/src/Cat/Diagram/Product/Solver.lagda.md b/src/Cat/Diagram/Product/Solver.lagda.md index 7c8592d15..020032d62 100644 --- a/src/Cat/Diagram/Product/Solver.lagda.md +++ b/src/Cat/Diagram/Product/Solver.lagda.md @@ -26,7 +26,7 @@ module NbE {o ℓ} (𝒞 : Precategory o ℓ) (cartesian : ∀ A B → Product -- Instead, what we want to do is perform type-directe open Cat.Reasoning 𝒞 - open Cartesian 𝒞 cartesian + open Binary-products 𝒞 cartesian ``` ## Expressions @@ -42,7 +42,7 @@ with un-quotiented syntax. ‶_‶ : Ob → ‶Ob‶ ⟦_⟧ₒ : ‶Ob‶ → Ob - ⟦ X ‶⊗‶ Y ⟧ₒ = ⟦ X ⟧ₒ ⊗ ⟦ Y ⟧ₒ + ⟦ X ‶⊗‶ Y ⟧ₒ = ⟦ X ⟧ₒ ⊗₀ ⟦ Y ⟧ₒ ⟦ ‶ X ‶ ⟧ₒ = X data Expr : ‶Ob‶ → ‶Ob‶ → Type (o ⊔ ℓ) where @@ -554,10 +554,10 @@ Wow, that was a lot of hard work! Let's marvel at the fruits of our labor. ```agda private module Tests {o ℓ} (𝒞 : Precategory o ℓ) (cartesian : ∀ X Y → Product 𝒞 X Y) where open Precategory 𝒞 - open Cartesian 𝒞 cartesian + open Binary-products 𝒞 cartesian open NbE 𝒞 cartesian - test-η : ∀ {X Y Z} → (f : Hom X (Y ⊗ Z)) + test-η : ∀ {X Y Z} → (f : Hom X (Y ⊗₀ Z)) → f ≡ ⟨ π₁ ∘ f , π₂ ∘ f ⟩ test-η f = products! 𝒞 cartesian @@ -579,7 +579,7 @@ private module Tests {o ℓ} (𝒞 : Precategory o ℓ) (cartesian : ∀ X Y → test-nested {W} {X} {Y} {Z} f g h = products! 𝒞 cartesian - test-big : ∀ {W X Y Z} → (f : Hom (W ⊗ X) (W ⊗ Y)) → (g : Hom (W ⊗ X) Z) + test-big : ∀ {W X Y Z} → (f : Hom (W ⊗₀ X) (W ⊗₀ Y)) → (g : Hom (W ⊗₀ X) Z) → (π₁ ∘ ⟨ f , g ⟩) ∘ id ≡ id ∘ ⟨ π₁ , π₂ ⟩ ∘ f test-big f g = products! 𝒞 cartesian ``` diff --git a/src/Cat/Diagram/Pullback.lagda.md b/src/Cat/Diagram/Pullback.lagda.md index eb3fd9eb9..27863f8ea 100644 --- a/src/Cat/Diagram/Pullback.lagda.md +++ b/src/Cat/Diagram/Pullback.lagda.md @@ -53,15 +53,15 @@ overall square has to commute. ~~~ ```agda - limiting : ∀ {P′} {p₁' : Hom P′ X} {p₂' : Hom P′ Y} + universal : ∀ {P′} {p₁' : Hom P′ X} {p₂' : Hom P′ Y} → f ∘ p₁' ≡ g ∘ p₂' → Hom P′ P - p₁∘limiting : {p : f ∘ p₁' ≡ g ∘ p₂'} → p₁ ∘ limiting p ≡ p₁' - p₂∘limiting : {p : f ∘ p₁' ≡ g ∘ p₂'} → p₂ ∘ limiting p ≡ p₂' + p₁∘universal : {p : f ∘ p₁' ≡ g ∘ p₂'} → p₁ ∘ universal p ≡ p₁' + p₂∘universal : {p : f ∘ p₁' ≡ g ∘ p₂'} → p₂ ∘ universal p ≡ p₂' unique : {p : f ∘ p₁' ≡ g ∘ p₂'} {lim' : Hom P′ P} → p₁ ∘ lim' ≡ p₁' → p₂ ∘ lim' ≡ p₂' - → lim' ≡ limiting p + → lim' ≡ universal p unique₂ : {p : f ∘ p₁' ≡ g ∘ p₂'} {lim' lim'' : Hom P′ P} @@ -109,9 +109,20 @@ record Pullback {X Y Z} (f : Hom X Z) (g : Hom Y Z) : Type (ℓ ⊔ ℓ′) wher open is-pullback has-is-pb public ``` - diff --git a/src/Cat/Diagram/Pullback/Properties.lagda.md b/src/Cat/Diagram/Pullback/Properties.lagda.md index cde29caff..13ba525a1 100644 --- a/src/Cat/Diagram/Pullback/Properties.lagda.md +++ b/src/Cat/Diagram/Pullback/Properties.lagda.md @@ -118,16 +118,16 @@ We thus have an induced map $x \to a$, which, since $a$ is a pullback, makes everything in sight commute, and does so uniquely. ```agda - pb .limiting {p₁' = P→b} {p₂' = P→d} p = - outer.limiting {p₁' = b→c ∘ P→b} {p₂' = P→d} (path p) + pb .universal {p₁' = P→b} {p₂' = P→d} p = + outer.universal {p₁' = b→c ∘ P→b} {p₂' = P→d} (path p) - pb .p₁∘limiting {p₁' = P→b} {p₂' = P→d} {p = p} = + pb .p₁∘universal {p₁' = P→b} {p₂' = P→d} {p = p} = right.unique₂ {p = pulll right.square ∙ pullr p} - (assoc _ _ _ ∙ outer.p₁∘limiting) - (pulll square ∙ pullr outer.p₂∘limiting) + (assoc _ _ _ ∙ outer.p₁∘universal) + (pulll square ∙ pullr outer.p₂∘universal) refl p - pb .p₂∘limiting {p₁' = P→b} {p₂' = P→d} {p = p} = outer.p₂∘limiting + pb .p₂∘universal {p₁' = P→b} {p₂' = P→d} {p = p} = outer.p₂∘universal pb .unique {p = p} q r = outer.unique (sym (ap (_ ∘_) (sym q) ∙ assoc _ _ _)) r @@ -151,11 +151,11 @@ then have a map $x \to a$, as we wanted. pb .is-pullback.square = c→f ∘ b→c ∘ a→b ≡⟨ square ⟩ (e→f ∘ d→e) ∘ a→d ∎ - pb .limiting {p₁' = P→c} {p₂' = P→d} x = - left.limiting {p₁' = right.limiting (x ∙ sym (assoc _ _ _))} {p₂' = P→d} - right.p₂∘limiting - pb .p₁∘limiting = pullr left.p₁∘limiting ∙ right.p₁∘limiting - pb .p₂∘limiting = left.p₂∘limiting + pb .universal {p₁' = P→c} {p₂' = P→d} x = + left.universal {p₁' = right.universal (x ∙ sym (assoc _ _ _))} {p₂' = P→d} + right.p₂∘universal + pb .p₁∘universal = pullr left.p₁∘universal ∙ right.p₁∘universal + pb .p₂∘universal = left.p₂∘universal pb .unique {p₁' = P→c} {P→d} {p = p} {lim'} q r = left.unique (right.unique (assoc _ _ _ ∙ q) s) r where @@ -187,13 +187,13 @@ is a monomorphism iff. the square below is a pullback. module _ {a b} {f : Hom a b} where is-monic→is-pullback : is-monic f → is-pullback id f id f is-monic→is-pullback mono .square = refl - is-monic→is-pullback mono .limiting {p₁' = p₁'} p = p₁' - is-monic→is-pullback mono .p₁∘limiting = idl _ - is-monic→is-pullback mono .p₂∘limiting {p = p} = idl _ ∙ mono _ _ p + is-monic→is-pullback mono .universal {p₁' = p₁'} p = p₁' + is-monic→is-pullback mono .p₁∘universal = idl _ + is-monic→is-pullback mono .p₂∘universal {p = p} = idl _ ∙ mono _ _ p is-monic→is-pullback mono .unique p q = introl refl ∙ p is-pullback→is-monic : is-pullback id f id f → is-monic f - is-pullback→is-monic pb f g p = sym (pb .p₁∘limiting {p = p}) ∙ pb .p₂∘limiting + is-pullback→is-monic pb f g p = sym (pb .p₁∘universal {p = p}) ∙ pb .p₂∘universal ``` Pullbacks additionally preserve monomorphisms, as shown below: @@ -228,9 +228,9 @@ rotate-pullback → is-pullback p1 f p2 g → is-pullback p2 g p1 f rotate-pullback pb .square = sym (pb .square) -rotate-pullback pb .limiting p = pb .limiting (sym p) -rotate-pullback pb .p₁∘limiting = pb .p₂∘limiting -rotate-pullback pb .p₂∘limiting = pb .p₁∘limiting +rotate-pullback pb .universal p = pb .universal (sym p) +rotate-pullback pb .p₁∘universal = pb .p₂∘universal +rotate-pullback pb .p₂∘universal = pb .p₁∘universal rotate-pullback pb .unique p q = pb .unique q p is-pullback-iso @@ -242,9 +242,9 @@ is-pullback-iso {f = f} {g} {p1} {p2} i pb = pb′ where module i = _≅_ i pb′ : is-pullback _ _ _ _ pb′ .square = extendl (pb .square) - pb′ .limiting p = i.to ∘ pb .limiting p - pb′ .p₁∘limiting = cancel-inner i.invr ∙ pb .p₁∘limiting - pb′ .p₂∘limiting = cancel-inner i.invr ∙ pb .p₂∘limiting + pb′ .universal p = i.to ∘ pb .universal p + pb′ .p₁∘universal = cancel-inner i.invr ∙ pb .p₁∘universal + pb′ .p₂∘universal = cancel-inner i.invr ∙ pb .p₂∘universal pb′ .unique p q = invertible→monic (iso→invertible (i Iso⁻¹)) _ _ $ sym $ cancell i.invr ∙ sym (pb .unique (assoc _ _ _ ∙ p) (assoc _ _ _ ∙ q)) @@ -257,15 +257,15 @@ pullback-unique pullback-unique {f = f} {g} {p1} {p2} {p1′} {p2′} pb pb′ = make-iso pb→pb′ pb′→pb il ir where - pb→pb′ = pb′ .limiting (pb .square) - pb′→pb = pb .limiting (pb′ .square) + pb→pb′ = pb′ .universal (pb .square) + pb′→pb = pb .universal (pb′ .square) il = unique₂ pb′ {p = pb′ .square} - (pulll (pb′ .p₁∘limiting) ∙ pb .p₁∘limiting) - (pulll (pb′ .p₂∘limiting) ∙ pb .p₂∘limiting) + (pulll (pb′ .p₁∘universal) ∙ pb .p₁∘universal) + (pulll (pb′ .p₂∘universal) ∙ pb .p₂∘universal) (idr _) (idr _) ir = unique₂ pb {p = pb .square} - (pulll (pb .p₁∘limiting) ∙ pb′ .p₁∘limiting) - (pulll (pb .p₂∘limiting) ∙ pb′ .p₂∘limiting) + (pulll (pb .p₁∘universal) ∙ pb′ .p₁∘universal) + (pulll (pb .p₂∘universal) ∙ pb′ .p₂∘universal) (idr _) (idr _) Pullback-unique @@ -280,17 +280,17 @@ Pullback-unique {x = X} {Y} {Z} {f} {g} c-cat x y = p where abstract p1s : PathP (λ i → Hom (apices i) X) x.p₁ y.p₁ - p1s = Univalent.Hom-pathp-refll-iso c-cat (x.p₁∘limiting) + p1s = Univalent.Hom-pathp-refll-iso c-cat (x.p₁∘universal) p2s : PathP (λ i → Hom (apices i) Y) x.p₂ y.p₂ - p2s = Univalent.Hom-pathp-refll-iso c-cat (x.p₂∘limiting) + p2s = Univalent.Hom-pathp-refll-iso c-cat (x.p₂∘universal) lims : ∀ {P′} {p1′ : Hom P′ X} {p2′ : Hom P′ Y} (p : f ∘ p1′ ≡ g ∘ p2′) - → PathP (λ i → Hom P′ (apices i)) (x.limiting p) (y.limiting p) + → PathP (λ i → Hom P′ (apices i)) (x.universal p) (y.universal p) lims p = Univalent.Hom-pathp-reflr-iso c-cat $ - y.unique (pulll y.p₁∘limiting ∙ x.p₁∘limiting) - (pulll y.p₂∘limiting ∙ x.p₂∘limiting) + y.unique (pulll y.p₁∘universal ∙ x.p₁∘universal) + (pulll y.p₂∘universal ∙ x.p₂∘universal) p : x ≡ y p i .apex = apices i @@ -299,13 +299,13 @@ Pullback-unique {x = X} {Y} {Z} {f} {g} c-cat x y = p where p i .has-is-pb .square = is-prop→pathp (λ i → Hom-set (apices i) Z (f ∘ p1s i) (g ∘ p2s i)) x.square y.square i - p i .has-is-pb .limiting p = lims p i - p i .has-is-pb .p₁∘limiting {p = p} = + p i .has-is-pb .universal p = lims p i + p i .has-is-pb .p₁∘universal {p = p} = is-prop→pathp (λ i → Hom-set _ X (p1s i ∘ lims p i) _) - x.p₁∘limiting y.p₁∘limiting i - p i .has-is-pb .p₂∘limiting {p = p} = + x.p₁∘universal y.p₁∘universal i + p i .has-is-pb .p₂∘universal {p = p} = is-prop→pathp (λ i → Hom-set _ _ (p2s i ∘ lims p i) _) - x.p₂∘limiting y.p₂∘limiting i + x.p₂∘universal y.p₂∘universal i p i .has-is-pb .unique {P′ = P′} {p₁' = p₁′} {p₂' = p₂′} {p = p′} {lim' = lim′} = is-prop→pathp (λ i → Π-is-hlevel {A = Hom P′ (apices i)} 1 diff --git a/src/Cat/Diagram/Pushout.lagda.md b/src/Cat/Diagram/Pushout.lagda.md index a2f714392..6df511617 100644 --- a/src/Cat/Diagram/Pushout.lagda.md +++ b/src/Cat/Diagram/Pushout.lagda.md @@ -56,15 +56,15 @@ The universal property ensures that we only perform the minimal number of identifications required to make the aforementioned square commute. ```agda - colimiting : ∀ {Q} {i₁′ : Hom Y Q} {i₂′ : Hom Z Q} + universal : ∀ {Q} {i₁′ : Hom Y Q} {i₂′ : Hom Z Q} → i₁′ ∘ f ≡ i₂′ ∘ g → Hom P Q - i₁∘colimiting : {p : i₁′ ∘ f ≡ i₂′ ∘ g} → colimiting p ∘ i₁ ≡ i₁′ - i₂∘colimiting : {p : i₁′ ∘ f ≡ i₂′ ∘ g} → colimiting p ∘ i₂ ≡ i₂′ + i₁∘universal : {p : i₁′ ∘ f ≡ i₂′ ∘ g} → universal p ∘ i₁ ≡ i₁′ + i₂∘universal : {p : i₁′ ∘ f ≡ i₂′ ∘ g} → universal p ∘ i₂ ≡ i₂′ unique : {p : i₁′ ∘ f ≡ i₂′ ∘ g} {colim′ : Hom P Q} → colim′ ∘ i₁ ≡ i₁′ → colim′ ∘ i₂ ≡ i₂′ - → colim′ ≡ colimiting p + → colim′ ≡ universal p unique₂ : {p : i₁′ ∘ f ≡ i₂′ ∘ g} {colim′ colim′′ : Hom P Q} diff --git a/src/Cat/Diagram/Pushout/Properties.lagda.md b/src/Cat/Diagram/Pushout/Properties.lagda.md index 2684a58c4..177e619a2 100644 --- a/src/Cat/Diagram/Pushout/Properties.lagda.md +++ b/src/Cat/Diagram/Pushout/Properties.lagda.md @@ -36,13 +36,13 @@ $f : A \to B$ is an epimorphism iff. the square below is a pushout module _ {a b} {f : Hom a b} where is-epic→is-pushout : is-epic f → is-pushout f id f id is-epic→is-pushout epi .square = refl - is-epic→is-pushout epi .colimiting {i₁′ = i₁′} p = i₁′ - is-epic→is-pushout epi .i₁∘colimiting = idr _ - is-epic→is-pushout epi .i₂∘colimiting {p = p} = idr _ ∙ epi _ _ p + is-epic→is-pushout epi .universal {i₁′ = i₁′} p = i₁′ + is-epic→is-pushout epi .i₁∘universal = idr _ + is-epic→is-pushout epi .i₂∘universal {p = p} = idr _ ∙ epi _ _ p is-epic→is-pushout epi .unique p q = intror refl ∙ p is-pushout→is-epic : is-pushout f id f id → is-epic f - is-pushout→is-epic pb g h p = sym (pb .i₁∘colimiting {p = p}) ∙ pb .i₂∘colimiting + is-pushout→is-epic pb g h p = sym (pb .i₁∘universal {p = p}) ∙ pb .i₂∘universal ``` Pushout additionally preserve epimorphisms, as shown below: diff --git a/src/Cat/Diagram/Zero.lagda.md b/src/Cat/Diagram/Zero.lagda.md index 0cf0db266..91af85a81 100644 --- a/src/Cat/Diagram/Zero.lagda.md +++ b/src/Cat/Diagram/Zero.lagda.md @@ -54,6 +54,12 @@ every hom set is inhabited! zero-∘r : ∀ {x y z} → (f : Hom x y) → zero→ {y} {z} ∘ f ≡ zero→ zero-∘r f = pullr (sym (!-unique (! ∘ f))) + + zero-comm : ∀ {x y z} → (f : Hom y z) → (g : Hom x y) → f ∘ zero→ ≡ zero→ ∘ g + zero-comm f g = zero-∘l f ∙ sym (zero-∘r g) + + zero-comm-sym : ∀ {x y z} → (f : Hom y z) → (g : Hom x y) → zero→ ∘ f ≡ g ∘ zero→ + zero-comm-sym f g = zero-∘r f ∙ sym (zero-∘l g) ``` ## Intuition diff --git a/src/Cat/Displayed/Bifibration.lagda.md b/src/Cat/Displayed/Bifibration.lagda.md index c220095ed..17af46a21 100644 --- a/src/Cat/Displayed/Bifibration.lagda.md +++ b/src/Cat/Displayed/Bifibration.lagda.md @@ -1,17 +1,17 @@ ```agda -open import Cat.Displayed.Base -open import Cat.Displayed.Fibre -open import Cat.Functor.Adjoint open import Cat.Functor.Adjoint.Hom open import Cat.Instances.Functor +open import Cat.Displayed.Fibre +open import Cat.Functor.Adjoint +open import Cat.Displayed.Base open import Cat.Prelude -import Cat.Displayed.Cartesian.Indexing import Cat.Displayed.Cocartesian.Indexing -import Cat.Displayed.Cocartesian +import Cat.Displayed.Cartesian.Indexing import Cat.Displayed.Cocartesian.Weak -import Cat.Displayed.Cartesian import Cat.Displayed.Cartesian.Weak +import Cat.Displayed.Cocartesian +import Cat.Displayed.Cartesian import Cat.Displayed.Reasoning import Cat.Reasoning @@ -35,10 +35,10 @@ open Displayed ℰ # Bifibrations -A displayed category $\cE \liesover \cB$ is a **bifibration** if it both -a fibration and an opfibration. This means that $\cE$ is equipped with -both [reindexing] and [opreindexing] functors, which allows us to both -restrict and extend along morphisms $X \to Y$ in the base. +A displayed category $\cE \liesover \cB$ is a **bifibration** if is it +both a fibration and an opfibration. This means that $\cE$ is equipped +with both [reindexing] and [opreindexing] functors, which allows us to +both restrict and extend along morphisms $X \to Y$ in the base. Note that a bifibration is *not* the same as a "profunctor of categories"; these are called **two-sided fibrations**, and are a distinct concept. @@ -64,9 +64,9 @@ record is-bifibration : Type (o ⊔ ℓ ⊔ o′ ⊔ ℓ′) where # Bifibrations and Adjoints -If $\cE$ is a bifibration, then its opreindexing functors are -[left adjoints] to reindexing functors. To see this, note that we need -to construct a natural isomorphism between $\cE_{y}(u_{*}(-),-)$ and +If $\cE$ is a bifibration, then its opreindexing functors are [left +adjoints] to its reindexing functors. To see this, note that we need to +construct a natural isomorphism between $\cE_{y}(u_{*}(-),-)$ and $\cE_{x}(-,u^{*}(-))$. However, we have already shown that $\cE_{y}(u_{*}(-),-)$ and $\cE_{x}(-,u^{*}(-))$ are both naturally isomorphic to $\cE_{u}(-,-)$ (see `opfibration→hom-iso`{.Agda} and @@ -110,7 +110,7 @@ module _ (fib : Cartesian-fibration) where left-adjoint-base-change→opfibration : (L : ∀ {x y} → (f : Hom x y) → Functor (Fibre ℰ x) (Fibre ℰ y)) - → (∀ {x y} → (f : Hom x y) → (L f ⊣ base-change f)) + → (∀ {x y} → (f : Hom x y) → (L f ⊣ base-change f)) → Cocartesian-fibration left-adjoint-base-change→opfibration L adj = cartesian+weak-opfibration→opfibration fib $ diff --git a/src/Cat/Displayed/Cartesian.lagda.md b/src/Cat/Displayed/Cartesian.lagda.md index b8876109d..eaae98d72 100644 --- a/src/Cat/Displayed/Cartesian.lagda.md +++ b/src/Cat/Displayed/Cartesian.lagda.md @@ -2,9 +2,9 @@ open import Cat.Displayed.Base open import Cat.Prelude -import Cat.Reasoning -import Cat.Displayed.Morphism import Cat.Displayed.Reasoning as DR +import Cat.Displayed.Morphism +import Cat.Reasoning module Cat.Displayed.Cartesian {o ℓ o′ ℓ′} {B : Precategory o ℓ} (E : Displayed B o′ ℓ′) where diff --git a/src/Cat/Displayed/Cartesian/Discrete.lagda.md b/src/Cat/Displayed/Cartesian/Discrete.lagda.md index 606ca144d..3e01663c8 100644 --- a/src/Cat/Displayed/Cartesian/Discrete.lagda.md +++ b/src/Cat/Displayed/Cartesian/Discrete.lagda.md @@ -13,9 +13,9 @@ open import Cat.Displayed.Base open import Cat.Displayed.Path open import Cat.Prelude -import Cat.Reasoning -import Cat.Displayed.Morphism import Cat.Displayed.Reasoning +import Cat.Displayed.Morphism +import Cat.Reasoning module Cat.Displayed.Cartesian.Discrete where ``` diff --git a/src/Cat/Displayed/Cartesian/Identity.lagda.md b/src/Cat/Displayed/Cartesian/Identity.lagda.md index 87cfb255c..dd10eb63f 100644 --- a/src/Cat/Displayed/Cartesian/Identity.lagda.md +++ b/src/Cat/Displayed/Cartesian/Identity.lagda.md @@ -5,8 +5,8 @@ open import Cat.Prelude import Cat.Displayed.Univalence.Reasoning import Cat.Displayed.Univalence -import Cat.Displayed.Morphism import Cat.Displayed.Reasoning as Dr +import Cat.Displayed.Morphism import Cat.Reasoning as Cr module Cat.Displayed.Cartesian.Identity diff --git a/src/Cat/Displayed/Cartesian/Right.lagda.md b/src/Cat/Displayed/Cartesian/Right.lagda.md index c5d673911..d7019fef5 100644 --- a/src/Cat/Displayed/Cartesian/Right.lagda.md +++ b/src/Cat/Displayed/Cartesian/Right.lagda.md @@ -3,15 +3,15 @@ description: | We define right fibrations, and characterize their fibre categories. --- ```agda -open import Cat.Displayed.Base open import Cat.Displayed.Cartesian.Discrete open import Cat.Displayed.Functor +open import Cat.Displayed.Base open import Cat.Prelude -import Cat.Reasoning import Cat.Displayed.Cartesian -import Cat.Displayed.Morphism import Cat.Displayed.Reasoning +import Cat.Displayed.Morphism +import Cat.Reasoning module Cat.Displayed.Cartesian.Right {o ℓ o′ ℓ′} diff --git a/src/Cat/Displayed/Cartesian/Weak.lagda.md b/src/Cat/Displayed/Cartesian/Weak.lagda.md index c2696eb45..82d54aa92 100644 --- a/src/Cat/Displayed/Cartesian/Weak.lagda.md +++ b/src/Cat/Displayed/Cartesian/Weak.lagda.md @@ -1,14 +1,18 @@ ```agda -open import Cat.Displayed.Base -open import Cat.Displayed.Fibre -open import Cat.Functor.Hom open import Cat.Functor.Hom.Displayed +open import Cat.Functor.Hom.Displayed +open import Cat.Instances.Functor open import Cat.Instances.Functor open import Cat.Instances.Product +open import Cat.Instances.Product +open import Cat.Displayed.Fibre +open import Cat.Displayed.Base +open import Cat.Functor.Hom open import Cat.Prelude -import Cat.Displayed.Cartesian as Cart import Cat.Displayed.Cartesian.Indexing as Indexing +import Cat.Displayed.Cartesian.Indexing as Indexing +import Cat.Displayed.Cartesian as Cart import Cat.Displayed.Reasoning as DR import Cat.Displayed.Morphism as DM import Cat.Reasoning as CR @@ -128,12 +132,9 @@ cartesian→weak-cartesian {f = f} {f′ = f′} cart = weak-cart where open is-cartesian cart weak-cart : is-weak-cartesian f f′ - weak-cart .universal g′ = - universalv g′ - weak-cart .commutes g′ = - commutesv g′ - weak-cart .unique h′ p = - uniquev h′ p + weak-cart .universal g′ = universalv g′ + weak-cart .commutes g′ = commutesv g′ + weak-cart .unique h′ p = uniquev h′ p ``` Furthermore, if $\cE$ is a fibration, weakly cartesian morphisms are @@ -296,7 +297,7 @@ module _ where → Cartesian-fibration weak-fibration→fibration weak-fib weak-∘ .has-lift {x = x} f y′ = f-lift where open is-weak-cartesian-fibration weak-fib - + module weak-∘ {x y z} (f : Hom y z) (g : Hom x y) (z′ : Ob[ z ]) = is-weak-cartesian (weak-∘ (weak-lift.weak-cartesian f z′) (weak-lift.weak-cartesian g _)) @@ -320,13 +321,13 @@ cartesian lift $f^{*}$ of $f$. ```agda x* : Ob[ x ] x* = weak-lift.x′ f y′ - + f* : Hom[ f ] x* y′ f* = weak-lift.lifting f y′ - + f*-weak-cartesian : is-weak-cartesian f f* f*-weak-cartesian = weak-lift.weak-cartesian f y′ - + module f* = is-weak-cartesian (f*-weak-cartesian) ``` @@ -364,13 +365,13 @@ to obtain the requisite map. where u* : Ob[ u ] u* = weak-lift.x′ m _ - + m* : Hom[ m ] u* x* m* = weak-lift.lifting m _ - + m*-weak-cartesian : is-weak-cartesian m m* m*-weak-cartesian = weak-lift.weak-cartesian m x* - + module m* = is-weak-cartesian m*-weak-cartesian module f*∘m* = is-weak-cartesian (weak-∘ f*-weak-cartesian m*-weak-cartesian) ``` @@ -392,7 +393,7 @@ $h'$ via $f^{*} \cdot m^{*}$ commutes. f*-cartesian .commutes {u = u} {u′ = u′} m h′ = path where open Morphisms m h′ - + abstract path : f* ∘′ hom[ idr m ] (m* ∘′ f*∘m*.universal h′) ≡ h′ path = @@ -413,7 +414,7 @@ maps. f*-cartesian .unique {u = u} {u′ = u′} {m = m} {h′ = h′} m′ p = path where open Morphisms m h′ - + abstract universal-path : (f* ∘′ m*) ∘′ m*.universal m′ ≡[ idr (f ∘ m) ] h′ universal-path = to-pathp $ @@ -422,7 +423,7 @@ maps. hom[] (f* ∘′ hom[] m′) ≡⟨ smashr _ _ ∙ liberate _ ⟩ f* ∘′ m′ ≡⟨ p ⟩ h′ ∎ - + path : m′ ≡ hom[ idr m ] (m* ∘′ f*∘m*.universal h′) path = m′ ≡˘⟨ from-pathp (m*.commutes m′) ⟩ @@ -531,7 +532,7 @@ between $\cE_{u}(-,y')$ and $\cE_{x}(-,u^{*}(y'))$. u*y′ : Ob[ x ] u*y′ = weak-lift.x′ u y′ - + mi : make-natural-iso (Hom-over-into ℰ u y′) (Hom-into (Fibre ℰ x) u*y′) mi .eta x u′ = weak-lift.universal u y′ u′ mi .inv x v′ = hom[ idr u ] (weak-lift.lifting u y′ ∘′ v′) diff --git a/src/Cat/Displayed/Cocartesian.lagda.md b/src/Cat/Displayed/Cocartesian.lagda.md index 8d07823a2..df954f95a 100644 --- a/src/Cat/Displayed/Cocartesian.lagda.md +++ b/src/Cat/Displayed/Cocartesian.lagda.md @@ -1,13 +1,13 @@ ```agda -open import Cat.Displayed.Base open import Cat.Displayed.Cartesian open import Cat.Displayed.Total.Op +open import Cat.Displayed.Base open import Cat.Prelude -import Cat.Reasoning -import Cat.Displayed.Morphism import Cat.Displayed.Morphism.Duality import Cat.Displayed.Reasoning as DR +import Cat.Displayed.Morphism +import Cat.Reasoning module Cat.Displayed.Cocartesian {o ℓ o′ ℓ′} {ℬ : Precategory o ℓ} (ℰ : Displayed ℬ o′ ℓ′) where @@ -102,9 +102,9 @@ to a unique universal factorisation of $h'$ through a map $b' \to_{m} u'$ → Hom[ m ] b′ u′ universal′ {u′ = u′} p h′ = universal _ (coe1→0 (λ i → Hom[ p i ] a′ u′) h′) - commutesp : ∀ {u u′} {m : Hom b u} {k : Hom a u} + commutesp : ∀ {u u′} {m : Hom b u} {k : Hom a u} → (p : m ∘ f ≡ k) (h′ : Hom[ k ] a′ u′) - → universal′ p h′ ∘′ f′ ≡[ p ] h′ + → universal′ p h′ ∘′ f′ ≡[ p ] h′ commutesp {u′ = u′} p h′ = to-pathp⁻ (commutes _ (coe1→0 (λ i → Hom[ p i ] a′ u′) h′)) diff --git a/src/Cat/Displayed/Cocartesian/Indexing.lagda.md b/src/Cat/Displayed/Cocartesian/Indexing.lagda.md index 555069ca9..47a0e8fc6 100644 --- a/src/Cat/Displayed/Cocartesian/Indexing.lagda.md +++ b/src/Cat/Displayed/Cocartesian/Indexing.lagda.md @@ -1,8 +1,8 @@ ```agda -open import Cat.Displayed.Base open import Cat.Displayed.Cocartesian -open import Cat.Displayed.Fibre open import Cat.Instances.Functor +open import Cat.Displayed.Fibre +open import Cat.Displayed.Base open import Cat.Prelude import Cat.Displayed.Reasoning diff --git a/src/Cat/Displayed/Cocartesian/Weak.lagda.md b/src/Cat/Displayed/Cocartesian/Weak.lagda.md index 1551a24ba..15ba729e3 100644 --- a/src/Cat/Displayed/Cocartesian/Weak.lagda.md +++ b/src/Cat/Displayed/Cocartesian/Weak.lagda.md @@ -1,20 +1,21 @@ ```agda -open import Cat.Displayed.Base -open import Cat.Displayed.Cartesian open import Cat.Displayed.Cartesian.Weak -open import Cat.Displayed.Fibre -open import Cat.Displayed.Total.Op -open import Cat.Functor.Hom open import Cat.Functor.Hom.Displayed +open import Cat.Displayed.Cartesian +open import Cat.Displayed.Total.Op open import Cat.Instances.Functor open import Cat.Instances.Product +open import Cat.Displayed.Fibre +open import Cat.Displayed.Base +open import Cat.Functor.Hom open import Cat.Prelude -import Cat.Displayed.Cocartesian as Cocart import Cat.Displayed.Cocartesian.Indexing as Indexing -import Cat.Displayed.Morphism import Cat.Displayed.Morphism.Duality +import Cat.Displayed.Cocartesian as Cocart +import Cat.Displayed.Cocartesian as Cocart import Cat.Displayed.Reasoning +import Cat.Displayed.Morphism import Cat.Reasoning as CR module Cat.Displayed.Cocartesian.Weak @@ -607,7 +608,7 @@ Furthermore, this equivalence is natural. mi .eta∘inv _ = funext λ v′ → sym $ weak-lift.unique u _ _ (to-pathp refl) mi .inv∘eta _ = funext λ u′ → - from-pathp $ weak-lift.commutes u _ _ + from-pathp $ weak-lift.commutes u _ _ mi .natural _ _ v′ = funext λ u′ → weak-lift.unique _ _ _ $ to-pathp $ smashl _ _ diff --git a/src/Cat/Displayed/Composition.lagda.md b/src/Cat/Displayed/Composition.lagda.md index 8400e0970..89592ab74 100644 --- a/src/Cat/Displayed/Composition.lagda.md +++ b/src/Cat/Displayed/Composition.lagda.md @@ -1,11 +1,9 @@ ```agda -open import Cat.Prelude - -open import Cat.Displayed.Base +open import Cat.Displayed.Cartesian open import Cat.Displayed.Functor open import Cat.Displayed.Total - -open import Cat.Displayed.Cartesian +open import Cat.Displayed.Base +open import Cat.Prelude import Cat.Displayed.Reasoning as DR diff --git a/src/Cat/Displayed/Instances/CT-Structure.lagda.md b/src/Cat/Displayed/Instances/CT-Structure.lagda.md index 483421ba9..af9a937a0 100644 --- a/src/Cat/Displayed/Instances/CT-Structure.lagda.md +++ b/src/Cat/Displayed/Instances/CT-Structure.lagda.md @@ -3,7 +3,7 @@ open import Cat.Displayed.Cartesian.Discrete open import Cat.Diagram.Product.Solver open import Cat.Displayed.Cartesian open import Cat.Displayed.Functor -open import Cat.Diagram.Product renaming (module Cartesian to Products) +open import Cat.Diagram.Product open import Cat.Displayed.Base open import Cat.Prelude @@ -16,7 +16,7 @@ module Cat.Displayed.Instances.CT-Structure where open Precategory B -open Products B has-prods +open Binary-products B has-prods open Simple B has-prods ``` @@ -54,8 +54,8 @@ CT-structure distinguishes as types. ```agda simple-ct : ∀ {s} → CT-Structure s → Displayed B (o ⊔ s) ℓ simple-ct ct .Displayed.Ob[_] Γ = Σ[ X ∈ Ob ] ∣ is-tp ct X ∣ -simple-ct ct .Displayed.Hom[_] {Γ} {Δ} u X Y = Hom (Γ ⊗ X .fst) (Y .fst) -simple-ct ct .Displayed.Hom[_]-set {Γ} {Δ} u X Y = Hom-set (Γ ⊗ X .fst) (Y .fst) +simple-ct ct .Displayed.Hom[_] {Γ} {Δ} u X Y = Hom (Γ ⊗₀ X .fst) (Y .fst) +simple-ct ct .Displayed.Hom[_]-set {Γ} {Δ} u X Y = Hom-set (Γ ⊗₀ X .fst) (Y .fst) simple-ct ct .Displayed.id′ = π₂ simple-ct ct .Displayed._∘′_ {f = u} {g = v} f g = f ∘ ⟨ v ∘ π₁ , g ⟩ simple-ct ct .Displayed.idr′ {f = u} f = diff --git a/src/Cat/Displayed/Instances/Simple.lagda.md b/src/Cat/Displayed/Instances/Simple.lagda.md index d2d93de98..abeda3c76 100644 --- a/src/Cat/Displayed/Instances/Simple.lagda.md +++ b/src/Cat/Displayed/Instances/Simple.lagda.md @@ -1,7 +1,7 @@ ```agda open import Cat.Diagram.Product.Solver open import Cat.Displayed.Cartesian -open import Cat.Diagram.Product renaming (module Cartesian to Products) +open import Cat.Diagram.Product open import Cat.Displayed.Base open import Cat.Prelude @@ -13,7 +13,7 @@ module Cat.Displayed.Instances.Simple where open Precategory B -open Products B has-prods +open Binary-products B has-prods ``` # Simple Fibrations @@ -66,8 +66,8 @@ can be entirely automated. ```agda simple : Displayed B o ℓ simple .Displayed.Ob[_] Γ = Ob -simple .Displayed.Hom[_] {Γ} {Δ} u X Y = Hom (Γ ⊗ X) Y -simple .Displayed.Hom[_]-set _ _ _ = Hom-set (_ ⊗ _) _ +simple .Displayed.Hom[_] {Γ} {Δ} u X Y = Hom (Γ ⊗₀ X) Y +simple .Displayed.Hom[_]-set _ _ _ = Hom-set (_ ⊗₀ _) _ simple .Displayed.id′ = π₂ simple .Displayed._∘′_ {f = u} {g = v} f g = f ∘ ⟨ v ∘ π₁ , g ⟩ simple .Displayed.idr′ f = diff --git a/src/Cat/Displayed/Instances/Slice.lagda.md b/src/Cat/Displayed/Instances/Slice.lagda.md index 36f53d5f4..a5deb3d7e 100644 --- a/src/Cat/Displayed/Instances/Slice.lagda.md +++ b/src/Cat/Displayed/Instances/Slice.lagda.md @@ -1,6 +1,7 @@ ```agda -open import Cat.Displayed.Cartesian open import Cat.Displayed.Cocartesian +open import Cat.Displayed.Cocartesian +open import Cat.Displayed.Cartesian open import Cat.Functor.Equivalence open import Cat.Diagram.Pullback open import Cat.Displayed.Fibre @@ -183,15 +184,15 @@ Codomain-fibration : (∀ {x y z} (f : Hom x y) (g : Hom z y) → Pullback B f g) → Cartesian-fibration Slices Codomain-fibration pullbacks .has-lift f y′ = lift-f where - open Pullback (pullbacks f (y′ .map)) + module pb = Pullback (pullbacks f (y′ .map)) lift-f : Cartesian-lift Slices f y′ - lift-f .x′ = cut p₁ - lift-f .lifting .to = p₂ - lift-f .lifting .commute = square - lift-f .cartesian .universal m h′ .to = limiting (assoc _ _ _ ∙ h′ .commute) - lift-f .cartesian .universal m h′ .commute = sym p₁∘limiting - lift-f .cartesian .commutes m h′ = Slice-pathp refl p₂∘limiting + lift-f .x′ = cut pb.p₁ + lift-f .lifting .to = pb.p₂ + lift-f .lifting .commute = pb.square + lift-f .cartesian .universal m h′ .to = pb.universal (assoc _ _ _ ∙ h′ .commute) + lift-f .cartesian .universal m h′ .commute = sym pb.p₁∘universal + lift-f .cartesian .commutes m h′ = Slice-pathp refl pb.p₂∘universal lift-f .cartesian .unique m′ x = Slice-pathp refl $ Pullback.unique (pullbacks f (y′ .map)) (sym (m′ .commute)) (ap to x) ``` @@ -212,20 +213,20 @@ Codomain-fibration→pullbacks Codomain-fibration→pullbacks f g lifts = pb where open Pullback open is-pullback - the-lift = lifts .has-lift f (cut g) + module the-lift = Cartesian-lift (lifts .has-lift f (cut g)) pb : Pullback B f g - pb .apex = the-lift .x′ .domain - pb .p₁ = the-lift .x′ .map - pb .p₂ = the-lift .lifting .to - pb .has-is-pb .square = the-lift .lifting .commute - pb .has-is-pb .limiting {p₁' = p₁'} {p₂'} p = - the-lift .cartesian .universal {u′ = cut id} + pb .apex = the-lift.x′ .domain + pb .p₁ = the-lift.x′ .map + pb .p₂ = the-lift.lifting .to + pb .has-is-pb .square = the-lift.lifting .commute + pb .has-is-pb .universal {p₁' = p₁'} {p₂'} p = + the-lift.cartesian .universal {u′ = cut id} p₁' (slice-hom p₂' (pullr (idr _) ∙ p)) .to - pb .has-is-pb .p₁∘limiting = - sym (the-lift .cartesian .universal _ _ .commute) ∙ idr _ - pb .has-is-pb .p₂∘limiting = ap to (the-lift .cartesian .commutes _ _) - pb .has-is-pb .unique p q = ap to $ the-lift .cartesian .unique + pb .has-is-pb .p₁∘universal = + sym (the-lift.universal _ _ .commute) ∙ idr _ + pb .has-is-pb .p₂∘universal = ap to (the-lift.cartesian .commutes _ _) + pb .has-is-pb .unique p q = ap to $ the-lift.cartesian .unique (slice-hom _ (idr _ ∙ sym p)) (Slice-pathp refl q) ``` diff --git a/src/Cat/Displayed/Instances/Trivial.lagda.md b/src/Cat/Displayed/Instances/Trivial.lagda.md index f841f2955..07f18b523 100644 --- a/src/Cat/Displayed/Instances/Trivial.lagda.md +++ b/src/Cat/Displayed/Instances/Trivial.lagda.md @@ -1,13 +1,13 @@ ```agda -open import Cat.Displayed.Base +open import Cat.Functor.Equivalence.Path +open import Cat.Instances.Shape.Terminal open import Cat.Displayed.Bifibration -open import Cat.Displayed.Cartesian open import Cat.Displayed.Cocartesian +open import Cat.Displayed.Cartesian +open import Cat.Functor.Equivalence open import Cat.Displayed.Fibre open import Cat.Displayed.Total -open import Cat.Instances.Shape.Terminal -open import Cat.Functor.Equivalence -open import Cat.Functor.Equivalence.Path +open import Cat.Displayed.Base open import Cat.Prelude module Cat.Displayed.Instances.Trivial @@ -29,7 +29,7 @@ Any category $\ca{C}$ can be regarded as being displayed over the terminal category $\top$. ```agda -Trivial : Displayed ⊤Cat o ℓ +Trivial : Displayed ⊤Cat o ℓ Trivial .Displayed.Ob[_] _ = Ob Trivial .Displayed.Hom[_] _ = Hom Trivial .Displayed.Hom[_]-set _ _ _ = Hom-set _ _ @@ -111,4 +111,3 @@ trivial→fibre-iso : is-precat-iso trivial-fibre trivial→fibre-iso .is-precat-iso.has-is-ff = id-equiv trivial→fibre-iso .is-precat-iso.has-is-iso = id-equiv ``` - diff --git a/src/Cat/Displayed/Morphism.lagda.md b/src/Cat/Displayed/Morphism.lagda.md index 94a394ad0..d53bcaa12 100644 --- a/src/Cat/Displayed/Morphism.lagda.md +++ b/src/Cat/Displayed/Morphism.lagda.md @@ -2,8 +2,8 @@ open import Cat.Displayed.Base open import Cat.Prelude -import Cat.Reasoning import Cat.Displayed.Reasoning +import Cat.Reasoning module Cat.Displayed.Morphism {o ℓ o′ ℓ′} @@ -380,7 +380,7 @@ make-iso[ inv ] f′ g′ p q .inverses′ .Inverses[_].invr′ = q make-vertical-iso : ∀ {x} {x′ x″ : Ob[ x ]} → (f′ : Hom[ id ] x′ x″) (g′ : Hom[ id ] x″ x′) - → f′ ∘′ g′ ≡[ idl _ ] id′ + → f′ ∘′ g′ ≡[ idl _ ] id′ → g′ ∘′ f′ ≡[ idl _ ] id′ → x′ ≅↓ x″ make-vertical-iso = make-iso[ id-iso ] diff --git a/src/Cat/Displayed/Morphism/Duality.lagda.md b/src/Cat/Displayed/Morphism/Duality.lagda.md index 6fe206d91..42000a3ad 100644 --- a/src/Cat/Displayed/Morphism/Duality.lagda.md +++ b/src/Cat/Displayed/Morphism/Duality.lagda.md @@ -1,12 +1,12 @@ ```agda -open import Cat.Displayed.Base open import Cat.Displayed.Total.Op +open import Cat.Displayed.Base open import Cat.Prelude -import Cat.Morphism -import Cat.Morphism.Duality -import Cat.Displayed.Morphism import Cat.Displayed.Reasoning +import Cat.Displayed.Morphism +import Cat.Morphism.Duality +import Cat.Morphism module Cat.Displayed.Morphism.Duality {o ℓ o′ ℓ′} @@ -123,7 +123,7 @@ co-section[]→retract[] f .retract′ = co-section[]→retract[] f .is-retract′ = ℰ^op.is-section′ f -retract[]→co-section[] +retract[]→co-section[] : ∀ {r : ℬ.has-retract f} → ℰ.has-retract[ r ] f′ → ℰ^op.has-section[ retract→co-section r ] f′ @@ -285,4 +285,3 @@ vertical-iso→vertical-co-iso f = (cast[] (ℰ.invr′ f)) (cast[] (ℰ.invl′ f)) ``` - diff --git a/src/Cat/Displayed/Total.lagda.md b/src/Cat/Displayed/Total.lagda.md index 828145514..daa87dce2 100644 --- a/src/Cat/Displayed/Total.lagda.md +++ b/src/Cat/Displayed/Total.lagda.md @@ -2,8 +2,8 @@ open import Cat.Displayed.Base open import Cat.Prelude -import Cat.Reasoning as CR import Cat.Displayed.Morphism as DM +import Cat.Reasoning as CR module Cat.Displayed.Total {o ℓ o′ ℓ′} {B : Precategory o ℓ} (E : Displayed B o′ ℓ′) where diff --git a/src/Cat/Displayed/Total/Op.lagda.md b/src/Cat/Displayed/Total/Op.lagda.md index 73c7b9058..80ec5434e 100644 --- a/src/Cat/Displayed/Total/Op.lagda.md +++ b/src/Cat/Displayed/Total/Op.lagda.md @@ -1,13 +1,18 @@ ```agda -open import Cat.Displayed.Base +open import 1Lab.Rewrite + +open import Cat.Functor.Equivalence.Path +open import Cat.Functor.Equivalence +open import Cat.Functor.Equivalence +open import Cat.Displayed.Fibre open import Cat.Displayed.Fibre open import Cat.Displayed.Total -open import Cat.Functor.Equivalence -open import Cat.Functor.Equivalence.Path +open import Cat.Displayed.Total +open import Cat.Displayed.Base +open import Cat.Displayed.Base open import Cat.Prelude -open import 1Lab.Rewrite - +import Cat.Displayed.Reasoning as DR import Cat.Displayed.Reasoning as DR module Cat.Displayed.Total.Op where @@ -162,7 +167,7 @@ fibre-functor-total-op {ℰ = ℰ} F .F-∘ f g = + @@ -46,148 +53,108 @@ then instantiate this theorem to the "canonical" shapes of limit: [pullbacks]: Cat.Diagram.Pullback.html [equalisers]: Cat.Diagram.Equaliser.html -```agda - module _ {od ℓd} {J : Precategory od ℓd} {F : Functor J D} where -``` - - - -## Passing cones along +This follows directly from the fact that [adjoints preserve Kan +extensions]. -The first thing we prove is that, given a cone over a diagram $F$ in -$\cD$, we can get a cone in $\cC$ over $R \circ F$, by passing -both the apex and the morphisms "over" using $R$. In reality, this is -just the canonically-defined action of $R$ on cones over $F$: - -```agda - cone-right-adjoint : Cone F → Cone (R F∘ F) - cone-right-adjoint = F-map-cone R -``` - -Conversely, if we have a cone over $R \circ F$, we can turn that into a -cone for $F$. In this direction, we use $L$ on the apex, but we must -additionally use the `adjunction counit`{.Agda ident=adj.counit.ε} to -"adjust" the cone maps (that's `ψ`{.Agda}). - -```agda - right-adjoint-cone : Cone (R F∘ F) → Cone F - right-adjoint-cone K .apex = L.₀ (K .apex) - right-adjoint-cone K .ψ x = adj.counit.ε _ D.∘ L.₁ (K .ψ x) - right-adjoint-cone K .commutes {x} {y} f = - F.₁ f D.∘ adj.counit.ε _ D.∘ L.₁ (K .ψ x) ≡⟨ D.extendl (sym (adj.counit.is-natural _ _ _)) ⟩ - adj.counit.ε (F.₀ y) D.∘ L.₁ (R.₁ (F.₁ f)) D.∘ L.₁ (K .ψ x) ≡˘⟨ ap (λ e → adj.counit.ε _ D.∘ e) (L.F-∘ _ _) ⟩ - adj.counit.ε (F.₀ y) D.∘ L.₁ (R.₁ (F.₁ f) C.∘ K .ψ x) ≡⟨ ap (λ e → adj.counit.ε _ D.∘ L.₁ e) (K .commutes f) ⟩ - adj.counit.ε _ D.∘ L.₁ _ ∎ -``` - -The key fact is that we can also pass _homomorphisms_ along, both ways! - -```agda - cone-hom-right-adjoint - : {K : Cone (R F∘ F)} {K′ : Cone F} - → Cone-hom F (right-adjoint-cone K) K′ - → Cone-hom (R F∘ F) K (cone-right-adjoint K′) - cone-hom-right-adjoint map .hom = R.₁ (map .hom) C.∘ adj.unit.η _ - cone-hom-right-adjoint {K = K} {K′ = K′} map .commutes o = - R.₁ (K′ .ψ o) C.∘ R.₁ (map .hom) C.∘ adj.unit.η _ ≡⟨ C.pulll (sym (R.F-∘ _ _)) ⟩ - R.₁ (K′ .ψ o D.∘ map .hom) C.∘ adj.unit.η _ ≡⟨ ap (λ e → R.₁ e C.∘ _) (map .commutes _) ⟩ - R.₁ (adj.counit.ε _ D.∘ L.₁ (Cone.ψ K o)) C.∘ adj.unit.η _ ≡⟨ C.pushl (R.F-∘ _ _) ⟩ - R.₁ (adj.counit.ε _) C.∘ R.₁ (L.₁ (Cone.ψ K o)) C.∘ adj.unit.η _ ≡˘⟨ C.pullr (adj.unit.is-natural _ _ _) ⟩ - (R.F₁ (adj.counit.ε _) C.∘ adj.unit.η _) C.∘ Cone.ψ K o ≡⟨ ap (λ e → e C.∘ Cone.ψ K _) adj.zag ⟩ - C.id C.∘ Cone.ψ K o ≡⟨ C.idl _ ⟩ - Cone.ψ K o ∎ - - right-adjoint-cone-hom - : {K : Cone (R F∘ F)} {K′ : Cone F} - → Cone-hom (R F∘ F) K (cone-right-adjoint K′) - → Cone-hom F (right-adjoint-cone K) K′ - right-adjoint-cone-hom map .hom = adj.counit.ε _ D.∘ L.₁ (map .hom) - right-adjoint-cone-hom {K = K} {K′ = K′} map .commutes o = - K′ .ψ o D.∘ adj.counit.ε _ D.∘ L.₁ (map .hom) ≡⟨ D.extendl (sym (adj.counit.is-natural _ _ _)) ⟩ - adj.counit.ε _ D.∘ (L.₁ (R.₁ (K′ .ψ o)) D.∘ L.₁ (map .hom)) ≡⟨ ap (λ e → _ D.∘ e) (sym (L.F-∘ _ _)) ⟩ - adj.counit.ε _ D.∘ (L.₁ (R.₁ (K′ .ψ o) C.∘ map .hom)) ≡⟨ ap (λ e → _ D.∘ L.₁ e) (map .commutes _) ⟩ - adj.counit.ε _ D.∘ L.₁ (K .ψ o) ∎ -``` - -Hence, if we have a limit for $F$, the two lemmas above (in the "towards -right adjoint" direction) already get us 66% of the way to having a -limit for $R \circ F$. The missing bit is a very short calculation: - -``` - right-adjoint-limit : Limit F → Limit (R F∘ F) - right-adjoint-limit lim .top = cone-right-adjoint (lim .top) - right-adjoint-limit lim .has⊤ cone = contr ! !-unique where - pre! = lim .has⊤ _ .centre - ! = cone-hom-right-adjoint pre! - - !-unique : ∀ x → ! ≡ x - !-unique x = Cone-hom-path _ ( - R.₁ (pre! .hom) C.∘ adj.unit.η _ ≡⟨ ap (λ e → R.₁ e C.∘ _) (ap hom (lim .has⊤ _ .paths (right-adjoint-cone-hom x))) ⟩ - R.₁ (adj.counit.ε _ D.∘ L.₁ (x .hom)) C.∘ adj.unit.η _ ≡˘⟨ C.pulll (sym (R.F-∘ _ _)) ⟩ - R.₁ (adj.counit.ε _) C.∘ R.₁ (L.₁ (x .hom)) C.∘ adj.unit.η _ ≡˘⟨ ap (R.₁ _ C.∘_) (adj.unit.is-natural _ _ _) ⟩ - R.₁ (adj.counit.ε _) C.∘ adj.unit.η _ C.∘ x .hom ≡⟨ C.cancell adj.zag ⟩ - x .hom ∎) -``` +[adjoints preserve Kan extensions]: Cat.Functor.Adjoint.Kan.html -We then have the promised theorem: right adjoints preserve limits. ```agda right-adjoint-is-continuous - : ∀ {os ℓs} → is-continuous {oshape = os} {hshape = ℓs} R - right-adjoint-is-continuous L x = Terminal.has⊤ (right-adjoint-limit (record { top = L ; has⊤ = x })) + : ∀ {os ℓs} → is-continuous os ℓs R + right-adjoint-is-continuous lim = + right-adjoint→right-extension lim L⊣R + + left-adjoint-is-cocontinuous + : ∀ {os ℓs} → is-cocontinuous os ℓs L + left-adjoint-is-cocontinuous colim = + left-adjoint→left-extension colim L⊣R + + module _ {od ℓd} {J : Precategory od ℓd} where + right-adjoint-limit : ∀ {F : Functor J D} → Limit F → Limit (R F∘ F) + right-adjoint-limit lim = + to-limit (right-adjoint-is-continuous (Limit.has-limit lim)) + + left-adjoint-colimit : ∀ {F : Functor J C} → Colimit F → Colimit (L F∘ F) + left-adjoint-colimit colim = + to-colimit (left-adjoint-is-cocontinuous (Colimit.has-colimit colim)) ``` ## Concrete limits -For establishing the preservation of "concrete limits", in addition to -the preexisting conversion functions (`Lim→Prod`{.Agda}, -`Limit→Pullback`{.Agda}, `Limit→Equaliser`{.Agda}, etc.), we must -establish results analogous to `canonical-functors`{.Agda}: Functors out -of shape categories are entirely determined by the "introduction forms" -`cospan→cospan-diagram`{.Agda} and `par-arrows→par-diagram`{.Agda}. +We now show that adjoint functors preserve "concrete limits". We could +show this using general abstract nonsense, but we can avoid transports +if we do it by hand. + - right-adjoint→product - : ∀ {A B} → Product D A B → Product C (R.₀ A) (R.₀ B) - right-adjoint→product {A = A} {B} prod = - Lim→Prod C (fixup (right-adjoint-limit (Prod→Lim D prod))) - where - fixup : Limit (R F∘ 2-object-diagram D {iss = Bool-is-set} A B) - → Limit (2-object-diagram C {iss = Bool-is-set} (R.₀ A) (R.₀ B)) - fixup = subst Limit (canonical-functors _ _) - - right-adjoint→pullback - : ∀ {A B c} {f : D.Hom A c} {g : D.Hom B c} - → Pullback D f g → Pullback C (R.₁ f) (R.₁ g) - right-adjoint→pullback {f = f} {g} pb = - Limit→Pullback C {x = lzero} {y = lzero} - (right-adjoint-limit (Pullback→Limit D pb)) - - right-adjoint→equaliser - : ∀ {A B} {f g : D.Hom A B} - → Equaliser D f g → Equaliser C (R.₁ f) (R.₁ g) - right-adjoint→equaliser {f = f} {g} eq = - Limit→Equaliser C (right-adjoint-limit - (Equaliser→Limit D {F = par-arrows→par-diagram f g} eq)) +```agda + right-adjoint→is-product + : ∀ {x a b} {p1 : D.Hom x a} {p2 : D.Hom x b} + → is-product D p1 p2 + → is-product C (R.₁ p1) (R.₁ p2) + right-adjoint→is-product {x = x} {a} {b} {p1} {p2} d-prod = c-prod where + open is-product + + c-prod : is-product C (R.₁ p1) (R.₁ p2) + c-prod .⟨_,_⟩ f g = + L-adjunct L⊣R (d-prod .⟨_,_⟩ (R-adjunct L⊣R f) (R-adjunct L⊣R g)) + c-prod .π₁∘factor = + R.pulll (d-prod .π₁∘factor) ∙ L-R-adjunct L⊣R _ + c-prod .π₂∘factor = + R.pulll (d-prod .π₂∘factor) ∙ L-R-adjunct L⊣R _ + c-prod .unique other p q = + sym (L-R-adjunct L⊣R other) + ∙ ap (L-adjunct L⊣R) + (d-prod .unique _ (R-adjunct-ap L⊣R p) (R-adjunct-ap L⊣R q)) + + right-adjoint→is-pullback + : ∀ {p x y z} + → {p1 : D.Hom p x} {f : D.Hom x z} {p2 : D.Hom p y} {g : D.Hom y z} + → is-pullback D p1 f p2 g + → is-pullback C (R.₁ p1) (R.₁ f) (R.₁ p2) (R.₁ g) + right-adjoint→is-pullback {p1 = p1} {f} {p2} {g} d-pb = c-pb where + open is-pullback + + c-pb : is-pullback C (R.₁ p1) (R.₁ f) (R.₁ p2) (R.₁ g) + c-pb .square = R.weave (d-pb .square) + c-pb .universal sq = + L-adjunct L⊣R (d-pb .universal (R-adjunct-square L⊣R sq)) + c-pb .p₁∘universal = + R.pulll (d-pb .p₁∘universal) ∙ L-R-adjunct L⊣R _ + c-pb .p₂∘universal = + R.pulll (d-pb .p₂∘universal) ∙ L-R-adjunct L⊣R _ + c-pb .unique {_} {p₁'} {p₂'} {sq} {other} p q = + sym (L-R-adjunct L⊣R other) + ∙ ap (L-adjunct L⊣R) + (d-pb .unique (R-adjunct-ap L⊣R p) (R-adjunct-ap L⊣R q)) + + right-adjoint→is-equaliser + : ∀ {e a b} {f g : D.Hom a b} {equ : D.Hom e a} + → is-equaliser D f g equ + → is-equaliser C (R.₁ f) (R.₁ g) (R.₁ equ) + right-adjoint→is-equaliser {f = f} {g} {equ} d-equal = c-equal where + open is-equaliser + + c-equal : is-equaliser C (R.₁ f) (R.₁ g) (R.₁ equ) + c-equal .equal = R.weave (d-equal .equal) + c-equal .universal sq = + L-adjunct L⊣R (d-equal .universal (R-adjunct-square L⊣R sq)) + c-equal .factors = + R.pulll (d-equal .factors) ∙ L-R-adjunct L⊣R _ + c-equal .unique p = + sym (L-R-adjunct L⊣R _) + ∙ ap (L-adjunct L⊣R) + (d-equal .unique (R-adjunct-ap L⊣R p)) right-adjoint→terminal - : ∀ {X} → is-terminal D X → is-terminal C (R.₀ X) + : ∀ {x} → is-terminal D x → is-terminal C (R.₀ x) right-adjoint→terminal term x = contr fin uniq where fin = L-adjunct L⊣R (term (L.₀ x) .centre) uniq : ∀ x → fin ≡ x @@ -196,40 +163,8 @@ of shape categories are entirely determined by the "introduction forms" (x , is-contr→is-prop (term _) _ _) right-adjoint→lex : is-lex R - right-adjoint→lex .is-lex.pres-⊤ = right-adjoint→terminal + right-adjoint→lex .is-lex.pres-⊤ = + right-adjoint→terminal right-adjoint→lex .is-lex.pres-pullback {f = f} {g = g} pb = - right-adjoint→pullback (record { p₁ = _ ; p₂ = _ ; has-is-pb = pb }) .Pullback.has-is-pb + right-adjoint→is-pullback pb ``` - - diff --git a/src/Cat/Functor/Adjoint/Hom.lagda.md b/src/Cat/Functor/Adjoint/Hom.lagda.md index d791bdfbd..9f6e98243 100644 --- a/src/Cat/Functor/Adjoint/Hom.lagda.md +++ b/src/Cat/Functor/Adjoint/Hom.lagda.md @@ -5,10 +5,12 @@ description: | isomorphism of Hom-sets. --- ```agda -open import Cat.Functor.Adjoint -open import Cat.Functor.Hom open import Cat.Instances.Functor +open import Cat.Instances.Functor +open import Cat.Instances.Product open import Cat.Instances.Product +open import Cat.Functor.Adjoint +open import Cat.Functor.Hom open import Cat.Prelude import Cat.Functor.Reasoning as Func @@ -69,7 +71,7 @@ terms. hom-iso-natural f = ∀ {a b c d} (g : C.Hom a b) (h : D.Hom c d) x → f (g C.∘ x C.∘ L.₁ h) ≡ R.₁ g D.∘ f x D.∘ h - + hom-iso→adjoints : (f : ∀ {x y} → C.Hom (L.₀ x) y → D.Hom x (R.₀ y)) → (eqv : ∀ {x y} → is-equiv (f {x} {y})) @@ -78,7 +80,7 @@ terms. hom-iso→adjoints f f-equiv natural = adj′ where f⁻¹ : ∀ {x y} → D.Hom x (R.₀ y) → C.Hom (L.₀ x) y f⁻¹ = equiv→inverse f-equiv - + inv-natural : ∀ {a b c d} (g : C.Hom a b) (h : D.Hom c d) x → f⁻¹ (R.₁ g D.∘ x D.∘ h) ≡ g C.∘ f⁻¹ x C.∘ L.₁ h inv-natural g h x = ap fst $ is-contr→is-prop (f-equiv .is-eqv _) diff --git a/src/Cat/Functor/Adjoint/Kan.lagda.md b/src/Cat/Functor/Adjoint/Kan.lagda.md new file mode 100644 index 000000000..fba886616 --- /dev/null +++ b/src/Cat/Functor/Adjoint/Kan.lagda.md @@ -0,0 +1,221 @@ +```agda +open import Cat.Functor.Kan.Duality +open import Cat.Functor.Coherence +open import Cat.Instances.Functor +open import Cat.Functor.Kan.Base +open import Cat.Functor.Adjoint +open import Cat.Functor.Base +open import Cat.Prelude + +import Cat.Functor.Reasoning as Fr +import Cat.Reasoning as Cr + +module Cat.Functor.Adjoint.Kan where +``` + +# Adjoints preserve Kan extensions + +Let $L \adj R$ be a pair of [adjoint functors]. It's well-known that +right adjoints preserve limits[^rapl], and dually that left adjoints +preserve colimits, but it turns out that this theorem can be made a bit +more general: If $G$ is a [left] (resp. [right]) extension of $F$ along +$p$, then $LG$ is a left extension of $LF$ along $p$: left adjoints +preserve left extensions. This dualises to right adjoints preserving +_right_ extensions. + +[adjoint functors]: Cat.Functor.Adjoint.html +[left]: Cat.Functor.Kan.Base.html#left-kan-extensions +[right]: Cat.Functor.Kan.Base.html#right-kan-extensions + +[^rapl]: Here on the 1Lab, we derive the proof that right (resp. left) +adjoints preserve limits (resp. colimits) from _this proof_ that +adjoints preserve Kan extensions. For a more concrete approach to that +proof, we recommend [the nLab's] + +[the nLab's]: https://ncatlab.org/nlab/show/adjoints+preserve+%28co-%29limits. + +The proof could be given in bicategorical generality: If the triangle +below is a left extension diagram, then the overall pasted diagram is +also a left extension. This is essentially because the $L \dashv R$ +adjunction provides a bunch of isomorphisms between natural +transformations, e.g. $(LF \to Mp) \cong (F \to RMp)$, which we can use +to "grow" the original extension diagram. + +~~~{.quiver .tall-15} +\[\begin{tikzcd} + {\mathcal{C}} && D && A \\ + \\ + {\mathcal{C'}} + \arrow["p"', from=1-1, to=3-1] + \arrow["F", from=1-1, to=1-3] + \arrow[""{name=0, anchor=center, inner sep=0}, "{\mathrm{Lan}_pF}"', from=3-1, to=1-3] + \arrow[""{name=1, anchor=center, inner sep=0}, "L"', shift right=2, from=1-3, to=1-5] + \arrow[""{name=2, anchor=center, inner sep=0}, "R"', shift right=3, color={rgb,255:red,137;green,133;blue,142}, from=1-5, to=1-3] + \arrow["{L\mathrm{Lan}_pF}"', curve={height=18pt}, from=3-1, to=1-5] + \arrow["\dashv"{anchor=center, rotate=90}, draw=none, from=1, to=2] + \arrow["\eta"', shorten <=6pt, Rightarrow, from=0, to=1-1] +\end{tikzcd}\] +~~~ + + + +```agda + left-adjoint→left-extension : preserves-lan L lan + left-adjoint→left-extension = pres where +``` + + + +It wouldn't fit in the diagram above, but the 2-cell of the larger +diagram is simply the whiskering $L\eta$. Unfortunately, proof +assistants; Our existing definition of whiskering lands in $L(Gp)$, but +we need a natural transformation onto $(LG)p$. + +```agda + pres : is-lan p (L F∘ F) (L F∘ G) (nat-assoc-to (L ▸ eta)) +``` + +Given a 2-cell $\alpha : LF \to Mp$, how do we factor it through $\eta$ +as a 2-cell $\sigma : LG \to M$? Well, since the original diagram was an +extension, if we can get our hands on a 2-cell $F \to M'p$, that'll give +us a cell $G \to M'$. Hm: choose $M' = RM$, let $\alpha' : F \to (RM)p$ +be the adjunct of $\alpha$, factor it through the original $\eta$ into a +cell $\sigma' : G \to RM$, and let $\sigma : LG \to M$ be its adjunct. + +```agda + pres .σ α .η x = R-adjunct adj (l.σ (fixup α) .η x) + pres .σ {M = M} α .is-natural x y f = + (counit.ε _ A.∘ L.₁ (l.σ (fixup α) .η y)) A.∘ LG.₁ f ≡⟨ A.pullr (L.weave (l.σ (fixup α) .is-natural x y f)) ⟩ + counit.ε _ A.∘ (L.₁ (RM.₁ f) A.∘ L.₁ (l.σ (fixup α) .η x)) ≡⟨ A.extendl (counit.is-natural _ _ _) ⟩ + M.₁ f A.∘ pres .σ α .η x ∎ + where module M = Functor M + module RM = Functor (R F∘ M) +``` + +And since every part of the process in the preceeding paragraph is an +isomorphism, this is indeed a factorisation, and it's unique to boot: +we've shown that $LG$ is the extension of $LF$ along $p$! + +
+The details of the calculation is just some more calculation +with natural transformations. We'll leave them here for the intrepid +reader but they will not be elaborated on. + + +```agda + pres .σ-comm {α = α} = Nat-path λ x → + (R-adjunct adj (l.σ (fixup α) .η _)) A.∘ L.₁ (eta .η _) ≡⟨ L.pullr (l.σ-comm {α = fixup α} ηₚ _) ⟩ + R-adjunct adj (L-adjunct adj (α .η x)) ≡⟨ equiv→unit (L-adjunct-is-equiv adj) (α .η x) ⟩ + α .η x ∎ + + pres .σ-uniq {M = M} {α = α} {σ′ = σ′} wit = Nat-path λ x → + R-adjunct adj (l.σ (fixup α) .η x) ≡⟨ A.refl⟩∘⟨ ap L.₁ (l.σ-uniq lemma ηₚ x) ⟩ + R-adjunct adj (L-adjunct adj (σ′ .η x)) ≡⟨ equiv→unit (L-adjunct-is-equiv adj) (σ′ .η x) ⟩ + σ′ .η x ∎ + where + module M = Functor M + + σ′′ : G => R F∘ M + σ′′ .η x = L-adjunct adj (σ′ .η x) + σ′′ .is-natural x y f = + (R.₁ (σ′ .η _) D.∘ unit.η _) D.∘ G.₁ f ≡⟨ D.pullr (unit.is-natural _ _ _) ⟩ + (R.₁ (σ′ .η _) D.∘ (RL.₁ (G.₁ f)) D.∘ unit.η _) ≡⟨ D.extendl (R.weave (σ′ .is-natural _ _ _)) ⟩ + R.₁ (M.₁ f) D.∘ R.₁ (σ′ .η x) D.∘ unit.η _ ∎ + + lemma : fixup α ≡ ((σ′′ ◂ p) ∘nt eta) + lemma = Nat-path λ x → + R.₁ (α .η x) D.∘ unit.η _ ≡⟨ ap R.₁ (wit ηₚ _) D.⟩∘⟨refl ⟩ + R.₁ (σ′ .η _ A.∘ L.₁ (eta .η _)) D.∘ unit.η _ ≡⟨ ap (D._∘ unit.η _) (R.F-∘ _ _) ∙ D.extendr (sym (unit.is-natural _ _ _)) ⟩ + (R.₁ (σ′ .η _) D.∘ unit.η _) D.∘ eta .η x ∎ +``` + +
+ +## Dually + +By duality, right adjoints preserve right extensions. + + + +```agda + right-adjoint→right-extension : preserves-ran R ran + right-adjoint→right-extension = fixed where + pres-lan = left-adjoint→left-extension + (is-ran→is-co-lan _ _ ran) + (opposite-adjunction adj) + + module p = is-lan pres-lan + open is-ran + open _=>_ + + fixed : is-ran p (R F∘ F) (R F∘ G) (nat-assoc-from (R ▸ eps)) + fixed .is-ran.σ {M = M} α = σ′ where + unquoteDecl α′ = dualise-into α′ + (Functor.op R F∘ Functor.op F => Functor.op M F∘ Functor.op p) + α + unquoteDecl σ′ = dualise-into σ′ (M => R F∘ G) (p.σ α′) + + fixed .is-ran.σ-comm = Nat-path λ x → p.σ-comm ηₚ _ + fixed .is-ran.σ-uniq {M = M} {σ′ = σ′} p = + Nat-path λ x → p.σ-uniq {σ′ = σ′′} (Nat-path λ x → p ηₚ x) ηₚ x where + unquoteDecl σ′′ = dualise-into σ′′ _ σ′ +``` diff --git a/src/Cat/Functor/Coherence.agda b/src/Cat/Functor/Coherence.agda new file mode 100644 index 000000000..d5b1613d5 --- /dev/null +++ b/src/Cat/Functor/Coherence.agda @@ -0,0 +1,96 @@ +open import 1Lab.Reflection + +open import Cat.Prelude + +import Cat.Instances.Functor.Compose + +module Cat.Functor.Coherence where + +open Cat.Instances.Functor.Compose public +private + variable + o ℓ : Level + C D E : Precategory o ℓ + f g h i : Functor C D + + _⊗_ = _F∘_ + infixr 30 _⊗_ + + _⇒_ = _=>_ + infix 20 _⇒_ + + {-# INLINE _⊗_ #-} + {-# INLINE _⇒_ #-} + +record Dualises {ℓ} (T : Type ℓ) : Type where + field + fields : List (Arg Name) + dualiser : Name + +instance + Dualises-nat-trans + : ∀ {o ℓ o′ ℓ′} {C : Precategory o ℓ} {D : Precategory o′ ℓ′} + {F G : Functor C D} + → Dualises (F => G) + Dualises-nat-trans .Dualises.fields = + argN (quote _=>_.η) + ∷ argN (quote _=>_.is-natural) + ∷ [] + Dualises-nat-trans .Dualises.dualiser = quote _=>_.op + +open Dualises + +cohere-dualise-into + : Bool → ∀ {ℓ ℓ′} {S : Type ℓ′} → Name → (T : Type ℓ) → ⦃ Dualises T ⦄ → S → TC ⊤ +cohere-dualise-into is-dual a T ⦃ du ⦄ tm = do + `tm ← quoteTC tm + clauses ← for (du .fields) λ where (arg ai prj) → do + t ← if is-dual + then (reduce (def prj [ argN (def (du .dualiser) [ argN `tm ]) ])) + else (reduce (def prj [ argN `tm ])) + pure $ clause [] [ arg ai (proj prj) ] t + + `T ← quoteTC T + declareDef (argN a) `T + defineFun a clauses + +define-coherator-dualiser : Bool → Name → TC ⊤ +define-coherator-dualiser is-dual nam = do + (fs , dual) ← runSpeculative do + ty ← inferType (def nam [ argN unknown ]) + (mv , _) ← new-meta′ (def (quote Dualises) [ argN ty ]) + getInstances mv >>= λ where + (qn ∷ []) -> do + fs ← normalise (def (quote Dualises.fields) [ argN qn ]) + >>= unquoteTC {A = List (Arg Name)} + du ← normalise (def (quote Dualises.dualiser) [ argN qn ]) + >>= unquoteTC {A = Name} + pure ((fs , du) , false) + _ → typeError [ strErr "Failed to determine how to define a coherence map for" + , termErr (def nam []) + ] + + clauses ← for fs λ where (arg ai prj) → do + pure $ if is-dual + then (clause (("α" , argN unknown) ∷ []) + [ argN (var 0) , arg ai (proj prj) ] + (def prj [ argN (def dual [ argN (var 0 []) ]) ])) + else (clause (("α" , argN unknown) ∷ []) + [ argN (var 0) , arg ai (proj prj) ] + (def prj [ argN (var 0 []) ])) + + defineFun nam clauses + +define-coherence = define-coherator-dualiser false +define-dualiser = define-coherator-dualiser true +cohere-into = cohere-dualise-into false +dualise-into = cohere-dualise-into true + +nat-assoc-to : f ⇒ g ⊗ h ⊗ i → f ⇒ (g ⊗ h) ⊗ i +nat-assoc-from : f ⊗ g ⊗ h ⇒ i → (f ⊗ g) ⊗ h ⇒ i +op-compose-into : f ⇒ Functor.op (g ⊗ h) → f ⇒ Functor.op g ⊗ Functor.op h + +unquoteDef nat-assoc-to nat-assoc-from op-compose-into = do + define-coherence nat-assoc-to + define-coherence nat-assoc-from + define-coherence op-compose-into diff --git a/src/Cat/Functor/Conservative.lagda.md b/src/Cat/Functor/Conservative.lagda.md index 30bdd1bc1..ed4483e16 100644 --- a/src/Cat/Functor/Conservative.lagda.md +++ b/src/Cat/Functor/Conservative.lagda.md @@ -5,6 +5,9 @@ open import Cat.Functor.Base open import Cat.Morphism open import Cat.Prelude hiding (J) +import Cat.Functor.Reasoning as Func +import Cat.Reasoning as Cat + module Cat.Functor.Conservative where ``` @@ -40,36 +43,45 @@ By the universal property of $L$, there exists a map $\eta$ from the apex of $K$ to the apex of $L$ in $C$. Furthermore, as $F(K)$ is a limit in $D$, $F(\eta)$ becomes an isomorphism in $D$. However, $F$ is conservative, which implies that $\eta$ was an isomorphism in $C$ all along! This means that $K$ must be a limit -in $C$ as well (see `apex-iso→is-limit`{.Agda}). +in $C$ as well (see `is-invertible→is-limitp`{.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) - conservative-reflects-limits {Dia = Dia} L-lim preserves K limits = - apex-iso→is-limit Dia K L-lim + private + open _=>_ + module C = Cat C + module D = Cat D + module F = Func F + + conservative-reflects-limits : ∀ {Dia : Functor J C} + → (L : Limit Dia) + → preserves-limit F Dia + → reflects-limit F Dia + conservative-reflects-limits L-lim preservesa {K} {eps} lim = + is-invertible→is-limitp + {K = Limit.Ext L-lim} {epsy = Limit.cone L-lim} (Limit.has-limit L-lim) + (eps .η) (λ f → sym (eps .is-natural _ _ f) ∙ C.elimr (K .F-id)) refl $ 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 - where - F∘L-lim : Limit (F F∘ Dia) - F∘L-lim .Terminal.top = F-map-cone F (Terminal.top L-lim) - F∘L-lim .Terminal.has⊤ = preserves (Terminal.top L-lim) (Terminal.has⊤ L-lim) + $ invert - K-lim : Limit (F F∘ Dia) - K-lim .Terminal.top = F-map-cone F K - K-lim .Terminal.has⊤ = limits + where + module L-lim = Limit L-lim + module FL-lim = is-limit (preservesa L-lim.has-limit) + module lim = is-limit lim - module L-lim = Terminal L-lim - module F∘L-lim = Terminal F∘L-lim - open Cone-hom + uinv : D.Hom (F .F₀ L-lim.apex) (F .F₀ (K .F₀ tt)) + uinv = + (lim.universal + (λ j → F .F₁ (L-lim.ψ j)) + (λ f → sym (F .F-∘ _ _) ∙ ap (F .F₁) (L-lim.commutes f))) - F-preserves-universal - : hom F∘L-lim.! ≡ F .F₁ (hom {x = K} L-lim.!) - 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.!) ∎ + invert : D.is-invertible (F .F₁ (L-lim.universal (eps .η) _)) + invert = + D.make-invertible uinv + (FL-lim.unique₂ _ (λ j → FL-lim.commutes j) + (λ j → F.pulll (L-lim.factors _ _) ∙ lim.factors _ _) + (λ j → D.idr _)) + (lim.unique₂ _ (λ j → lim.commutes j) + (λ j → D.pulll (lim.factors _ _) ∙ F.collapse (L-lim.factors _ _)) + (λ j → D.idr _)) ``` diff --git a/src/Cat/Functor/Dense.lagda.md b/src/Cat/Functor/Dense.lagda.md index a0d554028..a810a10d3 100644 --- a/src/Cat/Functor/Dense.lagda.md +++ b/src/Cat/Functor/Dense.lagda.md @@ -1,4 +1,5 @@ ```agda +open import Cat.Instances.Shape.Terminal open import Cat.Diagram.Colimit.Base open import Cat.Functor.Kan.Nerve open import Cat.Instances.Comma @@ -35,9 +36,7 @@ nadir of a cocone over $J$. module _ {o ℓ} {C : Precategory ℓ ℓ} {D : Precategory o ℓ} (F : Functor C D) where - open Cocone-hom open Functor - open Cocone open ↓Obj open ↓Hom open _=>_ @@ -50,11 +49,12 @@ module --> ```agda + dense-cocone : ∀ d → F F∘ Dom F (const! d) => Const d + dense-cocone d .η x = x .map + dense-cocone d .is-natural _ _ f = f .sq + is-dense : Type _ - is-dense = ∀ d → is-colimit {J = F ↘ d} (F F∘ Dom _ _) $ λ where - .coapex → d - .ψ x → x .map - .commutes f → f .sq ∙ D.idl _ + is-dense = ∀ d → is-colimit {J = F ↘ d} (F F∘ Dom _ _) d (dense-cocone d) ``` The functor $F$ is called _dense_ if this cocone is colimiting for every @@ -65,26 +65,21 @@ the induced [nerve] functor is fully faithful. ```agda is-dense→nerve-is-ff : is-dense → is-fully-faithful (Nerve F) - is-dense→nerve-is-ff is-colim = is-iso→is-equiv $ iso inv invr invl where - nt→cone : ∀ {x y} → (Nerve F .F₀ x => Nerve F .F₀ y) → Cocone _ - nt→cone {x} {y} nt .coapex = y - nt→cone {x} {y} nt .ψ o = nt .η _ (o .map) - nt→cone {x} {y} nt .commutes {γ} {δ} f = - nt .η _ (δ .map) D.∘ F.₁ (f .α) ≡˘⟨ nt .is-natural _ _ _ $ₚ _ ⟩ - nt .η _ ⌜ δ .map D.∘ F.₁ (f .α) ⌝ ≡⟨ ap! (f .sq ∙ D.idl _) ⟩ - nt .η _ (γ .map) ∎ + is-dense→nerve-is-ff is-dense = is-iso→is-equiv $ iso inv invr invl where + module is-dense d = is-colimit (is-dense d) inv : ∀ {x y} → (Nerve F .F₀ x => Nerve F .F₀ y) → D.Hom x y - inv {x} {y} nt = is-colim x (nt→cone nt) .centre .hom + inv nt = + is-dense.universal _ + (λ j → nt .η _ (j .map)) + λ f → sym (nt .is-natural _ _ _ $ₚ _) ∙ ap (nt .η _) (f .sq ∙ D.idl _) invr : ∀ {x y} (f : Nerve F .F₀ x => Nerve F .F₀ y) → Nerve F .F₁ (inv f) ≡ f invr f = Nat-path λ x → funext λ i → - is-colim _ (nt→cone f) .centre .commutes record { map = i } + is-dense.factors _ {j = ↓obj i} _ _ invl : ∀ {x y} (f : D.Hom x y) → inv (Nerve F .F₁ f) ≡ f - invl f = ap hom $ is-colim _ (nt→cone _) .paths $ λ where - .hom → f - .commutes o → refl + invl f = sym $ is-dense.unique _ _ _ f (λ _ → refl) ``` Another way of putting this is that probes by a dense subcategory are diff --git a/src/Cat/Functor/Equivalence/Complete.lagda.md b/src/Cat/Functor/Equivalence/Complete.lagda.md index 7355051ef..78fe4169f 100644 --- a/src/Cat/Functor/Equivalence/Complete.lagda.md +++ b/src/Cat/Functor/Equivalence/Complete.lagda.md @@ -34,6 +34,6 @@ naturally isomorphic to $K$, so $K$ also has a limit. ```agda equivalence→complete : ∀ {co cℓ} → is-complete co cℓ C → is-complete co cℓ D equivalence→complete ccomp K = - Limit-ap-iso (F∘-iso-id-l F∘F⁻¹≅Id) + natural-iso→limit (F∘-iso-id-l F∘F⁻¹≅Id) (subst Limit F∘-assoc (right-adjoint-limit F⁻¹⊣F (ccomp (F⁻¹ F∘ K)))) ``` diff --git a/src/Cat/Functor/Everything.agda b/src/Cat/Functor/Everything.agda index 90d76a3ab..45f19e588 100644 --- a/src/Cat/Functor/Everything.agda +++ b/src/Cat/Functor/Everything.agda @@ -16,9 +16,10 @@ open import Cat.Functor.Equivalence public open import Cat.Functor.FullSubcategory public open import Cat.Functor.Hom public open import Cat.Functor.Hom.Cocompletion public +open import Cat.Functor.Hom.Coyoneda public open import Cat.Functor.Hom.Representable public -open import Cat.Functor.Kan public -open import Cat.Functor.Kan.Right public +open import Cat.Functor.Kan.Base public open import Cat.Functor.Kan.Nerve public +open import Cat.Functor.Kan.Pointwise public open import Cat.Functor.Pullback public open import Cat.Functor.Slice public diff --git a/src/Cat/Functor/Final.lagda.md b/src/Cat/Functor/Final.lagda.md index c9ecf93d9..61306f92d 100644 --- a/src/Cat/Functor/Final.lagda.md +++ b/src/Cat/Functor/Final.lagda.md @@ -1,7 +1,9 @@ ```agda +open import Cat.Instances.Shape.Terminal open import Cat.Diagram.Colimit.Base open import Cat.Prelude +import Cat.Functor.Reasoning as Func import Cat.Reasoning as Cr module Cat.Functor.Final where @@ -24,9 +26,7 @@ module (F : Functor 𝒞 𝒟) where - open Cocone-hom open Functor - open Cocone private module 𝒞 = Cr 𝒞 @@ -84,8 +84,9 @@ commutes. where private module fin = is-final final - module D = Functor D + module D = Func D module ℰ = Cr ℰ + open _=>_ ``` --> @@ -97,18 +98,13 @@ D(x) \to DF(x_0)$ (where $x_!$ comes from the finality of $F$) defines a cocone $\{D(x) \to K\}$. ```agda - extend-cocone : Cocone (D F∘ F) → Cocone D - extend-cocone cone = cone′ where - open is-iso - module cone = Cocone cone - cone′ : Cocone D - cone′ .coapex = cone.coapex - cone′ .ψ x = cone.ψ _ ℰ.∘ D.₁ (fin.map x) - cone′ .commutes f = - (cone.ψ _ ℰ.∘ D.₁ (fin.map _)) ℰ.∘ D.₁ f ≡⟨ ℰ.pullr (sym (D.F-∘ _ _)) ⟩ - cone.ψ _ ℰ.∘ D.₁ (fin.map _ 𝒟.∘ f) ≡⟨ ℰ.pushl (sym (cone.commutes (fin.extend (fin.map _ 𝒟.∘ f) (fin.map _)))) ⟩ - cone.ψ _ ℰ.∘ _ ≡⟨ ℰ.refl⟩∘⟨ sym (D.F-∘ _ _) ∙ ap D.₁ (sym (fin.extend-commutes _ _)) ⟩ - cone.ψ _ ℰ.∘ D.₁ (fin.map _) ∎ + extend-cocone : ∀ {coapex} → D F∘ F => Const coapex → D => Const coapex + extend-cocone cone .η x = cone .η _ ℰ.∘ D.₁ (fin.map x) + extend-cocone cone .is-natural x y f = + ℰ.pullr (sym (D.F-∘ _ _)) + ·· ℰ.pushl (sym (cone .is-natural _ _ _ ∙ ℰ.idl _)) + ·· (ℰ.refl⟩∘⟨ D.collapse (sym (fin.extend-commutes _ _))) + ∙ sym (ℰ.idl _) ``` In the other direction, suppose that we have a cocone $\{D(x) \to K\}$ @@ -116,13 +112,9 @@ In the other direction, suppose that we have a cocone $\{D(x) \to K\}$ K\}$. ```agda - restrict-cocone : Cocone D → Cocone (D F∘ F) - restrict-cocone K = K′ where - module K = Cocone K - K′ : Cocone (D F∘ F) - K′ .coapex = K.coapex - K′ .ψ x = K.ψ (F.F₀ x) - K′ .commutes f = K.commutes (F.F₁ f) + restrict-cocone : ∀ {coapex} → D => Const coapex → D F∘ F => Const coapex + restrict-cocone K .η x = K .η (F.F₀ x) + restrict-cocone K .is-natural x y f = K .is-natural (F.F₀ x) (F.F₀ y) (F.F₁ f) ``` A computation using connectedness of the comma categories shows that @@ -130,18 +122,17 @@ these formulae are mutually inverse: ```agda open is-iso - extend-cocone-is-iso : is-iso extend-cocone + extend-cocone-is-iso : ∀ {coapex} → is-iso (extend-cocone {coapex}) extend-cocone-is-iso .inv = restrict-cocone - extend-cocone-is-iso .rinv x = Cocone-path _ refl $ λ o → - x .commutes _ - extend-cocone-is-iso .linv x = Cocone-path _ refl $ λ o → - x .ψ _ ℰ.∘ D.₁ (fin.map (F.F₀ o)) ≡˘⟨ x .commutes (fin.extend (fin.map (F.F₀ o)) 𝒟.id) ℰ.⟩∘⟨refl ⟩ - (x .ψ o ℰ.∘ D.₁ (F.₁ (fin.extend _ _))) ℰ.∘ D.₁ (fin.map _) ≡⟨ ℰ.pullr (sym (D.F-∘ _ _) ·· ap D.₁ (fin.extend-commutes _ _) ·· ap D.₁ (𝒟.idr _)) ⟩ - x .ψ o ℰ.∘ D.₁ (F.₁ (fin.extend _ _)) ≡⟨ x .commutes _ ⟩ - x .ψ o ∎ - - restriction-eqv : Cocone (D F∘ F) ≃ Cocone D - restriction-eqv = _ , is-iso→is-equiv extend-cocone-is-iso + extend-cocone-is-iso .rinv x = + Nat-path λ o → + x .is-natural _ _ _ ∙ ℰ.idl _ + extend-cocone-is-iso .linv x = + Nat-path λ o → + (sym (ℰ.idl _) ∙ sym (x .is-natural _ _ (fin.extend (fin.map (F.F₀ o)) 𝒟.id)) ℰ.⟩∘⟨refl) + ·· ℰ.pullr (D.collapse (fin.extend-commutes _ _ ∙ 𝒟.idr _)) + ·· x .is-natural _ _ _ + ∙ ℰ.idl _ ``` The most important conclusion that we get is the following: If you can @@ -153,8 +144,9 @@ the polarity mismatch. ```agda extend-is-colimit - : (K : Cocone (D F∘ F)) - → is-colimit (D F∘ F) K → is-colimit D (extend-cocone K) + : ∀ {coapex} (K : D F∘ F => Const coapex) + → is-colimit (D F∘ F) coapex K + → is-colimit D coapex (extend-cocone K) ```
@@ -164,40 +156,41 @@ it in this `
`{.html} tag for the curious reader only. ```agda - extend-is-colimit K colim x = contr x¡ x¡-unique where - module K = Cocone K - module x = Cocone x - x′ : Cocone (D F∘ F) - x′ = restrict-cocone x - - x′¡ = colim x′ - x¡ : Cocone-hom D (extend-cocone K) x - x¡ .hom = x′¡ .centre .hom - x¡ .commutes o = - x′¡ .centre .hom ℰ.∘ K.ψ _ ℰ.∘ D.₁ _ ≡⟨ ℰ.pulll (x′¡ .centre .commutes _) ⟩ - x′ .ψ _ ℰ.∘ D.₁ (fin.map o) ≡⟨ x .commutes _ ⟩ - x.ψ o ∎ - - x¡-unique : ∀ h′ → x¡ ≡ h′ - x¡-unique h′ = Cocone-hom-path D $ ap hom $ x′¡ .paths go where - go : Cocone-hom (D F∘ F) K x′ - go .hom = h′ .hom - go .commutes o = - h′ .hom ℰ.∘ K.ψ o ≡˘⟨ ℰ.refl⟩∘⟨ K.commutes (fin.extend 𝒟.id (fin.map _)) ⟩ - h′ .hom ℰ.∘ K.ψ _ ℰ.∘ D.₁ (F.₁ _) ≡⟨ ℰ.refl⟩∘⟨ ℰ.refl⟩∘⟨ ap D.₁ (𝒟.intror refl ∙ sym (fin.extend-commutes _ _)) ⟩ - h′ .hom ℰ.∘ K.ψ _ ℰ.∘ D.₁ (fin.map _) ≡⟨ h′ .commutes _ ⟩ - x.ψ (F.₀ o) ∎ + extend-is-colimit {coapex} K colim = + to-is-colimitp mc refl + module extend-is-colimit where + module colim = is-colimit colim + open make-is-colimit + + mc : make-is-colimit D coapex + mc .ψ x = extend-cocone K .η x + mc .commutes f = extend-cocone K .is-natural _ _ _ ∙ ℰ.idl _ + mc .universal eps p = + colim.universal (λ j → eps (F.F₀ j)) λ f → p (F.F₁ f) + mc .factors eps p = + ℰ.pulll (colim.factors _ _) + ∙ p (fin.map _) + mc .unique eps p other q = + colim.unique _ _ _ λ j → + other ℰ.∘ K .η j ≡⟨ ℰ.refl⟩∘⟨ (sym (ℰ.idl _) ∙ sym (K .is-natural _ _ _)) ⟩ + other ℰ.∘ K .η _ ℰ.∘ D.F₁ (F.F₁ (fin.extend _ _)) ≡⟨ ℰ.refl⟩∘⟨ ℰ.refl⟩∘⟨ ap D.₁ (sym (𝒟.idr _) ∙ sym (fin.extend-commutes _ _)) ⟩ + other ℰ.∘ K .η _ ℰ.∘ D.F₁ (fin.map _) ≡⟨ q (F.F₀ j) ⟩ + eps (F.F₀ j) ∎ ```
```agda is-colimit-restrict - : (K : Cocone D) - → is-colimit (D F∘ F) (restrict-cocone K) → is-colimit D K - is-colimit-restrict K colim = subst (is-colimit D) - (Equiv.ε restriction-eqv _) - (extend-is-colimit (restrict-cocone K) colim) + : ∀ {coapex} + → (K : D => Const coapex) + → is-colimit (D F∘ F) coapex (restrict-cocone K) + → is-colimit D coapex K + is-colimit-restrict {coapex} K colim = + to-is-colimitp + (extend-is-colimit.mc (restrict-cocone K) colim) + (extend-cocone-is-iso .rinv K ηₚ _) + where open is-iso ``` - -We can now prove that, if $f, g : X \To Y$ are two maps such that, for -every map with representable domain $h : \yo(A) \to X$, $fh = gh$, then -$f = g$. The quantifier structure of this sentence is a bit funky, so -watch out for the formalisation below: - -```agda - Representables-generate-presheaf - : {f g : X => Y} - → ( ∀ {A : Ob} (h : よ₀ A => X) → f PSh.∘ h ≡ g PSh.∘ h ) - → f ≡ g -``` - -A map $h : \yo(A) \To X$ can be seen as a "generalised element" of $X$, -so that the precondition for $f = g$ can be read as "$f$ and $g$ agree -for _all_ generalised elements with domain _any_ representable". The -proof is deceptively simple: Since $X$ is a colimit, it is an initial -object in the category of cocones under $\yo (\pi X)$. - -The construction `Map→cocone-under`{.Agda} lets us express $Y$ as a -cocone under $\yo (\pi X)$ in a way that $f$ becomes a cocone -homomorphism $X \to Y$; The condition that $g$ agrees with $f$ on all -generalised elements with representable domains ensures that $g$ is -_also_ a cocone homomorphism $X \to Y$; But $X$ is initial, so $f = g$! - -```agda - Representables-generate-presheaf {f} {g} sep = - ap hom $ is-contr→is-prop (coyoneda X (Map→cocone-under X f)) f′ g′ where - f′ : Cocone-hom (よ F∘ El.πₚ C X) (Reassemble X) (Map→cocone-under X f) - f′ .hom = f - f′ .commutes o = Nat-path (λ _ → refl) - - g′ : Cocone-hom (よ F∘ El.πₚ C X) (Reassemble X) (Map→cocone-under X f) - g′ .hom = g - g′ .commutes o = Nat-path λ x → sym (sep $ - NT (λ i a → P.₁ a (o .section)) λ x y h → - funext λ a → P.F-∘ _ _ $ₚ o .section) ηₚ x -``` - -An immediate consequence is that, since any pair of maps $f, g : X \to -Y$ in $\cC$ extend to maps $\yo(f), \yo(g) : \yo(X) \to \yo(Y)$, and -the functor $\yo(-)$ is fully faithful, two maps in $\cC$ are equal -iff. they agree on all generalised elements: - -```agda -private module _ where private - よ-cancelr - : ∀ {X Y : Ob} {f g : Hom X Y} - → (∀ {Z} (h : Hom Z X) → f ∘ h ≡ g ∘ h) - → f ≡ g - よ-cancelr sep = - fully-faithful→faithful {F = よ} よ-is-fully-faithful $ - Representables-generate-presheaf λ h → Nat-path λ x → funext λ a → - sep (h .η x a) -``` - -However note that we have eliminated a mosquito using a low-orbit ion -cannon: +よcov-is-fully-faithful : is-fully-faithful よcov +よcov-is-fully-faithful = is-iso→is-equiv isom where + open is-iso -```agda -よ-cancelr - : ∀ {X Y : Ob} {f g : Hom X Y} - → (∀ {Z} (h : Hom Z X) → f ∘ h ≡ g ∘ h) - → f ≡ g -よ-cancelr sep = sym (idr _) ∙ sep id ∙ idr _ + isom : is-iso よcov₁ + isom .inv nt = nt .η _ id + isom .rinv nt = Nat-path λ c → funext λ g → + sym (nt .is-natural _ _ _) $ₚ _ ∙ ap (nt .η c) (idr g) + isom .linv _ = idl _ ``` diff --git a/src/Cat/Functor/Hom/Cocompletion.lagda.md b/src/Cat/Functor/Hom/Cocompletion.lagda.md index a3cb65005..cf239606e 100644 --- a/src/Cat/Functor/Hom/Cocompletion.lagda.md +++ b/src/Cat/Functor/Hom/Cocompletion.lagda.md @@ -1,12 +1,14 @@ ```agda -open import Cat.Functor.Adjoint.Continuous +open import Cat.Functor.Kan.Pointwise open import Cat.Diagram.Colimit.Base -open import Cat.Functor.Kan.Nerve open import Cat.Instances.Functor +open import Cat.Functor.Kan.Base open import Cat.Functor.Hom -open import Cat.Functor.Kan open import Cat.Prelude +import Cat.Functor.Reasoning as Func +import Cat.Reasoning + module Cat.Functor.Hom.Cocompletion {κ o} (C : Precategory κ κ) (D : Precategory o κ) @@ -16,7 +18,13 @@ module @@ -72,17 +80,12 @@ risen above it: Everything claimed above follows from the general theory of Kan extensions. [Yoneda embedding]: Cat.Functor.Hom.html#the-yoneda-embedding -[left Kan extension]: Cat.Functor.Kan.html +[left Kan extension]: Cat.Functor.Kan.Base.html ```agda -extend : Functor C D → Functor (PSh κ C) D -extend F = Realisation colim F - -extend-cocontinuous - : ∀ {od ℓd} {J : Precategory od ℓd} {Dg : Functor J (PSh κ C)} (F : Functor C D) - → Colimit Dg → Colimit (extend F F∘ Dg) -extend-cocontinuous F = left-adjoint-colimit (Realisation⊣Nerve colim F) +よ-extension : (F : Functor C D) → Lan (よ C) F +よ-extension F = cocomplete→lan (よ C) F colim -extend-factors : (F : Functor C D) → (extend F F∘ よ C) ≅ F -extend-factors F = ff-lan-ext colim (よ C) F (よ-is-fully-faithful C) +extend-factors : (F : Functor C D) → (よ-extension F .Lan.Ext F∘ よ C) ≅ F +extend-factors F = ff→cocomplete-lan-ext (よ C) F colim (よ-is-fully-faithful C) ``` diff --git a/src/Cat/Functor/Hom/Coyoneda.lagda.md b/src/Cat/Functor/Hom/Coyoneda.lagda.md new file mode 100644 index 000000000..5a132b54a --- /dev/null +++ b/src/Cat/Functor/Hom/Coyoneda.lagda.md @@ -0,0 +1,247 @@ +```agda +open import Cat.Diagram.Colimit.Cocone +open import Cat.Diagram.Colimit.Base +open import Cat.Instances.Functor +open import Cat.Instances.Product +open import Cat.Diagram.Initial +open import Cat.Functor.Base +open import Cat.Prelude + +import Cat.Instances.Elements as El +import Cat.Functor.Hom +import Cat.Reasoning + +module Cat.Functor.Hom.Coyoneda {o h} (C : Precategory o h) where +``` + + + +## The Coyoneda Lemma + +The Coyoneda lemma is, like its dual, a statement about presheaves. It +states that "every presheaf is a colimit of representables", which, in +less abstract terms, means that every presheaf arises as some way of +gluing together a bunch of (things isomorphic to) hom functors! + +```agda +module _ (P : Functor (C ^op) (Sets h)) where + private + module P = Functor P + open El C P + open Element + open Element-hom +``` + +We start by fixing some presheaf $P$, and constructing a colimit +whose coapex is $P$. This involves a clever choice of diagram category: +specifically, the [category of elements] of $P$. This may seem like a +somewhat odd choice, but recall that the data contained in $\int P$ is +the _same_ data as $P$, just melted into a soup of points. The colimit +we construct will then glue all those points back together into $P$. + +[category of elements]: Cat.Instances.Elements.html + +```agda + coyoneda : is-colimit (よ F∘ πₚ) P _ + coyoneda = to-is-colimit colim where +``` + +This is done by projecting out of $\int P$ into $\cC$ via the +[canonical projection], and then embedding $\cC$ into the category of +presheaves over $\cC$ via the yoneda embedding. Concretely, what this +diagram gives us is a bunch of copies of the hom functor, one for each +$px : P(X)$. Then, to construct the injection map, we can just use the +(contravariant) functorial action of $P$ to take a $px : P(X)$ and a $f +: Hom(A, X)$ to a $P(A)$. This map is natural by functoriality of $P$. + +[canonical projection]: Cat.Instances.Elements.html#projection + +```agda + open make-is-colimit + module ∫ = Precategory ∫ + + colim : make-is-colimit (よ F∘ πₚ) P + colim .ψ x .η y f = P.F₁ f (x .section) + colim .ψ x .is-natural y z f = + funext (λ g → happly (P.F-∘ f g) (x .section)) + colim .commutes {x = x} {y = y} f = + Nat-path λ z → funext λ g → + P.F₁ (f .hom ∘ g) (y .section) ≡⟨ happly (P.F-∘ g (f .hom)) (y .section) ⟩ + P.F₁ g (P.F₁ (f .hom) (y .section)) ≡⟨ ap (P.F₁ g) (f .commute) ⟩ + P.F₁ g (x .section) ∎ +``` + +Now that we've constructed a cocone, all that remains is to see that +this is a _colimiting_ cocone. Intuitively, it makes sense that +`Reassemble`{.Agda} should be colimiting: all we've done is taken all +the data associated with $P$ and glued it back together. However, +proving this does involve futzing about with various naturality + cocone +commuting conditions. + +We start by constructing the universal map from $P$ into the coapex of +some other cocone $K$. The components of this natural transformation are +obtained in a similar manner to the yoneda lemma; we bundle up the data +to construct an object of $\int P$, and then apply the function we +construct to the identity morphism. Naturality follows from the fact +that $K$ is a cocone, and the components of $K$ are natural. + +```agda + colim .universal eps _ .η x px = eps (elem x px) .η x id + colim .universal {Q} eps comm .is-natural x y f = funext λ px → + eps (elem y (P.F₁ f px)) .η y id ≡˘⟨ (λ i → comm (induce f px) i .η y id) ⟩ + eps (elem x px) .η y (f ∘ id) ≡⟨ ap (eps (elem x px) .η y) id-comm ⟩ + eps (elem x px) .η y (id ∘ f) ≡⟨ happly (eps (elem x px) .is-natural x y f) id ⟩ + F₁ Q f (eps (elem x px) .η x id) ∎ +``` + +Next, we need to show that this morphism factors each of the components +of $K$. The tricky bit of the proof here is that we need to use +`induce`{.Agda} to regard `f` as a morphism in the category of elements. + +```agda + colim .factors {o} eps comm = Nat-path λ x → funext λ f → + eps (elem x (P.F₁ f (o .section))) .η x id ≡˘⟨ (λ i → comm (induce f (o .section)) i .η x id) ⟩ + eps o .η x (f ∘ id) ≡⟨ ap (eps o .η x) (idr f) ⟩ + eps o .η x f ∎ +``` + +Finally, uniqueness: This just follows by the commuting conditions on +`α`. + +```agda + colim .unique eps comm α p = Nat-path λ x → funext λ px → + α .η x px ≡˘⟨ ap (α .η x) (happly P.F-id px) ⟩ + α .η x (P.F₁ id px) ≡⟨ happly (p _ ηₚ x) id ⟩ + eps (elem x px) .η x id ∎ +``` + +And that's it! The important takeaway here is not the shuffling around +of natural transformations required to prove this lemma, but rather the +idea that, unlike Humpty Dumpty, if a presheaf falls off a wall, we +_can_ put it back together again. + +An important consequence of being able to disassemble presheaves into +colimits of representables is that **representables generate +$\psh(C)$**, in that if a pair $f, g$ of natural transformations that +agrees on all representables, then $f = g$ all along. + +```agda + module _ {Y} (f : P => Y) where + private + module Y = Functor Y + open Cocone +``` + +The first thing we prove is that any map $P \To Y$ of presheaves +expresses $Y$ as a cocone over $\yo (\pi P)$. The special case +`Reassemble`{.Agda} above is this procedure for the identity map --- +whence we see that `coyoneda`{.Agda} is essentially a restatement of the +fact that $\id$ is initial the coslice category under $P$. + +```agda + Map→cocone-under : Cocone (よ F∘ πₚ) + Map→cocone-under .coapex = Y + + Map→cocone-under .ψ (elem ob sect) .η x i = f .η x (P.₁ i sect) + Map→cocone-under .ψ (elem ob sect) .is-natural x y h = funext λ a → + f .η _ (P.₁ (a ∘ h) sect) ≡⟨ happly (f .is-natural _ _ _) _ ⟩ + Y.₁ (a ∘ h) (f .η _ sect) ≡⟨ happly (Y.F-∘ _ _) _ ⟩ + Y.₁ h (Y.₁ a (f .η _ sect)) ≡˘⟨ ap (Y .F₁ h) (happly (f .is-natural _ _ _) _) ⟩ + Y.₁ h (f .η _ (P.₁ a sect)) ∎ + + Map→cocone-under .commutes {x} {y} o = Nat-path λ i → funext λ a → ap (f .η _) $ + P.₁ (o .hom ∘ a) (y .section) ≡⟨ happly (P.F-∘ _ _) _ ⟩ + P.₁ a (P.₁ (o .hom) (y .section)) ≡⟨ ap (P.F₁ _) (o .commute) ⟩ + P.₁ a (x .section) ∎ +``` + + + +We can now prove that, if $f, g : X \To Y$ are two maps such that, for +every map with representable domain $h : \yo(A) \to X$, $fh = gh$, then +$f = g$. The quantifier structure of this sentence is a bit funky, so +watch out for the formalisation below: + +```agda + Representables-generate-presheaf + : {f g : X => Y} + → (∀ {A : Ob} (h : よ₀ A => X) → f PSh.∘ h ≡ g PSh.∘ h) + → f ≡ g +``` + +A map $h : \yo(A) \To X$ can be seen as a "generalised element" of $X$, +so that the precondition for $f = g$ can be read as "$f$ and $g$ agree +for _all_ generalised elements with domain _any_ representable". The +proof is deceptively simple: Since $X$ is a colimit, it is an initial +object in the category of cocones under $\yo (\pi X)$. + +The construction `Map→cocone-under`{.Agda} lets us express $Y$ as a +cocone under $\yo (\pi X)$ in a way that $f$ becomes a cocone +homomorphism $X \to Y$; The condition that $g$ agrees with $f$ on all +generalised elements with representable domains ensures that $g$ is +_also_ a cocone homomorphism $X \to Y$; But $X$ is initial, so $f = g$! + +```agda + Representables-generate-presheaf {f} {g} sep = + ap hom $ is-contr→is-prop + (is-colimit→is-initial-cocone _ (coyoneda X) (Map→cocone-under X f)) + f′ g′ + where + f′ : Cocone-hom (よ F∘ El.πₚ C X) _ (Map→cocone-under X f) + f′ .hom = f + f′ .commutes o = Nat-path (λ _ → refl) + + g′ : Cocone-hom (よ F∘ El.πₚ C X) _ (Map→cocone-under X f) + g′ .hom = g + g′ .commutes o = Nat-path λ x → sym (sep $ + NT (λ i a → P.₁ a (o .section)) λ x y h → + funext λ a → P.F-∘ _ _ $ₚ o .section) ηₚ x +``` + +An immediate consequence is that, since any pair of maps $f, g : X \to +Y$ in $\cC$ extend to maps $\yo(f), \yo(g) : \yo(X) \to \yo(Y)$, and +the functor $\yo(-)$ is fully faithful, two maps in $\cC$ are equal +iff. they agree on all generalised elements: + +```agda +private module _ where private + よ-cancelr + : ∀ {X Y : Ob} {f g : Hom X Y} + → (∀ {Z} (h : Hom Z X) → f ∘ h ≡ g ∘ h) + → f ≡ g + よ-cancelr sep = + fully-faithful→faithful {F = よ} よ-is-fully-faithful $ + Representables-generate-presheaf λ h → Nat-path λ x → funext λ a → + sep (h .η x a) +``` + +However note that we have eliminated a mosquito using a low-orbit ion +cannon: + +```agda +よ-cancelr + : ∀ {X Y : Ob} {f g : Hom X Y} + → (∀ {Z} (h : Hom Z X) → f ∘ h ≡ g ∘ h) + → f ≡ g +よ-cancelr sep = sym (idr _) ∙ sep id ∙ idr _ +``` diff --git a/src/Cat/Functor/Hom/Displayed.lagda.md b/src/Cat/Functor/Hom/Displayed.lagda.md index ec2029d50..79535d680 100644 --- a/src/Cat/Functor/Hom/Displayed.lagda.md +++ b/src/Cat/Functor/Hom/Displayed.lagda.md @@ -1,9 +1,9 @@ ```agda -open import Cat.Displayed.Base +open import Cat.Instances.Product +open import Cat.Displayed.Solver open import Cat.Displayed.Fibre open import Cat.Displayed.Total -open import Cat.Displayed.Solver -open import Cat.Instances.Product +open import Cat.Displayed.Base open import Cat.Prelude import Cat.Displayed.Reasoning diff --git a/src/Cat/Functor/Hom/Representable.lagda.md b/src/Cat/Functor/Hom/Representable.lagda.md index 53fc96f26..83fd7fa6b 100644 --- a/src/Cat/Functor/Hom/Representable.lagda.md +++ b/src/Cat/Functor/Hom/Representable.lagda.md @@ -1,7 +1,12 @@ ```agda +open import Cat.Univalent.Instances.Opposite +open import Cat.Diagram.Colimit.Base +open import Cat.Diagram.Limit.Base +open import Cat.Functor.Kan.Unique open import Cat.Instances.Elements open import Cat.Instances.Functor open import Cat.Diagram.Terminal +open import Cat.Morphism.Duality open import Cat.Instances.Sets open import Cat.Functor.Base open import Cat.Functor.Hom @@ -17,6 +22,7 @@ module Cat.Functor.Hom.Representable {o κ} {C : Precategory o κ} where private module C = Cat.Reasoning C module C^ = Cat.Reasoning Cat[ C ^op , Sets κ ] + module [C,Sets] = Cat.Reasoning Cat[ C , Sets κ ] module Sets = Cat.Reasoning (Sets κ) open Element-hom open Functor @@ -44,16 +50,17 @@ record Representation (F : Functor (C ^op) (Sets κ)) : Type (o ⊔ κ) where no-eta-equality field rep : C.Ob - represents : F C^.≅ よ₀ C rep + represents : natural-iso F (Hom-into C rep) + + module rep = natural-iso represents equiv : ∀ {a} → C.Hom a rep ≃ ∣ F .F₀ a ∣ equiv = Iso→Equiv λ where - .fst → represents .C^.from .η _ - .snd .is-iso.inv → represents .C^.to .η _ - .snd .is-iso.rinv x → represents .C^.invr ηₚ _ $ₚ x - .snd .is-iso.linv x → represents .C^.invl ηₚ _ $ₚ x + .fst → rep.from .η _ + .snd .is-iso.inv → rep.to .η _ + .snd .is-iso.rinv x → rep.invr ηₚ _ $ₚ x + .snd .is-iso.linv x → rep.invl ηₚ _ $ₚ x - module rep = C^._≅_ represents module Rep {a} = Equiv (equiv {a}) open Representation @@ -71,7 +78,7 @@ of representing objects for functors: [limits], [colimits], [coends], [limits]: Cat.Diagram.Limit.Base.html [coends]: Cat.Diagram.Coend.html [colimits]: Cat.Diagram.Colimit.Base.html -[Kan extensions]: Cat.Functor.Kan.html +[Kan extensions]: Cat.Functor.Kan.Base.html [adjoint functors]: Cat.Functor.Adjoint.html The first thing we will observe is an immediate consequence of the @@ -117,7 +124,7 @@ Representation-is-prop {F = F} c-cat x y = path where (Nat-pathp _ _ λ a → Hom-pathp-reflr (Sets _) {A = F .F₀ a} {q = λ i → el! (C.Hom a (objs i))} (funext λ x → - ap (λ e → e .Sets.to) (ap-F₀-iso c-cat (Hom[_,-] C a) _) $ₚ _ + ap (λ e → e .Sets.to) (ap-F₀-iso c-cat (Hom-from C a) _) $ₚ _ ·· sym (Y.rep.to .is-natural _ _ _) $ₚ _ ·· ap Y.Rep.from (sym (X.rep.from .is-natural _ _ _ $ₚ _) ·· ap X.Rep.to (C.idl _) @@ -192,3 +199,189 @@ representable-unit→terminal repr .Terminal.top = repr .rep representable-unit→terminal repr .Terminal.has⊤ ob = retract→is-contr (Rep.from repr) (λ _ → lift tt) (Rep.η repr) (hlevel 0) ``` + +## Corepresentable functors + +As noted earlier, we can dualise the definition of a representable +functor to the covariant setting to get **corepresentable** functors. + +```agda +record Corepresentation (F : Functor C (Sets κ)) : Type (o ⊔ κ) where + no-eta-equality + field + corep : C.Ob + corepresents : natural-iso F (Hom-from C corep) + + module corep = natural-iso corepresents + + coequiv : ∀ {a} → C.Hom corep a ≃ ∣ F .F₀ a ∣ + coequiv = Iso→Equiv λ where + .fst → corep.from .η _ + .snd .is-iso.inv → corep.to .η _ + .snd .is-iso.rinv x → corep.invr ηₚ _ $ₚ x + .snd .is-iso.linv x → corep.invl ηₚ _ $ₚ x + + module Corep {a} = Equiv (coequiv {a}) + +open Corepresentation +open Corepresentation using (module Corep) public +``` + +Much like their contravariant cousins, corepresenting objects are unique up to +isomorphism. + +```agda +corepresentation-unique + : {F : Functor C (Sets κ)} (X Y : Corepresentation F) + → X .corep C.≅ Y .corep +``` + +
+We omit the proof, as it is identical to the representable case. + + +```agda +corepresentation-unique X Y = + is-ff→essentially-injective {F = Functor.op (よcov C)} + (よcov-is-fully-faithful C) + (iso→co-iso (Cat[ C , Sets κ ]) ni) + where + ni : natural-iso (Hom-from C (Y .corep)) (Hom-from C (X .corep)) + ni = (Y .corepresents ni⁻¹) ni∘ X .corepresents +``` +
+ +This implies that the type of corepresentations is a proposition when +$\cC$ is univalent. + +```agda +Corepresentation-is-prop : ∀ {F} → is-category C → is-prop (Corepresentation F) +``` + +
+We opt to not show the proof, as it is even nastier than the +proof for representables due to the fact that the yoneda embedding +for covariant functors is itself contravariant. + + +```agda +Corepresentation-is-prop {F = F} c-cat X Y = path where + + objs : X .corep ≡ Y .corep + objs = c-cat .to-path (corepresentation-unique X Y) + + path : X ≡ Y + path i .corep = objs i + path i .corepresents = + [C,Sets].≅-pathp refl (ap (Hom-from C) objs) + {f = X .corepresents} {g = Y .corepresents} + (Nat-pathp _ _ λ a → Hom-pathp-reflr (Sets _) + {A = F .F₀ a} {q = λ i → el! (C.Hom (objs i) a)} + (funext λ x → + ap (λ e → e .Sets.to) (ap-F₀-iso (opposite-is-category c-cat) (Hom-into C a) _) $ₚ _ + ·· sym (corep.to Y .is-natural _ _ _ $ₚ _) + ·· ap (Corep.from Y) (sym (corep.from X .is-natural _ _ _ $ₚ _) + ·· ap (Corep.to X) (C.idr _) + ·· Corep.ε X _))) + i +``` +
+ +## Corepresentable functors preserve limits + +A useful fact about corepresentable functors is that they preserve +all limits. To show this, we first need to show that the covariant +hom functor $\cC(x,-)$ preserves limits. + +To get an intuition for why this is true, consider how the +functor $\cC(x,-)$ behaves on products. The set of morphisms +$\cC(x,a \times b)$ is equivalent to the set $\cC(x, a) \times \cC(x, b)$ +of pairs of morphisms (See [`product-repr`] for a proof of this +equivalence). + +[`product-repr`]: Cat.Diagram.Product.html#product-repr + +```agda +Hom-from-preserves-limits + : ∀ {o′ κ′} + → (c : C.Ob) + → is-continuous o′ κ′ (Hom-from C c) +Hom-from-preserves-limits c {Diagram = Dia} {K} {eps} lim = + to-is-limitp ml (funext λ _ → refl) where + open make-is-limit + module lim = is-limit lim + + ml : make-is-limit _ _ + ml .ψ j f = lim.ψ j C.∘ f + ml .commutes f = funext λ g → + C.pulll (sym (eps .is-natural _ _ _)) + ∙ (C.elimr (K .F-id) C.⟩∘⟨refl) + ml .universal eta p x = + lim.universal (λ j → eta j x) (λ f → p f $ₚ x) + ml .factors _ _ = funext λ _ → + lim.factors _ _ + ml .unique eps p other q = funext λ x → + lim.unique _ _ _ λ j → q j $ₚ x +``` + +Preservation of limits by corepresentable functors then follows from +a general fact about functors: if $F$ preserves limits, and $F$ is +naturally isomorphic to $F'$, then $F'$ must also preserve limits. + +```agda +corepresentable-preserves-limits + : ∀ {o′ κ′} {F} + → Corepresentation F + → is-continuous o′ κ′ F +corepresentable-preserves-limits F-corep lim = + natural-iso→preserves-limits + (F-corep .corepresents ni⁻¹) + (Hom-from-preserves-limits (F-corep .corep)) + lim +``` + +We can show a similar fact for representable functors, but with a twist: +they **reverse** colimits! This is due to the fact that a representable +functor $F : \cC\op \to \Sets$ is contravariant. Specifically, $F$ will +take limits in $\cC\op$ to limits in $\Sets$, but limits in $\cC\op$ +are colimits, so $F$ will take colimits in $\cC$ to limits in $\Sets$. + +A less formal perspective on this is that the collection of maps +out of a colimit is still defined as a limit in $\Sets$. For instance, +to give a $a + b \to x$ out of a coproduct, we are required to give +a pair of maps $a \to x$ and $b \to x$. + +```agda +よ-reverses-colimits + : ∀ {o′ κ′} + → (c : C.Ob) + → is-cocontinuous o′ κ′ (Functor.op (よ₀ C c)) +よ-reverses-colimits c {Diagram = Dia} {K} {eta} colim = + to-is-colimitp mc (funext λ _ → refl) where + open make-is-colimit + module colim = is-colimit colim + + mc : make-is-colimit _ _ + mc .ψ j f = f C.∘ colim.ψ j + mc .commutes f = funext λ g → + C.pullr (eta .is-natural _ _ _) + ∙ (C.refl⟩∘⟨ C.eliml (K .F-id)) + mc .universal eps p x = + colim.universal (λ j → eps j x) (λ f → p f $ₚ x) + mc .factors eps p = funext λ _ → + colim.factors _ _ + mc .unique eps p other q = funext λ x → + colim.unique _ _ _ λ j → q j $ₚ x + +representable-reverses-colimits + : ∀ {o′ κ′} {F} + → Representation F + → is-cocontinuous o′ κ′ (Functor.op F) +representable-reverses-colimits F-rep colim = + natural-iso→preserves-colimits + ((F-rep .represents ni^op) ni⁻¹) + (よ-reverses-colimits (F-rep .rep)) + colim +``` + + diff --git a/src/Cat/Functor/Kan.lagda.md b/src/Cat/Functor/Kan.lagda.md deleted file mode 100644 index 56f8d3e0e..000000000 --- a/src/Cat/Functor/Kan.lagda.md +++ /dev/null @@ -1,413 +0,0 @@ -```agda -open import Cat.Instances.Functor.Compose -open import Cat.Instances.Shape.Terminal -open import Cat.Diagram.Colimit.Base -open import Cat.Instances.Functor -open import Cat.Diagram.Initial -open import Cat.Functor.Adjoint -open import Cat.Instances.Comma -open import Cat.Functor.Base -open import Cat.Prelude - -import Cat.Functor.Reasoning as Func -import Cat.Reasoning as Cat - -module Cat.Functor.Kan where -``` - - - -# Kan extensions - -Suppose we have a functor $F : \cC \to \cD$, and a functor $p : -\cC \to \cC'$ --- perhaps to be thought of as a [full subcategory] -inclusion, where $\cC'$ is a completion of $\cC$, but the -situation applies just as well to any pair of functors --- which -naturally fit into a commutative diagram - -[full subcategory]: Cat.Functor.FullSubcategory.html - -~~~{.quiver} -\[\begin{tikzcd} - \mathcal{C} && \mathcal{D} \\ - \\ - {\mathcal{C}'} - \arrow["F", from=1-1, to=1-3] - \arrow["p"', from=1-1, to=3-1] -\end{tikzcd}\] -~~~ - -but as we can see this is a particularly sad commutative diagram; it's -crying out for a third edge $\cC' \to \cD$ - -~~~{.quiver} -\[\begin{tikzcd} - \mathcal{C} && \mathcal{D} \\ - \\ - {\mathcal{C}'} - \arrow["F", from=1-1, to=1-3] - \arrow["p"', from=1-1, to=3-1] - \arrow[dashed, from=3-1, to=1-3] -\end{tikzcd}\] -~~~ - -extending $F$ to a functor $\cC' \to \cD$. If there exists an -_universal_ such extension (we'll define what "universal" means in just -a second), we call it the **left Kan extension** of $F$ along $p$, and -denote it $\Lan_p F$. Such extensions do not come for free (in a sense -they're pretty hard to come by), but concept of Kan extension can be -used to rephrase the definition of both [limit] and [adjoint functor]. - -[limit]: Cat.Diagram.Limit.Base.html -[adjoint functor]: Cat.Functor.Adjoint.html - -A left Kan extension $\Lan_p F$ is equipped with a natural -transformation $\eta : F \To \Lan_p F \circ p$ witnessing the -("directed") commutativity of the triangle (so that it need not commute -on-the-nose) which is universal among such transformations; Meaning that -if $M : \ca{C'} \to \cD$ is another functor with a transformation -$\alpha : M \To M \circ p$, there is a unique natural transformation -$\sigma : \Lan_p F \To M$ which commutes with $\alpha$. - -Note that in general the triangle commutes "weakly", but when $p$ is -[fully faithful] and $\cD$ is [cocomplete], $\Lan_p F$ genuinely extends -$p$, in that $\eta$ is a natural isomorphism. - -[fully faithful]: Cat.Functor.Base.html#ff-functors -[cocomplete]: Cat.Diagram.Colimit.Base.html#cocompleteness - -```agda -record Lan (p : Functor C C′) (F : Functor C D) : Type (kan-lvl p F) where - field - Ext : Functor C′ D - eta : F => Ext F∘ p -``` - -Universality of `eta`{.Agda} is witnessed by the following fields, which -essentially say that, in the diagram below (assuming $M$ has a natural -transformation $\alpha$ witnessing the same "directed commutativity" -that $\eta$ does for $\Lan_p F$), the 2-cell exists and is unique. - -~~~{.quiver} -\[\begin{tikzcd} - C && D \\ - \\ - {C'} - \arrow["F", from=1-1, to=1-3] - \arrow["p"', from=1-1, to=3-1] - \arrow[""{name=0, anchor=center, inner sep=0}, "{\mathrm{Lan}_pF}"{description}, curve={height=-12pt}, from=3-1, to=1-3] - \arrow[""{name=1, anchor=center, inner sep=0}, "M"{description}, curve={height=12pt}, from=3-1, to=1-3] - \arrow[shorten <=6pt, shorten >=4pt, Rightarrow, from=0, to=1] -\end{tikzcd}\] -~~~ - -```agda - σ : {M : Functor C′ D} (α : F => M F∘ p) → Ext => M - σ-comm : {M : Functor C′ D} {α : F => M F∘ p} → (σ α ◂ p) ∘nt eta ≡ α - σ-uniq : {M : Functor C′ D} {α : F => M F∘ p} {σ′ : Ext => M} - → α ≡ (σ′ ◂ p) ∘nt eta - → σ α ≡ σ′ -``` - -## Ubiquity - -The elevator pitch for Kan extensions is that "all concepts are Kan -extensions". The example we will give here is that, if $F \dashv G$ is -an adjunction, then $(G, \eta)$ gives $\Lan_F(\rm{Id})$. This isn't -exactly enlightening: adjunctions and Kan extensions have very different -vibes, but the latter concept _is_ a legitimate generalisation. - - - -```agda - adjoint→lan : Lan F Id - adjoint→lan .Ext = G - adjoint→lan .eta = unit -``` - -The proof is mostly pushing symbols around, and the calculation is -available below unabridged. In components, $\sigma_x$ must give, -assuming a map $\alpha : \rm{Id} \To MFx$, a map $Gx \to Mx$. The -transformation we're looking for arises as the composite - -$$ -Gx \xto{\alpha} MFGx \xto{M\epsilon} Mx\text{,} -$$ - -where uniqueness and commutativity follows from the triangle identities -`zig`{.Agda} and `zag`{.Agda}. - -```agda - adjoint→lan .σ {M} α .η x = M .Functor.F₁ (counit.ε _) C.∘ α .η (G.₀ x) - adjoint→lan .σ {M} nt .is-natural x y f = - (M.₁ (counit.ε _) C.∘ nt .η _) C.∘ G.₁ f ≡⟨ C.pullr (nt .is-natural _ _ _) ⟩ - M.₁ (counit.ε _) C.∘ M.₁ (F.₁ (G.₁ f)) C.∘ nt .η _ ≡⟨ M.extendl (counit.is-natural _ _ _) ⟩ - M.₁ f C.∘ M.₁ (counit.ε _) C.∘ nt .η _ ∎ - where module M = Func M - - adjoint→lan .σ-comm {M} {α} = Nat-path λ _ → - (M.₁ (counit.ε _) C.∘ α.η _) C.∘ unit.η _ ≡⟨ C.pullr (α.is-natural _ _ _) ⟩ - M.₁ (counit.ε _) C.∘ M.₁ (F.F₁ (unit .η _)) C.∘ α.η _ ≡⟨ M.cancell zig ⟩ - α.η _ ∎ - where module α = _=>_ α - module M = Func M - - adjoint→lan .σ-uniq {M} {α} {σ'} p = Nat-path λ x → - M.₁ (counit.ε _) C.∘ α.η _ ≡⟨ ap (_ C.∘_) (p ηₚ _) ⟩ - M.₁ (counit.ε _) C.∘ σ' .η _ C.∘ unit.η _ ≡⟨ C.extendl (sym (σ' .is-natural _ _ _)) ⟩ - σ' .η _ C.∘ G.₁ (counit.ε _) C.∘ unit.η _ ≡⟨ C.elimr zag ⟩ - σ' .η x ∎ - where module α = _=>_ α - module M = Func M -``` - -# A formula - -In the cases where $\cC, \cD$ are "small enough" and $\cE$ is -"cocomplete enough", the left Kan extension of _any_ functor $F : \cC -\to \cE$ along _any_ functor $K : \cC \to \cD$ exists, and is -computed as a colimit in $\cE$. The size concerns here are -unavoidable, so let's be explicit about them: Suppose that $\cE$ -admits colimits of [$\kappa$-small diagrams], e.g. because it is -$\Sets_\kappa$. Then the category $\cC$ must be $\kappa$-small, and -$\cD$ must be locally $\kappa$-small, i.e. its Hom-sets must live in -the $\kappa$th universe. - -[$\kappa$-small diagrams]: 1Lab.intro.html#universes-and-size-issues - - - -The size restrictions on $\cC$ and $\cD$ ensure that the [comma -category] $K \searrow d$ is $\kappa$-small, so that $\cE$ has a -colimit for it. The objects of this category can be considered -"approximations of $d$ coming from $\cC$", so the colimit over this -category is a "best approximation of $d$"! The rest of the computation -is "straightforward" in the way that most category-theoretic -computations are: it looks mighty complicated from the outside, but when -you're actually working them out, there's only one step you can take at -each point. Agda's goal-and-context display guides you the whole way. - -[comma category]: Cat.Instances.Comma.html - -```agda - lan-approximate : ∀ {d e} (f : D.Hom d e) → Cocone (F F∘ Dom K (const! d)) - lan-approximate {e = e} f .coapex = colim (F F∘ Dom K (const! e)) .bot .coapex - lan-approximate {e = e} f .ψ x = - colim (F F∘ Dom K (const! e)) .bot .ψ (record { map = f D.∘ x .map }) - lan-approximate {e = e} f .commutes {x} {y} h = - colim (F F∘ Dom K (const! e)) .bot .commutes (record { sq = path }) - where abstract - path : (f D.∘ y .map) D.∘ K.₁ (h .α) ≡ D.id D.∘ (f D.∘ x .map) - path = - (f D.∘ y .map) D.∘ K.₁ (h .α) ≡⟨ D.pullr (h .sq) ⟩ - f D.∘ D.id D.∘ x .map ≡⟨ cat! D ⟩ - D.id D.∘ (f D.∘ x .map) ∎ - - cocomplete→lan : Lan K F - cocomplete→lan = lan module cocomplete→lan where - diagram : ∀ d → Functor (K ↘ d) E - diagram d = F F∘ Dom K (const! d) - - approx = lan-approximate -``` - -Our extension will associate to each object $d$ the colimit of - -$$ -(K \searrow d) \xto{\rm{Dom}} C \xto{F} E\text{,} -$$ - -where $\rm{Dom}$ is the functor which projects out the *dom*ain of each -object of $K \searrow d$. Now, we must also associate _arrows_ $f : d -\to e \in \cD$ to arrows between the respective colimits of $d$ and -$e$. What we note is that any arrow $f : d \to e$ displays (the colimit -associated with) $e$ as a cocone under $d$, as can be seen in the -computation of `approx`{.Agda} above. - -Our functor can then take an arrow $f : d \to e$ to the uniqueness arrow -from $\colim(d) \to \colim(e)$ (punning $d$ and $e$ for their respective -diagrams), which exists because $\colim(d)$ is initial. Uniqueness of -this arrow ensures that this assignment is functorial --- but as the -functoriality proof is (to use a technical term) goddamn nasty, we leave -it hidden from the page. - -```agda - func : Functor D E - func .F₀ d = colim (diagram d) .bot .coapex - func .F₁ {d} {e} f = colim (diagram d) .has⊥ (approx f) .centre .hom -``` - - - -It remains to show that our extension functor admits a natural -transformation (with components) $Fx \to \colim(Fx)$, but we can take -these arrows to be the colimit coprojections `ψ`{.Agda}; The factoring -natural transformation `σ`{.Agda} is given by eliminating the colimit, -which ensures commutativity and uniqueness. We leave the rest of the -computation in this `
`{.html} tag, for the interested reader. - -
- -Fair advance warning that the computation here doesn't have any further comments. - - -```agda - lan : Lan K F - lan .Ext = func - lan .eta .η x = colim (diagram (K.₀ x)) .bot .ψ (record { map = D.id }) - - lan .σ {M} α .η x = colim (diagram x) .has⊥ cocone′ .centre .hom - where - module M = Func M - - cocone′ : Cocone _ - cocone′ .coapex = M.₀ x - cocone′ .ψ ob = M.₁ (ob .map) E.∘ α .η _ - cocone′ .commutes {x} {y} f = - (M.₁ (y .map) E.∘ α .η _) E.∘ F.₁ (f .↓Hom.α) ≡⟨ E.pullr (α .is-natural _ _ _) ⟩ - M.₁ (y .map) E.∘ M.₁ (K.₁ (f .↓Hom.α)) E.∘ α .η _ ≡⟨ M.pulll (f .↓Hom.sq ∙ D.idl _) ⟩ - M.₁ (x .map) E.∘ α .η _ ∎ - - lan .eta .is-natural x y f = sym $ - colim (diagram (K.₀ x)) .has⊥ (approx (K.₁ f)) .centre .commutes _ - ∙ sym (colim (diagram (K.₀ y)) .bot .commutes - (record { sq = D.introl refl ∙ ap₂ D._∘_ refl (sym D.id-comm) })) - - lan .σ {M} α .is-natural x y f = - ap hom $ is-contr→is-prop (colim (diagram x) .has⊥ cocone′) - (cocone-hom _ λ o → E.pullr (colim (diagram x) .has⊥ (approx f) .centre .commutes _) - ∙ colim (diagram y) .has⊥ _ .centre .commutes _) - (cocone-hom _ λ o → E.pullr (colim (diagram x) .has⊥ _ .centre .commutes _) - ∙ M.pulll refl) - where - module M = Func M - - cocone′ : Cocone _ - cocone′ .coapex = M.₀ y - cocone′ .ψ x = _ - cocone′ .commutes {x} {y} f = - E.pullr (α .is-natural _ _ _) - ∙ M.pulll (D.pullr (f .↓Hom.sq ∙ D.idl _)) - - lan .σ-comm {M = M} = - Nat-path λ x → colim (diagram (K.₀ x)) .has⊥ _ .centre .commutes _ ∙ M.eliml refl - where module M = Func M - - lan .σ-uniq {M = M} {α} {σ'} path = Nat-path λ x → - ap hom $ colim (diagram _) .has⊥ _ .paths - (cocone-hom _ λ o → sym $ - ap₂ E._∘_ refl (path ηₚ ↓Obj.x o) - ∙ E.pulll (sym (σ' .is-natural _ _ _)) - ∙ E.pullr ( colim _ .has⊥ _ .centre .commutes _ - ∙ ap (colim (diagram x) .bot .ψ) - (↓Obj-path _ _ refl refl (D.idr _)))) -``` -
- -A useful result about this calculation of $\Lan_F(G)$ is that, if $F$ is -fully faithful, then $\Lan_F(G) \circ F \cong G$ --- the left Kan -extension along a fully-faithful functor does actually _extend_. - -```agda - private module Fn = Cat Cat[ C , E ] - open _=>_ - - ff-lan-ext : is-fully-faithful K → cocomplete→lan .Ext F∘ K Fn.≅ F - ff-lan-ext ff = Fn._Iso⁻¹ (Fn.invertible→iso (cocomplete→lan .eta) inv) where - module ff {x} {y} = Equiv (_ , ff {x} {y}) - inv′ : ∀ x → E.is-invertible (cocomplete→lan .eta .η x) - inv′ x = E.make-invertible to invl invr where - cocone′ : Cocone _ - cocone′ .coapex = F.₀ x - cocone′ .ψ ob = F.₁ (ff.from (ob .map)) - cocone′ .commutes {x = y} {z} f = - F.collapse (fully-faithful→faithful {F = K} ff - ( K .Functor.F-∘ _ _ - ·· ap₂ D._∘_ (ff.ε _) refl - ·· f .sq - ·· D.idl _ - ·· sym (ff.ε _))) - - to : E.Hom _ (F.₀ x) - to = colim _ .has⊥ cocone′ .centre .hom - - invl : cocomplete→lan .eta .η x E.∘ to ≡ E.id - invl = ap hom $ is-contr→is-prop - (colim _ .has⊥ (colim (F F∘ Dom K (const! _)) .bot)) - (cocone-hom _ λ o → - E.pullr (colim _ .has⊥ cocone′ .centre .commutes _) - ∙ colim _ .bot .commutes - (record { sq = ap (D.id D.∘_) (ff.ε _) })) - (cocone-hom _ λ o → E.idl _) - - invr : to E.∘ cocomplete→lan .eta .η x ≡ E.id - invr = colim _ .has⊥ cocone′ .centre .commutes _ - ∙ F.elim (fully-faithful→faithful {F = K} ff - (ff.ε _ ∙ sym K.F-id)) - - inv : Fn.is-invertible (cocomplete→lan .eta) - inv = componentwise-invertible→invertible (cocomplete→lan .eta) inv′ -``` diff --git a/src/Cat/Functor/Kan/Adjoint.lagda.md b/src/Cat/Functor/Kan/Adjoint.lagda.md new file mode 100644 index 000000000..e29f6d912 --- /dev/null +++ b/src/Cat/Functor/Kan/Adjoint.lagda.md @@ -0,0 +1,113 @@ +```agda +open import Cat.Functor.Kan.Duality +open import Cat.Functor.Kan.Base +open import Cat.Functor.Adjoint +open import Cat.Prelude + +import Cat.Functor.Reasoning as Func +import Cat.Reasoning as Cat + +module Cat.Functor.Kan.Adjoint where +``` + + + +# Adjoints are Kan extensions + +The elevator pitch for Kan extensions is that "all concepts are Kan +extensions". The example we will give here is that, if $F \dashv G$ is +an adjunction, then $(G, \eta)$ gives $\Lan_F(\rm{Id})$. This isn't +exactly enlightening: adjunctions and Kan extensions have very different +vibes, but the latter concept _is_ a legitimate generalisation. + + + +```agda + adjoint→is-lan : is-lan F Id G unit +``` + +The proof is mostly pushing symbols around, and the calculation is +available below unabridged. In components, $\sigma_x$ must give, +assuming a map $\alpha : \rm{Id} \To MFx$, a map $Gx \to Mx$. The +transformation we're looking for arises as the composite + +$$ +Gx \xto{\alpha} MFGx \xto{M\epsilon} Mx\text{,} +$$ + +where uniqueness and commutativity follows from the triangle identities +`zig`{.Agda} and `zag`{.Agda}. + +```agda + adjoint→is-lan .σ {M} α .η x = M .F₁ (counit.ε _) C.∘ α .η (G.₀ x) + adjoint→is-lan .σ {M} nt .is-natural x y f = + (M.₁ (counit.ε _) C.∘ nt .η _) C.∘ G.₁ f ≡⟨ C.pullr (nt .is-natural _ _ _) ⟩ + M.₁ (counit.ε _) C.∘ M.₁ (F.₁ (G.₁ f)) C.∘ nt .η _ ≡⟨ M.extendl (counit.is-natural _ _ _) ⟩ + M.₁ f C.∘ M.₁ (counit.ε _) C.∘ nt .η _ ∎ + where module M = Func M + + adjoint→is-lan .σ-comm {M} {α} = Nat-path λ _ → + (M.₁ (counit.ε _) C.∘ α.η _) C.∘ unit.η _ ≡⟨ C.pullr (α.is-natural _ _ _) ⟩ + M.₁ (counit.ε _) C.∘ M.₁ (F.F₁ (unit .η _)) C.∘ α.η _ ≡⟨ M.cancell zig ⟩ + α.η _ ∎ + where module α = _=>_ α + module M = Func M + + adjoint→is-lan .σ-uniq {M} {α} {σ'} p = Nat-path λ x → + M.₁ (counit.ε _) C.∘ α.η _ ≡⟨ ap (_ C.∘_) (p ηₚ _) ⟩ + M.₁ (counit.ε _) C.∘ σ' .η _ C.∘ unit.η _ ≡⟨ C.extendl (sym (σ' .is-natural _ _ _)) ⟩ + σ' .η _ C.∘ G.₁ (counit.ε _) C.∘ unit.η _ ≡⟨ C.elimr zag ⟩ + σ' .η x ∎ + where module α = _=>_ α + module M = Func M +``` + +As expected, adjoints also yield right Kan extensions. + +```agda + adjoint→is-ran : is-ran G Id F counit +``` + +
+The proof is the same as left adjoints, just dualized. + + +```agda + adjoint→is-ran .σ {M} β .η x = β .η _ D.∘ M .F₁ (unit.η _) + adjoint→is-ran .σ {M} β .is-natural _ _ _ = + M.extendr (unit.is-natural _ _ _) + ∙ D.pushl (β .is-natural _ _ _) + where module M = Func M + adjoint→is-ran .σ-comm {M} {β} = Nat-path λ _ → + D.pulll (sym $ β .is-natural _ _ _) + ∙ M.cancelr zag + where module M = Func M + adjoint→is-ran .σ-uniq {M} {β} {σ′} p = Nat-path λ _ → + ap (D._∘ _) (p ηₚ _) + ·· D.extendr (σ′ .is-natural _ _ _) + ·· D.eliml zig +``` +
diff --git a/src/Cat/Functor/Kan/Base.lagda.md b/src/Cat/Functor/Kan/Base.lagda.md new file mode 100644 index 000000000..9164f7623 --- /dev/null +++ b/src/Cat/Functor/Kan/Base.lagda.md @@ -0,0 +1,348 @@ +```agda +open import Cat.Instances.Functor.Compose +open import Cat.Instances.Shape.Terminal +open import Cat.Functor.Coherence +open import Cat.Instances.Functor +open import Cat.Prelude + +import Cat.Functor.Reasoning as Func +import Cat.Reasoning as Cat + +module Cat.Functor.Kan.Base where +``` + + + +# Left Kan extensions + +Suppose we have a functor $F : \cC \to \cD$, and a functor $p : +\cC \to \cC'$ --- perhaps to be thought of as a [full subcategory] +inclusion, where $\cC'$ is a completion of $\cC$, but the +situation applies just as well to any pair of functors --- which +naturally fit into a commutative diagram + +[full subcategory]: Cat.Functor.FullSubcategory.html + +~~~{.quiver} +\[\begin{tikzcd} + \mathcal{C} && \mathcal{D} \\ + \\ + {\mathcal{C}'} + \arrow["F", from=1-1, to=1-3] + \arrow["p"', from=1-1, to=3-1] +\end{tikzcd}\] +~~~ + +but as we can see this is a particularly sad commutative diagram; it's +crying out for a third edge $\cC' \to \cD$ + +~~~{.quiver} +\[\begin{tikzcd} + \mathcal{C} && \mathcal{D} \\ + \\ + {\mathcal{C}'} + \arrow["F", from=1-1, to=1-3] + \arrow["p"', from=1-1, to=3-1] + \arrow[dashed, from=3-1, to=1-3] +\end{tikzcd}\] +~~~ + +extending $F$ to a functor $\cC' \to \cD$. If there exists an +_universal_ such extension (we'll define what "universal" means in just +a second), we call it the **left Kan extension** of $F$ along $p$, and +denote it $\Lan_p F$. Such extensions do not come for free (in a sense +they're pretty hard to come by), but concept of Kan extension can be +used to rephrase the definition of both [limit] and [adjoint functor]. + +[limit]: Cat.Diagram.Limit.Base.html +[adjoint functor]: Cat.Functor.Adjoint.html + +A left Kan extension $\Lan_p F$ is equipped with a natural +transformation $\eta : F \To \Lan_p F \circ p$ witnessing the +("directed") commutativity of the triangle (so that it need not commute +on-the-nose) which is universal among such transformations; Meaning that +if $M : \ca{C'} \to \cD$ is another functor with a transformation +$\alpha : M \To M \circ p$, there is a unique natural transformation +$\sigma : \Lan_p F \To M$ which commutes with $\alpha$. + +Note that in general the triangle commutes "weakly", but when $p$ is +[fully faithful] and $\cD$ is [cocomplete], $\Lan_p F$ genuinely extends +$p$, in that $\eta$ is a natural isomorphism. + +[fully faithful]: Cat.Functor.Base.html#ff-functors +[cocomplete]: Cat.Diagram.Colimit.Base.html#cocompleteness + +```agda +record + is-lan (p : Functor C C′) (F : Functor C D) (L : Functor C′ D) (eta : F => L F∘ p) + : Type (kan-lvl p F) where + field +``` + +Universality of `eta`{.Agda} is witnessed by the following fields, which +essentially say that, in the diagram below (assuming $M$ has a natural +transformation $\alpha$ witnessing the same "directed commutativity" +that $\eta$ does for $\Lan_p F$), the 2-cell exists and is unique. + +~~~{.quiver} +\[\begin{tikzcd} + C && D \\ + \\ + {C'} + \arrow["F", from=1-1, to=1-3] + \arrow["p"', from=1-1, to=3-1] + \arrow[""{name=0, anchor=center, inner sep=0}, "{\mathrm{Lan}_pF}"{description}, curve={height=-12pt}, from=3-1, to=1-3] + \arrow[""{name=1, anchor=center, inner sep=0}, "M"{description}, curve={height=12pt}, from=3-1, to=1-3] + \arrow[shorten <=6pt, shorten >=4pt, Rightarrow, from=0, to=1] +\end{tikzcd}\] +~~~ + +```agda + σ : {M : Functor C′ D} (α : F => M F∘ p) → L => M + σ-comm : {M : Functor C′ D} {α : F => M F∘ p} → (σ α ◂ p) ∘nt eta ≡ α + σ-uniq : {M : Functor C′ D} {α : F => M F∘ p} {σ′ : L => M} + → α ≡ (σ′ ◂ p) ∘nt eta + → σ α ≡ σ′ + + σ-uniq₂ + : {M : Functor C′ D} (α : F => M F∘ p) {σ₁′ σ₂′ : L => M} + → α ≡ (σ₁′ ◂ p) ∘nt eta + → α ≡ (σ₂′ ◂ p) ∘nt eta + → σ₁′ ≡ σ₂′ + σ-uniq₂ β p q = sym (σ-uniq p) ∙ σ-uniq q + + σ-uniqp + : ∀ {M₁ M₂ : Functor C′ D} + → {α₁ : F => M₁ F∘ p} {α₂ : F => M₂ F∘ p} + → (q : M₁ ≡ M₂) + → PathP (λ i → F => q i F∘ p) α₁ α₂ + → PathP (λ i → L => q i) (σ α₁) (σ α₂) + σ-uniqp q r = Nat-pathp refl q λ c' i → + σ {M = q i} (r i) .η c' + + open _=>_ eta +``` + +We also provide a bundled form of this data. + +```agda +record Lan (p : Functor C C′) (F : Functor C D) : Type (kan-lvl p F) where + field + Ext : Functor C′ D + eta : F => Ext F∘ p + has-lan : is-lan p F Ext eta + + module Ext = Func Ext + open is-lan has-lan public +``` + +# Right Kan extensions + +Our choice of universal property in the section above isn't the only +choice; we could instead require a [terminal] solution to the lifting +problem, instead of an [initial] one. We can picture the terminal +situation using the following diagram. + +[terminal]: Cat.Diagram.Terminal.html +[initial]: Cat.Diagram.Initial.html + +~~~{.quiver} +\[\begin{tikzcd} + {\mathcal{C}} && {\mathcal{D}} \\ + \\ + {\mathcal{C}'} + \arrow["F", from=1-1, to=1-3] + \arrow["p"', from=1-1, to=3-1] + \arrow[""{name=0, anchor=center, inner sep=0}, "{\mathrm{Ran}_pF}"{description}, curve={height=-12pt}, from=3-1, to=1-3] + \arrow[""{name=1, anchor=center, inner sep=0}, "M"{description}, curve={height=12pt}, from=3-1, to=1-3] + \arrow[shorten <=4pt, shorten >=4pt, Rightarrow, from=1, to=0] +\end{tikzcd}\] +~~~ + +Note the same warnings about "weak, directed" commutativity as for left +Kan extensions apply here, too. Rather than either of the triangles +commuting on the nose, we have natural transformations $\eps$ witnessing +their commutativity. + +```agda +record is-ran + (p : Functor C C′) (F : Functor C D) (Ext : Functor C′ D) + (eps : Ext F∘ p => F) + : Type (kan-lvl p F) where + no-eta-equality + + field + σ : {M : Functor C′ D} (α : M F∘ p => F) → M => Ext + σ-comm : {M : Functor C′ D} {β : M F∘ p => F} → eps ∘nt (σ β ◂ p) ≡ β + σ-uniq : {M : Functor C′ D} {β : M F∘ p => F} {σ′ : M => Ext} + → β ≡ eps ∘nt (σ′ ◂ p) + → σ β ≡ σ′ + + open _=>_ eps renaming (η to ε) + + σ-uniq₂ + : {M : Functor C′ D} (β : M F∘ p => F) {σ₁′ σ₂′ : M => Ext} + → β ≡ eps ∘nt (σ₁′ ◂ p) + → β ≡ eps ∘nt (σ₂′ ◂ p) + → σ₁′ ≡ σ₂′ + σ-uniq₂ β p q = sym (σ-uniq p) ∙ σ-uniq q + +record Ran (p : Functor C C′) (F : Functor C D) : Type (kan-lvl p F) where + no-eta-equality + field + Ext : Functor C′ D + eps : Ext F∘ p => F + has-ran : is-ran p F Ext eps + + module Ext = Func Ext + open is-ran has-ran public +``` + + + +# Preservation and reflection of Kan extensions + +Let $(G : C' \to D, \eta : F \to G \circ p)$ be the left Kan extension +of $F : C \to D$ along $p : C \to C'$, and suppose that $H : D \to E$ is +a functor. We can “apply” $H$ to all the data of the Kan extension, +obtaining the following diagram. + +~~~{.quiver} +\begin{tikzcd} + C &&&& E \\ + \\ + && {C'} + \arrow["p"', from=1-1, to=3-3] + \arrow[""{name=0, anchor=center, inner sep=0}, "HF", from=1-1, to=1-5] + \arrow["HG"', from=3-3, to=1-5] + \arrow["H\eta", shorten <=4pt, shorten >=4pt, Rightarrow, from=0, to=3-3] +\end{tikzcd} +~~~ + +This looks like yet another Kan extension diagram, but it may not be +universal! If this diagram _is_ a left Kan extension, we say that $H$ +**preserves** $(G, \eta)$. + + + +```agda + preserves-lan : (H : Functor D E) → is-lan p F G eta → Type _ + preserves-lan H _ = + is-lan p (H F∘ F) (H F∘ G) (nat-assoc-to (H ▸ eta)) +``` + +In the diagram above, the 2-cell is simply the whiskering $H\eta$. +Unfortunately, proof assistants; our definition of whiskering lands in +$H(Gp)$, but we requires a natural transformation to $(HG)p$. + +It may also be the case that $(HG, H\eta)$ is already a left kan +extension of $HF$ along $p$. We say that $H$ reflects this Kan extension +if $G, \eta$ is a also a left extension of $F$ along $p$. + +```agda + reflects-lan + : (H : Functor D E) + → is-lan p (H F∘ F) (H F∘ G) (nat-assoc-to (H ▸ eta)) + → Type _ + reflects-lan _ _ = + is-lan p F G eta +``` + + + +We can define dual notions for right Kan extensions as well. + +```agda + preserves-ran : (H : Functor D E) → is-ran p F G eps → Type _ + preserves-ran H _ = + is-ran p (H F∘ F) (H F∘ G) (nat-assoc-from (H ▸ eps)) + + reflects-ran + : (H : Functor D E) + → is-ran p (H F∘ F) (H F∘ G) (nat-assoc-from (H ▸ eps)) + → Type _ + reflects-ran _ _ = + is-ran p F G eps +``` + + diff --git a/src/Cat/Functor/Kan/Duality.lagda.md b/src/Cat/Functor/Kan/Duality.lagda.md new file mode 100644 index 000000000..01541a244 --- /dev/null +++ b/src/Cat/Functor/Kan/Duality.lagda.md @@ -0,0 +1,186 @@ +```agda +open import Cat.Functor.Coherence +open import Cat.Functor.Kan.Base +open import Cat.Prelude + +module Cat.Functor.Kan.Duality where +``` + + + +# Duality for Kan extensions + +Left Kan extensions are dual to right Kan extensions. This is +straightforward enough to prove, but does involve some bureaucracy +involving opposite categories. + + + +```agda + is-co-lan'→is-ran + : ∀ {G : Functor (C′ ^op) (D ^op)} {eta : Functor.op F => G F∘ Functor.op p} + → is-lan (Functor.op p) (Functor.op F) G eta + → is-ran p F (Functor.op G) (co-unit→counit eta) + is-co-lan'→is-ran {G = G} {eta = eta} is-lan = ran where + module lan = is-lan is-lan + + ran : is-ran p F (Functor.op G) (co-unit→counit eta) + ran .σ {M = M} α = op (lan.σ α′) where + unquoteDecl α′ = dualise-into α′ + (Functor.op F => Functor.op M F∘ Functor.op p) + α + + ran .σ-comm = Nat-path λ x → lan.σ-comm ηₚ x + ran .σ-uniq {M = M} {σ′ = σ′} p = + Nat-path λ x → lan.σ-uniq {σ′ = σ′op} (Nat-path λ x → p ηₚ x) ηₚ x + where unquoteDecl σ′op = dualise-into σ′op _ σ′ +``` + + diff --git a/src/Cat/Functor/Kan/Global.lagda.md b/src/Cat/Functor/Kan/Global.lagda.md index bb9740e62..f42fd8420 100644 --- a/src/Cat/Functor/Kan/Global.lagda.md +++ b/src/Cat/Functor/Kan/Global.lagda.md @@ -1,11 +1,10 @@ ```agda open import Cat.Instances.Functor.Compose -open import Cat.Diagram.Colimit.Base open import Cat.Functor.Adjoint.Hom open import Cat.Instances.Functor open import Cat.Instances.Product +open import Cat.Functor.Kan.Base open import Cat.Functor.Adjoint -open import Cat.Functor.Kan open import Cat.Prelude import Cat.Reasoning @@ -37,7 +36,7 @@ is a universal solution $\Lan_p(F)$ to extending $F$ to a functor $C' **any** $F$; When this happens, the assignment $F \mapsto \Lan_p(F)$ *extends to a functor, which we call a **global Kan extension**. -[left Kan extension]: Cat.Functor.Kan.html +[left Kan extension]: Cat.Functor.Kan.Base.html [small]: 1Lab.intro.html#universes-and-size-issues [cocomplete]: Cat.Diagram.Colimit.Base.html#cocompleteness @@ -97,15 +96,15 @@ values generate Kan extensions: ```agda adjoint→Lan : (F : Functor Cat[ C , D ] Cat[ C′ , D ]) - → F ⊣ p ! - → (G : Functor C D) → Lan p G + → (F⊣p! : F ⊣ p !) + → (G : Functor C D) + → is-lan p G (F .F₀ G) (F⊣p! ._⊣_.unit .η G) adjoint→Lan F F⊣p! G = ext where open Lan + open is-lan module F⊣p! = _⊣_ F⊣p! - ext : Lan p G - ext .Ext = F .F₀ G - ext .eta = F⊣p!.unit .η G + ext : is-lan p G _ _ ext .σ α = R-adjunct F⊣p! α ext .σ-comm {M = M} {α = α} = Nat-path λ a → D.pullr (sym (F⊣p!.unit .is-natural _ _ _) ηₚ a) diff --git a/src/Cat/Functor/Kan/Nerve.lagda.md b/src/Cat/Functor/Kan/Nerve.lagda.md index 3d1d3f76b..e2496c5e1 100644 --- a/src/Cat/Functor/Kan/Nerve.lagda.md +++ b/src/Cat/Functor/Kan/Nerve.lagda.md @@ -6,18 +6,21 @@ description: | --- ```agda -{-# OPTIONS -vtc.decl:5 --lossy-unification -WnoEmptyWhere #-} +{-# OPTIONS -vtc.decl:5 -WnoEmptyWhere #-} +open import Cat.Instances.Functor.Compose open import Cat.Instances.Shape.Terminal +open import Cat.Functor.Kan.Pointwise open import Cat.Diagram.Colimit.Base open import Cat.Instances.Functor +open import Cat.Functor.Kan.Base open import Cat.Diagram.Initial open import Cat.Functor.Adjoint open import Cat.Instances.Comma +open import Cat.Functor.Base open import Cat.Functor.Hom -open import Cat.Functor.Kan open import Cat.Prelude -import Cat.Functor.Reasoning as Functor-kit +import Cat.Functor.Reasoning as Func import Cat.Reasoning module Cat.Functor.Kan.Nerve where @@ -27,8 +30,9 @@ module Cat.Functor.Kan.Nerve where ```agda private variable o κ : Level -open Functor +open Func open _=>_ +open is-lan ``` --> @@ -71,137 +75,264 @@ Kan extension] of $F$ along the [Yoneda embedding] $\yo$, which can be given by the _restricted_ Yoneda embedding functor $X \mapsto \hom(F(-), X)$. -[left Kan extension]: Cat.Functor.Kan.html +[left Kan extension]: Cat.Functor.Kan.Base.html [Yoneda embedding]: Cat.Functor.Hom.html -[computed]: Cat.Functor.Kan.html#a-formula +[computed]: Cat.Functor.Kan.Pointwise.html ```agda - Nerve : {D : Precategory κ κ} → Functor D C → Functor C (PSh κ D) - Nerve F .F₀ x .F₀ y = el (C.Hom (F .F₀ y) x) (C.Hom-set _ _) - Nerve F .F₀ x .F₁ f g = g C.∘ F .F₁ f - Nerve F .F₀ x .F-id = funext λ x → C.elimr (F .F-id) - Nerve F .F₀ x .F-∘ f g = funext λ x → C.pushr (F .F-∘ _ _) + Nerve : Functor D (PSh κ C) + Nerve .F₀ c = Hom-into D c F∘ Functor.op F + Nerve .F₁ f = よ₁ D f ◂ (Functor.op F) + Nerve .F-id = Nat-path (λ _ → funext λ _ → D.idl _) + Nerve .F-∘ _ _ = Nat-path (λ _ → funext λ _ → sym (D.assoc _ _ _)) +``` - Nerve F .F₁ f .η _ g = f C.∘ g - Nerve F .F₁ _ .is-natural _ _ _ = funext λ _ → C.assoc _ _ _ +The action of $F$ on morphisms assembles into a natural transformation +$\yo_\cC \To \rm{Nerve}(F)F$, which is universal in the following sense: +the nerve functor associated to $F$ is the [left Kan extension] of $\cC$'s +[yoneda embedding] along $F$. - Nerve F .F-id = Nat-path λ _ → funext λ _ → C.idl _ - Nerve F .F-∘ f g = Nat-path λ _ → funext λ _ → sym (C.assoc _ _ _) +```agda + coapprox : よ C => Nerve F∘ F + coapprox .η x .η y f = F.₁ f + coapprox .η x .is-natural _ _ _ = funext λ _ → F.F-∘ _ _ + coapprox .is-natural _ _ _ = Nat-path λ _ → funext λ _ → F.F-∘ _ _ ``` -The realisation left adjoint is constructed by general abstract -nonsense. +~~~{.quiver} +\[\begin{tikzcd} + {\cC} && {\psh(\cC)} \\ + \\ + {\cD} + \arrow["{\rm{yo}_\cC}", from=1-1, to=1-3] + \arrow["F"', from=1-1, to=3-1] + \arrow["{\rm{Nerve}(F)}"', from=3-1, to=1-3] +\end{tikzcd}\] +~~~ + +If we're given a functor $M : \cD \to \psh(\cC)$ and natural +transformation $\yo_\cC \to MF$, we can produce a natural transformation +$\alpha : \rm{Nerve}(F) \to M$ using $M$'s own action on morphisms, +since --- in components --- we have to transform an element of $f : +\hom(Fc,d)$ into one of $M(d)(c)$. Using $\alpha$, we obtain an element +of $M(Fc)(c)$, which maps under $M(f)$ to what we want. + +```agda + Nerve-is-lan : is-lan F (よ C) Nerve coapprox + Nerve-is-lan .σ {M = M} α .η d .η c f = + M .F₁ f .η c (α .η c .η c C.id) +``` + +
+ +The construction is concluded by straightforward, though tedious, +computations using naturality. It's not very enlightening. + + +```agda + Nerve-is-lan .σ {M = M} α .η d .is-natural x y f = + funext λ g → + M.₁ (g D.∘ F.₁ f) .η y (α .η y .η y C.id) ≡⟨ M.F-∘ g (F .F₁ f) ηₚ _ $ₚ _ ⟩ + M.₁ g .η y (M .F₁ (F.₁ f) .η y (α .η y .η y C.id)) ≡˘⟨ ap (M.F₁ g .η y) (α .is-natural _ _ _ ηₚ _ $ₚ _) ⟩ + M.₁ g .η y (α .η x .η y ⌜ f C.∘ C.id ⌝) ≡⟨ ap! C.id-comm ⟩ + M.₁ g .η y (α .η x .η y (C.id C.∘ f)) ≡⟨ ap (M.₁ g .η y) (α .η _ .is-natural _ _ _ $ₚ _) ⟩ + M.₁ g .η y (M.₀ (F.₀ x) .F₁ f (α .η x .η x C.id)) ≡⟨ M.₁ g .is-natural _ _ _ $ₚ _ ⟩ + M.₀ d .F₁ f (M.₁ g .η x (α .η x .η x C.id)) ∎ + where module M = Functor M + + Nerve-is-lan .σ {M = M} α .is-natural x y f = + Nat-path λ z → funext λ g → M .F-∘ f g ηₚ _ $ₚ _ + + Nerve-is-lan .σ-comm {M = M} {α = α} = + Nat-path λ x → Nat-path λ y → funext λ f → + M.₁ (F.₁ f) .η y (α .η y .η y C.id) ≡˘⟨ α .is-natural _ _ _ ηₚ _ $ₚ _ ⟩ + α .η x .η y (f C.∘ C.id) ≡⟨ ap (α .η x .η y) (C.idr _) ⟩ + α .η x .η y f ∎ + where module M = Functor M + + Nerve-is-lan .σ-uniq {M = M} {α = α} {σ′ = σ′} p = + Nat-path λ x → Nat-path λ y → funext λ f → + M.₁ f .η y (α .η y .η y C.id) ≡⟨ ap (M.₁ f .η y) (p ηₚ _ ηₚ _ $ₚ _) ⟩ + M.₁ f .η y (σ′ .η _ .η y ⌜ F.₁ C.id ⌝) ≡⟨ ap! F.F-id ⟩ + M.₁ f .η y (σ′ .η _ .η y D.id) ≡˘⟨ σ′ .is-natural _ _ _ ηₚ _ $ₚ _ ⟩ + σ′ .η x .η y (f D.∘ D.id) ≡⟨ ap (σ′ .η x .η y) (D.idr _) ⟩ + σ′ .η x .η y f ∎ + where module M = Functor M +``` + +
+Since we have assumed $\cC$ is $\kappa$-small and $\cD$ is +$\kappa$-cocomplete, any functor $F : \cC \to \cD$ admits an extension +$\Lan_{\yo}(F)$. This instance of generalised abstract nonsense is the +left adjoint we were after, the "realisation" functor. + ```agda - Realisation : {D : Precategory κ κ} → Functor D C → Functor (PSh κ D) C - Realisation {D} F = cocomplete→lan cocompl (よ D) F .Lan.Ext + Realisation : Functor (PSh κ C) D + Realisation = Lan.Ext (cocomplete→lan (よ C) F cocompl) - Realisation⊣Nerve - : {D : Precategory κ κ} (F : Functor D C) - → Realisation F ⊣ Nerve F -``` + approx : F => Realisation F∘ よ C + approx = Lan.eta (cocomplete→lan (よ C) F cocompl) -The construction of the nerve-realisation adjunction is done below in -components, but understanding it is not necessary: Either ponder the -$\Delta \mono \strcat$ example from above, or take it as a foundational -assumption. However, if you're feeling particularly brave, feel free to -look at the code. Godspeed. + Realisation-is-lan : is-lan (よ C) F Realisation approx + Realisation-is-lan = Lan.has-lan (cocomplete→lan (よ C) F cocompl) +``` + + +```agda + Realisation⊣Nerve : Realisation F cocompl ⊣ Nerve F + Realisation⊣Nerve = adj where open _⊣_ open ↓Obj open ↓Hom - module F = Functor-kit F +``` + +Recall that $\Lan_{\yo}(F)(P)$ is defined pointwise as the colimit under +the diagram + +$$ +(\yo_\cC \downarrow P) \to C \xto{F} D\text{,} +$$ - hom′ : (P : Functor (D ^op) (Sets κ)) (i : D.Ob) +where we can identify $\yo_\cC \downarrow P$ with the [category of +elements] of the presheaf $P$. Any object $i : \cC$ with +section $P(i)$ defines an object of this category, which in the proof +below we denote `elem`{.Agda}. + +[category of elements]: Cat.Instances.Elements.html + +```agda + elem : (P : Functor (C ^op) (Sets κ)) (i : C.Ob) → (arg : ∣ P .F₀ i ∣) - → ↓Obj (よ D) _ - hom′ P i arg .x = i - hom′ P i arg .y = tt - hom′ P i arg .map .η j h = P .F₁ h arg - hom′ P i arg .map .is-natural _ _ f = funext λ _ → happly (P .F-∘ _ _) _ - - Shape : (c : C.Ob) → Functor (よ D ↘ Nerve F .F₀ c) C - Shape c = F F∘ Dom (よ D) (const! (Nerve F .F₀ c)) - - cocone′ : ∀ ob → Cocone (Shape ob) - cocone′ ob .coapex = ob - cocone′ ob .ψ obj = obj .map .η _ D.id - cocone′ ob .commutes {x} {y} f = - sym (y .map .is-natural _ _ _) $ₚ _ - ∙ ap (y .map .η (x .↓Obj.x)) D.id-comm-sym - ∙ f .sq ηₚ _ $ₚ _ -``` - -Before proceeding, take a moment to appreciate the beauty of the -adjunction unit and counit, and you'll see that it _makes sense_ that -nerve and realisation are adjoints: The unit is given by the -coprojections defining the left Kan extension as a colimit, and the -counit is given by the unique "colimiting" map _from_ that colimit. + → ↓Obj (よ C) (const! P) + elem P i arg .x = i + elem P i arg .y = tt + elem P i arg .map .η j h = P .F₁ h arg + elem P i arg .map .is-natural _ _ f = funext λ _ → happly (P .F-∘ _ _) _ +``` + +The adjunction unit is easiest to describe: we must come up with a map +$F(i) \to \Lan_{\yo}(F)(P)$, in the context of an object $i : \cC$ and +section $a : P(i)$. This is an object in the diagram that we took a +colimit over, so the coprojection $\psi$ has a component expressing +precisely what we need. Naturality is omitted from this page for +brevity, but it follows from $\psi$'s universal property. ```agda - adj : Realisation F ⊣ Nerve F - adj .unit .η P .η i arg = cocompl _ .bot .ψ (hom′ P i arg) - adj .counit .η ob = cocompl _ .has⊥ (cocone′ ob) .centre .hom - - adj .unit .η P .is-natural x y f = funext λ arg → - sym $ cocompl (F F∘ Dom (よ D) (Const P)) .bot .commutes - (record { sq = Nat-path (λ _ → funext λ _ → P .F-∘ _ _ $ₚ _) }) - - adj .unit .is-natural x y f = Nat-path λ i → funext λ arg → sym $ - cocompl (F F∘ Dom (よ D) (Const x)) .has⊥ (lan-approximate cocompl (よ D) F f) - .centre .commutes (hom′ x i arg) - ∙ ap (cocompl (F F∘ Dom (よ D) (Const y)) .bot .ψ) - (↓Obj-path _ _ refl refl (Nat-path λ _ → funext λ _ → f .is-natural _ _ _ $ₚ _)) - - adj .counit .is-natural x y f = ap hom $ - is-contr→is-prop (cocompl _ .has⊥ cocone₂) - (cocone-hom _ λ o → - C.pullr (cocompl (Shape x) .has⊥ _ .centre .commutes o) - ∙ cocompl (Shape y) .has⊥ (cocone′ _) .centre .commutes _) - (cocone-hom _ λ o → - C.pullr (cocompl _ .has⊥ (cocone′ x) .centre .commutes _)) - where - cocone₂ : Cocone (Shape x) - cocone₂ .coapex = y - cocone₂ .ψ ob = f C.∘ ob .map .η _ D.id - cocone₂ .commutes {x₂} {y₂} f = - C.pullr ( sym (happly (y₂ .map .is-natural _ _ _) _) - ∙ ap (y₂ .map .η _) (sym D.id-comm)) - ∙ ap (_ C.∘_) (f .sq ηₚ _ $ₚ D.id) - - adj .zig {A} = ap hom $ - is-contr→is-prop (cocompl (F F∘ Dom (よ D) (const! A)) .has⊥ (cocompl _ .bot)) - (cocone-hom _ λ o → C.pullr ( - cocompl (F F∘ Dom (よ D) (Const A)) .has⊥ _ .centre .commutes o) - ·· cocompl (Shape (Realisation F .F₀ A)) .has⊥ _ .centre .commutes _ - ·· ap (cocompl (F F∘ Dom (よ D) (const! A)) .bot .ψ) - {x = hom′ A (o .x) (o .map .η (o .x) D.id)} - (↓Obj-path _ _ _ refl (Nat-path λ x → funext λ _ → - sym (o .map .is-natural _ _ _ $ₚ _) ∙ ap (o .map .η x) (D.idl _)))) - (cocone-hom _ λ o → C.idl _) - - adj .zag {B} = Nat-path λ x → funext λ a → - cocompl (F F∘ Dom (よ D) (const! (Nerve F .F₀ B))) .has⊥ (cocone′ B) - .centre .commutes (hom′ _ x a) - ∙ F.elimr refl + adj : Realisation F cocompl ⊣ Nerve F + adj .unit .η P .η i a = ↓colim.ψ P (elem P i a) ``` + +In the other direction, we're mapping _out_ of a colimit, into an +arbitrary object $o : \cD$. It will suffice to come up with a family of +compatible maps $F(j) \to o$, where $j$ is an object in the comma +category + +$$ +\yo_\cC \downarrow (\yo_\cD(o) \circ F)\text{.} +$$ + +These, by definition, come with (natural) functions $\hom_\cC(-, j) \to +\hom_\cD(F-, o)$, which we can evaluate at the identity morphism to get +the map we want. That this assembles into a map from the colimit follows +from that same naturality: + +``` + adj .counit .η ob = ↓colim.universal _ (λ j → j .map .η _ C.id) comm + where abstract + comm + : ∀ {x y} (f : ↓Hom (よ C) (const! (Nerve F .F₀ ob)) x y) + → y .map .η _ C.id D.∘ F.₁ (f .α) ≡ x .map .η _ C.id + comm {x} {y} f = + y .map .η _ C.id D.∘ F.₁ (f .α) ≡˘⟨ y .map .is-natural _ _ _ $ₚ _ ⟩ + y .map .η _ (C.id C.∘ f .α) ≡⟨ ap (y .map .η _) C.id-comm-sym ⟩ + y .map .η _ (f .α C.∘ C.id) ≡⟨ f .sq ηₚ _ $ₚ _ ⟩ + x .map .η (x .↓Obj.x) C.id ∎ +``` + +Naturality of this putative counit follows from the uniqueness of maps +from a colimit, and the triangle identities follow because, essentially, +"matching" on a colimit after applying one of the coprojections does the +obvious thing. + + diff --git a/src/Cat/Functor/Kan/Pointwise.lagda.md b/src/Cat/Functor/Kan/Pointwise.lagda.md new file mode 100644 index 000000000..0d99daf1e --- /dev/null +++ b/src/Cat/Functor/Kan/Pointwise.lagda.md @@ -0,0 +1,621 @@ +```agda +open import Cat.Diagram.Colimit.Representable +open import Cat.Functor.Hom.Representable +open import Cat.Functor.Kan.Representable +open import Cat.Instances.Functor.Compose +open import Cat.Instances.Functor.Compose +open import Cat.Instances.Shape.Terminal +open import Cat.Diagram.Colimit.Base +open import Cat.Diagram.Limit.Base +open import Cat.Functor.Kan.Unique +open import Cat.Functor.Coherence +open import Cat.Instances.Functor +open import Cat.Functor.Kan.Base +open import Cat.Instances.Comma +open import Cat.Functor.Base +open import Cat.Functor.Hom +open import Cat.Prelude + +import Cat.Functor.Reasoning as Func +import Cat.Reasoning + +module Cat.Functor.Kan.Pointwise where +``` + +# Pointwise Kan Extensions + +One useful perspective on [Kan extensions] is that they are +generalizations of (co)limits; in fact, we have defined (co)limits as a +special case of Kan extensions! This means that many theorems involving +(co)limits can be directly generalized to theorems of Kan extensions. +A concrete example of this phenomena is the fact that right adjoints don't +just preserve limits, they preserve *all* right extensions! + +[Kan extensions]: Cat.Functor.Kan.Base.html + +However, this pattern of generalization fails in one critical way: +[corepresentable functors preserve limits], but corepresentable functors +do _not_ necessarily preserve Kan extensions! This seemingly slight +issue has far-reaching consequences, to the point that some authors +write off general Kan extensions entirely. + +Instead of throwing the whole thing out, an alternative is to focus on +only the Kan extensions that _are_ preserved by arbitrary +(co)representables; we call such extensions **pointwise**. + +[corepresentable functors preserve limits]: Cat.Functor.Hom.Representable.html#corepresentable-functors-preserve-limits + + + + +```agda + is-pointwise-lan : ∀ {eta : G => E F∘ F} → is-lan F G E eta → Type _ + is-pointwise-lan lan = + ∀ (x : D.Ob) → preserves-lan (Functor.op (Hom-into D x)) lan + + is-pointwise-ran : ∀ {eps : E F∘ F => G} → is-ran F G E eps → Type _ + is-pointwise-ran ran = + ∀ (x : D.Ob) → preserves-ran (Hom-from D x) ran +``` + + + +As noted earlier, limits and colimits are pointwise Kan extensions. + +```agda + limit→pointwise + : {eps : Const x => Dia} + → (lim : is-limit Dia x eps) + → is-pointwise-ran lim + limit→pointwise lim x = Hom-from-preserves-limits x lim + + colimit→pointwise + : {eta : Dia => Const x} + → (colim : is-colimit Dia x eta) + → is-pointwise-lan colim + colimit→pointwise colim x = よ-reverses-colimits x colim +``` + +## Computing Pointwise Extensions + +One useful fact about pointwise left Kan extensions (resp. right) is +that they can be computed via colimits (resp. limits). We will focus on +the left extensions for the moment. Given functors $F : \cC \to \cC'$ +and $G : \cC \to \cD$, if $\cC$ is a [$\kappa$-small] category, $\cC'$ +is _locally_ $\kappa$-small, and $\cD$ has $\kappa$-small colimits, then +$\Lan_F(G)$ exists _and_ is pointwise. + +[$\kappa$-small]: 1Lab.intro.html#universes-and-size-issues + + + +The big idea of our construction is to approximate $G$ by gluing +together "all the ways that $C$ is able to see $C'$ through $G$". +Concretely, this means looking at the [comma category] $F \swarrow c'$ +for a given $c' : \cC'$. Objects in this category are pairs of an object +$c: \cC$ and morphisms $\cC'(F(c), c')$. If we apply the “domain” +projection $F \swarrow c' \to \cC$, we obtain a diagram in $\cC$ +encoding the hand-wavy intuition from before. + +[comma category]: Cat.Instances.Comma.html + +To finish the construction, we take our diagram in $\cC$ and +post-compose with $G$ to obtain a diagram + +$$ +F \swarrow c' \to \cC \xto{G} \cD\text{.} +$$ + +Since $F \swarrow c'$ is put together out of objects in $\cC$ and +morphisms in $\cC'$, it is also a $\kappa$-small category, so these +diagrams all have colimits in $\cD$. + +```agda + ↓Dia : (c' : C'.Ob) → Functor (F ↘ c') D + ↓Dia c' = G F∘ Dom F (Const c') +``` + +In fact, we can weaken the precondition from cocompleteness of $\cD$ to +having colimits of these comma-category-shaped diagrams. + +```agda + comma-colimits→lan + : (∀ (c' : C'.Ob) → Colimit (↓Dia c')) + → Lan F G + comma-colimits→lan ↓colim = lan module comma-colimits→lan where + module ↓colim c' = Colimit (↓colim c') +``` + +Taking the colimit at each $c' : \cC'$ gives an assignment of objects +$F'(c') : \cD$, so we're left with extending it to a proper functor +$\cC' \to \cD$. Coming up with morphisms in the image of $F'$ isn't too +complicated, and universality of colimits guarantees the functoriality +constraints are satisfied. + +```agda + F′ : Functor C' D + F′ .F₀ c' = ↓colim.coapex c' + F′ .F₁ f = ↓colim.universal _ + (λ j → ↓colim.ψ _ (↓obj (f C'.∘ j .map))) + (λ f → ↓colim.commutes _ (↓hom {β = f .β} (C'.pullr (f .sq) + ·· C'.elim-inner refl + ·· sym (C'.idl _)))) +``` + + +Next, we need to show that the functor we constructed actually is a left +extension. The first step will be to construct a "unit" natural +transformation $\eta : G \to F'F$, which, in components, means we're +looking for maps $G(x) \to F'(F(x))$. Since each $F'(-)$ is a colimit, +we can use the coprojections! + +```agda + eta : G => F′ F∘ F + eta .η c = ↓colim.ψ (F .F₀ c) (↓obj C'.id) +``` + +This "type checks" because the colimit coprojections for our $F +\swarrow c'$-colimits, essentially, convert maps $C'(F(X), Y)$ into +maps $G(X) \to F'(Y)$. If we take the identity $C'(F(c), F(c))$, we get +what we wanted: a map $G(c) \to F'(F(c))$. + +```agda + eta .is-natural x y f = + ↓colim.commutes (F₀ F y) (↓hom (ap (C'.id C'.∘_) (sym (C'.idr _)))) + ∙ sym (↓colim.factors _ _ _) +``` + +For the other obligation, suppose we're given some $M : \cC' \to \cD$ +and natural transformation $\alpha : G \to MF$. How do we extend it to a +transformation $F' \to M$? By "matching" on the colimit, with a slight +adjustment to $\alpha$: + +```agda + has-lan : is-lan F G F′ eta + has-lan .σ {M = M} α .η c' = ↓colim.universal c' + (λ j → M .F₁ (j .map) D.∘ α .η (j .x)) + (λ f → D.pullr (α .is-natural _ _ _) + ∙ pulll M ((f .sq) ∙ C'.idl _)) + has-lan .σ {M = M} α .is-natural x y f = ↓colim.unique₂ _ _ + (λ f → D.pullr (α .is-natural _ _ _) + ∙ pulll M (C'.pullr (f .sq) ∙ C'.elim-inner refl)) + (λ j → D.pullr (↓colim.factors _ _ _) + ∙ ↓colim.factors _ _ _) + (λ j → D.pullr (↓colim.factors _ _ _) + ∙ D.pulll (sym (M .F-∘ _ _))) + +``` + + +Finally, commutativity and uniqueness follow from the corresponding +properties of colimits. + +```agda + has-lan .σ-comm {M = M} = Nat-path λ c → + ↓colim.factors _ _ _ ∙ D.eliml (M .F-id) + has-lan .σ-uniq {M = M} {α = α} {σ′ = σ′} p = Nat-path λ c' → + sym $ ↓colim.unique _ _ _ _ λ j → + σ′ .η c' D.∘ ↓colim.ψ c' j ≡⟨ ap (λ ϕ → σ′ .η c' D.∘ ↓colim.ψ c' ϕ) (↓Obj-path _ _ refl refl (sym (C'.idr _))) ⟩ + (σ′ .η c' D.∘ ↓colim.ψ c' (↓obj (j .map C'.∘ C'.id))) ≡⟨ D.pushr (sym $ ↓colim.factors _ _ _) ⟩ + (σ′ .η c' D.∘ ↓colim.universal _ _ _) D.∘ ↓colim.ψ _ _ ≡⟨ D.pushl (σ′ .is-natural _ _ _) ⟩ + M .F₁ (j .map) D.∘ (σ′ .η _ D.∘ ↓colim.ψ _ (↓obj C'.id)) ≡˘⟨ (D.refl⟩∘⟨ (p ηₚ j .x)) ⟩ + M .F₁ (j .map) D.∘ α .η (j .x) ∎ +``` + +All that remains is to bundle up the data! + +```agda + lan : Lan F G + lan .Lan.Ext = F′ + lan .Lan.eta = eta + lan .Lan.has-lan = has-lan +``` + +And, if $\cD$ is $\kappa$-cocomplete, then it certainly has the required +colimits: we can "un-weaken" our result. + +```agda + cocomplete→lan : is-cocomplete ℓ ℓ D → Lan F G + cocomplete→lan colimits = comma-colimits→lan (λ c' → colimits (↓Dia c')) +``` + + +Next, we shall show that the left extension we just constructed is +pointwise. To do this, we will show a more general fact: if $H : D \to +E$ is $\kappa$-cocontinuous, then it also preserves extensions "built +from" colimits. + + + +The mathematical content of the proof is quite straightforward: +$H$ preserves the colimits used to construct the extension, so we +can perform the exact same process in $\cE$ to obtain a left extension. +However, the mechanical content of this proof is less pleasant: we +end up being off by a bunch of natural isomorphisms. + +```agda + preserves-colimits→preserves-pointwise-lan + : ∀ {o″ ℓ″} {E : Precategory o″ ℓ″} + → (colimits : is-cocomplete ℓ ℓ D) + → (H : Functor D E) + → is-cocontinuous ℓ ℓ H + → preserves-lan H (Lan.has-lan (cocomplete→lan F G colimits)) + preserves-colimits→preserves-pointwise-lan {E = E} colimits H cocont = + natural-isos→is-lan idni idni HF′-cohere fixup $ + comma-colimits→lan.has-lan F (H F∘ G) H-↓colim + where + module E = Cat.Reasoning E + open make-natural-iso + open Func + + ↓colim : (c' : C'.Ob) → Colimit (G F∘ Dom F (Const c')) + ↓colim c' = colimits (G F∘ Dom F (Const c')) + + module ↓colim c' = Colimit (↓colim c') + + H-↓colim : (c' : C'.Ob) → Colimit ((H F∘ G) F∘ Dom F (Const c')) + H-↓colim c' = + natural-iso→colimit ni-assoc $ + preserves-colimit.colimit cocont (↓colim c') + + module H-↓colim c' = Colimit (H-↓colim c') +``` + +
+ +Unfortunately, proof assistants. By "a bunch", we really do mean _a +bunch_. This fold contains all the required coherence data, which ends +up not being very interesting. + + +```agda + F′ : Functor C' D + F′ = comma-colimits→lan.F′ F G λ c' → colimits (G F∘ Dom F (Const c')) + + HF′ : Functor C' E + HF′ = comma-colimits→lan.F′ F (H F∘ G) H-↓colim + + HF′-cohere : natural-iso HF′ (H F∘ F′) + HF′-cohere = to-natural-iso mi where + mi : make-natural-iso HF′ (H F∘ F′) + mi .eta c' = E.id + mi .inv c' = E.id + mi .eta∘inv _ = E.idl _ + mi .inv∘eta _ = E.idl _ + mi .natural _ _ _ = + E.idr _ + ∙ H-↓colim.unique _ _ _ _ (λ j → pulll H (↓colim.factors _ _ _)) + ∙ sym (E.idl _) + + module HF′-cohere = natural-iso HF′-cohere + + abstract + fixup : (HF′-cohere.to ◆ idnt) ∘nt comma-colimits→lan.eta F (H F∘ G) _ ∘nt idnt ≡ nat-assoc-to (H ▸ comma-colimits→lan.eta F G _) + fixup = Nat-path λ j → + (H .F₁ (↓colim.universal _ _ _) E.∘ E.id) E.∘ (H-↓colim.ψ _ _ E.∘ E.id) ≡⟨ ap₂ E._∘_ (E.idr _) (E.idr _) ⟩ + H .F₁ (↓colim.universal _ _ _) E.∘ H-↓colim.ψ _ _ ≡⟨ pulll H (↓colim.factors _ _ _) ⟩ + H .F₁ (↓colim.ψ _ _) E.∘ E.id ≡⟨ E.idr _ ⟩ + H .F₁ (↓colim.ψ _ (↓obj ⌜ C'.id C'.∘ C'.id ⌝)) ≡⟨ ap! (C'.idl _) ⟩ + H .F₁ (↓colim.ψ _ (↓obj (C'.id))) ∎ +``` +
+ +Hom functors take colimits in $\cD$ to colimits in $\Sets^op$, so by +the previous lemma, they must preserve the left extension. In other +words, the extension we constructed is pointwise. + +```agda + cocomplete→pointwise-lan + : (colim : is-cocomplete ℓ ℓ D) + → is-pointwise-lan (Lan.has-lan (cocomplete→lan F G colim)) + cocomplete→pointwise-lan colim d = + preserves-colimits→preserves-pointwise-lan + colim (Functor.op (Hom-into D d)) + (よ-reverses-colimits d) +``` + +## All Pointwise Extensions are Computed via (Co)limits + +As we've seen earlier, we can compute the extension of $F : \cC \to \cD$ +along $p : \cC \to \cC'$ when $\cD$ has enough colimits, and that this +extension is pointwise. It turns out that this is an exact +characterization of the pointwise extensions: if $L$ is a pointwise +extension of $F$ along $p$, then $\cD$ must have colimits of all +diagrams of the form $F \circ \mathrm{Dom} : p \swarrow c' \to C \to D$, +and $L$ must be computed via these colimits. This is where the name +"pointwise extension" comes from --- they really _are_ computed +pointwise! + + + +We begin by constructing a cocone for every object $c' : \cC'$. + +```agda + ↓cocone : ∀ (c' : C'.Ob) → F F∘ Dom p (const! c') => Const (L .F₀ c') + ↓cocone c' .η j = L .F₁ (j .map) D.∘ eta .η _ + ↓cocone c' .is-natural _ _ f = + D.pullr (eta .is-natural _ _ _ ) + ∙ pulll L (f .sq ∙ C'.idl _) + ∙ sym (D.idl _) +``` + +To show that the extension is computed pointwise by these extensions, +we shall appeal to the fact that [colimits are representable]. + +[colimits are representable]: Cat.Diagram.Colimit.Representable.html + +```agda + pointwise-lan→has-comma-colimits + : ∀ (c' : C'.Ob) + → is-colimit (F F∘ Dom p (const! c')) (L .F₀ c') (↓cocone c') + pointwise-lan→has-comma-colimits c' = + represents→is-colimit $ + [D,Sets].make-invertible inv invl invr + where +``` + +As $(L,\eta)$ is pointwise, we can represent every cocone $F \circ p +\searrow c' \to d$ as a natural transformation $\cC'(-,c') \to +\cD(L(-),d)$, though we do need to pass through some abstract +representability nonsense to get there. + +```agda + represent-↓cocone + : ∀ (d : D.Ob) + → F F∘ Dom p (const! c') => Const d + → Functor.op (よ₀ D d) F∘ F => Functor.op (よ₀ C' c') F∘ p + represent-↓cocone d α .η c f = α .η (↓obj f) + represent-↓cocone d α .is-natural _ _ f = funext λ g → + α .is-natural (↓obj (g C'.∘ p .F₁ f)) (↓obj g) (↓hom (sym (C'.idl _))) + ∙ D.idl _ + + pointwise-↓cocone + : ∀ (d : D.Ob) + → (α : F F∘ Dom p (const! c') => Const d) + → Functor.op (Hom-into D d) F∘ L => Functor.op (Hom-into C' c') + pointwise-↓cocone d α = pointwise.σ d (represent-↓cocone d α) +``` + +We can use this representation to construct the required inverse, via +the usual Yoneda-like argument. + +```agda + inv : Lim[C[F-,=]] => Hom-from D (L .F₀ c') + inv .η d α = + pointwise-↓cocone d α .η c' C'.id + inv .is-natural x y f = funext λ α → + pointwise.σ-uniq y {σ′ = pointwise-↓cocone x α ∘nt (_=>_.op (よ₁ D f) ◂ L)} + (Nat-path λ c → funext λ g → D.pushr (sym (pointwise.σ-comm x ηₚ _ $ₚ _))) ηₚ c' $ₚ C'.id +``` + +
+ +To show that we've constructed an isomorphism, we'll forget the +_pointwise_, and remember that we're working with a Kan extension. + + +```agda + invl : Hom-into-inj (↓cocone c') ∘nt inv ≡ idnt + invl = Nat-path λ d → funext λ α → Nat-path λ p↓c' → + pointwise-↓cocone d α .η _ C'.id D.∘ L .Functor.F₁ (p↓c' .map) D.∘ eta .η _ ≡⟨ D.pulll (pointwise.σ d (represent-↓cocone d α) .is-natural _ _ _ $ₚ _) ⟩ + pointwise-↓cocone d α .η _ ⌜ C'.id C'.∘ p↓c' .map ⌝ D.∘ eta .η _ ≡⟨ ap! (C'.idl _) ⟩ + pointwise-↓cocone d α .η _ (p↓c' .map) D.∘ eta .η (x p↓c') ≡⟨ pointwise.σ-comm d ηₚ _ $ₚ p↓c' .map ⟩ + α .η (↓obj (p↓c' .map)) ≡⟨ ap (α .η) (↓Obj-path _ _ refl refl refl) ⟩ + α .η p↓c' ∎ + + vaguely-yoneda + : ∀ {d : D.Ob} (α : D.Hom (L .F₀ c') d) + → Functor.op (Hom-into D d) F∘ L => Functor.op (Hom-into C' c') + vaguely-yoneda α .η c'' f = α D.∘ L .F₁ f + vaguely-yoneda α .is-natural x y f = + funext λ g → D.pullr (sym (L .F-∘ _ _)) + + invr : inv ∘nt Hom-into-inj (↓cocone c') ≡ idnt + invr = Nat-path λ d → funext λ α → + pointwise.σ-uniq d {σ′ = vaguely-yoneda α} + (Nat-path λ c → funext λ f → D.assoc _ _ _) ηₚ c' $ₚ C'.id + ∙ D.elimr (L .F-id) +``` +
+ +A corollary is that if $(L, \eta)$ is a pointwise left extension along a +fully faithful functor, then $\eta$ is a natural isomorphism. + +```agda + ff→pointwise-lan-ext : is-fully-faithful p → is-natural-invertible eta +``` + +The idea is to use the fact that $L$ is computed via colimits to +construct an inverse to $\eta$. In particular, we use the universal map +out of each colimit, relying on the full faithfulness of $p$ to +construct the requisite cocone. + +```agda + ff→pointwise-lan-ext p-ff = + componentwise-invertible→invertible eta λ c → + D.make-invertible (inv c) + (pointwise-colim.unique₂ _ _ + (λ f → D.pullr (eta .is-natural _ _ _) + ∙ pulll L (sym (p .F-∘ _ _) ∙ path f)) + (λ j → D.pullr (pointwise-colim.factors _ {j = j} _ _) + ∙ eta .is-natural _ _ _) + (λ j → D.idl _ + ∙ ap₂ D._∘_ (ap (L .F₁) (sym (equiv→counit p-ff (j .map)))) refl)) + (pointwise.σ-comm _ ηₚ c $ₚ C'.id + ∙ elimr F (ap (equiv→inverse p-ff) (sym (p .F-id)) ∙ equiv→unit p-ff _)) +``` + + diff --git a/src/Cat/Functor/Kan/Representable.lagda.md b/src/Cat/Functor/Kan/Representable.lagda.md new file mode 100644 index 000000000..589893fca --- /dev/null +++ b/src/Cat/Functor/Kan/Representable.lagda.md @@ -0,0 +1,159 @@ +```agda +open import Cat.Functor.Hom.Representable +open import Cat.Instances.Functor.Compose +open import Cat.Instances.Functor +open import Cat.Functor.Kan.Base +open import Cat.Functor.Hom +open import Cat.Prelude + +import Cat.Reasoning + +module Cat.Functor.Kan.Representable where +``` + +## Representability of Kan Extensions + +Like most constructions with a universal property, we can phrase the +definition of [Kan extensions] in terms of an equivalence of $\hom$ +functors. This rephrasing lets us construct extensions in terms of +chains of natural isomorphisms, which can be very handy! + +[Kan extensions]: Cat.Functor.Kan.Base.html + + + +Let $p : \cC \to \cC'$, and $F : \cC \to \cD$, and $(G, \eta)$ be +a candidate for a left extension, as in the following diagram. + +~~~{.quiver} +\begin{tikzcd} + && {\cC'} \\ + \\ + \cC &&&& \cD + \arrow[""{name=0, anchor=center, inner sep=0}, "F"', from=3-1, to=3-5] + \arrow["p"', from=3-1, to=1-3] + \arrow["G"', from=1-3, to=3-5] + \arrow["\eta"', shorten <=4pt, shorten >=7pt, Rightarrow, from=1-3, to=0] +\end{tikzcd} +~~~ + +Any such pair $(G, \eta)$ induces a natural transformation +$D^{C'}(G, -) \to D^{C}(F, p(-))$. + +```agda + Hom-from-! + : F => G F∘ p + → Hom-from Cat[ C' , D ] G => Hom-from Cat[ C , D ] F F∘ (p !) + Hom-from-! eta .η H α = (α ◂ p) ∘nt eta + Hom-from-! eta .is-natural H K α = funext λ β → [C,D].pushl ◂-distribl +``` + +If this natural transformation is an isomorphism, then $(G, \eta)$ is a +left Kan extension of $F$ along $p$. + +```agda + represents→is-lan + : (eta : F => G F∘ p) + → is-natural-invertible (Hom-from-! eta) + → is-lan p F G eta + represents→is-lan eta nat-inv = lan where +``` + +This may seem somewhat difficult to prove at first glance, but it ends +up being an exercise in shuffling data around. We can use the inverse +to `Hom-from-! eta` to obtain the universal map of the extension, and +factorization/uniqueness follow directly from the fact that we have +a natural isomorphism. + +```agda + module nat-inv = is-natural-invertible nat-inv + + lan : is-lan p F G eta + lan .σ {M} α = nat-inv.inv .η M α + lan .σ-comm {M} {α} = nat-inv.invl ηₚ M $ₚ α + lan .σ-uniq {M} {α} {σ′} q = ap (nat-inv.inv .η M) q ∙ nat-inv.invr ηₚ M $ₚ σ′ +``` + +Furthermore, if $(G, \eta)$ is a left extension, then we can show that +`Hom-from-! eta` is a natural isomorphism. The proof is yet another +exercise in moving data around. + +```agda + is-lan→represents + : {eta : F => G F∘ p} + → is-lan p F G eta + → is-natural-invertible (Hom-from-! eta) + is-lan→represents {eta} lan = + to-is-natural-invertible inv + (λ M → funext λ α → lan .σ-comm) + (λ M → funext λ α → lan .σ-uniq refl) + where + inv : Hom-from Cat[ C , D ] F F∘ (p !) => Hom-from Cat[ C' , D ] G + inv .η M α = lan .σ α + inv .is-natural M N α = funext λ β → + lan .σ-uniq (Nat-path λ _ → D.pushr (sym $ lan .σ-comm ηₚ _)) +``` + + + +```agda + lan→represents : Lan p F → Corepresentation (Hom-from Cat[ C , D ] F F∘ (p !)) + lan→represents lan .corep = lan .Ext + lan→represents lan .corepresents = + (is-natural-invertible→natural-iso (is-lan→represents (lan .has-lan))) ni⁻¹ + + represents→lan : Corepresentation (Hom-from Cat[ C , D ] F F∘ (p !)) → Lan p F + represents→lan has-corep .Ext = has-corep .corep + represents→lan has-corep .eta = has-corep .corepresents .from .η _ idnt + represents→lan has-corep .has-lan = + represents→is-lan (Corep.to has-corep idnt) $ + to-is-natural-invertible (has-corep .corepresents .to) + (λ M → funext λ α → + (Corep.from has-corep α ◂ p) ∘nt Corep.to has-corep idnt ≡˘⟨ has-corep .corepresents .from .is-natural _ _ _ $ₚ idnt ⟩ + Corep.to has-corep (Corep.from has-corep α ∘nt idnt) ≡⟨ ap (Corep.to has-corep) ([C',D].idr _) ⟩ + Corep.to has-corep (Corep.from has-corep α) ≡⟨ Corep.ε has-corep α ⟩ + α ∎) + (λ M → funext λ α → + Corep.from has-corep ((α ◂ p) ∘nt Corep.to has-corep idnt) ≡⟨ has-corep .corepresents .to .is-natural _ _ _ $ₚ _ ⟩ + α ∘nt Corep.from has-corep (Corep.to has-corep idnt) ≡⟨ [C',D].elimr (Corep.η has-corep idnt) ⟩ + α ∎) +``` diff --git a/src/Cat/Functor/Kan/Right.lagda.md b/src/Cat/Functor/Kan/Right.lagda.md deleted file mode 100644 index 733fe93f4..000000000 --- a/src/Cat/Functor/Kan/Right.lagda.md +++ /dev/null @@ -1,133 +0,0 @@ -```agda -open import Cat.Instances.Functor.Compose -open import Cat.Diagram.Limit.Base -open import Cat.Instances.Functor -open import Cat.Diagram.Duals -open import Cat.Functor.Kan -open import Cat.Prelude - -module Cat.Functor.Kan.Right where -``` - - - -# Right Kan extensions - -Dually to our setup for a [left Kan extension], we have **right Kan -extensions**: The (suitably weakly) [terminal] solution to the problem of -lifting a functor $F : \cC \to \cD$ along a functor $p : \cC' -\to \cC$. We picture the situation as in the following commutative -diagram: - -[left Kan extension]: Cat.Functor.Kan.html -[terminal]: Cat.Diagram.Terminal.html - -~~~{.quiver} -\[\begin{tikzcd} - {\mathcal{C}} && {\mathcal{D}} \\ - \\ - {\mathcal{C}'} - \arrow["F", from=1-1, to=1-3] - \arrow["p"', from=1-1, to=3-1] - \arrow[""{name=0, anchor=center, inner sep=0}, "{\mathrm{Ran}_pF}"{description}, curve={height=-12pt}, from=3-1, to=1-3] - \arrow[""{name=1, anchor=center, inner sep=0}, "M"{description}, curve={height=12pt}, from=3-1, to=1-3] - \arrow[shorten <=4pt, shorten >=4pt, Rightarrow, from=1, to=0] -\end{tikzcd}\] -~~~ - -Note the same warnings about "weak, directed" commutativity as for [left -Kan extensions] apply here, too. Rather than either of the triangles -commuting on the nose, we have natural transformations $\eps$ witnessing -their commutativity. - -```agda -record Ran (p : Functor C C′) (F : Functor C D) : Type (kan-lvl p F) where - field - Ext : Functor C′ D - eps : Ext F∘ p => F - - σ : {M : Functor C′ D} (α : M F∘ p => F) → M => Ext - σ-comm : {M : Functor C′ D} {β : M F∘ p => F} → eps ∘nt (σ β ◂ p) ≡ β - σ-uniq : {M : Functor C′ D} {β : M F∘ p => F} {σ′ : M => Ext} - → β ≡ eps ∘nt (σ′ ◂ p) - → σ β ≡ σ′ - - σ-uniq₂ - : {M : Functor C′ D} (β : M F∘ p => F) {σ₁′ σ₂′ : M => Ext} - → β ≡ eps ∘nt (σ₁′ ◂ p) - → β ≡ eps ∘nt (σ₂′ ◂ p) - → σ₁′ ≡ σ₂′ - σ-uniq₂ β p q = sym (σ-uniq p) ∙ σ-uniq q -``` - -The first thing we'll verify is that this construction is indeed dual to -the left Kan extension. This is straightforward enough to do, but we -have some administrative noise from all the opposite categories getting -in the way. - -```agda -module _ (p : Functor C C′) (F : Functor C D) where - open Ran - open _=>_ - - Co-lan→Ran : Lan (Functor.op p) (Functor.op F) -> Ran p F - Co-lan→Ran lan = ran where - module lan = Lan lan - ran : Ran p F - ran .Ext = Functor.op lan.Ext - ran .eps .η x = lan.eta .η x - ran .eps .is-natural x y f = sym (lan.eta .is-natural y x f) - - ran .σ {M = M} α = op (lan.σ α′) where - α′ : Functor.op F => Functor.op M F∘ Functor.op p - α′ .η x = α .η x - α′ .is-natural x y f = sym (α .is-natural y x f) - - ran .σ-comm = Nat-path λ x → lan.σ-comm ηₚ x - ran .σ-uniq {M = M} {σ′ = σ′} p = - Nat-path λ x → lan.σ-uniq {σ′ = σ′op} (Nat-path λ x → p ηₚ x) ηₚ x - where - σ′op : lan.Ext => Functor.op M - σ′op .η x = σ′ .η x - σ′op .is-natural x y f = sym (σ′ .is-natural y x f) -``` - -## Computation - -Using the helper `Co-lan→Ran`{.Agda} defined above and the formula for -computing _left_ Kan extensions, we can formulate a condition for the -existence of right Kan extensions based on the size and completeness of -the categories involved. If $\cE$ admits limits of [$\kappa$-small -diagrams], $\cC$ is $\kappa$-small, and $\cD$ is locally -$\kappa$-small, then for any $p : \cC \to \cD$ and $F : \cD \to -\cE$, the right Kan extension $\Ran_p F$ exists. - -[$\kappa$-small diagrams]: 1Lab.intro.html#universes-and-size-issues - -```agda -module _ - {o o′ ℓ κ} {C : Precategory κ κ} {D : Precategory o′ κ} {E : Precategory o ℓ} - (lims : is-complete κ κ E) (p : Functor C D) (F : Functor C E) - where - - complete→ran : Ran p F - complete→ran = - Co-lan→Ran p F $ - cocomplete→lan - (λ F → Co-limit→Colimit (E ^op) (lims (Functor.op F))) - (Functor.op p) (Functor.op F) -``` - -As before, it's impossible to cheat the size limitations for computing -Kan extensions as (co)limits, but this does not preclude the existence -of extensions in other situations. diff --git a/src/Cat/Functor/Kan/Unique.lagda.md b/src/Cat/Functor/Kan/Unique.lagda.md index 39d3a8584..08ffdb12a 100644 --- a/src/Cat/Functor/Kan/Unique.lagda.md +++ b/src/Cat/Functor/Kan/Unique.lagda.md @@ -1,161 +1,563 @@ ```agda open import Cat.Instances.Functor.Compose open import Cat.Instances.Functor +open import Cat.Functor.Kan.Base open import Cat.Functor.Base -open import Cat.Functor.Kan open import Cat.Prelude +import Cat.Functor.Reasoning import Cat.Reasoning -module Cat.Functor.Kan.Unique - +module Cat.Functor.Kan.Unique where ``` +# Uniqueness of Kan extensions + +[Kan extensions] (both left and [right]) are [universal constructions], +so they are [unique when they exist]. To get a theorem out of this +intuition, we must be careful about how the structure and the properties +are separated: Informally, we refer to the _functor_ as "the Kan +extension", but in reality, the data associated with "the Kan extension +of $F$ along $p$" also includes the natural transformation. For +accuracy, using the setup from the diagram below, we should say “$(G, +\eta)$ is the Kan extension of $F$ along $p$". + +[Kan extensions]: Cat.Functor.Kan.Base.html +[right]: Cat.Functor.Kan.Base.html#right-kan-extensions +[universal constructions]: Cat.Functor.Hom.Representable.html +[unique when they exist]: 1Lab.HLevel.html#is-prop + +~~~{.quiver .tall-15} +\[\begin{tikzcd} + C && D \\ + \\ + {C'} + \arrow["F", from=1-1, to=1-3] + \arrow["p"', from=1-1, to=3-1] + \arrow[""{name=0, anchor=center, inner sep=0}, "G"', from=3-1, to=1-3] + \arrow["\eta"', shorten <=6pt, Rightarrow, from=0, to=1-1] +\end{tikzcd}\] +~~~ + -# Uniqueness of Kan extensions +To show uniqueness, suppose that $(G_1, \eta_1)$ and $(G_2, \eta_2)$ and +both left extensions of $F$ along $p$. Diagramming this with both +natural transformations shown is a bit of a nightmare: the option which +doesn't result in awful crossed arrows is to duplicate the span $\cC' +\ot \cC \to \cD$. So, to be clear: The upper triangle and the lower +triangle _are the same_. -Since [(left) Kan extensions] are defined by a universal property (they -are partial values of a specific [adjunction]), they are unique: When -$\cD$ is a category, then the type $\Lan_p(F)$ of left Kan extensions -of $F$ along $p$ is a proposition. +~~~{.quiver} +\[\begin{tikzcd} + && \cC \\ + \\ + {\cC'} &&&& \cD \\ + \\ + && \cC + \arrow["F", from=1-3, to=3-5] + \arrow["p"', from=1-3, to=3-1] + \arrow[""{name=0, anchor=center, inner sep=0}, "{G_1}", curve={height=-12pt}, from=3-1, to=3-5] + \arrow["p", from=5-3, to=3-1] + \arrow["F"', from=5-3, to=3-5] + \arrow[""{name=1, anchor=center, inner sep=0}, "{G_2}"', curve={height=12pt}, from=3-1, to=3-5] + \arrow["{\eta_1}", shorten <=9pt, Rightarrow, from=0, to=1-3] + \arrow["{\eta_2}"', shorten <=9pt, Rightarrow, from=1, to=5-3] +\end{tikzcd}\] +~~~ -[(left) Kan extensions]: Cat.Functor.Kan.html -[adjunction]: Cat.Functor.Kan.Global.html +Recall that $(G_1, \eta_1)$ being a left extension means we can +(uniquely) factor natural transformations $F \to Mp$ through +transformations $G_1 \to M$. We want a map $G_1 \to G_2$, for which it +will suffice to find a map $F \to G_2p$ --- but $\eta_2$ is right there! +In the other direction, we can factor $\eta_1$ to get a map $G_2 \to +G_1$. Since these factorisations are unique, we have a natural +isomorphism. ```agda -Lan-is-prop : is-prop (Lan p F) -Lan-is-prop L₁ L₂ = Lan-unique module Lan-unique where + σ-inversesp + : ∀ {α : G₁ => G₂} {β : G₂ => G₁} + → (α ◂ p) ∘nt η₁ ≡ η₂ + → (β ◂ p) ∘nt η₂ ≡ η₁ + → natural-inverses α β + σ-inversesp α-factor β-factor = C′D.make-inverses + (l₂.σ-uniq₂ η₂ + (Nat-path λ j → sym (D.pullr (β-factor ηₚ j) ∙ α-factor ηₚ j)) + (Nat-path λ j → sym (D.idl _))) + (l₁.σ-uniq₂ η₁ + (Nat-path λ j → sym (D.pullr (α-factor ηₚ j) ∙ β-factor ηₚ j)) + (Nat-path λ j → sym (D.idl _))) ``` -We diagram the situation generically as follows: $p$ and $F$ are fixed, -with $G$ and $G'$ being both left extensions of $F$ along $p$. The -functor $G$ (resp. $G'$) is equipped with a natural transformation $\eta -: F \to Gp$ (resp. $\eta' : F \to G'p$), and if $M$ is equipped with -$\theta : F \to Mp$, then $G$ can cough up a _unique_ natural -transformation $\sigma_\theta : G \to M$ (resp. $\sigma'_\theta : G' \to -M$) making everything in sight commute. + + +It's immediate from the construction that this isomorphism "sends +$\eta_1$ to $\eta_2$". + +```agda + unit : (l₁.σ η₂ ◂ p) ∘nt η₁ ≡ η₂ + unit = l₁.σ-comm +``` -The isomorphism $G \cong G'$ is constructed as follows: Since $G'$ is -equipped with $\eta'$, $G$ can produce $\sigma_{\eta'} : G \to G'$; -Since $G$ is equipped with $\eta$, $G'$ can produce $\sigma'_\eta : G' -\to G$. To show $\sigma_{\eta'}\sigma'_\eta : G' \to G'$ is the -identity, note that both make "everything in sight commute", so they -inhabit a contractible space since $G'$ is an extension. The argument -for $\sigma'_\eta\sigma_{\eta'}$ is analogous. +Another _uniqueness-like_ result we can state for left extensions is the +following: Given any functor $G' : \cC' \to \cD$ and candidate "unit" +transformation $\eta' : F \to G'p$, if a left extension $\Lan_p(F)$ +sends $\eta'$ to a natural isomorphism, then $(G', \eta')$ is also a +left extension of $F$ along $p$. ```agda - Ext-unique : [C′,D].Isomorphism L₁.Ext L₂.Ext - Ext-unique = [C′,D].make-iso (L₁.σ L₂.eta) (L₂.σ L₁.eta) - ( sym (L₂.σ-uniq {α = L₂.eta} - (Nat-path λ _ → sym ( D.pullr (L₂.σ-comm ηₚ _) - ∙ L₁.σ-comm ηₚ _))) - ∙ L₂.σ-uniq (Nat-path λ _ → D.introl refl)) - ( sym (L₁.σ-uniq {α = L₁.eta} - (Nat-path λ _ → sym ( D.pullr (L₁.σ-comm ηₚ _) - ∙ L₂.σ-comm ηₚ _))) - ∙ L₁.σ-uniq (Nat-path λ _ → D.introl refl)) - - Ext-uniqueₚ : L₁.Ext ≡ L₂.Ext - Ext-uniqueₚ = cd-cat .to-path Ext-unique + is-invertible→is-lan + : ∀ {G' : Functor C′ D} {eta' : F => G' F∘ p} + → is-natural-invertible (lan.σ eta') + → is-lan p F G' eta' + is-invertible→is-lan {G' = G'} {eta'} invert = lan' where + open is-lan + open C′D.is-invertible invert + + lan' : is-lan p F G' eta' + lan' .σ α = lan.σ α ∘nt inv + lan' .σ-comm {M} {α} = Nat-path λ j → + (lan.σ α .η _ D.∘ inv .η _) D.∘ eta' .η j ≡˘⟨ D.refl⟩∘⟨ (lan.σ-comm ηₚ _) ⟩ + (lan.σ α .η _ D.∘ inv .η _) D.∘ (lan.σ eta' .η _ D.∘ eta .η j) ≡⟨ D.cancel-inner (invr ηₚ _) ⟩ + lan.σ α .η _ D.∘ eta .η j ≡⟨ lan.σ-comm ηₚ _ ⟩ + α .η j ∎ + lan' .σ-uniq {M} {α} {σ′} p = Nat-path λ j → + lan.σ α .η j D.∘ inv .η j ≡⟨ (lan.σ-uniq {σ′ = σ′ ∘nt lan.σ eta'} (Nat-path λ j → p ηₚ j ∙ D.pushr (sym (lan.σ-comm ηₚ j))) ηₚ j) D.⟩∘⟨refl ⟩ + (σ′ .η j D.∘ lan.σ eta' .η j) D.∘ inv .η _ ≡⟨ D.cancelr (invl ηₚ _) ⟩ + σ′ .η j ∎ +``` + + -The functor is not the only data associated with a left extension, -though: we must also verify that, under the identification $G \equiv G'$ -we just produced, the natural transformations $\eta$ and $\eta'$ are -also identified. This boils down to verifying, in components, that -$\sigma_{\eta'}\eta = \eta'$, but that is immediate by the specification -for $\sigma$. +Left Kan extensions are also invariant under arbitrary natural +isomorphisms. To get better definitional control, we allow “adjusting” +the resulting construction to talk about any natural transformation +which is propositionally equal to the whiskering: ```agda - eta-uniqueₚ : PathP (λ i → F => Ext-uniqueₚ i F∘ p) L₁.eta L₂.eta - eta-uniqueₚ = Nat-pathp refl _ λ _ → - Univalent.Hom-pathp-reflr-iso dcat (L₁.σ-comm ηₚ _) + natural-isos→is-lan + : (p-iso : natural-iso p p') + → (F-iso : natural-iso F F') + → (G-iso : natural-iso G G') + → (natural-iso.to G-iso ◆ natural-iso.to p-iso) ∘nt eps ∘nt natural-iso.from F-iso ≡ eps' + → is-lan p F G eps + → is-lan p' F' G' eps' ``` -
-A similar argument shows that $\sigma_j$ and $\sigma'_j$ are -also identified. + + +## Into univalent categories -Now $(G, \eta, \sigma)$ _is_ all the data of a left extension: The other -two fields are propositions, and so they are automatically identified ---- regardless of the specific isomorphism we would have exhibited. +As traditional with universal constructions, if $F : \cC \to \cD$ takes +values in a [univalent category], we can sharpen our result: the type of +left extensions of $F$ along $p$ is a proposition. + +[univalent category]: Cat.Univalent.html#univalent-categories ```agda +Lan-is-prop + : ∀ {p : Functor C C′} {F : Functor C D} → is-category D → is-prop (Lan p F) +Lan-is-prop {C = C} {C′ = C′} {D = D} {p = p} {F = F} d-cat L₁ L₂ = path where +``` + + + +That's because if $\cD$ is univalent, then [so is $[\cC', +\cD]$][Functor-is-category], so our natural isomorphism $i : G_1 \cong +G_2$ is equivalent to an identification $i' : G_1 \equiv G_2$. Then, our +tiny lemma stating that this isomorphism "sends $\eta_1$ to $\eta_2$" is +precisely the data of a dependent identification $\eta_1 \equiv \eta_2$ +over $i'$. + +[Functor-is-category]: Cat.Instances.Functor.html#Functor-is-category + +```agda + functor-path : L₁.Ext ≡ L₂.Ext + functor-path = c′d-cat .to-path Lu.unique + + eta-path : PathP (λ i → F => functor-path i F∘ p) L₁.eta L₂.eta + eta-path = Nat-pathp _ _ λ x → + Univalent.Hom-pathp-reflr-iso d-cat (Lu.unit ηₚ _) +``` + +Since being a left extension is always a proposition when applied to +$(G, \eta)$, even when the categories are not univalent, we can finish +our proof. + +```agda + path : L₁ ≡ L₂ + path i .Ext = functor-path i + path i .eta = eta-path i + path i .has-lan = + is-prop→pathp (λ i → is-lan-is-prop {p = p} {F} {functor-path i} {eta-path i}) + L₁.has-lan L₂.has-lan i +``` + + diff --git a/src/Cat/Functor/Monadic/Beck.lagda.md b/src/Cat/Functor/Monadic/Beck.lagda.md index 40399826a..8664fd1bf 100644 --- a/src/Cat/Functor/Monadic/Beck.lagda.md +++ b/src/Cat/Functor/Monadic/Beck.lagda.md @@ -157,9 +157,9 @@ can be seen below. Uniqueness of this map follows almost immediately from the algebra laws. ```agda - algebra-is-coequaliser .coequalise {F = F} {e'} p .morphism = + algebra-is-coequaliser .universal {F = F} {e'} p .morphism = e' .morphism C.∘ T.unit .η A - algebra-is-coequaliser .coequalise {F = F} {e'} p .commutes = + algebra-is-coequaliser .universal {F = F} {e'} p .commutes = (e' .morphism C.∘ unit.η A) C.∘ A.ν ≡⟨ C.extendr (unit.is-natural _ _ _) ⟩ (e' .morphism C.∘ T.M₁ A.ν) C.∘ unit.η (T.M₀ A) ≡˘⟨ ap morphism p C.⟩∘⟨refl ⟩ (e' .morphism C.∘ T.mult .η A) C.∘ unit.η (T.M₀ A) ≡⟨ C.cancelr T.right-ident ⟩ @@ -168,14 +168,14 @@ from the algebra laws. (e' .morphism C.∘ T.mult .η A) C.∘ T.M₁ (unit.η A) ≡⟨ C.pushl (e' .commutes) ⟩ F .snd .ν C.∘ T.M₁ (e' .morphism) C.∘ T.M₁ (unit.η A) ≡˘⟨ C.refl⟩∘⟨ T.M-∘ _ _ ⟩ F .snd .ν C.∘ T.M₁ (e' .morphism C.∘ unit.η A) ∎ - algebra-is-coequaliser .universal {F = F} {e'} {p} = Algebra-hom-path C $ + algebra-is-coequaliser .factors {F = F} {e'} {p} = Algebra-hom-path C $ (e' .morphism C.∘ unit.η _) C.∘ A.ν ≡⟨ C.extendr (unit.is-natural _ _ _) ⟩ (e' .morphism C.∘ T.M₁ A.ν) C.∘ unit.η _ ≡˘⟨ ap morphism p C.⟩∘⟨refl ⟩ (e' .morphism C.∘ T.mult .η _) C.∘ unit.η _ ≡⟨ C.cancelr T.right-ident ⟩ e' .morphism ∎ algebra-is-coequaliser .unique {F = F} {e'} {p} {colim} q = Algebra-hom-path C $ sym $ - e' .morphism C.∘ unit.η A ≡⟨ ap morphism q C.⟩∘⟨refl ⟩ + e' .morphism C.∘ unit.η A ≡⟨ ap morphism (sym q) C.⟩∘⟨refl ⟩ (colim .morphism C.∘ A.ν) C.∘ unit.η A ≡⟨ C.cancelr A.ν-unit ⟩ colim .morphism ∎ ``` @@ -218,7 +218,7 @@ far $\cD$ is from being the category of $T$-algebras. Comparison⁻¹ : Functor (Eilenberg-Moore C T) D Comparison⁻¹ .F₀ = coapex ⊙ has-coeq Comparison⁻¹ .F₁ {X} {Y} alg-map = - has-coeq X .coequalise {e′ = e′} path where + has-coeq X .universal {e′ = e′} path where e′ : D.Hom (F.F₀ (X .fst)) (Comparison⁻¹ .F₀ Y) e′ = has-coeq Y .coeq D.∘ F.₁ (alg-map .morphism) ``` @@ -231,10 +231,10 @@ far $\cD$ is from being the category of $T$-algebras. has-coeq Y .coeq D.∘ F.₁ (Y .snd .ν) D.∘ F.₁ (T.M₁ (alg-map .morphism)) ≡⟨ D.extendl (has-coeq Y .coequal) ⟩ has-coeq Y .coeq D.∘ counit.ε _ D.∘ F.₁ (T.M₁ (alg-map .morphism)) ≡⟨ D.pushr (counit.is-natural _ _ _) ⟩ (has-coeq Y .coeq D.∘ F.₁ (alg-map .morphism)) D.∘ counit.ε _ ∎ - Comparison⁻¹ .F-id {X} = sym $ has-coeq X .unique (F.elimr refl ∙ D.introl refl) - Comparison⁻¹ .F-∘ {X} f g = sym $ has-coeq X .unique $ sym $ - D.pullr (has-coeq X .universal) - ·· D.pulll (has-coeq _ .universal) + Comparison⁻¹ .F-id {X} = sym $ has-coeq X .unique (D.idl _ ∙ D.intror F.F-id) + Comparison⁻¹ .F-∘ {X} f g = sym $ has-coeq X .unique $ + D.pullr (has-coeq X .factors) + ·· D.pulll (has-coeq _ .factors) ·· F.pullr refl open _⊣_ @@ -252,7 +252,7 @@ readers. Comparison⁻¹⊣Comparison .unit .η x .morphism = G.₁ (has-coeq _ .coeq) C.∘ T.unit .η _ Comparison⁻¹⊣Comparison .counit .η x = - has-coeq _ .coequalise (F⊣G .counit.is-natural _ _ _) + has-coeq _ .universal (F⊣G .counit.is-natural _ _ _) ```
@@ -268,25 +268,27 @@ readers. Comparison⁻¹⊣Comparison .unit .is-natural x y f = Algebra-hom-path C $ (G.₁ (has-coeq y .coeq) C.∘ T.unit.η _) C.∘ f .morphism ≡⟨ C.pullr (T.unit.is-natural _ _ _) ⟩ G.₁ (has-coeq y .coeq) C.∘ T.M₁ (f .morphism) C.∘ T.unit .η (x .fst) ≡⟨ C.pulll (sym (G.F-∘ _ _)) ⟩ - G.₁ (has-coeq y .coeq D.∘ F.₁ (f .morphism)) C.∘ T.unit .η (x .fst) ≡⟨ ap G.₁ (sym (has-coeq _ .universal)) C.⟩∘⟨refl ⟩ - G.₁ (has-coeq x .coequalise _ D.∘ has-coeq x .coeq) C.∘ T.unit .η (x .fst) ≡⟨ C.pushl (G.F-∘ _ _) ⟩ - G.₁ (has-coeq x .coequalise _) C.∘ G.₁ (has-coeq x .coeq) C.∘ T.unit.η _ ∎ + G.₁ (has-coeq y .coeq D.∘ F.₁ (f .morphism)) C.∘ T.unit .η (x .fst) ≡⟨ ap G.₁ (sym (has-coeq _ .factors)) C.⟩∘⟨refl ⟩ + G.₁ (has-coeq x .universal _ D.∘ has-coeq x .coeq) C.∘ T.unit .η (x .fst) ≡⟨ C.pushl (G.F-∘ _ _) ⟩ + G.₁ (has-coeq x .universal _) C.∘ G.₁ (has-coeq x .coeq) C.∘ T.unit.η _ ∎ Comparison⁻¹⊣Comparison .counit .is-natural x y f = has-coeq (F₀ (Comparison F⊣G) x) .unique {p = ap₂ D._∘_ (F⊣G .counit.is-natural _ _ _) refl ·· D.pullr (F⊣G .counit.is-natural _ _ _) ·· D.pulll (sym (F⊣G .counit.is-natural _ _ _))} - (sym (D.pullr (has-coeq _ .universal) ∙ D.pulll (has-coeq _ .universal))) - ∙ sym (has-coeq _ .unique (F⊣G .counit.is-natural _ _ _ ∙ D.pushr (sym (has-coeq _ .universal)))) + (D.pullr (has-coeq _ .factors) ∙ D.pulll (has-coeq _ .factors)) + ∙ sym (has-coeq _ .unique (D.pullr (has-coeq _ .factors) ∙ sym (F⊣G .counit.is-natural _ _ _))) Comparison⁻¹⊣Comparison .zig = - has-coeq _ .unique {e′ = has-coeq _ .coeq} {p = has-coeq _ .coequal} - (sym (D.pullr (has-coeq _ .universal) - ∙ D.pulll (has-coeq _ .universal) - ∙ ap₂ D._∘_ refl (F.F-∘ _ _) ∙ D.pulll (F⊣G .counit.is-natural _ _ _) - ∙ D.cancelr (F⊣G .zig))) - ∙ sym (has-coeq _ .unique (D.introl refl)) + unique₂ (has-coeq _) + (has-coeq _ .coequal) + (D.pullr (has-coeq _ .factors) + ∙ D.pulll (has-coeq _ .factors) + ∙ ap₂ D._∘_ refl (F.F-∘ _ _) + ∙ D.pulll (F⊣G .counit.is-natural _ _ _) + ∙ D.cancelr (F⊣G .zig)) + (D.idl _) Comparison⁻¹⊣Comparison .zag = Algebra-hom-path C $ - G.pulll (has-coeq _ .universal) ∙ F⊣G .zag + G.pulll (has-coeq _ .factors) ∙ F⊣G .zag ```
diff --git a/src/Cat/Functor/Monadic/Crude.lagda.md b/src/Cat/Functor/Monadic/Crude.lagda.md index 742f750e8..c47439ce2 100644 --- a/src/Cat/Functor/Monadic/Crude.lagda.md +++ b/src/Cat/Functor/Monadic/Crude.lagda.md @@ -179,17 +179,18 @@ we're seeking. η⁻¹η : adj.unit.η _ .morphism C.∘ η⁻¹ ≡ C.id ηη⁻¹ : η⁻¹ C.∘ adj.unit.η _ .morphism ≡ C.id - η⁻¹ = preserved .coequalise {e′ = o .snd .ν} (o .snd .ν-mult) + η⁻¹ = preserved .universal {e′ = o .snd .ν} (o .snd .ν-mult) η⁻¹η = is-coequaliser.unique₂ preserved - {e′ = U.₁ (has-coeq o .coeq)} {p = preserved .coequal} - (sym (C.pullr (preserved .universal) - ∙ C.pullr (unit.is-natural _ _ _) - ∙ C.pulll (preserved .coequal) - ∙ C.cancelr zag)) - (C.introl refl) - - ηη⁻¹ = C.pulll (preserved .universal) ∙ o .snd .ν-unit + {e′ = U.₁ (has-coeq o .coeq)} + (preserved .coequal) + (C.pullr (preserved .factors) + ∙ C.pullr (unit.is-natural _ _ _) + ∙ C.pulll (preserved .coequal) + ∙ C.cancelr zag) + (C.idl _) + + ηη⁻¹ = C.pulll (preserved .factors) ∙ o .snd .ν-unit ``` It remains to show that $\eta^{-1}$ is a homomorphism of algebras. This @@ -203,7 +204,7 @@ is a calculation reusing the established proof that $\eta^{-1}\eta = η⁻¹ C.∘ U.₁ (counit.ε _) ≡⟨ C.refl⟩∘⟨ ap U.₁ (D.intror (F.annihilate (C.assoc _ _ _ ∙ η⁻¹η))) ⟩ η⁻¹ C.∘ U.₁ (counit.ε _ D.∘ F.₁ (U.₁ (has-coeq o .coeq)) D.∘ F.₁ (unit.η _ C.∘ η⁻¹)) ≡⟨ C.refl⟩∘⟨ ap U.₁ (D.extendl (counit.is-natural _ _ _)) ⟩ η⁻¹ C.∘ U.₁ (has-coeq o .coeq D.∘ counit.ε _ D.∘ F.₁ (unit.η _ C.∘ η⁻¹)) ≡⟨ C.refl⟩∘⟨ U.F-∘ _ _ ⟩ - η⁻¹ C.∘ U.₁ (has-coeq o .coeq) C.∘ U.₁ (counit.ε _ D.∘ F.₁ (unit.η _ C.∘ η⁻¹)) ≡⟨ C.pulll (preserved .universal) ⟩ + η⁻¹ C.∘ U.₁ (has-coeq o .coeq) C.∘ U.₁ (counit.ε _ D.∘ F.₁ (unit.η _ C.∘ η⁻¹)) ≡⟨ C.pulll (preserved .factors) ⟩ o .snd .ν C.∘ U.₁ (counit.ε _ D.∘ F.₁ (unit.η _ C.∘ η⁻¹)) ≡⟨ C.refl⟩∘⟨ ap U.₁ (ap (counit.ε _ D.∘_) (F.F-∘ _ _) ∙ D.cancell zig) ⟩ o .snd .ν C.∘ UF.₁ η⁻¹ ∎ ``` @@ -244,7 +245,7 @@ $U$ is conservative, so $\eps$ is an isomorphism, as desired. : ∀ {o} → is-conservative U → D.is-invertible (adj.counit.ε o) conservative-prcoeq→counit-is-iso {o} reflect-iso = reflect-iso $ C.make-invertible - (U.₁ (coequ .coeq) C.∘ unit.η _) (U.pulll (coequ .universal) ∙ zag) + (U.₁ (coequ .coeq) C.∘ unit.η _) (U.pulll (coequ .factors) ∙ zag) inversel where @@ -260,12 +261,13 @@ $U$ is conservative, so $\eps$ is an isomorphism, as desired. inversel = is-coequaliser.unique₂ preserved - {e′ = U.₁ (coequ .coeq)} {p = preserved .coequal} - (sym (C.pullr (U.collapse (coequ .universal)) + {e′ = U.₁ (coequ .coeq)} + (preserved .coequal) + (C.pullr (U.collapse (coequ .factors)) ∙ C.pullr (unit.is-natural _ _ _) ∙ C.pulll (preserved .coequal) - ∙ C.cancelr zag)) - (C.introl refl) + ∙ C.cancelr zag) + (C.idl _) ``` Packaging it all up, we get the claimed theorem: If $\cD$ has Beck's diff --git a/src/Cat/Functor/Pullback.lagda.md b/src/Cat/Functor/Pullback.lagda.md index 8bce318b2..f0277aaae 100644 --- a/src/Cat/Functor/Pullback.lagda.md +++ b/src/Cat/Functor/Pullback.lagda.md @@ -85,9 +85,9 @@ diagram below is a cone over $K' \to X \ot Y$. module ypb = Pullback (pullbacks (y .map) f) module xpb = Pullback (pullbacks (x .map) f) dh′ : /-Hom _ _ - dh′ .map = ypb.limiting {p₁' = dh .map ∘ xpb.p₁} + dh′ .map = ypb.universal {p₁' = dh .map ∘ xpb.p₁} (pulll (dh .commutes) ∙ xpb.square) - dh′ .commutes = ypb.p₂∘limiting + dh′ .commutes = ypb.p₂∘universal ```
@@ -100,8 +100,8 @@ functorial, but the details are not particularly enlightening. Base-change .F-∘ {x} {y} {z} am bm = /-Hom-path (sym (zpb.unique - (pulll zpb.p₁∘limiting ∙ pullr ypb.p₁∘limiting ∙ assoc _ _ _) - (pulll zpb.p₂∘limiting ∙ ypb.p₂∘limiting))) + (pulll zpb.p₁∘universal ∙ pullr ypb.p₁∘universal ∙ assoc _ _ _) + (pulll zpb.p₂∘universal ∙ ypb.p₂∘universal))) where module ypb = Pullback (pullbacks (y .map) f) module zpb = Pullback (pullbacks (z .map) f) @@ -179,15 +179,15 @@ module _ (pullbacks : ∀ {X Y Z} f g → Pullback C {X} {Y} {Z} f g) {X Y : Ob} Σf⊣f* .unit .η obj = dh where module pb = Pullback (pullbacks (f ∘ obj .map) f) dh : /-Hom _ _ - dh .map = pb.limiting {p₁' = id} {p₂' = obj .map} (idr _) - dh .commutes = pb.p₂∘limiting + dh .map = pb.universal {p₁' = id} {p₂' = obj .map} (idr _) + dh .commutes = pb.p₂∘universal Σf⊣f* .unit .is-natural x y g = /-Hom-path (pb.unique₂ {p = (f ∘ y .map) ∘ id ∘ g .map ≡⟨ cat! C ⟩ f ∘ y .map ∘ g .map ∎} - (pulll pb.p₁∘limiting) - (pulll pb.p₂∘limiting) - (pulll pb.p₁∘limiting ∙ pullr pb′.p₁∘limiting ∙ id-comm) - (pulll pb.p₂∘limiting ∙ pb′.p₂∘limiting ∙ sym (g .commutes))) + (pulll pb.p₁∘universal) + (pulll pb.p₂∘universal) + (pulll pb.p₁∘universal ∙ pullr pb′.p₁∘universal ∙ id-comm) + (pulll pb.p₂∘universal ∙ pb′.p₂∘universal ∙ sym (g .commutes))) where module pb = Pullback (pullbacks (f ∘ y .map) f) module pb′ = Pullback (pullbacks (f ∘ x .map) f) @@ -197,17 +197,17 @@ module _ (pullbacks : ∀ {X Y Z} f g → Pullback C {X} {Y} {Z} f g) {X Y : Ob} dh : /-Hom _ _ dh .map = pb.p₁ dh .commutes = pb.square - Σf⊣f* .counit .is-natural x y g = /-Hom-path pb.p₁∘limiting + Σf⊣f* .counit .is-natural x y g = /-Hom-path pb.p₁∘universal where module pb = Pullback (pullbacks (y .map) f) - Σf⊣f* .zig {A} = /-Hom-path pb.p₁∘limiting + Σf⊣f* .zig {A} = /-Hom-path pb.p₁∘universal where module pb = Pullback (pullbacks (f ∘ A .map) f) Σf⊣f* .zag {B} = /-Hom-path (sym (pb.unique₂ {p = pb.square} (idr _) (idr _) - (pulll pb.p₁∘limiting ∙ pullr pb′.p₁∘limiting ∙ idr _) - (pulll pb.p₂∘limiting ∙ pb′.p₂∘limiting))) where + (pulll pb.p₁∘universal ∙ pullr pb′.p₁∘universal ∙ idr _) + (pulll pb.p₂∘universal ∙ pb′.p₂∘universal))) where module pb = Pullback (pullbacks (B .map) f) module pb′ = Pullback (pullbacks (f ∘ pb.p₂) f) ``` diff --git a/src/Cat/Instances/Comma.lagda.md b/src/Cat/Instances/Comma.lagda.md index 126676f65..064cc1647 100644 --- a/src/Cat/Instances/Comma.lagda.md +++ b/src/Cat/Instances/Comma.lagda.md @@ -71,6 +71,7 @@ $x : \cA$, $y : \cB$, and $f : F(x) \to G(y)$. ```agda record ↓Obj : Type (h ⊔ ao ⊔ bo) where no-eta-equality + constructor ↓obj field {x} : Ob A {y} : Ob B @@ -99,6 +100,7 @@ component of a [naturality square]. ```agda record ↓Hom (a b : ↓Obj) : Type (h ⊔ bh ⊔ ah) where no-eta-equality + constructor ↓hom private module a = ↓Obj a module b = ↓Obj b diff --git a/src/Cat/Instances/Discrete.lagda.md b/src/Cat/Instances/Discrete.lagda.md index f03f3b6e5..0362aee9a 100644 --- a/src/Cat/Instances/Discrete.lagda.md +++ b/src/Cat/Instances/Discrete.lagda.md @@ -4,6 +4,8 @@ open import Cat.Prelude open import Data.Dec +import Cat.Reasoning + module Cat.Instances.Discrete where ``` @@ -16,6 +18,7 @@ private variable open Precategory open Functor +open _=>_ ``` --> @@ -84,18 +87,16 @@ adjoint] to the `Ob`{.Agda} functor. ```agda Disc-diagram - : ∀ {iss : is-set X} (disc : Discrete X) - → (X → Ob C) - → Functor (Disc′ (el X iss)) C -Disc-diagram {X = X} {C = C} disc f = F where + : ∀ {X : Set ℓ} (disc : Discrete ∣ X ∣) + → (∣ X ∣ → Ob C) + → Functor (Disc′ X) C +Disc-diagram {C = C} {X = X} disc f = F where module C = Precategory C - set : is-set X - set = Discrete→is-set disc - P : X → X → Type _ + P : ∣ X ∣ → ∣ X ∣ → Type _ P x y = C.Hom (f x) (f y) - map : ∀ {x y : X} → x ≡ y → Dec (x ≡ y) → P x y + map : ∀ {x y : ∣ X ∣} → x ≡ y → Dec (x ≡ y) → P x y map {x} {y} p = Dec-elim (λ _ → P x y) (λ q → subst (P x) q C.id) @@ -123,7 +124,7 @@ computations with equalities and a whole waterfall of absurd cases: F .F-id {x} with inspect (disc x x) ... | yes p , q = map refl (disc x x) ≡⟨ ap (map refl) q ⟩ - map refl (yes p) ≡⟨ ap (map refl ⊙ yes) (set _ _ p refl) ⟩ + map refl (yes p) ≡⟨ ap (map refl ⊙ yes) (X .is-tr _ _ p refl) ⟩ map refl (yes refl) ≡⟨⟩ subst (P x) refl C.id ≡⟨ transport-refl _ ⟩ C.id ∎ @@ -132,7 +133,7 @@ computations with equalities and a whole waterfall of absurd cases: F .F-∘ {x} {y} {z} f g with inspect (disc x y) | inspect (disc x z) | inspect (disc y z) ... | yes x=y , p1 | yes x=z , p2 | yes y=z , p3 = map (g ∙ f) (disc x z) ≡⟨ ap (map (g ∙ f)) p2 ⟩ - map (g ∙ f) (yes ⌜ x=z ⌝) ≡⟨ ap! (set _ _ _ _) ⟩ + map (g ∙ f) (yes ⌜ x=z ⌝) ≡⟨ ap! (X .is-tr _ _ _ _) ⟩ map (g ∙ f) (yes (x=y ∙ y=z)) ≡⟨⟩ subst (P x) (x=y ∙ y=z) C.id ≡⟨ subst-∙ (P x) _ _ _ ⟩ subst (P x) y=z (subst (P _) x=y C.id) ≡⟨ from-pathp (Hom-pathp-reflr C refl) ⟩ @@ -170,3 +171,19 @@ Disc-adjunct {C = C} {iss = iss} F .F-∘ {x} {y} {z} f g = path where f {x} g ``` --> + + diff --git a/src/Cat/Instances/Functor.lagda.md b/src/Cat/Instances/Functor.lagda.md index d9c6d29b2..efa75c550 100644 --- a/src/Cat/Instances/Functor.lagda.md +++ b/src/Cat/Instances/Functor.lagda.md @@ -14,7 +14,7 @@ module Cat.Instances.Functor where private variable o h o₁ h₁ o₂ h₂ : Level - C D E : Precategory o h + B C D E : Precategory o h F G : Functor C D ``` @@ -61,7 +61,8 @@ _∘nt_ {C = C} {D = D} {F} {G} {H} f g = nat @@ -442,10 +443,23 @@ module module D = Cat.Reasoning D module C = Cat.Reasoning C + natural-inverses : {F G : Functor C D} → F => G → G => F → Type _ + natural-inverses = CD.Inverses + + is-natural-invertible : {F G : Functor C D} → F => G → Type _ + is-natural-invertible = CD.is-invertible + natural-iso : (F G : Functor C D) → Type _ natural-iso F G = F CD.≅ G - module natural-iso {F G : Functor C D} (eta : F CD.≅ G) = CD._≅_ eta + module natural-inverses {F G : Functor C D} {α : F => G} {β : G => F} (inv : natural-inverses α β) = + CD.Inverses inv + module is-natural-invertible {F G : Functor C D} {α : F => G} (inv : is-natural-invertible α) = + CD.is-invertible inv + module natural-iso {F G : Functor C D} (eta : natural-iso F G) = CD._≅_ eta + + idni : natural-iso F F + idni = CD.id-iso _ni∘_ : ∀ {F G H : Functor C D} → natural-iso F G → natural-iso G H @@ -460,6 +474,7 @@ module → F DD.≅ Id → (F F∘ G) CD.≅ G F∘-iso-id-l {F} {G} isom = subst ((F F∘ G) CD.≅_) F∘-idl (F∘-iso-l isom) + record make-natural-iso (F G : Functor C D) : Type (o ⊔ ℓ ⊔ ℓ′) where no-eta-equality field @@ -469,6 +484,22 @@ module inv∘eta : ∀ x → inv x D.∘ eta x ≡ D.id natural : ∀ x y f → G .F₁ f D.∘ eta x ≡ eta y D.∘ F .F₁ f + to-natural-inverses + : {F G : Functor C D} {α : F => G} {β : G => F} + → (∀ x → α .η x D.∘ β .η x ≡ D.id) + → (∀ x → β .η x D.∘ α .η x ≡ D.id) + → natural-inverses α β + to-natural-inverses p q = + CD.make-inverses (Nat-path p) (Nat-path q) + + to-is-natural-invertible + : {F G : Functor C D} {α : F => G} + → (β : G => F) + → (∀ x → α .η x D.∘ β .η x ≡ D.id) + → (∀ x → β .η x D.∘ α .η x ≡ D.id) + → is-natural-invertible α + to-is-natural-invertible β p q = CD.make-invertible β (Nat-path p) (Nat-path q) + to-natural-iso : {F G : Functor C D} → make-natural-iso F G → F CD.≅ G to-natural-iso {F = F} {G = G} x = isom where open CD._≅_ @@ -489,8 +520,48 @@ module isom .inverses .invl = Nat-path eta∘inv isom .inverses .invr = Nat-path inv∘eta + natural-inverses→inverses + : ∀ {α : F => G} {β : G => F} + → natural-inverses α β + → ∀ x → D.Inverses (α .η x) (β .η x) + natural-inverses→inverses inv x = + D.make-inverses + (CD.Inverses.invl inv ηₚ x) + (CD.Inverses.invr inv ηₚ x) + + is-natural-invertible→invertible + : ∀ {α : F => G} + → is-natural-invertible α + → ∀ x → D.is-invertible (α .η x) + is-natural-invertible→invertible inv x = + D.make-invertible + (CD.is-invertible.inv inv .η x) + (CD.is-invertible.invl inv ηₚ x) + (CD.is-invertible.invr inv ηₚ x) + + is-natural-invertible→natural-iso + : ∀ {α : F => G} + → is-natural-invertible α + → natural-iso F G + is-natural-invertible→natural-iso nat-inv = + CD.invertible→iso _ nat-inv + + natural-iso→is-natural-invertible + : (i : natural-iso F G) + → is-natural-invertible (natural-iso.to i) + natural-iso→is-natural-invertible i = + CD.iso→invertible i + open _=>_ +_ni^op : natural-iso F G → natural-iso (Functor.op F) (Functor.op G) +_ni^op α = + Cat.Reasoning.make-iso _ + (_=>_.op (natural-iso.from α)) + (_=>_.op (natural-iso.to α)) + (Nat-path λ j → natural-iso.invl α ηₚ _) + (Nat-path λ j → natural-iso.invr α ηₚ _) + module _ {o ℓ o′ ℓ′ o₂ ℓ₂} {C : Precategory o ℓ} @@ -562,5 +633,19 @@ module _ {o ℓ κ} {C : Precategory o ℓ} where natural-iso→equiv eta x = natural-iso.to eta .η x , natural-iso-to-is-equiv eta x + +module _ where + open Cat.Reasoning + + -- [TODO: Reed M, 14/03/2023] Extend the coherence machinery to handle natural + -- isos. + ni-assoc : {F : Functor D E} {G : Functor C D} {H : Functor B C} + → natural-iso (F F∘ G F∘ H) ((F F∘ G) F∘ H) + ni-assoc {E = E} = to-natural-iso λ where + .make-natural-iso.eta _ → E .id + .make-natural-iso.inv _ → E .id + .make-natural-iso.eta∘inv _ → E .idl _ + .make-natural-iso.inv∘eta _ → E .idl _ + .make-natural-iso.natural _ _ _ → E .idr _ ∙ sym (E .idl _) ``` --> diff --git a/src/Cat/Instances/Functor/Compose.lagda.md b/src/Cat/Instances/Functor/Compose.lagda.md index d3e3c7e95..0f714e955 100644 --- a/src/Cat/Instances/Functor/Compose.lagda.md +++ b/src/Cat/Instances/Functor/Compose.lagda.md @@ -5,6 +5,7 @@ open import Cat.Prelude import Cat.Functor.Reasoning as Fr import Cat.Reasoning +import Cat.Morphism open Functor open _=>_ @@ -38,7 +39,8 @@ opposite direction to $p$. private variable o ℓ : Level A B C C′ D E : Precategory o ℓ - F G H : Functor C D + F G H K : Functor C D + α β γ : F => G ``` --> @@ -85,6 +87,7 @@ _▸_ H nt .is-natural x y f = sym (H .F-∘ _ _) ∙ ap (H .F₁) (nt .is-natural _ _ _) ∙ H .F-∘ _ _ ``` + With the whiskerings already defined, defining $p_!$ and $p^*$ is easy: ```agda @@ -101,3 +104,59 @@ module _ (p : Functor C C′) where _^* .F-id = Nat-path λ _ → p .F-id _^* .F-∘ f g = Nat-path λ _ → p .F-∘ _ _ ``` + +Whiskerings are instances of a more general form of composition for +natural transformations, known as **horizontal composition**. + +```agda +_◆_ : ∀ {F G : Functor D E} {H K : Functor C D} + → F => G → H => K → F F∘ H => G F∘ K +_◆_ {E = E} {F = F} {G} {H} {K} α β = nat module horizontal-comp where + private module E = Cat.Reasoning E + open Fr + nat : F F∘ H => G F∘ K + nat .η x = G .F₁ (β .η _) E.∘ α .η _ + nat .is-natural x y f = + E.pullr (α .is-natural _ _ _) + ∙ E.extendl (weave G (β .is-natural _ _ _)) +``` + + + + + + + + diff --git a/src/Cat/Instances/Functor/Limits.lagda.md b/src/Cat/Instances/Functor/Limits.lagda.md index b8991401b..d9939c7cd 100644 --- a/src/Cat/Instances/Functor/Limits.lagda.md +++ b/src/Cat/Instances/Functor/Limits.lagda.md @@ -35,10 +35,8 @@ module _ diff --git a/src/Cat/Instances/Sets/Cocomplete.lagda.md b/src/Cat/Instances/Sets/Cocomplete.lagda.md index e56b5dff6..6f6ad3d18 100644 --- a/src/Cat/Instances/Sets/Cocomplete.lagda.md +++ b/src/Cat/Instances/Sets/Cocomplete.lagda.md @@ -12,9 +12,7 @@ module Cat.Instances.Sets.Cocomplete where @@ -86,9 +84,11 @@ set-coequalisers already includes a truncation. ```agda Sets-is-cocomplete : ∀ {ι κ o} → is-cocomplete ι κ (Sets (ι ⊔ κ ⊔ o)) -Sets-is-cocomplete {D = D} F = colim where +Sets-is-cocomplete {ι} {κ} {o} {D = D} F = to-colimit (to-is-colimit colim) where module D = Precategory D module F = Functor F + open _=>_ + open make-is-colimit sum : Type _ sum = Σ[ d ∈ D.Ob ] ∣ F.₀ d ∣ @@ -104,33 +104,32 @@ that $F(f)(x) = y$. [quotient]: Data.Set.Coequaliser.html#quotients -```agda - apex : Cocone F - apex .coapex = el (sum / rel) squash - apex .ψ x p = inc (x , p) - apex .commutes f = funext (λ i → sym (quot (f , refl))) -``` - By the same truncation nonsense as above, we can apply `Coeq-rec`{.Agda} to eliminate from our quotient to the coapex of any other cocone over $F$; The family of maps $\psi$ respects the quotient essentially by definition. ```agda - colim : Initial _ - colim .bot = apex - colim .has⊥ other = contr map unique where - map : Cocone-hom F apex other - map .hom = Coeq-rec (other .coapex .is-tr) (λ { (x , p) → other .ψ x p }) - λ { ((X , x) , (Y , y) , f , p) → - sym (happly (other .commutes f) x) ∙ ap (other .ψ Y) p - } - map .commutes o = refl - - unique : ∀ x → map ≡ x - unique hom′ = Cocone-hom-path _ - (funext (Coeq-elim-prop (λ x → other .coapex .is-tr _ _) - λ y → sym (happly (hom′ .commutes _) _))) + univ : ∀ {A : Set (ι ⊔ κ ⊔ o)} + → (eps : ∀ j → ∣ F.F₀ j ∣ → ∣ A ∣) + → (∀ {x y} (f : D.Hom x y) → ∀ Fx → eps y (F.F₁ f Fx) ≡ eps x Fx) + → sum / rel + → ∣ A ∣ + univ {A} eps p = + Coeq-rec (A .is-tr) + (λ { (x , p) → eps x p }) + (λ { ((X , x) , (Y , y) , f , q) → sym (p f x) ∙ ap (eps _) q}) + + colim : make-is-colimit F (el! (sum / rel)) + colim .ψ x p = inc (x , p) + colim .commutes f = funext λ _ → sym (quot (f , refl)) + colim .universal {A} eps p x = univ {A} eps (λ f → happly (p f)) x + colim .factors eps p = refl + colim .unique {A} eps p other q = funext λ x → + Coeq-elim-prop + (λ x → A .is-tr (other x) (univ {A} eps (λ f → happly (p f)) x)) + (λ x → happly (q (x .fst)) (x .snd)) + x ``` # Coproducts are disjoint diff --git a/src/Cat/Instances/Sets/Complete.lagda.md b/src/Cat/Instances/Sets/Complete.lagda.md index 22eb8a83d..c26bbc160 100644 --- a/src/Cat/Instances/Sets/Complete.lagda.md +++ b/src/Cat/Instances/Sets/Complete.lagda.md @@ -19,12 +19,10 @@ build a limit for it! ```agda Sets-is-complete : ∀ {ι κ o} → is-complete ι κ (Sets (ι ⊔ κ ⊔ o)) -Sets-is-complete {D = D} F = lim where +Sets-is-complete {D = D} F = to-limit (to-is-limit lim) module Sets-is-complete where module D = Precategory D module F = Functor F - - comm-prop : ∀ f → is-prop (∀ x y (g : D.Hom x y) → F.₁ g (f x) ≡ (f y)) - comm-prop f = hlevel! + open make-is-limit ``` Since `Set`{.Agda} is closed under (arbitrary) products, we can build @@ -35,8 +33,8 @@ those for which $F(g) \circ f(x) = f(y)$, i.e., those which are cones over $F$. ```agda - f-apex : Set _ - f-apex = el! $ + apex : Set _ + apex = el! $ Σ[ f ∈ ((j : D.Ob) → ∣ F.₀ j ∣) ] (∀ x y (g : D.Hom x y) → F.₁ g (f x) ≡ (f y)) ``` @@ -47,14 +45,6 @@ of the type underlying `f-apex`{.Agda}, we must cough up (for us. Similarly, since $p$ witnesses that $\psi$ `commutes`{.Agda}, we can project it directly. -```agda - open Cone - cone : Cone F - cone .Cone.apex = f-apex - cone .ψ x = λ { (f , p) → f x } - cone .commutes o = funext λ { (_ , p) → p _ _ o } -``` - Given some other cone $K$, to build a cone homomorphism $K \to \lim F$, recall that $K$ comes equipped with its own function $\psi : \prod_{x : \cD} K \to F(x)$, which we can simply flip around to get a function @@ -63,19 +53,16 @@ out by $\lim F$ since $K$ is a cone, hence $F(f) \circ \psi(x) = \psi(y)$, as required. ```agda - open Terminal - lim : Limit F - lim .top = cone - lim .has⊤ K = contr map map-unique where - module K = Cone K - open Cone-hom - map : Cone-hom F K cone - map .hom x = (λ j → K.ψ j x) , λ x y f → happly (K.commutes f) _ - map .commutes _ = refl - - map-unique : ∀ m → map ≡ m - map-unique m = Cone-hom-path _ (funext λ x → - Σ-prop-path comm-prop (funext λ y i → m .commutes y (~ i) x)) + -- open Terminal + lim : make-is-limit F apex + lim .ψ x (f , p) = f x + lim .commutes f = funext λ where + (_ , p) → p _ _ f + lim .universal eta p x = + (λ j → eta j x) , λ x y f → p f $ₚ _ + lim .factors _ _ = refl + lim .unique eta p other q = funext λ x → + Σ-prop-path hlevel! (funext λ j → q j $ₚ x) ``` + +The **initial category** is the category with no objects. + +```agda +⊥Cat : Precategory lzero lzero +⊥Cat .Ob = ⊥ +⊥Cat .Hom _ _ = ⊥ +``` + +This category is notable for the existence of a unique functor from +it to any other category. + +```agda +module _ {o h} {A : Precategory o h} where + private module A = Precategory A + open Functor + + + ¡F : Functor ⊥Cat A + ¡F .F₀ () + + ¡F-unique : ∀ (F G : Functor ⊥Cat A) → F ≡ G + ¡F-unique F G i .F₀ () + + ¡nt : ∀ {F G : Functor ⊥Cat A} → F => G + ¡nt ._=>_.η () + ¡nt ._=>_.is-natural () +``` + diff --git a/src/Cat/Instances/Shape/Parallel.lagda.md b/src/Cat/Instances/Shape/Parallel.lagda.md index 66dc33986..23a16d5aa 100644 --- a/src/Cat/Instances/Shape/Parallel.lagda.md +++ b/src/Cat/Instances/Shape/Parallel.lagda.md @@ -3,6 +3,8 @@ open import Cat.Prelude open import Data.Bool +import Cat.Reasoning + module Cat.Instances.Shape.Parallel where ``` @@ -57,11 +59,12 @@ parallel arrows between them. It is the shape of [equaliser] and ```agda module _ {o ℓ} {C : Precategory o ℓ} where - open Precategory C + open Cat.Reasoning C open Functor + open _=>_ - par-arrows→par-diagram : ∀ {a b} (f g : Hom a b) → Functor ·⇉· C - par-arrows→par-diagram f g = funct where + Fork : ∀ {a b} (f g : Hom a b) → Functor ·⇉· C + Fork f g = funct where funct : Functor _ _ funct .F₀ false = _ funct .F₀ true = _ @@ -75,4 +78,23 @@ module _ {o ℓ} {C : Precategory o ℓ} where funct .F-∘ {false} {false} {true} f g = sym (idr _) funct .F-∘ {false} {true} {true} tt _ = sym (idl _) funct .F-∘ {true} {true} {true} tt tt = sym (idl _) + + forkl : (F : Functor ·⇉· C) → Hom (F .F₀ false) (F .F₀ true) + forkl F = F .F₁ {false} {true} false + + forkr : (F : Functor ·⇉· C) → Hom (F .F₀ false) (F .F₀ true) + forkr F = F .F₁ {false} {true} true + + Fork→Cone + : ∀ {e} (F : Functor ·⇉· C) {equ : Hom e (F .F₀ false)} + → forkl F ∘ equ ≡ forkr F ∘ equ + → Const e => F + Fork→Cone {e = e} F {equ = equ} equal = nt where + nt : Const e => F + nt .η true = forkl F ∘ equ + nt .η false = equ + nt .is-natural true true tt = idr _ ∙ introl (F .F-id) + nt .is-natural false true true = idr _ ∙ equal + nt .is-natural false true false = idr _ + nt .is-natural false false tt = idr _ ∙ introl (F .F-id) ``` diff --git a/src/Cat/Instances/Shape/Terminal.lagda.md b/src/Cat/Instances/Shape/Terminal.lagda.md index 7bf03149c..8fc500c1a 100644 --- a/src/Cat/Instances/Shape/Terminal.lagda.md +++ b/src/Cat/Instances/Shape/Terminal.lagda.md @@ -1,8 +1,11 @@ ```agda open import 1Lab.Prelude +open import Cat.Instances.Functor open import Cat.Base +import Cat.Reasoning + module Cat.Instances.Shape.Terminal where ``` @@ -32,4 +35,63 @@ module _ {o h} {A : Precategory o h} where const! : Ob A → Functor ⊤Cat A const! = Const + + !F : Functor A ⊤Cat + !F .F₀ _ = tt + !F .F₁ _ = tt + !F .F-id = refl + !F .F-∘ _ _ = refl + + const-η : ∀ (F G : Functor ⊤Cat A) → F .F₀ tt ≡ G .F₀ tt → F ≡ G + const-η F G p = + Functor-path + (λ _ → p) + (λ _ i → hcomp (∂ i) λ where + j (i = i0) → F .F-id (~ j) + j (i = i1) → G .F-id (~ j) + j (j = i0) → A.id) ``` + + +Natural isomorphisms between functors $\top \to \cC$ +correspond to isomorphisms in $\cC$. + +```agda +module _ {o ℓ} {𝒞 : Precategory o ℓ} {F G : Functor ⊤Cat 𝒞} where + private + module 𝒞 = Cat.Reasoning 𝒞 + open Functor + open _=>_ + + hom→⊤-natural-trans : 𝒞.Hom (F .F₀ tt) (G .F₀ tt) → F => G + hom→⊤-natural-trans f .η _ = f + hom→⊤-natural-trans f .is-natural _ _ _ = 𝒞.elimr (F .F-id) ∙ 𝒞.introl (G .F-id) + + iso→⊤-natural-iso : F .F₀ tt 𝒞.≅ G .F₀ tt → natural-iso F G + iso→⊤-natural-iso i = to-natural-iso mi where + open make-natural-iso + open 𝒞._≅_ + + mi : make-natural-iso F G + mi .eta _ = i .to + mi .inv _ = i .from + mi .eta∘inv _ = i .invl + mi .inv∘eta _ = i .invr + mi .natural _ _ _ = 𝒞.eliml (G .F-id) ∙ 𝒞.intror (F .F-id) +``` + + diff --git a/src/Cat/Instances/Slice.lagda.md b/src/Cat/Instances/Slice.lagda.md index 45103a280..034927ce4 100644 --- a/src/Cat/Instances/Slice.lagda.md +++ b/src/Cat/Instances/Slice.lagda.md @@ -2,6 +2,7 @@ open import Cat.Instances.Shape.Terminal open import Cat.Instances.Shape.Join open import Cat.Diagram.Limit.Base +open import Cat.Diagram.Limit.Base open import Cat.Instances.Discrete open import Cat.Instances.Functor open import Cat.Diagram.Pullback @@ -311,14 +312,14 @@ $\cC$, as below: module g = /-Hom /g factor : C/c.Hom Q _ - factor .map = pb.limiting (f.commutes ∙ sym g.commutes) + factor .map = pb.universal (f.commutes ∙ sym g.commutes) factor .commutes = - (f .map C.∘ π₁) C.∘ pb.limiting _ ≡⟨ C.pullr pb.p₁∘limiting ⟩ + (f .map C.∘ π₁) C.∘ pb.universal _ ≡⟨ C.pullr pb.p₁∘universal ⟩ f .map C.∘ f.map ≡⟨ f.commutes ⟩ Q .map ∎ - is-pullback→is-fibre-product .π₁∘factor = /-Hom-path pb.p₁∘limiting - is-pullback→is-fibre-product .π₂∘factor = /-Hom-path pb.p₂∘limiting + is-pullback→is-fibre-product .π₁∘factor = /-Hom-path pb.p₁∘universal + is-pullback→is-fibre-product .π₂∘factor = /-Hom-path pb.p₂∘universal is-pullback→is-fibre-product .unique other p q = /-Hom-path (pb.unique (ap map p) (ap map q)) @@ -340,7 +341,7 @@ consisting of their underlying objects and maps is a square in $\cC$. The "if" direction (what is `pullback-above→pullback-below`{.Agda}) in the code is essentially an immediate consequence of the fact that equality of morphisms in $\cC/X$ may be tested in $\cC$, but we do -have to take some care when extending the "limiting" morphism back down +have to take some care when extending the "universal" morphism back down to the slice category (see the calculation marked `{- * -}`{.Agda}). ```agda @@ -360,14 +361,14 @@ module _ where pb′ : is-pullback (Slice _ _) _ _ _ _ pb′ .square = /-Hom-path (pb .square) - pb′ .limiting p .map = pb .limiting (ap map p) - pb′ .limiting {P'} {p₁' = p₁'} p .commutes = - (c .map ∘ pb .limiting (ap map p)) ≡˘⟨ (pulll (p1 .commutes)) ⟩ - (P .map ∘ p1 .map ∘ pb .limiting (ap map p)) ≡⟨ ap (_ ∘_) (pb .p₁∘limiting) ⟩ + pb′ .universal p .map = pb .universal (ap map p) + pb′ .universal {P'} {p₁' = p₁'} p .commutes = + (c .map ∘ pb .universal (ap map p)) ≡˘⟨ (pulll (p1 .commutes)) ⟩ + (P .map ∘ p1 .map ∘ pb .universal (ap map p)) ≡⟨ ap (_ ∘_) (pb .p₁∘universal) ⟩ (P .map ∘ p₁' .map) ≡⟨ p₁' .commutes ⟩ P' .map ∎ {- * -} - pb′ .p₁∘limiting = /-Hom-path (pb .p₁∘limiting) - pb′ .p₂∘limiting = /-Hom-path (pb .p₂∘limiting) + pb′ .p₁∘universal = /-Hom-path (pb .p₁∘universal) + pb′ .p₂∘universal = /-Hom-path (pb .p₂∘universal) pb′ .unique p q = /-Hom-path (pb .unique (ap map p) (ap map q)) pullback-below→pullback-above @@ -380,8 +381,8 @@ module _ where open is-pullback pb′ : is-pullback _ _ _ _ _ pb′ .square = ap map (pb .square) - pb′ .limiting {P′ = P'} {p₁'} {p₂'} p = - pb .limiting {P′ = cut (P .map ∘ p₁')} + pb′ .universal {P′ = P'} {p₁'} {p₂'} p = + pb .universal {P′ = cut (P .map ∘ p₁')} {p₁' = record { map = p₁' ; commutes = refl }} {p₂' = record { map = p₂' ; commutes = sym (pulll (g .commutes)) ∙ sym (ap (_ ∘_) p) @@ -389,8 +390,8 @@ module _ where }} (/-Hom-path p) .map - pb′ .p₁∘limiting = ap map $ pb .p₁∘limiting - pb′ .p₂∘limiting = ap map $ pb .p₂∘limiting + pb′ .p₁∘universal = ap map $ pb .p₁∘universal + pb′ .p₂∘universal = ap map $ pb .p₂∘universal pb′ .unique p q = ap map $ pb .unique {lim' = record { map = _ ; commutes = sym (pulll (p1 .commutes)) ∙ ap (_ ∘_) p }} (/-Hom-path p) (/-Hom-path q) @@ -573,13 +574,12 @@ module where open Terminal - open Cone-hom - open Cone open /-Obj open /-Hom private module C = Cat.Reasoning C + module J = Cat.Reasoning J module C/o = Cat.Reasoning (Slice C o) module F = Functor F ``` @@ -604,42 +604,55 @@ in $\cC$, then pass back to the slice category. F′ .F-∘ {inr x} {inr y} {inr z} (lift f) (lift g) = C.introl refl limit-above→limit-in-slice : Limit F′ → Limit F - limit-above→limit-in-slice lims = lim where - module lim = Terminal lims - module limob = Cone lim.top - - nadir : Cone F - nadir .apex .domain = limob.apex - nadir .apex .map = limob.ψ (inr tt) - nadir .ψ x .map = limob.ψ (inl x) - nadir .ψ x .commutes = limob.commutes (lift tt) - nadir .commutes f = /-Hom-path (limob.commutes (lift f)) - - lim : Limit F - lim .top = nadir - lim .has⊤ other = contr ch cont where - other′ : Cone F′ - other′ .apex = other .apex .domain - other′ .ψ (inl x) = other .ψ x .map - other′ .ψ (inr tt) = other .apex .map - other′ .commutes {inl x} {inl y} (lift f) = ap map (other .commutes f) - other′ .commutes {inl x} {inr y} (lift f) = other .ψ x .commutes - other′ .commutes {inr x} {inr y} (lift f) = C.idl _ - - module cont = is-contr (lim.has⊤ other′) - ch : Cone-hom F other nadir - ch .hom .map = cont.centre .hom - ch .hom .commutes = cont.centre .commutes (inr tt) - ch .commutes o = /-Hom-path (cont.centre .commutes (inl o)) - - cont : ∀ c → ch ≡ c - cont c = Cone-hom-path _ (/-Hom-path (ap hom uniq)) where - c′ : Cone-hom F′ other′ lim.top - c′ .hom = c .hom .map - c′ .commutes (inl x) = ap map (c .commutes x) - c′ .commutes (inr tt) = c .hom .commutes - - uniq = cont.paths c′ + limit-above→limit-in-slice lims = to-limit (to-is-limit lim) where + module lims = Limit lims + open make-is-limit + + apex : C/o.Ob + apex = cut (lims.ψ (inr tt)) + + nadir : (j : J.Ob) → /-Hom apex (F .F₀ j) + nadir j .map = lims.ψ (inl j) + nadir j .commutes = lims.commutes (lift tt) + + module Cone + {x : C/o.Ob} + (eta : (j : J.Ob) → C/o.Hom x (F .F₀ j)) + (p : ∀ {i j : J.Ob} → (f : J.Hom i j) → F .F₁ f C/o.∘ eta i ≡ eta j) + where + + ϕ : (j : J.Ob ⊎ ⊤) → C.Hom (x .domain) (F′ .F₀ j) + ϕ (inl j) = eta j .map + ϕ (inr _) = x .map + + ϕ-commutes + : ∀ {i j : J.Ob ⊎ ⊤} + → (f : ⋆Hom J ⊤Cat i j) + → F′ .F₁ f C.∘ ϕ i ≡ ϕ j + ϕ-commutes {inl i} {inl j} (lift f) = ap map (p f) + ϕ-commutes {inl i} {inr j} (lift f) = eta i .commutes + ϕ-commutes {inr i} {inr x} (lift f) = C.idl _ + + ϕ-factor + : ∀ (other : /-Hom x apex) + → (∀ j → nadir j C/o.∘ other ≡ eta j) + → (j : J.Ob ⊎ ⊤) + → lims.ψ j C.∘ other .map ≡ ϕ j + ϕ-factor other q (inl j) = ap map (q j) + ϕ-factor other q (inr tt) = other .commutes + + lim : make-is-limit F apex + lim .ψ = nadir + lim .commutes f = + /-Hom-path (lims.commutes (lift f)) + lim .universal {x} eta p .map = + lims.universal (Cone.ϕ eta p) (Cone.ϕ-commutes eta p) + lim .universal eta p .commutes = + lims.factors _ _ + lim .factors eta p = + /-Hom-path (lims.factors _ _) + lim .unique eta p other q = + /-Hom-path $ lims.unique _ _ (other .map) (Cone.ϕ-factor eta p other q) ``` In particular, if a category $\cC$ is complete, then so are its slices: diff --git a/src/Cat/Instances/Twisted.lagda.md b/src/Cat/Instances/Twisted.lagda.md index 963f9c6d8..852fdd812 100644 --- a/src/Cat/Instances/Twisted.lagda.md +++ b/src/Cat/Instances/Twisted.lagda.md @@ -34,6 +34,7 @@ $f \to g$ in $\rm{Tw}(\cC)$ as a factorisation of $g$ through $f$. module _ {o ℓ} {C : Precategory o ℓ} where open Precategory C record Twist {a₀ a₁ b₀ b₁} (f : Hom a₀ a₁) (g : Hom b₀ b₁) : Type ℓ where + constructor twist no-eta-equality field before : Hom b₀ a₀ diff --git a/src/Cat/Monoidal/Instances/Cartesian.lagda.md b/src/Cat/Monoidal/Instances/Cartesian.lagda.md index ea23c182a..cff983f58 100644 --- a/src/Cat/Monoidal/Instances/Cartesian.lagda.md +++ b/src/Cat/Monoidal/Instances/Cartesian.lagda.md @@ -37,7 +37,7 @@ module _ {o ℓ} {C : Precategory o ℓ} where ```agda Cartesian-monoidal : (∀ A B → Product A B) → Terminal C → Monoidal-category C Cartesian-monoidal prods term = mon where - open Cartesian prods + open Binary-products prods open Terminal term mon : Monoidal-category C mon .-⊗- = ×-functor diff --git a/src/Cat/Morphism.lagda.md b/src/Cat/Morphism.lagda.md index bdb258c7d..bd5f60550 100644 --- a/src/Cat/Morphism.lagda.md +++ b/src/Cat/Morphism.lagda.md @@ -192,6 +192,10 @@ We note that the identity morphism is always iso, and that isos compose: