From 1806521c23119e038bd1b909346c5622e7e7e7b8 Mon Sep 17 00:00:00 2001 From: fouche Date: Tue, 30 Nov 2021 12:03:25 +0200 Subject: [PATCH 01/46] Create Adjunctions.agda The category of adjunctions splitting a given monad, and proof that `Kleisli` and `EilenbergMoore` are initial and terminal in it. --- .../Adjoint/Construction/Adjunctions.agda | 62 +++++++++++++++++++ 1 file changed, 62 insertions(+) create mode 100644 Categories/Adjoint/Construction/Adjunctions.agda diff --git a/Categories/Adjoint/Construction/Adjunctions.agda b/Categories/Adjoint/Construction/Adjunctions.agda new file mode 100644 index 000000000..85cd24a07 --- /dev/null +++ b/Categories/Adjoint/Construction/Adjunctions.agda @@ -0,0 +1,62 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level + +open import Categories.Category.Core using (Category) +open import Categories.Category +open import Categories.Monad + +module Categories.Adjoint.Construction.Adjunctions where + +open import Categories.Category.Construction.Kleisli +open import Categories.Category.Construction.EilenbergMoore +open import Categories.Adjoint +open import Categories.Functor +open import Categories.Morphism +open import Categories.Functor.Properties +open import Categories.NaturalTransformation.Core +open import Categories.NaturalTransformation.NaturalIsomorphism using (_≃_) +open import Categories.Morphism.Reasoning as MR +open import Categories.Tactic.Category + +-- three things: +-- 1. the category of adjunctions splitting a given Monad +-- 2. the proof that EM(M) is the terminal object here +-- 3. the proof that KL(M) is the terminal object here + +private + variable + o ℓ e : Level + +record SplitObj {C : Category o ℓ e} (M : Monad C) : Set (suc o ⊔ suc ℓ ⊔ suc e) where + field + D : Category o ℓ e + F : Functor C D + G : Functor D C + adj : F ⊣ G + eqM : G ∘F F ≃ Monad.F M + +record Split⇒ {C : Category o ℓ e} (M : Monad C) (X Y : SplitObj M) : Set (suc o ⊔ suc ℓ ⊔ suc e) where + private + module X = SplitObj X + module Y = SplitObj Y + field + H : Functor X.D Y.D + HF≃F' : H ∘F X.F ≃ Y.F + G'H≃G : Y.G ∘F H ≃ X.G + +Split : {𝒞 : Category o ℓ e} → Monad 𝒞 → Category _ _ _ +Split {𝒞 = 𝒞} M = record + { Obj = SplitObj M + ; _⇒_ = Split⇒ M + ; _≈_ = {! !} + ; id = {! !} + ; _∘_ = {! !} + ; assoc = {! !} + ; sym-assoc = {! !} + ; identityˡ = {! !} + ; identityʳ = {! !} + ; identity² = {! !} + ; equiv = {! !} + ; ∘-resp-≈ = {! !} + } From c6f1d21caf1315667e23044a08aac656d2a27b3c Mon Sep 17 00:00:00 2001 From: fouche Date: Tue, 30 Nov 2021 12:16:53 +0200 Subject: [PATCH 02/46] whoops, file was in the wrong place --- .../Categories}/Adjoint/Construction/Adjunctions.agda | 0 1 file changed, 0 insertions(+), 0 deletions(-) rename {Categories => src/Categories}/Adjoint/Construction/Adjunctions.agda (100%) diff --git a/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda similarity index 100% rename from Categories/Adjoint/Construction/Adjunctions.agda rename to src/Categories/Adjoint/Construction/Adjunctions.agda From 4f125284c6bcfb15bf2208febb3b28860c871a66 Mon Sep 17 00:00:00 2001 From: fouche Date: Tue, 30 Nov 2021 13:15:21 +0200 Subject: [PATCH 03/46] stub one thing at a time... --- src/Categories/Adjoint/Construction/Adjunctions.agda | 10 +++++++++- 1 file changed, 9 insertions(+), 1 deletion(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 85cd24a07..a3ead2388 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -49,7 +49,7 @@ Split : {𝒞 : Category o ℓ e} → Monad 𝒞 → Category _ _ _ Split {𝒞 = 𝒞} M = record { Obj = SplitObj M ; _⇒_ = Split⇒ M - ; _≈_ = {! !} + ; _≈_ = λ H K → {! !} ; id = {! !} ; _∘_ = {! !} ; assoc = {! !} @@ -60,3 +60,11 @@ Split {𝒞 = 𝒞} M = record ; equiv = {! !} ; ∘-resp-≈ = {! !} } + where + open NaturalTransformation + split-id : {A : SplitObj M} → Split⇒ M A A + split-id = record + { H = Categories.Functor.id + ; HF≃F' = record { F⇒G = {! !} ; F⇐G = {! !} ; iso = {! !} } + ; G'H≃G = record { F⇒G = {! !} ; F⇐G = {! !} ; iso = {! !} } + } From 5973e7fb122c461ed9086cdae4e7a21947a80532 Mon Sep 17 00:00:00 2001 From: fouche Date: Tue, 30 Nov 2021 18:11:07 +0100 Subject: [PATCH 04/46] little progress --- .../Adjoint/Construction/Adjunctions.agda | 29 +++++++++++++++---- 1 file changed, 24 insertions(+), 5 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index a3ead2388..860323f77 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -15,7 +15,7 @@ open import Categories.Functor open import Categories.Morphism open import Categories.Functor.Properties open import Categories.NaturalTransformation.Core -open import Categories.NaturalTransformation.NaturalIsomorphism using (_≃_) +open import Categories.NaturalTransformation.NaturalIsomorphism -- using (_≃_; unitorʳ; unitorˡ) open import Categories.Morphism.Reasoning as MR open import Categories.Tactic.Category @@ -50,8 +50,8 @@ Split {𝒞 = 𝒞} M = record { Obj = SplitObj M ; _⇒_ = Split⇒ M ; _≈_ = λ H K → {! !} - ; id = {! !} - ; _∘_ = {! !} + ; id = split-id + ; _∘_ = comp ; assoc = {! !} ; sym-assoc = {! !} ; identityˡ = {! !} @@ -65,6 +65,25 @@ Split {𝒞 = 𝒞} M = record split-id : {A : SplitObj M} → Split⇒ M A A split-id = record { H = Categories.Functor.id - ; HF≃F' = record { F⇒G = {! !} ; F⇐G = {! !} ; iso = {! !} } - ; G'H≃G = record { F⇒G = {! !} ; F⇐G = {! !} ; iso = {! !} } + ; HF≃F' = unitorˡ + ; G'H≃G = unitorʳ } + comp : {A B X : SplitObj M} → Split⇒ M B X → Split⇒ M A B → Split⇒ M A X + comp U V = record + { H = H U ∘F H V + ; HF≃F' = {! !} + ; G'H≃G = {! !} + } + where + module U = Split⇒ U + module V = Split⇒ V + open U + open V + + -- comp record { H = H ; HF≃F' = record { F⇒G = F⇒G₁ ; F⇐G = F⇐G₁ ; iso = iso₁ } ; G'H≃G = isoGH } + -- record { H = K ; HF≃F' = record { F⇒G = F⇒G ; F⇐G = F⇐G ; iso = iso } ; G'H≃G = isoGK } + -- = record + -- { H = H ∘F K + -- ; HF≃F' = ≃.trans {! !} {! !} + -- ; G'H≃G = {! !} + -- } From f755c48c6984f8ee1d0b84aef30948a8d950b4de Mon Sep 17 00:00:00 2001 From: iwilare Date: Wed, 1 Dec 2021 13:41:19 -0500 Subject: [PATCH 05/46] Finish Split category definition --- .../Adjoint/Construction/Adjunctions.agda | 83 +++++++++---------- 1 file changed, 41 insertions(+), 42 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 860323f77..f12df15cd 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -1,34 +1,29 @@ {-# OPTIONS --without-K --safe #-} open import Level - open import Categories.Category.Core using (Category) open import Categories.Category open import Categories.Monad -module Categories.Adjoint.Construction.Adjunctions where +module Categories.Adjoint.Construction.Adjunctions {o ℓ e} {C : Category o ℓ e} (M : Monad C) where + +open Category C -open import Categories.Category.Construction.Kleisli -open import Categories.Category.Construction.EilenbergMoore open import Categories.Adjoint open import Categories.Functor open import Categories.Morphism open import Categories.Functor.Properties open import Categories.NaturalTransformation.Core -open import Categories.NaturalTransformation.NaturalIsomorphism -- using (_≃_; unitorʳ; unitorˡ) +open import Categories.NaturalTransformation.NaturalIsomorphism open import Categories.Morphism.Reasoning as MR open import Categories.Tactic.Category -- three things: -- 1. the category of adjunctions splitting a given Monad -- 2. the proof that EM(M) is the terminal object here --- 3. the proof that KL(M) is the terminal object here - -private - variable - o ℓ e : Level +-- 3. the proof that KL(M) is the initial object here -record SplitObj {C : Category o ℓ e} (M : Monad C) : Set (suc o ⊔ suc ℓ ⊔ suc e) where +record SplitObj : Set (suc o ⊔ suc ℓ ⊔ suc e) where field D : Category o ℓ e F : Functor C D @@ -36,7 +31,8 @@ record SplitObj {C : Category o ℓ e} (M : Monad C) : Set (suc o ⊔ suc ℓ adj : F ⊣ G eqM : G ∘F F ≃ Monad.F M -record Split⇒ {C : Category o ℓ e} (M : Monad C) (X Y : SplitObj M) : Set (suc o ⊔ suc ℓ ⊔ suc e) where +record Split⇒ (X Y : SplitObj) : Set (suc o ⊔ suc ℓ ⊔ suc e) where + constructor Splitc⇒ private module X = SplitObj X module Y = SplitObj Y @@ -45,45 +41,48 @@ record Split⇒ {C : Category o ℓ e} (M : Monad C) (X Y : SplitObj M) : Set (s HF≃F' : H ∘F X.F ≃ Y.F G'H≃G : Y.G ∘F H ≃ X.G -Split : {𝒞 : Category o ℓ e} → Monad 𝒞 → Category _ _ _ -Split {𝒞 = 𝒞} M = record - { Obj = SplitObj M - ; _⇒_ = Split⇒ M - ; _≈_ = λ H K → {! !} +Split : Monad C → Category _ _ _ +Split M = record + { Obj = SplitObj + ; _⇒_ = Split⇒ + ; _≈_ = λ U V → Split⇒.H U ≃ Split⇒.H V ; id = split-id ; _∘_ = comp - ; assoc = {! !} - ; sym-assoc = {! !} - ; identityˡ = {! !} - ; identityʳ = {! !} - ; identity² = {! !} - ; equiv = {! !} - ; ∘-resp-≈ = {! !} + ; assoc = λ { {f = f} {g = g} {h = h} → associator (Split⇒.H f) (Split⇒.H g) (Split⇒.H h) } + ; sym-assoc = λ { {f = f} {g = g} {h = h} → sym-associator (Split⇒.H f) (Split⇒.H g) (Split⇒.H h) } + ; identityˡ = unitorˡ + ; identityʳ = unitorʳ + ; identity² = unitor² + ; equiv = record { refl = refl ; sym = sym ; trans = trans } + ; ∘-resp-≈ = _ⓘₕ_ } where open NaturalTransformation - split-id : {A : SplitObj M} → Split⇒ M A A + split-id : {A : SplitObj} → Split⇒ A A split-id = record { H = Categories.Functor.id ; HF≃F' = unitorˡ ; G'H≃G = unitorʳ } - comp : {A B X : SplitObj M} → Split⇒ M B X → Split⇒ M A B → Split⇒ M A X - comp U V = record - { H = H U ∘F H V - ; HF≃F' = {! !} - ; G'H≃G = {! !} + comp : {A B X : SplitObj} → Split⇒ B X → Split⇒ A B → Split⇒ A X + comp {A = A} {B = B} {X = X} (Splitc⇒ Hᵤ HF≃F'ᵤ G'H≃Gᵤ) (Splitc⇒ Hᵥ HF≃F'ᵥ G'H≃Gᵥ) = record + { H = Hᵤ ∘F Hᵥ + ; HF≃F' = HF≃F'ᵤ ⓘᵥ (Hᵤ ⓘˡ HF≃F'ᵥ) ⓘᵥ associator (SplitObj.F A) Hᵥ Hᵤ + ; G'H≃G = G'H≃Gᵥ ⓘᵥ (G'H≃Gᵤ ⓘʳ Hᵥ) ⓘᵥ sym-associator Hᵥ Hᵤ (SplitObj.G X) } - where - module U = Split⇒ U - module V = Split⇒ V - open U - open V - -- comp record { H = H ; HF≃F' = record { F⇒G = F⇒G₁ ; F⇐G = F⇐G₁ ; iso = iso₁ } ; G'H≃G = isoGH } - -- record { H = K ; HF≃F' = record { F⇒G = F⇒G ; F⇐G = F⇐G ; iso = iso } ; G'H≃G = isoGK } - -- = record - -- { H = H ∘F K - -- ; HF≃F' = ≃.trans {! !} {! !} - -- ; G'H≃G = {! !} - -- } +open import Categories.Object.Terminal (Split M) +open import Categories.Object.Initial (Split M) +open import Categories.Category.Construction.EilenbergMoore +open import Categories.Category.Construction.Kleisli + +EM-object : SplitObj +EM-object = record { D = {! EilenbergMoore M !} + ; F = {! !} + ; G = {! !} + ; adj = {! !} + ; eqM = {! !} + } + +EM-terminal : IsTerminal EM-object +EM-terminal = {! !} From 5700870792e2beda5830f5e773ddc122215af1f5 Mon Sep 17 00:00:00 2001 From: fouche Date: Wed, 1 Dec 2021 22:55:46 +0200 Subject: [PATCH 06/46] def of kleisli-object --- agda-categories.agda-lib | 2 +- .../Adjoint/Construction/Adjunctions.agda | 32 +++++++++++++++---- 2 files changed, 27 insertions(+), 7 deletions(-) diff --git a/agda-categories.agda-lib b/agda-categories.agda-lib index 91c6c57d7..3f89afb89 100644 --- a/agda-categories.agda-lib +++ b/agda-categories.agda-lib @@ -1,3 +1,3 @@ name: agda-categories -depend: standard-library-1.7 +depend: standard-library-2.0 include: src/ diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index f12df15cd..1cdd9a1fc 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -75,14 +75,34 @@ open import Categories.Object.Terminal (Split M) open import Categories.Object.Initial (Split M) open import Categories.Category.Construction.EilenbergMoore open import Categories.Category.Construction.Kleisli +open import Categories.Adjoint.Construction.Kleisli M as KL +open import Categories.Adjoint.Construction.EilenbergMoore M as EM +-- there's a problem with levels here... EM-object : SplitObj -EM-object = record { D = {! EilenbergMoore M !} - ; F = {! !} - ; G = {! !} - ; adj = {! !} - ; eqM = {! !} - } +EM-object = record + { D = {! !} + ; F = {! !} + ; G = {! !} + ; adj = {! !} + ; eqM = {! !} + } EM-terminal : IsTerminal EM-object EM-terminal = {! !} + + +Kl-object : SplitObj +Kl-object = record + { D = Kleisli M + ; F = KL.Free + ; G = KL.Forgetful + ; adj = KL.Free⊣Forgetful + ; eqM = KL.FF≃F + } + +Kl-initial : IsInitial Kl-object +Kl-initial = record + { ! = {! !} + ; !-unique = {! !} + } \ No newline at end of file From fa0b87ce574868b6b1f353ba54821ad8b6d6b360 Mon Sep 17 00:00:00 2001 From: fouche Date: Thu, 2 Dec 2021 07:28:40 +0200 Subject: [PATCH 07/46] revert to 1.7 --- agda-categories.agda-lib | 2 +- .../AdjunctionTerminalInitial.agda | 53 +++++++++++++++++++ .../Adjoint/Construction/Adjunctions.agda | 46 ++++++++-------- 3 files changed, 77 insertions(+), 24 deletions(-) create mode 100644 src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda diff --git a/agda-categories.agda-lib b/agda-categories.agda-lib index 3f89afb89..91c6c57d7 100644 --- a/agda-categories.agda-lib +++ b/agda-categories.agda-lib @@ -1,3 +1,3 @@ name: agda-categories -depend: standard-library-2.0 +depend: standard-library-1.7 include: src/ diff --git a/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda b/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda new file mode 100644 index 000000000..0aba9d20d --- /dev/null +++ b/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda @@ -0,0 +1,53 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level +open import Categories.Category.Core using (Category) +open import Categories.Category +open import Categories.Monad + +module Categories.Adjoint.Construction.AdjunctionTerminalInitial {o ℓ e} {C : Category (o ⊔ ℓ ⊔ e) (ℓ ⊔ e) e} (M : Monad C) where + +open Category C + +open import Categories.Adjoint +open import Categories.Functor +open import Categories.NaturalTransformation.Core +open import Categories.NaturalTransformation.NaturalIsomorphism +open import Categories.Morphism.Reasoning as MR + +open import Categories.Adjoint.Construction.Adjunctions M +open import Categories.Object.Terminal (Split M) +open import Categories.Object.Initial (Split M) +open import Categories.Category.Construction.EilenbergMoore +open import Categories.Category.Construction.Kleisli +open import Categories.Adjoint.Construction.Kleisli M as KL +open import Categories.Adjoint.Construction.EilenbergMoore M as EM + + +EM-object : SplitObj +EM-object = record + { D = {! !} + ; F = {! !} + ; G = {! !} + ; adj = {! !} + ; eqM = {! !} + } + +EM-terminal : IsTerminal EM-object +EM-terminal = {! !} + + +Kl-object : SplitObj +Kl-object = record + { D = Kleisli M + ; F = KL.Free + ; G = KL.Forgetful + ; adj = KL.Free⊣Forgetful + ; eqM = KL.FF≃F + } + +Kl-initial : IsInitial Kl-object +Kl-initial = record + { ! = {! !} + ; !-unique = {! !} + } \ No newline at end of file diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 1cdd9a1fc..8a8ace8dd 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -79,30 +79,30 @@ open import Categories.Adjoint.Construction.Kleisli M as KL open import Categories.Adjoint.Construction.EilenbergMoore M as EM -- there's a problem with levels here... -EM-object : SplitObj -EM-object = record - { D = {! !} - ; F = {! !} - ; G = {! !} - ; adj = {! !} - ; eqM = {! !} - } +-- EM-object : SplitObj +-- EM-object = record +-- { D = {! !} +-- ; F = {! !} +-- ; G = {! !} +-- ; adj = {! !} +-- ; eqM = {! !} +-- } -EM-terminal : IsTerminal EM-object -EM-terminal = {! !} +-- EM-terminal : IsTerminal EM-object +-- EM-terminal = {! !} -Kl-object : SplitObj -Kl-object = record - { D = Kleisli M - ; F = KL.Free - ; G = KL.Forgetful - ; adj = KL.Free⊣Forgetful - ; eqM = KL.FF≃F - } +-- Kl-object : SplitObj +-- Kl-object = record +-- { D = Kleisli M +-- ; F = KL.Free +-- ; G = KL.Forgetful +-- ; adj = KL.Free⊣Forgetful +-- ; eqM = KL.FF≃F +-- } -Kl-initial : IsInitial Kl-object -Kl-initial = record - { ! = {! !} - ; !-unique = {! !} - } \ No newline at end of file +-- Kl-initial : IsInitial Kl-object +-- Kl-initial = record +-- { ! = {! !} +-- ; !-unique = {! !} +-- } \ No newline at end of file From f1bc8e7bbe272356dfad6b6ac929519e9e30fe00 Mon Sep 17 00:00:00 2001 From: fouche Date: Thu, 2 Dec 2021 09:01:30 +0200 Subject: [PATCH 08/46] progress --- .../AdjunctionTerminalInitial.agda | 33 +++++++++++++---- .../Adjoint/Construction/Adjunctions.agda | 37 +------------------ 2 files changed, 26 insertions(+), 44 deletions(-) diff --git a/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda b/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda index 0aba9d20d..b36f74282 100644 --- a/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda +++ b/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda @@ -26,15 +26,18 @@ open import Categories.Adjoint.Construction.EilenbergMoore M as EM EM-object : SplitObj EM-object = record - { D = {! !} - ; F = {! !} - ; G = {! !} - ; adj = {! !} - ; eqM = {! !} + { D = EilenbergMoore M + ; F = EM.Free + ; G = EM.Forgetful + ; adj = EM.Free⊣Forgetful + ; eqM = EM.FF≃F } EM-terminal : IsTerminal EM-object -EM-terminal = {! !} +EM-terminal = record + { ! = {! !} + ; !-unique = {! !} + } Kl-object : SplitObj @@ -48,6 +51,20 @@ Kl-object = record Kl-initial : IsInitial Kl-object Kl-initial = record - { ! = {! !} + { ! = bang ; !-unique = {! !} - } \ No newline at end of file + } + where + -- vanno aperte un po' di cose per non diventar matti + bang : {A : SplitObj} → Split⇒ Kl-object A + bang {Splitc D F G adj eq} = record + { H = record + { F₀ = λ T → Functor.F₀ F T + ; F₁ = λ f → {! !} + ; identity = {! !} + ; homomorphism = {! !} + ; F-resp-≈ = {! !} + } + ; HF≃F' = {! !} + ; G'H≃G = {! !} + } \ No newline at end of file diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 8a8ace8dd..2eae287bb 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -24,6 +24,7 @@ open import Categories.Tactic.Category -- 3. the proof that KL(M) is the initial object here record SplitObj : Set (suc o ⊔ suc ℓ ⊔ suc e) where + constructor Splitc field D : Category o ℓ e F : Functor C D @@ -70,39 +71,3 @@ Split M = record ; HF≃F' = HF≃F'ᵤ ⓘᵥ (Hᵤ ⓘˡ HF≃F'ᵥ) ⓘᵥ associator (SplitObj.F A) Hᵥ Hᵤ ; G'H≃G = G'H≃Gᵥ ⓘᵥ (G'H≃Gᵤ ⓘʳ Hᵥ) ⓘᵥ sym-associator Hᵥ Hᵤ (SplitObj.G X) } - -open import Categories.Object.Terminal (Split M) -open import Categories.Object.Initial (Split M) -open import Categories.Category.Construction.EilenbergMoore -open import Categories.Category.Construction.Kleisli -open import Categories.Adjoint.Construction.Kleisli M as KL -open import Categories.Adjoint.Construction.EilenbergMoore M as EM - --- there's a problem with levels here... --- EM-object : SplitObj --- EM-object = record --- { D = {! !} --- ; F = {! !} --- ; G = {! !} --- ; adj = {! !} --- ; eqM = {! !} --- } - --- EM-terminal : IsTerminal EM-object --- EM-terminal = {! !} - - --- Kl-object : SplitObj --- Kl-object = record --- { D = Kleisli M --- ; F = KL.Free --- ; G = KL.Forgetful --- ; adj = KL.Free⊣Forgetful --- ; eqM = KL.FF≃F --- } - --- Kl-initial : IsInitial Kl-object --- Kl-initial = record --- { ! = {! !} --- ; !-unique = {! !} --- } \ No newline at end of file From a0b6dba7218b3916e2b68192172d3e9c0a82257d Mon Sep 17 00:00:00 2001 From: fouche Date: Thu, 2 Dec 2021 17:40:39 +0200 Subject: [PATCH 09/46] proof of Kl-initial --- .../AdjunctionTerminalInitial.agda | 32 +++++++++++++++---- 1 file changed, 25 insertions(+), 7 deletions(-) diff --git a/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda b/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda index b36f74282..859974eb9 100644 --- a/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda +++ b/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda @@ -11,8 +11,8 @@ open Category C open import Categories.Adjoint open import Categories.Functor -open import Categories.NaturalTransformation.Core -open import Categories.NaturalTransformation.NaturalIsomorphism +open import Categories.NaturalTransformation.Core as NT +open import Categories.NaturalTransformation.NaturalIsomorphism as NI open import Categories.Morphism.Reasoning as MR open import Categories.Adjoint.Construction.Adjunctions M @@ -60,11 +60,29 @@ Kl-initial = record bang {Splitc D F G adj eq} = record { H = record { F₀ = λ T → Functor.F₀ F T - ; F₁ = λ f → {! !} - ; identity = {! !} - ; homomorphism = {! !} - ; F-resp-≈ = {! !} + ; F₁ = λ {A} {B} f → Adjoint.counit.η adj (Functor.F₀ F B) D.∘ (Functor.F₁ F (⇐.η B) D.∘ Functor.F₁ F f) + ; identity = λ {A} → begin + NaturalTransformation.η (Adjoint.counit adj) (Functor.F₀ F A) D.∘ Functor.F₁ F (⇐.η A) D.∘ Functor.F₁ F (M.η.η A) ≈⟨ (refl⟩∘⟨ D.Equiv.sym (Functor.homomorphism F)) ⟩ + NaturalTransformation.η (Adjoint.counit adj) (Functor.F₀ F A) D.∘ Functor.F₁ F (⇐.η A C.∘ M.η.η A) ≈⟨ (refl⟩∘⟨ Functor.F-resp-≈ F lemma) ⟩ + NaturalTransformation.η (Adjoint.counit adj) (Functor.F₀ F A) D.∘ Functor.F₁ F (NaturalTransformation.η (Adjoint.unit adj) A) ≈⟨ Adjoint.zig adj ⟩ + D.id ∎ + ; homomorphism = λ {X} {Y} {Z} {f} {g} → {! !} + ; F-resp-≈ = λ {A} {B} {f} {g} x → D.∘-resp-≈ʳ (D.∘-resp-≈ʳ (Functor.F-resp-≈ F x)) } ; HF≃F' = {! !} ; G'H≃G = {! !} - } \ No newline at end of file + } + where module D = Category D + module C = Category C + open C + open D + open D.HomReasoning + module M = Monad M + module eq = NaturalIsomorphism eq + open eq + lemma : {A : C.Obj} → ⇐.η A C.∘ M.η.η A C.≈ NaturalTransformation.η (Adjoint.unit adj) A + lemma {A} = C.HomReasoning.begin + ⇐.η A C.∘ M.η.η A C.HomReasoning.≈⟨ {! !} ⟩ + NaturalTransformation.η (Adjoint.unit adj) A + C.HomReasoning.∎ + where open C.HomReasoning From 59d6ef6a49fa1efc2c08f880ae99d397b9009416 Mon Sep 17 00:00:00 2001 From: fouche Date: Thu, 2 Dec 2021 17:49:24 +0200 Subject: [PATCH 10/46] cose --- .../Adjoint/Construction/AdjunctionTerminalInitial.agda | 3 +++ src/Categories/Adjoint/Construction/Adjunctions.agda | 2 ++ 2 files changed, 5 insertions(+) diff --git a/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda b/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda index 859974eb9..4e441dabe 100644 --- a/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda +++ b/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda @@ -23,6 +23,9 @@ open import Categories.Category.Construction.Kleisli open import Categories.Adjoint.Construction.Kleisli M as KL open import Categories.Adjoint.Construction.EilenbergMoore M as EM +open Category (Split M) +open SplitObj +open Split⇒ EM-object : SplitObj EM-object = record diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 2eae287bb..9b3b6c977 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -42,6 +42,8 @@ record Split⇒ (X Y : SplitObj) : Set (suc o ⊔ suc ℓ ⊔ suc e) where HF≃F' : H ∘F X.F ≃ Y.F G'H≃G : Y.G ∘F H ≃ X.G +open Split⇒ + Split : Monad C → Category _ _ _ Split M = record { Obj = SplitObj From b9f877c6d3b535c503bcadfd5fafa756321a4af0 Mon Sep 17 00:00:00 2001 From: iwilare Date: Thu, 2 Dec 2021 12:00:42 -0500 Subject: [PATCH 11/46] Clean modules and definitions --- .../AdjunctionTerminalInitial.agda | 69 +++++++++---------- .../Adjoint/Construction/Adjunctions.agda | 8 +-- 2 files changed, 34 insertions(+), 43 deletions(-) diff --git a/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda b/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda index 4e441dabe..45b965b60 100644 --- a/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda +++ b/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda @@ -7,8 +7,6 @@ open import Categories.Monad module Categories.Adjoint.Construction.AdjunctionTerminalInitial {o ℓ e} {C : Category (o ⊔ ℓ ⊔ e) (ℓ ⊔ e) e} (M : Monad C) where -open Category C - open import Categories.Adjoint open import Categories.Functor open import Categories.NaturalTransformation.Core as NT @@ -16,6 +14,7 @@ open import Categories.NaturalTransformation.NaturalIsomorphism as NI open import Categories.Morphism.Reasoning as MR open import Categories.Adjoint.Construction.Adjunctions M + open import Categories.Object.Terminal (Split M) open import Categories.Object.Initial (Split M) open import Categories.Category.Construction.EilenbergMoore @@ -23,10 +22,6 @@ open import Categories.Category.Construction.Kleisli open import Categories.Adjoint.Construction.Kleisli M as KL open import Categories.Adjoint.Construction.EilenbergMoore M as EM -open Category (Split M) -open SplitObj -open Split⇒ - EM-object : SplitObj EM-object = record { D = EilenbergMoore M @@ -38,7 +33,7 @@ EM-object = record EM-terminal : IsTerminal EM-object EM-terminal = record - { ! = {! !} + { ! = {!!} ; !-unique = {! !} } @@ -58,34 +53,32 @@ Kl-initial = record ; !-unique = {! !} } where - -- vanno aperte un po' di cose per non diventar matti - bang : {A : SplitObj} → Split⇒ Kl-object A - bang {Splitc D F G adj eq} = record - { H = record - { F₀ = λ T → Functor.F₀ F T - ; F₁ = λ {A} {B} f → Adjoint.counit.η adj (Functor.F₀ F B) D.∘ (Functor.F₁ F (⇐.η B) D.∘ Functor.F₁ F f) - ; identity = λ {A} → begin - NaturalTransformation.η (Adjoint.counit adj) (Functor.F₀ F A) D.∘ Functor.F₁ F (⇐.η A) D.∘ Functor.F₁ F (M.η.η A) ≈⟨ (refl⟩∘⟨ D.Equiv.sym (Functor.homomorphism F)) ⟩ - NaturalTransformation.η (Adjoint.counit adj) (Functor.F₀ F A) D.∘ Functor.F₁ F (⇐.η A C.∘ M.η.η A) ≈⟨ (refl⟩∘⟨ Functor.F-resp-≈ F lemma) ⟩ - NaturalTransformation.η (Adjoint.counit adj) (Functor.F₀ F A) D.∘ Functor.F₁ F (NaturalTransformation.η (Adjoint.unit adj) A) ≈⟨ Adjoint.zig adj ⟩ - D.id ∎ - ; homomorphism = λ {X} {Y} {Z} {f} {g} → {! !} - ; F-resp-≈ = λ {A} {B} {f} {g} x → D.∘-resp-≈ʳ (D.∘-resp-≈ʳ (Functor.F-resp-≈ F x)) - } - ; HF≃F' = {! !} - ; G'H≃G = {! !} - } - where module D = Category D - module C = Category C - open C - open D - open D.HomReasoning - module M = Monad M - module eq = NaturalIsomorphism eq - open eq - lemma : {A : C.Obj} → ⇐.η A C.∘ M.η.η A C.≈ NaturalTransformation.η (Adjoint.unit adj) A - lemma {A} = C.HomReasoning.begin - ⇐.η A C.∘ M.η.η A C.HomReasoning.≈⟨ {! !} ⟩ - NaturalTransformation.η (Adjoint.unit adj) A - C.HomReasoning.∎ - where open C.HomReasoning + -- vanno aperte un po' di cose per non diventar matti + bang : {A : SplitObj} → Split⇒ Kl-object A + bang {splitobj D F G adj eq} = record + { H = record + { F₀ = λ T → Functor.F₀ F T + ; F₁ = λ {A} {B} f → adj.counit.η (F.₀ B) ∘ F.₁ (eq.⇐.η B C.∘ f) + ; identity = λ {A} → begin + adj.counit.η (F.₀ A) ∘ F.₁ (eq.⇐.η A C.∘ M.η.η A) ≈⟨ refl⟩∘⟨ Functor.F-resp-≈ F lemma ⟩ + adj.counit.η (F.₀ A) ∘ F.₁ (adj.unit.η A) ≈⟨ adj.zig ⟩ + D.id ∎ + ; homomorphism = λ {X} {Y} {Z} {f} {g} → {! !} + ; F-resp-≈ = λ x → D.∘-resp-≈ʳ (Functor.F-resp-≈ F (C.∘-resp-≈ʳ x)) + } + ; HF≃F' = {! !} + ; G'H≃G = {! !} + } + where module adj = Adjoint adj + module F = Functor F + module M = Monad M + module eq = NaturalIsomorphism eq + open module D = Category D + open D.HomReasoning + module C = Category C + + lemma : ∀ {A} → eq.⇐.η A C.∘ M.η.η A C.≈ adj.unit.η A + lemma {A} = C.HomReasoning.begin + eq.⇐.η A C.∘ M.η.η A C.HomReasoning.≈⟨ {! !} ⟩ + adj.unit.η A C.HomReasoning.∎ + where open C.HomReasoning \ No newline at end of file diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 9b3b6c977..3fb3da76e 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -24,7 +24,7 @@ open import Categories.Tactic.Category -- 3. the proof that KL(M) is the initial object here record SplitObj : Set (suc o ⊔ suc ℓ ⊔ suc e) where - constructor Splitc + constructor splitobj field D : Category o ℓ e F : Functor C D @@ -33,7 +33,7 @@ record SplitObj : Set (suc o ⊔ suc ℓ ⊔ suc e) where eqM : G ∘F F ≃ Monad.F M record Split⇒ (X Y : SplitObj) : Set (suc o ⊔ suc ℓ ⊔ suc e) where - constructor Splitc⇒ + constructor split⇒ private module X = SplitObj X module Y = SplitObj Y @@ -42,8 +42,6 @@ record Split⇒ (X Y : SplitObj) : Set (suc o ⊔ suc ℓ ⊔ suc e) where HF≃F' : H ∘F X.F ≃ Y.F G'H≃G : Y.G ∘F H ≃ X.G -open Split⇒ - Split : Monad C → Category _ _ _ Split M = record { Obj = SplitObj @@ -68,7 +66,7 @@ Split M = record ; G'H≃G = unitorʳ } comp : {A B X : SplitObj} → Split⇒ B X → Split⇒ A B → Split⇒ A X - comp {A = A} {B = B} {X = X} (Splitc⇒ Hᵤ HF≃F'ᵤ G'H≃Gᵤ) (Splitc⇒ Hᵥ HF≃F'ᵥ G'H≃Gᵥ) = record + comp {A = A} {B = B} {X = X} (split⇒ Hᵤ HF≃F'ᵤ G'H≃Gᵤ) (split⇒ Hᵥ HF≃F'ᵥ G'H≃Gᵥ) = record { H = Hᵤ ∘F Hᵥ ; HF≃F' = HF≃F'ᵤ ⓘᵥ (Hᵤ ⓘˡ HF≃F'ᵥ) ⓘᵥ associator (SplitObj.F A) Hᵥ Hᵤ ; G'H≃G = G'H≃Gᵥ ⓘᵥ (G'H≃Gᵤ ⓘʳ Hᵥ) ⓘᵥ sym-associator Hᵥ Hᵤ (SplitObj.G X) From f6168a815a88cd8c85783069623f460c61e9e7cc Mon Sep 17 00:00:00 2001 From: fouche Date: Thu, 9 Dec 2021 20:32:40 +0200 Subject: [PATCH 12/46] lazily think about it --- agda-categories.agda-lib | 2 +- .../AdjunctionTerminalInitial.agda | 42 ++++++++++++------- .../Adjoint/Construction/Adjunctions.agda | 7 +--- 3 files changed, 29 insertions(+), 22 deletions(-) diff --git a/agda-categories.agda-lib b/agda-categories.agda-lib index 91c6c57d7..3f89afb89 100644 --- a/agda-categories.agda-lib +++ b/agda-categories.agda-lib @@ -1,3 +1,3 @@ name: agda-categories -depend: standard-library-1.7 +depend: standard-library-2.0 include: src/ diff --git a/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda b/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda index 45b965b60..9141ae857 100644 --- a/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda +++ b/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda @@ -33,7 +33,7 @@ EM-object = record EM-terminal : IsTerminal EM-object EM-terminal = record - { ! = {!!} + { ! = {! !} ; !-unique = {! !} } @@ -53,7 +53,6 @@ Kl-initial = record ; !-unique = {! !} } where - -- vanno aperte un po' di cose per non diventar matti bang : {A : SplitObj} → Split⇒ Kl-object A bang {splitobj D F G adj eq} = record { H = record @@ -66,19 +65,30 @@ Kl-initial = record ; homomorphism = λ {X} {Y} {Z} {f} {g} → {! !} ; F-resp-≈ = λ x → D.∘-resp-≈ʳ (Functor.F-resp-≈ F (C.∘-resp-≈ʳ x)) } - ; HF≃F' = {! !} + ; HF≃F' = niHelper (record + { η = λ X → Functor.F₁ F (Category.id C) + ; η⁻¹ = λ X → Functor.F₁ F (Category.id C) + ; commute = λ {X} {Y} f → begin + F.F₁ C.id ∘ adj.counit.η (F.F₀ Y) ∘ F.F₁ (eq.⇐.η Y C.∘ M.η.η Y C.∘ f) ≈⟨ {! !} ⟩ + adj.counit.η (F.F₀ Y) ∘ F.F₁ (eq.⇐.η Y C.∘ M.η.η Y C.∘ f) ≈⟨ {! !} ⟩ + adj.counit.η (F.F₀ Y) ∘ F.F₁ (eq.⇐.η Y) D.∘ F.F₁ (M.η.η Y) D.∘ F.F₁ f ≈⟨ {! !} ⟩ + adj.counit.η (F.F₀ Y) ∘ F.F₁ (eq.⇐.η Y) D.∘ F.F₁ (M.F.F₁ f) D.∘ F.F₁ (M.η.η X) ≈⟨ {! !} ⟩ + adj.counit.η (F.F₀ Y) ∘ F.F₁ (eq.⇐.η Y) D.∘ F.F₁ (M.F.F₁ f) D.∘ F.F₁ (M.η.η X) ≈⟨ {! !} ⟩ + F.F₁ f ∘ F.F₁ C.id ∎ + ; iso = {! !} + }) ; G'H≃G = {! !} } - where module adj = Adjoint adj - module F = Functor F - module M = Monad M - module eq = NaturalIsomorphism eq - open module D = Category D - open D.HomReasoning - module C = Category C - - lemma : ∀ {A} → eq.⇐.η A C.∘ M.η.η A C.≈ adj.unit.η A - lemma {A} = C.HomReasoning.begin - eq.⇐.η A C.∘ M.η.η A C.HomReasoning.≈⟨ {! !} ⟩ - adj.unit.η A C.HomReasoning.∎ - where open C.HomReasoning \ No newline at end of file + where + module adj = Adjoint adj + module F = Functor F + module M = Monad M + module eq = NaturalIsomorphism eq + open module D = Category D + open D.HomReasoning + module C = Category C + lemma : ∀ {A} → eq.⇐.η A C.∘ M.η.η A C.≈ adj.unit.η A + lemma {A} = C.HomReasoning.begin + eq.⇐.η A C.∘ M.η.η A C.HomReasoning.≈⟨ {! !} ⟩ + adj.unit.η A C.HomReasoning.∎ + where open C.HomReasoning \ No newline at end of file diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 3fb3da76e..77ade6c23 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -18,11 +18,7 @@ open import Categories.NaturalTransformation.NaturalIsomorphism open import Categories.Morphism.Reasoning as MR open import Categories.Tactic.Category --- three things: --- 1. the category of adjunctions splitting a given Monad --- 2. the proof that EM(M) is the terminal object here --- 3. the proof that KL(M) is the initial object here - +-- the category of adjunctions splitting a given monad record SplitObj : Set (suc o ⊔ suc ℓ ⊔ suc e) where constructor splitobj field @@ -31,6 +27,7 @@ record SplitObj : Set (suc o ⊔ suc ℓ ⊔ suc e) where G : Functor D C adj : F ⊣ G eqM : G ∘F F ≃ Monad.F M + -- probably this is asking too little... record Split⇒ (X Y : SplitObj) : Set (suc o ⊔ suc ℓ ⊔ suc e) where constructor split⇒ From fd94bd7df3634b716787261d9840b8d56af90e5f Mon Sep 17 00:00:00 2001 From: fouche Date: Thu, 9 Dec 2021 20:33:04 +0200 Subject: [PATCH 13/46] revert agda-lib to 1.7 --- agda-categories.agda-lib | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/agda-categories.agda-lib b/agda-categories.agda-lib index 3f89afb89..91c6c57d7 100644 --- a/agda-categories.agda-lib +++ b/agda-categories.agda-lib @@ -1,3 +1,3 @@ name: agda-categories -depend: standard-library-2.0 +depend: standard-library-1.7 include: src/ From 38c00aa513b4ac7c4225ee07c32eebde4dc7acd5 Mon Sep 17 00:00:00 2001 From: fouche Date: Fri, 14 Jan 2022 11:12:27 +0200 Subject: [PATCH 14/46] a blind, automatic commit --- agda-categories.agda-lib | 2 +- src/Categories/Adjoint/Construction/Adjunctions.agda | 4 ++++ 2 files changed, 5 insertions(+), 1 deletion(-) diff --git a/agda-categories.agda-lib b/agda-categories.agda-lib index 91c6c57d7..3f89afb89 100644 --- a/agda-categories.agda-lib +++ b/agda-categories.agda-lib @@ -1,3 +1,3 @@ name: agda-categories -depend: standard-library-1.7 +depend: standard-library-2.0 include: src/ diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 77ade6c23..d10b22dad 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -27,6 +27,10 @@ record SplitObj : Set (suc o ⊔ suc ℓ ⊔ suc e) where G : Functor D C adj : F ⊣ G eqM : G ∘F F ≃ Monad.F M + unit-eq : ∀ A → eqM.⇐.η A C.∘ M.η.η A C.≈ adj.unit.η A + -- mult-eq : + -- unit M = unit adj + -- mult M = G . counit . F -- probably this is asking too little... record Split⇒ (X Y : SplitObj) : Set (suc o ⊔ suc ℓ ⊔ suc e) where From 44603fcf9475fb0f2cf698cbb84cbeb8c9ed9bd1 Mon Sep 17 00:00:00 2001 From: iwilare Date: Sun, 29 May 2022 22:20:35 +0200 Subject: [PATCH 15/46] =?UTF-8?q?Add=20=CE=BC-eq,=20move=20to=20Adjunction?= =?UTF-8?q?.Properties,=20fill=20some=20definitions?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- agda-categories.agda-lib | 2 +- .../AdjunctionTerminalInitial.agda | 94 ----------- .../Adjoint/Construction/Adjunctions.agda | 43 ++++- .../Construction/Adjunctions/Properties.agda | 149 ++++++++++++++++++ 4 files changed, 186 insertions(+), 102 deletions(-) delete mode 100644 src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda create mode 100644 src/Categories/Adjoint/Construction/Adjunctions/Properties.agda diff --git a/agda-categories.agda-lib b/agda-categories.agda-lib index 3f89afb89..91c6c57d7 100644 --- a/agda-categories.agda-lib +++ b/agda-categories.agda-lib @@ -1,3 +1,3 @@ name: agda-categories -depend: standard-library-2.0 +depend: standard-library-1.7 include: src/ diff --git a/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda b/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda deleted file mode 100644 index 9141ae857..000000000 --- a/src/Categories/Adjoint/Construction/AdjunctionTerminalInitial.agda +++ /dev/null @@ -1,94 +0,0 @@ -{-# OPTIONS --without-K --safe #-} - -open import Level -open import Categories.Category.Core using (Category) -open import Categories.Category -open import Categories.Monad - -module Categories.Adjoint.Construction.AdjunctionTerminalInitial {o ℓ e} {C : Category (o ⊔ ℓ ⊔ e) (ℓ ⊔ e) e} (M : Monad C) where - -open import Categories.Adjoint -open import Categories.Functor -open import Categories.NaturalTransformation.Core as NT -open import Categories.NaturalTransformation.NaturalIsomorphism as NI -open import Categories.Morphism.Reasoning as MR - -open import Categories.Adjoint.Construction.Adjunctions M - -open import Categories.Object.Terminal (Split M) -open import Categories.Object.Initial (Split M) -open import Categories.Category.Construction.EilenbergMoore -open import Categories.Category.Construction.Kleisli -open import Categories.Adjoint.Construction.Kleisli M as KL -open import Categories.Adjoint.Construction.EilenbergMoore M as EM - -EM-object : SplitObj -EM-object = record - { D = EilenbergMoore M - ; F = EM.Free - ; G = EM.Forgetful - ; adj = EM.Free⊣Forgetful - ; eqM = EM.FF≃F - } - -EM-terminal : IsTerminal EM-object -EM-terminal = record - { ! = {! !} - ; !-unique = {! !} - } - - -Kl-object : SplitObj -Kl-object = record - { D = Kleisli M - ; F = KL.Free - ; G = KL.Forgetful - ; adj = KL.Free⊣Forgetful - ; eqM = KL.FF≃F - } - -Kl-initial : IsInitial Kl-object -Kl-initial = record - { ! = bang - ; !-unique = {! !} - } - where - bang : {A : SplitObj} → Split⇒ Kl-object A - bang {splitobj D F G adj eq} = record - { H = record - { F₀ = λ T → Functor.F₀ F T - ; F₁ = λ {A} {B} f → adj.counit.η (F.₀ B) ∘ F.₁ (eq.⇐.η B C.∘ f) - ; identity = λ {A} → begin - adj.counit.η (F.₀ A) ∘ F.₁ (eq.⇐.η A C.∘ M.η.η A) ≈⟨ refl⟩∘⟨ Functor.F-resp-≈ F lemma ⟩ - adj.counit.η (F.₀ A) ∘ F.₁ (adj.unit.η A) ≈⟨ adj.zig ⟩ - D.id ∎ - ; homomorphism = λ {X} {Y} {Z} {f} {g} → {! !} - ; F-resp-≈ = λ x → D.∘-resp-≈ʳ (Functor.F-resp-≈ F (C.∘-resp-≈ʳ x)) - } - ; HF≃F' = niHelper (record - { η = λ X → Functor.F₁ F (Category.id C) - ; η⁻¹ = λ X → Functor.F₁ F (Category.id C) - ; commute = λ {X} {Y} f → begin - F.F₁ C.id ∘ adj.counit.η (F.F₀ Y) ∘ F.F₁ (eq.⇐.η Y C.∘ M.η.η Y C.∘ f) ≈⟨ {! !} ⟩ - adj.counit.η (F.F₀ Y) ∘ F.F₁ (eq.⇐.η Y C.∘ M.η.η Y C.∘ f) ≈⟨ {! !} ⟩ - adj.counit.η (F.F₀ Y) ∘ F.F₁ (eq.⇐.η Y) D.∘ F.F₁ (M.η.η Y) D.∘ F.F₁ f ≈⟨ {! !} ⟩ - adj.counit.η (F.F₀ Y) ∘ F.F₁ (eq.⇐.η Y) D.∘ F.F₁ (M.F.F₁ f) D.∘ F.F₁ (M.η.η X) ≈⟨ {! !} ⟩ - adj.counit.η (F.F₀ Y) ∘ F.F₁ (eq.⇐.η Y) D.∘ F.F₁ (M.F.F₁ f) D.∘ F.F₁ (M.η.η X) ≈⟨ {! !} ⟩ - F.F₁ f ∘ F.F₁ C.id ∎ - ; iso = {! !} - }) - ; G'H≃G = {! !} - } - where - module adj = Adjoint adj - module F = Functor F - module M = Monad M - module eq = NaturalIsomorphism eq - open module D = Category D - open D.HomReasoning - module C = Category C - lemma : ∀ {A} → eq.⇐.η A C.∘ M.η.η A C.≈ adj.unit.η A - lemma {A} = C.HomReasoning.begin - eq.⇐.η A C.∘ M.η.η A C.HomReasoning.≈⟨ {! !} ⟩ - adj.unit.η A C.HomReasoning.∎ - where open C.HomReasoning \ No newline at end of file diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index d10b22dad..c7cfefa44 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -7,7 +7,10 @@ open import Categories.Monad module Categories.Adjoint.Construction.Adjunctions {o ℓ e} {C : Category o ℓ e} (M : Monad C) where -open Category C +private + module C = Category C + module M = Monad M + open C open import Categories.Adjoint open import Categories.Functor @@ -26,12 +29,38 @@ record SplitObj : Set (suc o ⊔ suc ℓ ⊔ suc e) where F : Functor C D G : Functor D C adj : F ⊣ G - eqM : G ∘F F ≃ Monad.F M - unit-eq : ∀ A → eqM.⇐.η A C.∘ M.η.η A C.≈ adj.unit.η A - -- mult-eq : - -- unit M = unit adj - -- mult M = G . counit . F - -- probably this is asking too little... + GF≃M : G ∘F F ≃ M.F + + module D = Category D + module F = Functor F + module G = Functor G + module adj = Adjoint adj + module GF≃M = NaturalIsomorphism GF≃M + + field + η-eq : ∀ {A} → GF≃M.⇐.η A ∘ M.η.η A ≈ adj.unit.η A + μ-eq : ∀ {A} → GF≃M.⇐.η A ∘ M.μ.η A ≈ G.₁ (adj.counit.η _) ∘ GF≃M.⇐.η _ ∘ M.F.F₁ (GF≃M.⇐.η A) + + {- + η-eq: + adj.unit.η A + A ------------------+ + M.η.η A | | + v v + MA --------------> GFA + GF≃M.⇐.η A + + μ-eq: + + M.F.F₁ (GF≃M.⇐.η A) GF≃M.⇐.η + MMA ---------------------> GFM ----------> GFGFA + | | + | M.μ A | G.₁ (adj.counit.η A) + v v + MA --------------------------------------> GFA + GF≃M.⇐.η A + -} + record Split⇒ (X Y : SplitObj) : Set (suc o ⊔ suc ℓ ⊔ suc e) where constructor split⇒ diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda new file mode 100644 index 000000000..617ec850e --- /dev/null +++ b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda @@ -0,0 +1,149 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level +open import Categories.Category.Core using (Category) +open import Categories.Category +open import Categories.Monad + +module Categories.Adjoint.Construction.Adjunctions.Properties {o ℓ e} {C : Category (o ⊔ ℓ ⊔ e) (ℓ ⊔ e) e} (M : Monad C) where + +module C = Category C +module M = Monad M + +open import Categories.Adjoint +open import Categories.Functor +open import Categories.Morphism.Reasoning +open import Categories.NaturalTransformation.Core as NT +open import Categories.NaturalTransformation.NaturalIsomorphism using (NaturalIsomorphism; niHelper) + +open import Categories.Adjoint.Construction.Adjunctions M + +open import Categories.Object.Terminal (Split M) +open import Categories.Object.Initial (Split M) +open import Categories.Category.Construction.EilenbergMoore +open import Categories.Category.Construction.Kleisli +open import Categories.Adjoint.Construction.Kleisli M as KL +open import Categories.Adjoint.Construction.EilenbergMoore M as EM + +EM-object : SplitObj +EM-object = record + { D = EilenbergMoore M + ; F = EM.Free + ; G = EM.Forgetful + ; adj = EM.Free⊣Forgetful + ; GF≃M = EM.FF≃F + ; η-eq = begin + M.F.F₁ C.id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ + C.id ∘ M.η.η _ ≈⟨ identityˡ ⟩ + M.η.η _ ∎ + ; μ-eq = begin + M.F.F₁ C.id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ + C.id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ + M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ + M.μ.η _ ∘ C.id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + M.μ.η _ ∘ M.F.F₁ C.id ≈⟨ Equiv.sym identityʳ ⟩ + (M.μ.η _ ∘ M.F.F₁ C.id) ∘ C.id ≈⟨ assoc ⟩ + M.μ.η _ ∘ M.F.F₁ C.id ∘ C.id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + M.μ.η _ ∘ M.F.F₁ C.id ∘ M.F.F₁ C.id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ + M.μ.η _ ∘ M.F.F₁ C.id ∘ M.F.F₁ (M.F.F₁ C.id) ∎ + } where open C + open HomReasoning + +Kl-object : SplitObj +Kl-object = record + { D = Kleisli M + ; F = KL.Free + ; G = KL.Forgetful + ; adj = KL.Free⊣Forgetful + ; GF≃M = KL.FF≃F + ; η-eq = begin + M.F.F₁ C.id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ + C.id ∘ M.η.η _ ≈⟨ identityˡ ⟩ + M.η.η _ ∎ + ; μ-eq = begin + M.F.F₁ C.id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ + C.id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ + M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ + M.μ.η _ ∘ C.id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + M.μ.η _ ∘ M.F.F₁ C.id ≈⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ + M.μ.η _ ∘ M.F.F₁ (M.F.F₁ C.id) ≈⟨ Equiv.sym identityʳ ⟩ + (M.μ.η _ ∘ M.F.F₁ (M.F.F₁ C.id)) ∘ C.id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + (M.μ.η _ ∘ M.F.F₁ (M.F.F₁ C.id)) ∘ M.F.F₁ C.id ≈⟨ Equiv.sym identityʳ ⟩ + ((M.μ.η _ ∘ M.F.F₁ (M.F.F₁ C.id)) ∘ M.F.F₁ C.id) ∘ C.id ≈⟨ assoc ⟩ + (M.μ.η _ ∘ M.F.F₁ (M.F.F₁ C.id)) ∘ M.F.F₁ C.id ∘ C.id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + (M.μ.η _ ∘ M.F.F₁ (M.F.F₁ C.id)) ∘ M.F.F₁ C.id ∘ M.F.F₁ C.id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ + (M.μ.η _ ∘ M.F.F₁ (M.F.F₁ C.id)) ∘ M.F.F₁ C.id ∘ M.F.F₁ (M.F.F₁ C.id) ∎ + } where open C + open HomReasoning + +EM-terminal : IsTerminal EM-object +EM-terminal = record + { ! = {! !} + ; !-unique = {! !} + } + +Kl-initial : IsInitial Kl-object +Kl-initial = record + { ! = bang + ; !-unique = λ { {A} f → niHelper ( + let open SplitObj A + open Split⇒ f in + record + { η = λ X → {! !} + ; η⁻¹ = {! !} + ; commute = {! !} + ; iso = {! !} + }) + } + } + where + bang : {A : SplitObj} → Split⇒ Kl-object A + bang {splitobj D F G adj GF≃M η-eq μ-eq} = record + { H = record + { F₀ = F.F₀ + ; F₁ = λ f → adj.counit.η _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) + ; identity = λ {A} → begin + adj.counit.η (F.₀ A) ∘ F.₁ (GF≃M.⇐.η A C.∘ M.η.η A) ≈⟨ refl⟩∘⟨ F.F-resp-≈ η-eq ⟩ + adj.counit.η (F.₀ A) ∘ F.₁ (adj.unit.η A) ≈⟨ adj.zig ⟩ + D.id ∎ + ; homomorphism = λ {X} {Y} {_} {f} {g} → + begin + adj.counit.η _ ∘ F.F₁ (GF≃M.⇐.η _ C.∘ (M.μ.η _ C.∘ M.F.F₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ C.assoc) ⟩ + adj.counit.η _ ∘ F.F₁ (GF≃M.⇐.η _ C.∘ M.μ.η _ C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C μ-eq) ⟩ + adj.counit.η _ ∘ F.F₁ ((G.F₁ (adj.counit.η _) C.∘ GF≃M.⇐.η _ C.∘ M.F.F₁ (GF≃M.⇐.η _)) C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ CMR.assoc²' ⟩ + adj.counit.η _ ∘ F.F₁ (G.F₁ (adj.counit.η _) C.∘ GF≃M.⇐.η _ C.∘ M.F.F₁ (GF≃M.⇐.η _) C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ pullˡ C (GF≃M.⇐.commute _)) ⟩ + adj.counit.η _ ∘ F.F₁ (G.F₁ (adj.counit.η _) C.∘ (G.F₁ (F.F₁ (GF≃M.⇐.η _)) C.∘ GF≃M.⇐.η _) C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ C.assoc) ⟩ + adj.counit.η _ ∘ F.F₁ (G.F₁ (adj.counit.η _) C.∘ G.F₁ (F.F₁ (GF≃M.⇐.η _)) C.∘ GF≃M.⇐.η _ C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ CHR.refl⟩∘⟨ pullˡ C (GF≃M.⇐.commute _)) ⟩ + adj.counit.η _ ∘ F.F₁ (G.F₁ (adj.counit.η _) C.∘ G.F₁ (F.F₁ (GF≃M.⇐.η _)) C.∘ (G.F₁ (F.F₁ g) C.∘ GF≃M.⇐.η Y) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ {! !} ⟩ + adj.counit.η _ ∘ F.F₁ ((G.F₁ (adj.counit.η _) C.∘ G.F₁ (F.F₁ (GF≃M.⇐.η _)) C.∘ G.F₁ (F.F₁ g)) C.∘ GF≃M.⇐.η Y C.∘ f) ≈⟨ refl⟩∘⟨ F.homomorphism ⟩ + adj.counit.η _ ∘ F.F₁ (G.F₁ (adj.counit.η _) C.∘ G.F₁ (F.F₁ (GF≃M.⇐.η _)) C.∘ G.F₁ (F.F₁ g)) ∘ F.F₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C (C.Equiv.sym G.homomorphism)) ⟩∘⟨refl ⟩ + adj.counit.η _ ∘ F.F₁ (G.F₁ (adj.counit.η _ ∘ F.F₁ (GF≃M.⇐.η _)) C.∘ G.F₁ (F.F₁ g)) ∘ F.F₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ (refl⟩∘⟨ F.F-resp-≈ (G.F-resp-≈ {! !} CHR.⟩∘⟨refl) ⟩∘⟨refl) ⟩ + {! !} ≈⟨ {! !} ⟩ + ((adj.counit.η _ ∘ F.F₁ (GF≃M.⇐.η _ C.∘ g)) ∘ adj.counit.η _ ∘ F.F₁ (GF≃M.⇐.η Y C.∘ f)) ∎ + ; F-resp-≈ = λ x → D.∘-resp-≈ʳ (F.F-resp-≈ (C.∘-resp-≈ʳ x)) + } + ; HF≃F' = niHelper (record + { η = λ _ → D.id + ; η⁻¹ = λ _ → D.id + ; commute = λ f → begin + D.id ∘ adj.counit.η _ ∘ F.F₁ (GF≃M.⇐.η _ C.∘ M.η.η _ C.∘ f) ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C η-eq)) ⟩ + (D.id ∘ adj.counit.η _ ∘ F.F₁ (adj.unit.η _ C.∘ f)) ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ F.homomorphism) ⟩ + (D.id ∘ adj.counit.η _ ∘ F.F₁ (adj.unit.η _) ∘ F.F₁ f) ≈⟨ (refl⟩∘⟨ pullˡ D adj.zig) ⟩ + (D.id ∘ D.id ∘ F.F₁ f) ≈⟨ identityˡ ⟩ + (D.id ∘ F.F₁ f) ≈⟨ identityˡ ⟩ + F.F₁ f ≈⟨ D.Equiv.sym identityʳ ⟩ + (F.F₁ f ∘ D.id) ∎ + ; iso = λ X → record { isoˡ = identityˡ ; isoʳ = identityˡ } + }) + ; G'H≃G = {! !} + } where + module adj = Adjoint adj + module F = Functor F + module G = Functor G + module GF≃M = NaturalIsomorphism GF≃M + module D = Category D + open D + open D.HomReasoning + module CHR = C.HomReasoning + module DMR = Categories.Morphism.Reasoning D + module CMR = Categories.Morphism.Reasoning C From 116611831e406dbdcbf33e60bef5d05208ee55bb Mon Sep 17 00:00:00 2001 From: iwilare Date: Sat, 23 Jul 2022 17:54:45 +0200 Subject: [PATCH 16/46] Refactor with simpler names --- .../Adjoint/Construction/Adjunctions.agda | 13 ++-- .../Construction/Adjunctions/Properties.agda | 65 +++++++++++++++---- 2 files changed, 59 insertions(+), 19 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index c7cfefa44..19ed73b98 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -52,12 +52,13 @@ record SplitObj : Set (suc o ⊔ suc ℓ ⊔ suc e) where μ-eq: - M.F.F₁ (GF≃M.⇐.η A) GF≃M.⇐.η - MMA ---------------------> GFM ----------> GFGFA - | | - | M.μ A | G.₁ (adj.counit.η A) - v v - MA --------------------------------------> GFA + GF≃M.⇐.η G.F₁ (F.F₁ (GF≃M.⇐.η)) + MMA ---------------------> GFMA --------------------> GFGFA + ---------------------> MGFA --------------------> + | M.F.F₁ (GF≃M.⇐.η A) GF≃M.⇐.η | + | M.μ A | G.₁ (adj.counit.η A) + v v + MA -------------------------------------------------> GFA GF≃M.⇐.η A -} diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda index 617ec850e..941b402f4 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda @@ -106,22 +106,61 @@ Kl-initial = record adj.counit.η (F.₀ A) ∘ F.₁ (GF≃M.⇐.η A C.∘ M.η.η A) ≈⟨ refl⟩∘⟨ F.F-resp-≈ η-eq ⟩ adj.counit.η (F.₀ A) ∘ F.₁ (adj.unit.η A) ≈⟨ adj.zig ⟩ D.id ∎ - ; homomorphism = λ {X} {Y} {_} {f} {g} → + ; homomorphism = λ {X} {Y} {Z} {f} {g} → + let FF = F.F₁ in + let GG = G.F₁ in + let ε x = adj.counit.η x in + let 𝑀 x = M.μ.η x in + let Γ x = GF≃M.⇐.η x in + let lemma : ∀ {A} → G.F₁ (F.F₁ (GF≃M.⇐.η A)) C.∘ GF≃M.⇐.η _ + C.≈ GF≃M.⇐.η _ C.∘ M.F.F₁ (GF≃M.⇐.η A) + lemma {A} = C.Equiv.sym (GF≃M.⇐.commute (GF≃M.⇐.η A)) in + let super : FF (GG (ε _)) ∘ FF (GG (FF (Γ _))) + ≈ FF (Γ _) ∘ ε _ + super = + begin + FF (GG (ε (F.F₀ Z))) ∘ FF (GG (FF (Γ Z))) ≈⟨ {! !} ⟩ + FF (GG (ε (F.F₀ Z)) C.∘ GG (FF (Γ Z))) ≈⟨ {! !} ⟩ + {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ {! !} ⟩ + FF (Γ Z) ∘ ε (F.F₀ (M.F.F₀ Z)) ∎ + in begin - adj.counit.η _ ∘ F.F₁ (GF≃M.⇐.η _ C.∘ (M.μ.η _ C.∘ M.F.F₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ C.assoc) ⟩ - adj.counit.η _ ∘ F.F₁ (GF≃M.⇐.η _ C.∘ M.μ.η _ C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C μ-eq) ⟩ - adj.counit.η _ ∘ F.F₁ ((G.F₁ (adj.counit.η _) C.∘ GF≃M.⇐.η _ C.∘ M.F.F₁ (GF≃M.⇐.η _)) C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ CMR.assoc²' ⟩ - adj.counit.η _ ∘ F.F₁ (G.F₁ (adj.counit.η _) C.∘ GF≃M.⇐.η _ C.∘ M.F.F₁ (GF≃M.⇐.η _) C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ pullˡ C (GF≃M.⇐.commute _)) ⟩ - adj.counit.η _ ∘ F.F₁ (G.F₁ (adj.counit.η _) C.∘ (G.F₁ (F.F₁ (GF≃M.⇐.η _)) C.∘ GF≃M.⇐.η _) C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ C.assoc) ⟩ - adj.counit.η _ ∘ F.F₁ (G.F₁ (adj.counit.η _) C.∘ G.F₁ (F.F₁ (GF≃M.⇐.η _)) C.∘ GF≃M.⇐.η _ C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ CHR.refl⟩∘⟨ pullˡ C (GF≃M.⇐.commute _)) ⟩ - adj.counit.η _ ∘ F.F₁ (G.F₁ (adj.counit.η _) C.∘ G.F₁ (F.F₁ (GF≃M.⇐.η _)) C.∘ (G.F₁ (F.F₁ g) C.∘ GF≃M.⇐.η Y) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ {! !} ⟩ - adj.counit.η _ ∘ F.F₁ ((G.F₁ (adj.counit.η _) C.∘ G.F₁ (F.F₁ (GF≃M.⇐.η _)) C.∘ G.F₁ (F.F₁ g)) C.∘ GF≃M.⇐.η Y C.∘ f) ≈⟨ refl⟩∘⟨ F.homomorphism ⟩ - adj.counit.η _ ∘ F.F₁ (G.F₁ (adj.counit.η _) C.∘ G.F₁ (F.F₁ (GF≃M.⇐.η _)) C.∘ G.F₁ (F.F₁ g)) ∘ F.F₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C (C.Equiv.sym G.homomorphism)) ⟩∘⟨refl ⟩ - adj.counit.η _ ∘ F.F₁ (G.F₁ (adj.counit.η _ ∘ F.F₁ (GF≃M.⇐.η _)) C.∘ G.F₁ (F.F₁ g)) ∘ F.F₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ (refl⟩∘⟨ F.F-resp-≈ (G.F-resp-≈ {! !} CHR.⟩∘⟨refl) ⟩∘⟨refl) ⟩ - {! !} ≈⟨ {! !} ⟩ - ((adj.counit.η _ ∘ F.F₁ (GF≃M.⇐.η _ C.∘ g)) ∘ adj.counit.η _ ∘ F.F₁ (GF≃M.⇐.η Y C.∘ f)) ∎ + ε _ ∘ FF (Γ _ C.∘ (𝑀 _ C.∘ M.F.F₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ C.assoc) ⟩ + ε _ ∘ FF (Γ _ C.∘ 𝑀 _ C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C μ-eq) ⟩ + ε _ ∘ FF ((GG (ε _) C.∘ Γ _ C.∘ M.F.F₁ (Γ _)) C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ CMR.assoc²' ⟩ + ε _ ∘ FF (GG (ε _) C.∘ Γ _ C.∘ M.F.F₁ (Γ _) C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ pullˡ C (GF≃M.⇐.commute _)) ⟩ + ε _ ∘ FF (GG (ε _) C.∘ (GG (FF (Γ _)) C.∘ Γ _) C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ C.assoc) ⟩ + ε _ ∘ FF (GG (ε _) C.∘ GG (FF (Γ _)) C.∘ Γ _ C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ CHR.refl⟩∘⟨ pullˡ C (GF≃M.⇐.commute _)) ⟩ + ε _ ∘ FF (GG (ε _) C.∘ GG (FF (Γ _)) C.∘ (GG (FF g) C.∘ Γ _) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ {! !} ⟩ + ε _ ∘ FF ((GG (ε _) C.∘ GG (FF (Γ _)) C.∘ GG (FF g)) C.∘ Γ _ C.∘ f) ≈⟨ refl⟩∘⟨ F.homomorphism ⟩ + ε _ ∘ FF (GG (ε _) C.∘ GG (FF (Γ _)) C.∘ GG (FF g)) ∘ FF (Γ _ C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C (C.Equiv.sym G.homomorphism)) ⟩∘⟨refl ⟩ + ε _ ∘ FF (GG (ε _ ∘ FF (Γ _)) C.∘ GG (FF g)) ∘ FF (Γ _ C.∘ f) ≈⟨ (refl⟩∘⟨ F.homomorphism ⟩∘⟨refl) ⟩ + ε _ ∘ (FF (GG (ε _ ∘ FF (Γ _))) ∘ FF (GG (FF g))) ∘ FF (Γ _ C.∘ f) ≈⟨ (refl⟩∘⟨ (F.F-resp-≈ G.homomorphism ⟩∘⟨refl) ⟩∘⟨refl) ⟩ + ε _ ∘ ((FF (GG (ε _) C.∘ GG (FF (Γ _)))) ∘ FF (GG (FF g))) ∘ FF (Γ _ C.∘ f) ≈⟨ (refl⟩∘⟨ (F.homomorphism ⟩∘⟨refl) ⟩∘⟨refl) ⟩ + ε _ ∘ ((FF (GG (ε _)) ∘ FF (GG (FF (Γ _)))) ∘ FF (GG (FF g))) ∘ FF (Γ _ C.∘ f) ≈⟨ (refl⟩∘⟨ (super ⟩∘⟨refl) ⟩∘⟨refl) ⟩ + ε _ ∘ ((FF (Γ _) ∘ ε _) ∘ FF (GG (FF g))) ∘ FF (Γ _ C.∘ f) ≈⟨ (refl⟩∘⟨ assoc ⟩∘⟨refl) ⟩ + ε _ ∘ (FF (Γ _) ∘ (ε _ ∘ FF (GG (FF g)))) ∘ FF (Γ _ C.∘ f) ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ adj.counit.commute _) ⟩∘⟨refl) ⟩ + ε _ ∘ (FF (Γ _) ∘ FF g ∘ ε _) ∘ FF (Γ _ C.∘ f) ≈⟨ (refl⟩∘⟨ Equiv.sym assoc ⟩∘⟨refl) ⟩ + ε _ ∘ ((FF (Γ _) ∘ FF g) ∘ ε _) ∘ FF (Γ _ C.∘ f) ≈⟨ (refl⟩∘⟨ (Equiv.sym F.homomorphism ⟩∘⟨refl) ⟩∘⟨refl) ⟩ + ε _ ∘ (FF (Γ _ C.∘ g) ∘ ε _) ∘ FF (Γ _ C.∘ f) ≈⟨ DMR.assoc²'' ⟩ + (ε _ ∘ FF (Γ _ C.∘ g)) ∘ ε _ ∘ FF (Γ _ C.∘ f) ∎ ; F-resp-≈ = λ x → D.∘-resp-≈ʳ (F.F-resp-≈ (C.∘-resp-≈ʳ x)) } +{- + + F (G (ε (F Z) ∘ F (Γ Z)) ∘ G (F g)) + F (G (ε (F Z) ∘ F (Γ Z))) ∘ F (G (F g)) + F (G (ε (F Z)) ∘ G (F (Γ Z))) ∘ F (G (F g)) + F (G (ε (F Z))) ∘ F (G (F (Γ Z))) ∘ F (G (F g)) + F (G (ε (F Z))) ∘ F (G (F (Γ Z))) ∘ F (G (F g)) + ≈ F (Γ Z ∘ g) ∘ ε (F Y) + +-} + + + ; HF≃F' = niHelper (record { η = λ _ → D.id ; η⁻¹ = λ _ → D.id From f1ad63f8f892459c63f27b5e985503d7b8e89590 Mon Sep 17 00:00:00 2001 From: iwilare Date: Sat, 6 Aug 2022 17:04:14 +0200 Subject: [PATCH 17/46] Simplify homomorphism, refactor into lemma on adjunction --- .../Adjoint/Construction/Adjunctions.agda | 11 +- .../Construction/Adjunctions/Properties.agda | 118 ++++++++---------- 2 files changed, 56 insertions(+), 73 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 19ed73b98..469ae17f4 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -38,8 +38,8 @@ record SplitObj : Set (suc o ⊔ suc ℓ ⊔ suc e) where module GF≃M = NaturalIsomorphism GF≃M field - η-eq : ∀ {A} → GF≃M.⇐.η A ∘ M.η.η A ≈ adj.unit.η A - μ-eq : ∀ {A} → GF≃M.⇐.η A ∘ M.μ.η A ≈ G.₁ (adj.counit.η _) ∘ GF≃M.⇐.η _ ∘ M.F.F₁ (GF≃M.⇐.η A) + η-eq : ∀ {A} → GF≃M.⇐.η _ ∘ M.η.η A ≈ adj.unit.η _ + μ-eq : ∀ {A} → GF≃M.⇐.η _ ∘ M.μ.η A ≈ G.₁ (adj.counit.η (F.₀ A)) ∘ GF≃M.⇐.η _ ∘ M.F.F₁ (GF≃M.⇐.η A) {- η-eq: @@ -53,10 +53,11 @@ record SplitObj : Set (suc o ⊔ suc ℓ ⊔ suc e) where μ-eq: GF≃M.⇐.η G.F₁ (F.F₁ (GF≃M.⇐.η)) - MMA ---------------------> GFMA --------------------> GFGFA - ---------------------> MGFA --------------------> + ---------------------> GFMA --------------------> + MMA ---------------------> MGFA --------------------> GFGFA | M.F.F₁ (GF≃M.⇐.η A) GF≃M.⇐.η | - | M.μ A | G.₁ (adj.counit.η A) + | | + | M.μ A | G.₁ (adj.counit.η (F.₀ A)) v v MA -------------------------------------------------> GFA GF≃M.⇐.η A diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda index 941b402f4..725d3f5a7 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda @@ -25,6 +25,8 @@ open import Categories.Category.Construction.Kleisli open import Categories.Adjoint.Construction.Kleisli M as KL open import Categories.Adjoint.Construction.EilenbergMoore M as EM +open import Categories.Tactic.Category + EM-object : SplitObj EM-object = record { D = EilenbergMoore M @@ -60,8 +62,7 @@ Kl-object = record M.F.F₁ C.id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ C.id ∘ M.η.η _ ≈⟨ identityˡ ⟩ M.η.η _ ∎ - ; μ-eq = begin - M.F.F₁ C.id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ + ; μ-eq = begin M.F.F₁ C.id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ C.id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ M.μ.η _ ∘ C.id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ @@ -102,76 +103,57 @@ Kl-initial = record { H = record { F₀ = F.F₀ ; F₁ = λ f → adj.counit.η _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) - ; identity = λ {A} → begin - adj.counit.η (F.₀ A) ∘ F.₁ (GF≃M.⇐.η A C.∘ M.η.η A) ≈⟨ refl⟩∘⟨ F.F-resp-≈ η-eq ⟩ - adj.counit.η (F.₀ A) ∘ F.₁ (adj.unit.η A) ≈⟨ adj.zig ⟩ - D.id ∎ + ; identity = λ {A} → + begin + adj.counit.η (F.₀ A) ∘ F.₁ (GF≃M.⇐.η A C.∘ M.η.η A) ≈⟨ refl⟩∘⟨ F.F-resp-≈ η-eq ⟩ + adj.counit.η (F.₀ A) ∘ F.₁ (adj.unit.η A) ≈⟨ adj.zig ⟩ + D.id ∎ ; homomorphism = λ {X} {Y} {Z} {f} {g} → - let FF = F.F₁ in - let GG = G.F₁ in - let ε x = adj.counit.η x in - let 𝑀 x = M.μ.η x in - let Γ x = GF≃M.⇐.η x in - let lemma : ∀ {A} → G.F₁ (F.F₁ (GF≃M.⇐.η A)) C.∘ GF≃M.⇐.η _ - C.≈ GF≃M.⇐.η _ C.∘ M.F.F₁ (GF≃M.⇐.η A) - lemma {A} = C.Equiv.sym (GF≃M.⇐.commute (GF≃M.⇐.η A)) in - let super : FF (GG (ε _)) ∘ FF (GG (FF (Γ _))) - ≈ FF (Γ _) ∘ ε _ - super = - begin - FF (GG (ε (F.F₀ Z))) ∘ FF (GG (FF (Γ Z))) ≈⟨ {! !} ⟩ - FF (GG (ε (F.F₀ Z)) C.∘ GG (FF (Γ Z))) ≈⟨ {! !} ⟩ - {! !} ≈⟨ {! !} ⟩ - {! !} ≈⟨ {! !} ⟩ - {! !} ≈⟨ {! !} ⟩ - FF (Γ Z) ∘ ε (F.F₀ (M.F.F₀ Z)) ∎ - in - begin - ε _ ∘ FF (Γ _ C.∘ (𝑀 _ C.∘ M.F.F₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ C.assoc) ⟩ - ε _ ∘ FF (Γ _ C.∘ 𝑀 _ C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C μ-eq) ⟩ - ε _ ∘ FF ((GG (ε _) C.∘ Γ _ C.∘ M.F.F₁ (Γ _)) C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ CMR.assoc²' ⟩ - ε _ ∘ FF (GG (ε _) C.∘ Γ _ C.∘ M.F.F₁ (Γ _) C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ pullˡ C (GF≃M.⇐.commute _)) ⟩ - ε _ ∘ FF (GG (ε _) C.∘ (GG (FF (Γ _)) C.∘ Γ _) C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ C.assoc) ⟩ - ε _ ∘ FF (GG (ε _) C.∘ GG (FF (Γ _)) C.∘ Γ _ C.∘ M.F.F₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ CHR.refl⟩∘⟨ pullˡ C (GF≃M.⇐.commute _)) ⟩ - ε _ ∘ FF (GG (ε _) C.∘ GG (FF (Γ _)) C.∘ (GG (FF g) C.∘ Γ _) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ {! !} ⟩ - ε _ ∘ FF ((GG (ε _) C.∘ GG (FF (Γ _)) C.∘ GG (FF g)) C.∘ Γ _ C.∘ f) ≈⟨ refl⟩∘⟨ F.homomorphism ⟩ - ε _ ∘ FF (GG (ε _) C.∘ GG (FF (Γ _)) C.∘ GG (FF g)) ∘ FF (Γ _ C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C (C.Equiv.sym G.homomorphism)) ⟩∘⟨refl ⟩ - ε _ ∘ FF (GG (ε _ ∘ FF (Γ _)) C.∘ GG (FF g)) ∘ FF (Γ _ C.∘ f) ≈⟨ (refl⟩∘⟨ F.homomorphism ⟩∘⟨refl) ⟩ - ε _ ∘ (FF (GG (ε _ ∘ FF (Γ _))) ∘ FF (GG (FF g))) ∘ FF (Γ _ C.∘ f) ≈⟨ (refl⟩∘⟨ (F.F-resp-≈ G.homomorphism ⟩∘⟨refl) ⟩∘⟨refl) ⟩ - ε _ ∘ ((FF (GG (ε _) C.∘ GG (FF (Γ _)))) ∘ FF (GG (FF g))) ∘ FF (Γ _ C.∘ f) ≈⟨ (refl⟩∘⟨ (F.homomorphism ⟩∘⟨refl) ⟩∘⟨refl) ⟩ - ε _ ∘ ((FF (GG (ε _)) ∘ FF (GG (FF (Γ _)))) ∘ FF (GG (FF g))) ∘ FF (Γ _ C.∘ f) ≈⟨ (refl⟩∘⟨ (super ⟩∘⟨refl) ⟩∘⟨refl) ⟩ - ε _ ∘ ((FF (Γ _) ∘ ε _) ∘ FF (GG (FF g))) ∘ FF (Γ _ C.∘ f) ≈⟨ (refl⟩∘⟨ assoc ⟩∘⟨refl) ⟩ - ε _ ∘ (FF (Γ _) ∘ (ε _ ∘ FF (GG (FF g)))) ∘ FF (Γ _ C.∘ f) ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ adj.counit.commute _) ⟩∘⟨refl) ⟩ - ε _ ∘ (FF (Γ _) ∘ FF g ∘ ε _) ∘ FF (Γ _ C.∘ f) ≈⟨ (refl⟩∘⟨ Equiv.sym assoc ⟩∘⟨refl) ⟩ - ε _ ∘ ((FF (Γ _) ∘ FF g) ∘ ε _) ∘ FF (Γ _ C.∘ f) ≈⟨ (refl⟩∘⟨ (Equiv.sym F.homomorphism ⟩∘⟨refl) ⟩∘⟨refl) ⟩ - ε _ ∘ (FF (Γ _ C.∘ g) ∘ ε _) ∘ FF (Γ _ C.∘ f) ≈⟨ DMR.assoc²'' ⟩ - (ε _ ∘ FF (Γ _ C.∘ g)) ∘ ε _ ∘ FF (Γ _ C.∘ f) ∎ - ; F-resp-≈ = λ x → D.∘-resp-≈ʳ (F.F-resp-≈ (C.∘-resp-≈ʳ x)) - } -{- - - F (G (ε (F Z) ∘ F (Γ Z)) ∘ G (F g)) - F (G (ε (F Z) ∘ F (Γ Z))) ∘ F (G (F g)) - F (G (ε (F Z)) ∘ G (F (Γ Z))) ∘ F (G (F g)) - F (G (ε (F Z))) ∘ F (G (F (Γ Z))) ∘ F (G (F g)) - F (G (ε (F Z))) ∘ F (G (F (Γ Z))) ∘ F (G (F g)) - ≈ F (Γ Z ∘ g) ∘ ε (F Y) - --} - - - + let 𝐹 = F.F₁ in + let 𝐺 = G.F₁ in + let ε x = adj.counit.η x in + let μ x = M.μ.η x in + let M⇒GF x = GF≃M.⇐.η x in + let GF⇒M x = GF≃M.⇒.η x in + let 𝑓 = M⇒GF _ C.∘ f in + let 𝑔 = M⇒GF _ C.∘ g in + let lemma : 𝐹 (𝐺 (ε (F.F₀ Z))) ≈ ε (F.F₀ (G.F₀ (F.F₀ Z))) + lemma = {! !} in + begin + -- Push f externally + ε _ ∘ 𝐹 (M⇒GF _ C.∘ (μ _ C.∘ M.F.F₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ {! DMR.assoc²'' !} ⟩ + ε _ ∘ 𝐹 ((M⇒GF _ C.∘ μ _ C.∘ M.F.F₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.homomorphism ⟩ + ε _ ∘ 𝐹 (M⇒GF _ C.∘ μ _ C.∘ M.F.F₁ g) ∘ 𝐹 f ≈⟨ refl⟩∘⟨ F.F-resp-≈ C.sym-assoc ⟩∘⟨refl ⟩ + ε _ ∘ 𝐹 ((M⇒GF _ C.∘ μ _) C.∘ M.F.F₁ g) ∘ 𝐹 f ≈⟨ refl⟩∘⟨ F.F-resp-≈ (μ-eq CHR.⟩∘⟨refl) ⟩∘⟨refl ⟩ + -- Switch from M to GF as close to g as possible + ε _ ∘ 𝐹 ((𝐺 (ε (F.₀ _)) C.∘ M⇒GF _ C.∘ M.F.F₁ (M⇒GF _)) C.∘ M.F.F₁ g) ∘ 𝐹 f ≈⟨ refl⟩∘⟨ F.F-resp-≈ {! !} ⟩∘⟨refl ⟩ + ε _ ∘ 𝐹 (𝐺 (ε (F.₀ _)) C.∘ M⇒GF _ C.∘ (M.F.F₁ (M⇒GF _) C.∘ M.F.F₁ g)) ∘ 𝐹 f ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ CHR.refl⟩∘⟨ C.Equiv.sym M.F.homomorphism) ⟩∘⟨refl ⟩ + ε _ ∘ 𝐹 (𝐺 (ε (F.₀ _)) C.∘ M⇒GF _ C.∘ (M.F.F₁ 𝑔)) ∘ 𝐹 f ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ GF≃M.⇐.commute _) ⟩∘⟨refl ⟩ + -- Switch from M to GF as close to f as possible + ε _ ∘ 𝐹 (𝐺 (ε (F.₀ _)) C.∘ 𝐺 (𝐹 𝑔) C.∘ M⇒GF _) ∘ 𝐹 f ≈⟨ ((refl⟩∘⟨ F.F-resp-≈ C.sym-assoc ⟩∘⟨refl)) ⟩ + ε _ ∘ 𝐹 ((𝐺 (ε (F.₀ _)) C.∘ 𝐺 (𝐹 𝑔)) C.∘ M⇒GF _) ∘ 𝐹 f ≈⟨ (refl⟩∘⟨ F.homomorphism ⟩∘⟨refl) ⟩ + ε _ ∘ (𝐹 ((𝐺 (ε (F.₀ _)) C.∘ 𝐺 (𝐹 𝑔))) ∘ 𝐹 (M⇒GF _)) ∘ 𝐹 f ≈⟨ (refl⟩∘⟨ assoc) ⟩ + ε _ ∘ 𝐹 ((𝐺 (ε (F.₀ _)) C.∘ 𝐺 (𝐹 𝑔))) ∘ 𝐹 (M⇒GF _) ∘ 𝐹 f ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ Equiv.sym F.homomorphism)) ⟩ + ε _ ∘ 𝐹 ((𝐺 (ε (F.₀ _)) C.∘ 𝐺 (𝐹 𝑔))) ∘ 𝐹 𝑓 ≈⟨ (refl⟩∘⟨ F.homomorphism ⟩∘⟨refl) ⟩ + -- Solve the adjunction + ε _ ∘ (𝐹 (𝐺 (ε (F.₀ _))) ∘ 𝐹 (𝐺 (𝐹 𝑔))) ∘ 𝐹 𝑓 ≈⟨ (refl⟩∘⟨ (lemma ⟩∘⟨refl) ⟩∘⟨refl) ⟩ + ε _ ∘ ((ε _) ∘ 𝐹 (𝐺 (𝐹 𝑔))) ∘ 𝐹 𝑓 ≈⟨ (refl⟩∘⟨ adj.counit.commute _ ⟩∘⟨refl) ⟩ + ε _ ∘ (𝐹 𝑔 ∘ ε _) ∘ 𝐹 𝑓 ≈⟨ DMR.assoc²'' ⟩ + (ε _ ∘ 𝐹 𝑔) ∘ ε _ ∘ 𝐹 𝑓 ∎ + ; F-resp-≈ = λ x → D.∘-resp-≈ʳ (F.F-resp-≈ (C.∘-resp-≈ʳ x)) + } ; HF≃F' = niHelper (record { η = λ _ → D.id ; η⁻¹ = λ _ → D.id - ; commute = λ f → begin - D.id ∘ adj.counit.η _ ∘ F.F₁ (GF≃M.⇐.η _ C.∘ M.η.η _ C.∘ f) ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C η-eq)) ⟩ - (D.id ∘ adj.counit.η _ ∘ F.F₁ (adj.unit.η _ C.∘ f)) ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ F.homomorphism) ⟩ - (D.id ∘ adj.counit.η _ ∘ F.F₁ (adj.unit.η _) ∘ F.F₁ f) ≈⟨ (refl⟩∘⟨ pullˡ D adj.zig) ⟩ - (D.id ∘ D.id ∘ F.F₁ f) ≈⟨ identityˡ ⟩ - (D.id ∘ F.F₁ f) ≈⟨ identityˡ ⟩ - F.F₁ f ≈⟨ D.Equiv.sym identityʳ ⟩ - (F.F₁ f ∘ D.id) ∎ + ; commute = λ f → + begin + D.id ∘ adj.counit.η _ ∘ F.F₁ (GF≃M.⇐.η _ C.∘ M.η.η _ C.∘ f) ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C η-eq)) ⟩ + (D.id ∘ adj.counit.η _ ∘ F.F₁ (adj.unit.η _ C.∘ f)) ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ F.homomorphism) ⟩ + (D.id ∘ adj.counit.η _ ∘ F.F₁ (adj.unit.η _) ∘ F.F₁ f) ≈⟨ (refl⟩∘⟨ pullˡ D adj.zig) ⟩ + (D.id ∘ D.id ∘ F.F₁ f) ≈⟨ identityˡ ⟩ + (D.id ∘ F.F₁ f) ≈⟨ identityˡ ⟩ + F.F₁ f ≈⟨ D.Equiv.sym identityʳ ⟩ + (F.F₁ f ∘ D.id) ∎ ; iso = λ X → record { isoˡ = identityˡ ; isoʳ = identityˡ } }) ; G'H≃G = {! !} From 0439ad0a79c2c14df56e0b33b254dbd759d08f79 Mon Sep 17 00:00:00 2001 From: iwilare Date: Thu, 18 Aug 2022 19:40:06 +0200 Subject: [PATCH 18/46] Complete homomorphism for bang --- .../Construction/Adjunctions/Properties.agda | 40 +++++-------------- 1 file changed, 11 insertions(+), 29 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda index 725d3f5a7..2a04c7f1f 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda @@ -109,37 +109,19 @@ Kl-initial = record adj.counit.η (F.₀ A) ∘ F.₁ (adj.unit.η A) ≈⟨ adj.zig ⟩ D.id ∎ ; homomorphism = λ {X} {Y} {Z} {f} {g} → - let 𝐹 = F.F₁ in - let 𝐺 = G.F₁ in let ε x = adj.counit.η x in - let μ x = M.μ.η x in - let M⇒GF x = GF≃M.⇐.η x in - let GF⇒M x = GF≃M.⇒.η x in - let 𝑓 = M⇒GF _ C.∘ f in - let 𝑔 = M⇒GF _ C.∘ g in - let lemma : 𝐹 (𝐺 (ε (F.F₀ Z))) ≈ ε (F.F₀ (G.F₀ (F.F₀ Z))) - lemma = {! !} in begin - -- Push f externally - ε _ ∘ 𝐹 (M⇒GF _ C.∘ (μ _ C.∘ M.F.F₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ {! DMR.assoc²'' !} ⟩ - ε _ ∘ 𝐹 ((M⇒GF _ C.∘ μ _ C.∘ M.F.F₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.homomorphism ⟩ - ε _ ∘ 𝐹 (M⇒GF _ C.∘ μ _ C.∘ M.F.F₁ g) ∘ 𝐹 f ≈⟨ refl⟩∘⟨ F.F-resp-≈ C.sym-assoc ⟩∘⟨refl ⟩ - ε _ ∘ 𝐹 ((M⇒GF _ C.∘ μ _) C.∘ M.F.F₁ g) ∘ 𝐹 f ≈⟨ refl⟩∘⟨ F.F-resp-≈ (μ-eq CHR.⟩∘⟨refl) ⟩∘⟨refl ⟩ - -- Switch from M to GF as close to g as possible - ε _ ∘ 𝐹 ((𝐺 (ε (F.₀ _)) C.∘ M⇒GF _ C.∘ M.F.F₁ (M⇒GF _)) C.∘ M.F.F₁ g) ∘ 𝐹 f ≈⟨ refl⟩∘⟨ F.F-resp-≈ {! !} ⟩∘⟨refl ⟩ - ε _ ∘ 𝐹 (𝐺 (ε (F.₀ _)) C.∘ M⇒GF _ C.∘ (M.F.F₁ (M⇒GF _) C.∘ M.F.F₁ g)) ∘ 𝐹 f ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ CHR.refl⟩∘⟨ C.Equiv.sym M.F.homomorphism) ⟩∘⟨refl ⟩ - ε _ ∘ 𝐹 (𝐺 (ε (F.₀ _)) C.∘ M⇒GF _ C.∘ (M.F.F₁ 𝑔)) ∘ 𝐹 f ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ GF≃M.⇐.commute _) ⟩∘⟨refl ⟩ - -- Switch from M to GF as close to f as possible - ε _ ∘ 𝐹 (𝐺 (ε (F.₀ _)) C.∘ 𝐺 (𝐹 𝑔) C.∘ M⇒GF _) ∘ 𝐹 f ≈⟨ ((refl⟩∘⟨ F.F-resp-≈ C.sym-assoc ⟩∘⟨refl)) ⟩ - ε _ ∘ 𝐹 ((𝐺 (ε (F.₀ _)) C.∘ 𝐺 (𝐹 𝑔)) C.∘ M⇒GF _) ∘ 𝐹 f ≈⟨ (refl⟩∘⟨ F.homomorphism ⟩∘⟨refl) ⟩ - ε _ ∘ (𝐹 ((𝐺 (ε (F.₀ _)) C.∘ 𝐺 (𝐹 𝑔))) ∘ 𝐹 (M⇒GF _)) ∘ 𝐹 f ≈⟨ (refl⟩∘⟨ assoc) ⟩ - ε _ ∘ 𝐹 ((𝐺 (ε (F.₀ _)) C.∘ 𝐺 (𝐹 𝑔))) ∘ 𝐹 (M⇒GF _) ∘ 𝐹 f ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ Equiv.sym F.homomorphism)) ⟩ - ε _ ∘ 𝐹 ((𝐺 (ε (F.₀ _)) C.∘ 𝐺 (𝐹 𝑔))) ∘ 𝐹 𝑓 ≈⟨ (refl⟩∘⟨ F.homomorphism ⟩∘⟨refl) ⟩ - -- Solve the adjunction - ε _ ∘ (𝐹 (𝐺 (ε (F.₀ _))) ∘ 𝐹 (𝐺 (𝐹 𝑔))) ∘ 𝐹 𝑓 ≈⟨ (refl⟩∘⟨ (lemma ⟩∘⟨refl) ⟩∘⟨refl) ⟩ - ε _ ∘ ((ε _) ∘ 𝐹 (𝐺 (𝐹 𝑔))) ∘ 𝐹 𝑓 ≈⟨ (refl⟩∘⟨ adj.counit.commute _ ⟩∘⟨refl) ⟩ - ε _ ∘ (𝐹 𝑔 ∘ ε _) ∘ 𝐹 𝑓 ≈⟨ DMR.assoc²'' ⟩ - (ε _ ∘ 𝐹 𝑔) ∘ ε _ ∘ 𝐹 𝑓 ∎ + ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ (M.μ.η _ C.∘ M.₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CMR.assoc²'') ⟩ + ε _ ∘ F.₁ ((GF≃M.⇐.η _ C.∘ M.μ.η _) C.∘ M.₁ g C.∘ f) ≈⟨ (refl⟩∘⟨ F.F-resp-≈ (μ-eq CHR.⟩∘⟨refl)) ⟩ + ε _ ∘ F.₁ ((G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.₁ (GF≃M.⇐.η _)) C.∘ M.₁ g C.∘ f) ≈⟨ (refl⟩∘⟨ F.F-resp-≈ CMR.assoc²') ⟩ + ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.₁ (GF≃M.⇐.η _) C.∘ M.₁ g C.∘ f) ≈⟨ (refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ CHR.refl⟩∘⟨ CMR.pullˡ (C.Equiv.sym M.F.homomorphism))) ⟩ + ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.₁ (GF≃M.⇐.η _ C.∘ g) C.∘ f) ≈⟨ (refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ CMR.extendʳ (GF≃M.⇐.commute _))) ⟩ + ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ (refl⟩∘⟨ F.homomorphism) ⟩ + ε _ ∘ F.₁ (G.₁ (ε (F.₀ _))) ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ DMR.extendʳ (adj.counit.commute _) ⟩ + ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ F.homomorphism) ⟩ + ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g))) ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ (refl⟩∘⟨ DMR.extendʳ (adj.counit.commute _)) ⟩ + ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ sym-assoc ⟩ + (ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g)) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ∎ ; F-resp-≈ = λ x → D.∘-resp-≈ʳ (F.F-resp-≈ (C.∘-resp-≈ʳ x)) } ; HF≃F' = niHelper (record From f876929bc69eb1803df01abac436ebdf7620faf9 Mon Sep 17 00:00:00 2001 From: iwilare Date: Fri, 19 Aug 2022 03:53:32 +0200 Subject: [PATCH 19/46] Progress on initiality of Kleisli --- .../Construction/Adjunctions/Properties.agda | 208 ++++++++++-------- 1 file changed, 119 insertions(+), 89 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda index 2a04c7f1f..68d3ae569 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda @@ -10,11 +10,10 @@ module Categories.Adjoint.Construction.Adjunctions.Properties {o ℓ e} {C : Cat module C = Category C module M = Monad M -open import Categories.Adjoint -open import Categories.Functor +open import Categories.Adjoint using (Adjoint) +open import Categories.Functor using (Functor) open import Categories.Morphism.Reasoning -open import Categories.NaturalTransformation.Core as NT -open import Categories.NaturalTransformation.NaturalIsomorphism using (NaturalIsomorphism; niHelper) +open import Categories.NaturalTransformation.NaturalIsomorphism using (NaturalIsomorphism; niHelper; sym) open import Categories.Adjoint.Construction.Adjunctions M @@ -35,20 +34,20 @@ EM-object = record ; adj = EM.Free⊣Forgetful ; GF≃M = EM.FF≃F ; η-eq = begin - M.F.F₁ C.id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ - C.id ∘ M.η.η _ ≈⟨ identityˡ ⟩ - M.η.η _ ∎ + M.F.₁ id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ + id ∘ M.η.η _ ≈⟨ identityˡ ⟩ + M.η.η _ ∎ ; μ-eq = begin - M.F.F₁ C.id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ - C.id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ - M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ - M.μ.η _ ∘ C.id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ - M.μ.η _ ∘ M.F.F₁ C.id ≈⟨ Equiv.sym identityʳ ⟩ - (M.μ.η _ ∘ M.F.F₁ C.id) ∘ C.id ≈⟨ assoc ⟩ - M.μ.η _ ∘ M.F.F₁ C.id ∘ C.id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ - M.μ.η _ ∘ M.F.F₁ C.id ∘ M.F.F₁ C.id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ - M.μ.η _ ∘ M.F.F₁ C.id ∘ M.F.F₁ (M.F.F₁ C.id) ∎ - } where open C + M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ + id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ + M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ + M.μ.η _ ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + M.μ.η _ ∘ M.F.₁ id ≈⟨ Equiv.sym identityʳ ⟩ + (M.μ.η _ ∘ M.F.₁ id) ∘ id ≈⟨ assoc ⟩ + M.μ.η _ ∘ M.F.₁ id ∘ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + M.μ.η _ ∘ M.F.₁ id ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ + M.μ.η _ ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ + } where open Category C open HomReasoning Kl-object : SplitObj @@ -59,22 +58,22 @@ Kl-object = record ; adj = KL.Free⊣Forgetful ; GF≃M = KL.FF≃F ; η-eq = begin - M.F.F₁ C.id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ - C.id ∘ M.η.η _ ≈⟨ identityˡ ⟩ - M.η.η _ ∎ - ; μ-eq = begin M.F.F₁ C.id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ - C.id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ - M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ - M.μ.η _ ∘ C.id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ - M.μ.η _ ∘ M.F.F₁ C.id ≈⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ - M.μ.η _ ∘ M.F.F₁ (M.F.F₁ C.id) ≈⟨ Equiv.sym identityʳ ⟩ - (M.μ.η _ ∘ M.F.F₁ (M.F.F₁ C.id)) ∘ C.id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ - (M.μ.η _ ∘ M.F.F₁ (M.F.F₁ C.id)) ∘ M.F.F₁ C.id ≈⟨ Equiv.sym identityʳ ⟩ - ((M.μ.η _ ∘ M.F.F₁ (M.F.F₁ C.id)) ∘ M.F.F₁ C.id) ∘ C.id ≈⟨ assoc ⟩ - (M.μ.η _ ∘ M.F.F₁ (M.F.F₁ C.id)) ∘ M.F.F₁ C.id ∘ C.id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ - (M.μ.η _ ∘ M.F.F₁ (M.F.F₁ C.id)) ∘ M.F.F₁ C.id ∘ M.F.F₁ C.id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ - (M.μ.η _ ∘ M.F.F₁ (M.F.F₁ C.id)) ∘ M.F.F₁ C.id ∘ M.F.F₁ (M.F.F₁ C.id) ∎ - } where open C + M.F.₁ id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ + id ∘ M.η.η _ ≈⟨ identityˡ ⟩ + M.η.η _ ∎ + ; μ-eq = begin M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ + id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ + M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ + M.μ.η _ ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + M.μ.η _ ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ + M.μ.η _ ∘ M.F.₁ (M.F.₁ id) ≈⟨ Equiv.sym identityʳ ⟩ + (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ≈⟨ Equiv.sym identityʳ ⟩ + ((M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id) ∘ id ≈⟨ assoc ⟩ + (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ + (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ + } where open Category C open HomReasoning EM-terminal : IsTerminal EM-object @@ -85,68 +84,99 @@ EM-terminal = record Kl-initial : IsInitial Kl-object Kl-initial = record - { ! = bang + { ! = ! ; !-unique = λ { {A} f → niHelper ( let open SplitObj A - open Split⇒ f in + open Split⇒ f + module G'H≃G = NaturalIsomorphism G'H≃G + module HF≃F' = NaturalIsomorphism HF≃F' in record - { η = λ X → {! !} - ; η⁻¹ = {! !} - ; commute = {! !} - ; iso = {! !} + { η = HF≃F'.⇐.η + ; η⁻¹ = HF≃F'.⇒.η + ; commute = + let open D + open D.HomReasoning in λ f → + {! !} + ; iso = NaturalIsomorphism.iso (sym HF≃F') }) } } where - bang : {A : SplitObj} → Split⇒ Kl-object A - bang {splitobj D F G adj GF≃M η-eq μ-eq} = record - { H = record - { F₀ = F.F₀ - ; F₁ = λ f → adj.counit.η _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) - ; identity = λ {A} → - begin - adj.counit.η (F.₀ A) ∘ F.₁ (GF≃M.⇐.η A C.∘ M.η.η A) ≈⟨ refl⟩∘⟨ F.F-resp-≈ η-eq ⟩ - adj.counit.η (F.₀ A) ∘ F.₁ (adj.unit.η A) ≈⟨ adj.zig ⟩ - D.id ∎ - ; homomorphism = λ {X} {Y} {Z} {f} {g} → - let ε x = adj.counit.η x in - begin - ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ (M.μ.η _ C.∘ M.₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CMR.assoc²'') ⟩ - ε _ ∘ F.₁ ((GF≃M.⇐.η _ C.∘ M.μ.η _) C.∘ M.₁ g C.∘ f) ≈⟨ (refl⟩∘⟨ F.F-resp-≈ (μ-eq CHR.⟩∘⟨refl)) ⟩ - ε _ ∘ F.₁ ((G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.₁ (GF≃M.⇐.η _)) C.∘ M.₁ g C.∘ f) ≈⟨ (refl⟩∘⟨ F.F-resp-≈ CMR.assoc²') ⟩ - ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.₁ (GF≃M.⇐.η _) C.∘ M.₁ g C.∘ f) ≈⟨ (refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ CHR.refl⟩∘⟨ CMR.pullˡ (C.Equiv.sym M.F.homomorphism))) ⟩ - ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.₁ (GF≃M.⇐.η _ C.∘ g) C.∘ f) ≈⟨ (refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ CMR.extendʳ (GF≃M.⇐.commute _))) ⟩ - ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ (refl⟩∘⟨ F.homomorphism) ⟩ - ε _ ∘ F.₁ (G.₁ (ε (F.₀ _))) ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ DMR.extendʳ (adj.counit.commute _) ⟩ - ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ F.homomorphism) ⟩ - ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g))) ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ (refl⟩∘⟨ DMR.extendʳ (adj.counit.commute _)) ⟩ - ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ sym-assoc ⟩ - (ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g)) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ∎ - ; F-resp-≈ = λ x → D.∘-resp-≈ʳ (F.F-resp-≈ (C.∘-resp-≈ʳ x)) - } - ; HF≃F' = niHelper (record - { η = λ _ → D.id - ; η⁻¹ = λ _ → D.id - ; commute = λ f → + ! : {A : SplitObj} → Split⇒ Kl-object A + ! {splitobj D F G adj GF≃M η-eq μ-eq} = record + { H = + let open D + open D.HomReasoning in + record + { F₀ = F.₀ + ; F₁ = λ f → adj.counit.η _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) + ; identity = λ {A} → begin - D.id ∘ adj.counit.η _ ∘ F.F₁ (GF≃M.⇐.η _ C.∘ M.η.η _ C.∘ f) ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C η-eq)) ⟩ - (D.id ∘ adj.counit.η _ ∘ F.F₁ (adj.unit.η _ C.∘ f)) ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ F.homomorphism) ⟩ - (D.id ∘ adj.counit.η _ ∘ F.F₁ (adj.unit.η _) ∘ F.F₁ f) ≈⟨ (refl⟩∘⟨ pullˡ D adj.zig) ⟩ - (D.id ∘ D.id ∘ F.F₁ f) ≈⟨ identityˡ ⟩ - (D.id ∘ F.F₁ f) ≈⟨ identityˡ ⟩ - F.F₁ f ≈⟨ D.Equiv.sym identityʳ ⟩ - (F.F₁ f ∘ D.id) ∎ - ; iso = λ X → record { isoˡ = identityˡ ; isoʳ = identityˡ } - }) - ; G'H≃G = {! !} + adj.counit.η (F.₀ A) ∘ F.₁ (GF≃M.⇐.η A C.∘ M.η.η A) ≈⟨ refl⟩∘⟨ F.F-resp-≈ η-eq ⟩ + adj.counit.η (F.₀ A) ∘ F.₁ (adj.unit.η A) ≈⟨ adj.zig ⟩ + D.id ∎ + ; homomorphism = λ {X} {Y} {Z} {f} {g} → + let ε x = adj.counit.η x in + begin + ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ (M.μ.η _ C.∘ M.F.₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (assoc²'' C) ⟩ + ε _ ∘ F.₁ ((GF≃M.⇐.η _ C.∘ M.μ.η _) C.∘ M.F.₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (μ-eq CHR.⟩∘⟨refl) ⟩ + ε _ ∘ F.₁ ((G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.F.₁ (GF≃M.⇐.η _)) C.∘ M.F.₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (assoc²' C) ⟩ + ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.F.₁ (GF≃M.⇐.η _) C.∘ M.F.₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ CHR.refl⟩∘⟨ pullˡ C (C.Equiv.sym M.F.homomorphism)) ⟩ + ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.F.₁ (GF≃M.⇐.η _ C.∘ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ extendʳ C (GF≃M.⇐.commute _)) ⟩ + ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ F.homomorphism ⟩ + ε _ ∘ F.₁ (G.₁ (ε (F.₀ _))) ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ DMR.extendʳ (adj.counit.commute _) ⟩ + ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.homomorphism ⟩ + ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g))) ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ DMR.extendʳ (adj.counit.commute _) ⟩ + ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ sym-assoc ⟩ + (ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g)) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ∎ + ; F-resp-≈ = λ x → D.∘-resp-≈ʳ (F.F-resp-≈ (C.∘-resp-≈ʳ x)) + } + ; HF≃F' = + let open D + open D.HomReasoning in + niHelper (record + { η = λ _ → D.id + ; η⁻¹ = λ _ → D.id + ; commute = λ f → + begin + D.id ∘ adj.counit.η _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ M.η.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C η-eq) ⟩ + D.id ∘ adj.counit.η _ ∘ F.₁ (adj.unit.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.homomorphism ⟩ + D.id ∘ adj.counit.η _ ∘ F.₁ (adj.unit.η _) ∘ F.₁ f ≈⟨ refl⟩∘⟨ pullˡ D adj.zig ⟩ + D.id ∘ D.id ∘ F.₁ f ≈⟨ identityˡ ⟩ + D.id ∘ F.₁ f ≈⟨ identityˡ ⟩ + F.₁ f ≈⟨ D.Equiv.sym identityʳ ⟩ + (F.₁ f ∘ D.id) ∎ + ; iso = λ X → record { isoˡ = identityˡ ; isoʳ = identityˡ } + }) + ; G'H≃G = + let open C + open C.HomReasoning in + niHelper (record + { η = GF≃M.⇒.η + ; η⁻¹ = GF≃M.⇐.η + ; commute = λ f → + begin + GF≃M.⇒.η _ ∘ G.₁ (adj.counit.η (F.₀ _) D.∘ F.₁ (GF≃M.⇐.η _ ∘ f)) ≈⟨ refl⟩∘⟨ G.F-resp-≈ (DHR.refl⟩∘⟨ F.homomorphism) ⟩ + GF≃M.⇒.η _ ∘ G.₁ (adj.counit.η (F.₀ _) D.∘ F.₁ (GF≃M.⇐.η _) D.∘ F.₁ f) ≈⟨ refl⟩∘⟨ G.homomorphism ⟩ + GF≃M.⇒.η _ ∘ G.₁ (adj.counit.η (F.₀ _)) ∘ G.₁ (F.₁ (GF≃M.⇐.η _) D.∘ F.₁ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ G.homomorphism ⟩ + GF≃M.⇒.η _ ∘ G.₁ (adj.counit.η (F.₀ _)) ∘ (G.₁ (F.₁ (GF≃M.⇐.η _))) ∘ G.₁ (F.₁ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ introʳ C (GF≃M.iso.isoˡ _) ⟩∘⟨refl ⟩ + GF≃M.⇒.η _ ∘ G.₁ (adj.counit.η (F.₀ _)) ∘ (G.₁ (F.₁ (GF≃M.⇐.η _)) ∘ GF≃M.⇐.η _ ∘ GF≃M.⇒.η _) ∘ G.₁ (F.₁ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ extendʳ C (C.Equiv.sym (GF≃M.⇐.commute _)) ⟩∘⟨refl ⟩ + GF≃M.⇒.η _ ∘ G.₁ (adj.counit.η (F.₀ _)) ∘ (GF≃M.⇐.η _ ∘ M.F.₁ (GF≃M.⇐.η _) ∘ GF≃M.⇒.η _) ∘ G.₁ (F.₁ f) ≈⟨ solve C ⟩ -- TODO: remove this + GF≃M.⇒.η _ ∘ (G.₁ (adj.counit.η (F.₀ _)) ∘ GF≃M.⇐.η _ ∘ M.F.₁ (GF≃M.⇐.η _)) ∘ GF≃M.⇒.η _ ∘ G.₁ (F.₁ f) ≈⟨ (refl⟩∘⟨ C.Equiv.sym μ-eq ⟩∘⟨refl) ⟩ + GF≃M.⇒.η _ ∘ (GF≃M.⇐.η _ ∘ M.μ.η _) ∘ GF≃M.⇒.η _ ∘ G.₁ (F.₁ f) ≈⟨ solve C ⟩ -- TODO: remove this + GF≃M.⇒.η _ ∘ GF≃M.⇐.η _ ∘ M.μ.η _ ∘ GF≃M.⇒.η _ ∘ G.₁ (F.₁ f) ≈⟨ pullˡ C (GF≃M.iso.isoʳ _) ⟩ + C.id ∘ M.μ.η _ ∘ GF≃M.⇒.η _ ∘ G.₁ (F.₁ f) ≈⟨ identityˡ ⟩ + M.μ.η _ ∘ GF≃M.⇒.η _ ∘ G.₁ (F.₁ f) ≈⟨ pushʳ C (GF≃M.⇒.commute _) ⟩ + (M.μ.η _ ∘ M.F.₁ f) ∘ GF≃M.⇒.η _ ∎ + ; iso = GF≃M.iso + }) } where - module adj = Adjoint adj - module F = Functor F - module G = Functor G - module GF≃M = NaturalIsomorphism GF≃M - module D = Category D - open D - open D.HomReasoning - module CHR = C.HomReasoning - module DMR = Categories.Morphism.Reasoning D - module CMR = Categories.Morphism.Reasoning C + module adj = Adjoint adj + module F = Functor F + module G = Functor G + module GF≃M = NaturalIsomorphism GF≃M + module D = Category D + module CHR = C.HomReasoning + module DHR = D.HomReasoning + module DMR = Categories.Morphism.Reasoning D + module CMR = Categories.Morphism.Reasoning C From a23a13276d56d83b1276ee05ea5be78b963cd7f4 Mon Sep 17 00:00:00 2001 From: iwilare Date: Mon, 2 Jan 2023 15:32:19 +0100 Subject: [PATCH 20/46] WIP --- .../Construction/Adjunctions/Properties.agda | 28 +++++++++++-------- 1 file changed, 17 insertions(+), 11 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda index 68d3ae569..0f807e92a 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda @@ -82,22 +82,28 @@ EM-terminal = record ; !-unique = {! !} } +{- + HF≃F'.⇐.η Y ∘ A.adj.counit.η (A.F.F₀ Y) ∘ A.F.F₁ (A.GF≃M.⇐.η Y C.∘ f) + +≈ Functor.F₁ H f ∘ Γ.⇐.η X +-} + Kl-initial : IsInitial Kl-object Kl-initial = record { ! = ! - ; !-unique = λ { {A} f → niHelper ( - let open SplitObj A - open Split⇒ f - module G'H≃G = NaturalIsomorphism G'H≃G - module HF≃F' = NaturalIsomorphism HF≃F' in + ; !-unique = λ { {A} H → niHelper ( + let module A = SplitObj A + module K = SplitObj Kl-object + module H = Split⇒ H + module G'H≃G = NaturalIsomorphism H.G'H≃G + module HF≃F' = NaturalIsomorphism H.HF≃F' + open Category A.D + open HomReasoning in record - { η = HF≃F'.⇐.η + { η = HF≃F'.⇐.η ; η⁻¹ = HF≃F'.⇒.η - ; commute = - let open D - open D.HomReasoning in λ f → - {! !} - ; iso = NaturalIsomorphism.iso (sym HF≃F') + ; commute = λ f → let open A.D.HomReasoning in {! HomReasoning._∎ !} + ; iso = NaturalIsomorphism.iso (sym H.HF≃F') }) } } From f7319266ff12d2ed35947906965238f5cba37322 Mon Sep 17 00:00:00 2001 From: fouche Date: Tue, 3 Jan 2023 13:14:45 +0100 Subject: [PATCH 21/46] gnagna --- agda-categories.agda-lib | 2 +- .../Adjoint/Construction/Adjunctions.agda | 31 +- .../NaturalTransformation/Core.agda | 2 + src/Categories/Tactic/Category.agda | 325 +++++++++--------- 4 files changed, 188 insertions(+), 172 deletions(-) diff --git a/agda-categories.agda-lib b/agda-categories.agda-lib index 91c6c57d7..3f89afb89 100644 --- a/agda-categories.agda-lib +++ b/agda-categories.agda-lib @@ -1,3 +1,3 @@ name: agda-categories -depend: standard-library-1.7 +depend: standard-library-2.0 include: src/ diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 469ae17f4..6468c06a6 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -51,7 +51,6 @@ record SplitObj : Set (suc o ⊔ suc ℓ ⊔ suc e) where GF≃M.⇐.η A μ-eq: - GF≃M.⇐.η G.F₁ (F.F₁ (GF≃M.⇐.η)) ---------------------> GFMA --------------------> MMA ---------------------> MGFA --------------------> GFGFA @@ -69,10 +68,26 @@ record Split⇒ (X Y : SplitObj) : Set (suc o ⊔ suc ℓ ⊔ suc e) where private module X = SplitObj X module Y = SplitObj Y + open X + open Y field H : Functor X.D Y.D - HF≃F' : H ∘F X.F ≃ Y.F - G'H≃G : Y.G ∘F H ≃ X.G + Γ : H ∘F X.F ≃ Y.F + + private + module Γ = NaturalIsomorphism Γ + + Γ' : NaturalTransformation X.G (Y.G ∘F H) + Γ' = F∘id⇒F ∘ᵥ + (((Y.G ∘F H) ∘ˡ X.adj.counit) ∘ᵥ zip ∘ᵥ Y.G ∘ˡ zop) + ∘ᵥ (Y.G ∘ˡ Γ.F⇐G ∘ʳ X.G) + ∘ᵥ (zap ∘ᵥ Y.adj.unit ∘ʳ X.G) + ∘ᵥ F⇒id∘F + where + open NaturalIsomorphism + zip = F⇐G (associator (X.F ∘F X.G) H Y.G) + zop = F⇒G (associator X.G X.F H) + zap = F⇒G (associator X.G Y.F Y.G) Split : Monad C → Category _ _ _ Split M = record @@ -94,12 +109,12 @@ Split M = record split-id : {A : SplitObj} → Split⇒ A A split-id = record { H = Categories.Functor.id - ; HF≃F' = unitorˡ - ; G'H≃G = unitorʳ + ; Γ = unitorˡ + -- ; G'H≃G = unitorʳ } comp : {A B X : SplitObj} → Split⇒ B X → Split⇒ A B → Split⇒ A X - comp {A = A} {B = B} {X = X} (split⇒ Hᵤ HF≃F'ᵤ G'H≃Gᵤ) (split⇒ Hᵥ HF≃F'ᵥ G'H≃Gᵥ) = record + comp {A = A} {B = B} {X = X} (split⇒ Hᵤ HF≃F'ᵤ) (split⇒ Hᵥ HF≃F'ᵥ) = record { H = Hᵤ ∘F Hᵥ - ; HF≃F' = HF≃F'ᵤ ⓘᵥ (Hᵤ ⓘˡ HF≃F'ᵥ) ⓘᵥ associator (SplitObj.F A) Hᵥ Hᵤ - ; G'H≃G = G'H≃Gᵥ ⓘᵥ (G'H≃Gᵤ ⓘʳ Hᵥ) ⓘᵥ sym-associator Hᵥ Hᵤ (SplitObj.G X) + ; Γ = HF≃F'ᵤ ⓘᵥ (Hᵤ ⓘˡ HF≃F'ᵥ) ⓘᵥ associator (SplitObj.F A) Hᵥ Hᵤ + -- ; G'H≃G = G'H≃Gᵥ ⓘᵥ (G'H≃Gᵤ ⓘʳ Hᵥ) ⓘᵥ sym-associator Hᵥ Hᵤ (SplitObj.G X) } diff --git a/src/Categories/NaturalTransformation/Core.agda b/src/Categories/NaturalTransformation/Core.agda index db06af31e..77a842aa6 100644 --- a/src/Categories/NaturalTransformation/Core.agda +++ b/src/Categories/NaturalTransformation/Core.agda @@ -102,6 +102,7 @@ _∘ₕ_ {E = E} {F} {I = I} Y X = ntHelper record open Functor I renaming (F₀ to I₀; F₁ to I₁) open MR E +-- F * α _∘ˡ_ : ∀ {G H : Functor C D} (F : Functor D E) → NaturalTransformation G H → NaturalTransformation (F ∘F G) (F ∘F H) _∘ˡ_ F α = record { η = λ X → F₁ (η X) @@ -111,6 +112,7 @@ _∘ˡ_ F α = record where open Functor F open NaturalTransformation α +-- α * F _∘ʳ_ : ∀ {G H : Functor D E} → NaturalTransformation G H → (F : Functor C D) → NaturalTransformation (G ∘F F) (H ∘F F) _∘ʳ_ {D = D} {E = E} {G = G} {H = H} α F = record { η = λ X → η (F₀ X) diff --git a/src/Categories/Tactic/Category.agda b/src/Categories/Tactic/Category.agda index 1ab24c52b..5c0a2e052 100644 --- a/src/Categories/Tactic/Category.agda +++ b/src/Categories/Tactic/Category.agda @@ -19,166 +19,165 @@ open import Data.List as List using (List; _∷_; []) open import Data.Product as Product using (_×_; _,_) open import Agda.Builtin.Reflection -open import Reflection.Argument -open import Reflection.Term using (getName; _⋯⟅∷⟆_) -open import Reflection.TypeChecking.Monad.Syntax - -module _ {o ℓ e} (𝒞 : Category o ℓ e) where - - open Category 𝒞 - open HomReasoning - open Equiv - - private - variable - A B C : Obj - f g : A ⇒ B - - -------------------------------------------------------------------------------- - -- An 'Expr' reifies the parentheses/identity morphisms of some series of - -- compositions of morphisms into a data structure. In fact, this is also - -- a category! - -------------------------------------------------------------------------------- - data Expr : Obj → Obj → Set (o ⊔ ℓ) where - _∘′_ : ∀ {A B C} → Expr B C → Expr A B → Expr A C - id′ : ∀ {A} → Expr A A - [_↑] : ∀ {A B} → A ⇒ B → Expr A B - - -- Embed a morphism in 'Expr' back into '𝒞' without normalizing. - [_↓] : Expr A B → A ⇒ B - [ f ∘′ g ↓] = [ f ↓] ∘ [ g ↓] - [ id′ ↓] = id - [ [ f ↑] ↓] = f - - -- Convert an 'Expr' back into a morphism, while normalizing - -- - -- This actually embeds the morphism into the category of copresheaves - -- on 𝒞, which obeys the category laws up to beta-eta equality. - -- This lets us normalize away all the associations/identity morphisms. - embed : Expr B C → A ⇒ B → A ⇒ C - embed (f ∘′ g) h = embed f (embed g h) - embed id′ h = h - embed [ f ↑] h = f ∘ h - - - preserves-≈′ : ∀ (f : Expr B C) → (h : A ⇒ B) → embed f id ∘ h ≈ embed f h - preserves-≈′ id′ f = identityˡ - preserves-≈′ [ x ↑] f = ∘-resp-≈ˡ identityʳ - preserves-≈′ (f ∘′ g) h = begin - embed (f ∘′ g) id ∘ h ≡⟨⟩ - embed f (embed g id) ∘ h ≈˘⟨ preserves-≈′ f (embed g id) ⟩∘⟨refl ⟩ - (embed f id ∘ embed g id) ∘ h ≈⟨ assoc ⟩ - embed f id ∘ embed g id ∘ h ≈⟨ refl⟩∘⟨ preserves-≈′ g h ⟩ - embed f id ∘ embed g h ≈⟨ preserves-≈′ f (embed g h) ⟩ - embed (f ∘′ g) h ∎ - - preserves-≈ : ∀ (f : Expr A B) → embed f id ≈ [ f ↓] - preserves-≈ id′ = refl - preserves-≈ [ x ↑] = identityʳ - preserves-≈ (f ∘′ g) = begin - embed (f ∘′ g) id ≈˘⟨ preserves-≈′ f (embed g id) ⟩ - embed f id ∘ embed g id ≈⟨ preserves-≈ f ⟩∘⟨ preserves-≈ g ⟩ - [ f ↓] ∘ [ g ↓] ≡⟨⟩ - [ f ∘′ g ↓] ∎ - --------------------------------------------------------------------------------- --- Reflection Helpers --------------------------------------------------------------------------------- - -_==_ = primQNameEquality -{-# INLINE _==_ #-} - -getArgs : Term → Maybe (Term × Term) -getArgs (def _ xs) = go xs - where - go : List (Arg Term) → Maybe (Term × Term) - go (vArg x ∷ vArg y ∷ []) = just (x , y) - go (x ∷ xs) = go xs - go _ = nothing -getArgs _ = nothing - --------------------------------------------------------------------------------- --- Getting Category Names --------------------------------------------------------------------------------- - -record CategoryNames : Set where - field - is-∘ : Name → Bool - is-id : Name → Bool - -buildMatcher : Name → Maybe Name → Name → Bool -buildMatcher n nothing x = n == x -buildMatcher n (just m) x = n == x ∨ m == x - -findCategoryNames : Term → TC CategoryNames -findCategoryNames cat = do - ∘-altName ← normalise (def (quote Category._∘_) (3 ⋯⟅∷⟆ cat ⟨∷⟩ [])) - id-altName ← normalise (def (quote Category.id) (3 ⋯⟅∷⟆ cat ⟨∷⟩ [])) - returnTC record - { is-∘ = buildMatcher (quote Category._∘_) (getName ∘-altName) - ; is-id = buildMatcher (quote Category.id) (getName id-altName) - } - --------------------------------------------------------------------------------- --- Constructing an Expr --------------------------------------------------------------------------------- - -″id″ : Term -″id″ = quote id′ ⟨ con ⟩ [] - -″[_↑]″ : Term → Term -″[ t ↑]″ = quote [_↑] ⟨ con ⟩ (t ⟨∷⟩ []) - -module _ (names : CategoryNames) where - - open CategoryNames names - - mutual - ″∘″ : List (Arg Term) → Term - ″∘″ (x ⟨∷⟩ y ⟨∷⟩ xs) = quote _∘′_ ⟨ con ⟩ buildExpr x ⟨∷⟩ buildExpr y ⟨∷⟩ [] - ″∘″ (x ∷ xs) = ″∘″ xs - ″∘″ _ = unknown - - buildExpr : Term → Term - buildExpr t@(def n xs) = - if (is-∘ n) - then ″∘″ xs - else if (is-id n) - then ″id″ - else - ″[ t ↑]″ - buildExpr t@(con n xs) = - if (is-∘ n) - then ″∘″ xs - else if (is-id n) - then ″id″ - else - ″[ t ↑]″ - buildExpr t = ″[ t ↑]″ - --------------------------------------------------------------------------------- --- Constructing the Solution --------------------------------------------------------------------------------- - -constructSoln : Term → CategoryNames → Term → Term → Term -constructSoln cat names lhs rhs = - quote Category.Equiv.trans ⟨ def ⟩ 3 ⋯⟅∷⟆ cat ⟨∷⟩ - (quote Category.Equiv.sym ⟨ def ⟩ 3 ⋯⟅∷⟆ cat ⟨∷⟩ - (quote preserves-≈ ⟨ def ⟩ 3 ⋯⟅∷⟆ cat ⟨∷⟩ buildExpr names lhs ⟨∷⟩ []) ⟨∷⟩ []) - ⟨∷⟩ - (quote preserves-≈ ⟨ def ⟩ 3 ⋯⟅∷⟆ cat ⟨∷⟩ buildExpr names rhs ⟨∷⟩ []) - ⟨∷⟩ [] - -solve-macro : Term → Term → TC _ -solve-macro mon hole = do - hole′ ← inferType hole >>= normalise - names ← findCategoryNames mon - just (lhs , rhs) ← returnTC (getArgs hole′) - where nothing → typeError (termErr hole′ ∷ []) - let soln = constructSoln mon names lhs rhs - unify hole soln - -macro - solve : Term → Term → TC _ - solve = solve-macro - +-- open import Reflection.Argument +-- open import Reflection.Term using (getName; _⋯⟅∷⟆_) +-- open import Reflection.TypeChecking.Monad.Syntax + +-- module _ {o ℓ e} (𝒞 : Category o ℓ e) where + +-- open Category 𝒞 +-- open HomReasoning +-- open Equiv + +-- private +-- variable +-- A B C : Obj +-- f g : A ⇒ B + +-- -------------------------------------------------------------------------------- +-- -- An 'Expr' reifies the parentheses/identity morphisms of some series of +-- -- compositions of morphisms into a data structure. In fact, this is also +-- -- a category! +-- -------------------------------------------------------------------------------- +-- data Expr : Obj → Obj → Set (o ⊔ ℓ) where +-- _∘′_ : ∀ {A B C} → Expr B C → Expr A B → Expr A C +-- id′ : ∀ {A} → Expr A A +-- [_↑] : ∀ {A B} → A ⇒ B → Expr A B + +-- -- Embed a morphism in 'Expr' back into '𝒞' without normalizing. +-- [_↓] : Expr A B → A ⇒ B +-- [ f ∘′ g ↓] = [ f ↓] ∘ [ g ↓] +-- [ id′ ↓] = id +-- [ [ f ↑] ↓] = f + +-- -- Convert an 'Expr' back into a morphism, while normalizing +-- -- +-- -- This actually embeds the morphism into the category of copresheaves +-- -- on 𝒞, which obeys the category laws up to beta-eta equality. +-- -- This lets us normalize away all the associations/identity morphisms. +-- embed : Expr B C → A ⇒ B → A ⇒ C +-- embed (f ∘′ g) h = embed f (embed g h) +-- embed id′ h = h +-- embed [ f ↑] h = f ∘ h + + +-- preserves-≈′ : ∀ (f : Expr B C) → (h : A ⇒ B) → embed f id ∘ h ≈ embed f h +-- preserves-≈′ id′ f = identityˡ +-- preserves-≈′ [ x ↑] f = ∘-resp-≈ˡ identityʳ +-- preserves-≈′ (f ∘′ g) h = begin +-- embed (f ∘′ g) id ∘ h ≡⟨⟩ +-- embed f (embed g id) ∘ h ≈˘⟨ preserves-≈′ f (embed g id) ⟩∘⟨refl ⟩ +-- (embed f id ∘ embed g id) ∘ h ≈⟨ assoc ⟩ +-- embed f id ∘ embed g id ∘ h ≈⟨ refl⟩∘⟨ preserves-≈′ g h ⟩ +-- embed f id ∘ embed g h ≈⟨ preserves-≈′ f (embed g h) ⟩ +-- embed (f ∘′ g) h ∎ + +-- preserves-≈ : ∀ (f : Expr A B) → embed f id ≈ [ f ↓] +-- preserves-≈ id′ = refl +-- preserves-≈ [ x ↑] = identityʳ +-- preserves-≈ (f ∘′ g) = begin +-- embed (f ∘′ g) id ≈˘⟨ preserves-≈′ f (embed g id) ⟩ +-- embed f id ∘ embed g id ≈⟨ preserves-≈ f ⟩∘⟨ preserves-≈ g ⟩ +-- [ f ↓] ∘ [ g ↓] ≡⟨⟩ +-- [ f ∘′ g ↓] ∎ + +-- -------------------------------------------------------------------------------- +-- -- Reflection Helpers +-- -------------------------------------------------------------------------------- + +-- _==_ = primQNameEquality +-- {-# INLINE _==_ #-} + +-- getArgs : Term → Maybe (Term × Term) +-- getArgs (def _ xs) = go xs +-- where +-- go : List (Arg Term) → Maybe (Term × Term) +-- go (vArg x ∷ vArg y ∷ []) = just (x , y) +-- go (x ∷ xs) = go xs +-- go _ = nothing +-- getArgs _ = nothing + +-- -------------------------------------------------------------------------------- +-- -- Getting Category Names +-- -------------------------------------------------------------------------------- + +-- record CategoryNames : Set where +-- field +-- is-∘ : Name → Bool +-- is-id : Name → Bool + +-- buildMatcher : Name → Maybe Name → Name → Bool +-- buildMatcher n nothing x = n == x +-- buildMatcher n (just m) x = n == x ∨ m == x + +-- findCategoryNames : Term → TC CategoryNames +-- findCategoryNames cat = do +-- ∘-altName ← normalise (def (quote Category._∘_) (3 ⋯⟅∷⟆ cat ⟨∷⟩ [])) +-- id-altName ← normalise (def (quote Category.id) (3 ⋯⟅∷⟆ cat ⟨∷⟩ [])) +-- returnTC record +-- { is-∘ = buildMatcher (quote Category._∘_) (getName ∘-altName) +-- ; is-id = buildMatcher (quote Category.id) (getName id-altName) +-- } + +-- -------------------------------------------------------------------------------- +-- -- Constructing an Expr +-- -------------------------------------------------------------------------------- + +-- ″id″ : Term +-- ″id″ = quote id′ ⟨ con ⟩ [] + +-- ″[_↑]″ : Term → Term +-- ″[ t ↑]″ = quote [_↑] ⟨ con ⟩ (t ⟨∷⟩ []) + +-- module _ (names : CategoryNames) where + +-- open CategoryNames names + +-- mutual +-- ″∘″ : List (Arg Term) → Term +-- ″∘″ (x ⟨∷⟩ y ⟨∷⟩ xs) = quote _∘′_ ⟨ con ⟩ buildExpr x ⟨∷⟩ buildExpr y ⟨∷⟩ [] +-- ″∘″ (x ∷ xs) = ″∘″ xs +-- ″∘″ _ = unknown + +-- buildExpr : Term → Term +-- buildExpr t@(def n xs) = +-- if (is-∘ n) +-- then ″∘″ xs +-- else if (is-id n) +-- then ″id″ +-- else +-- ″[ t ↑]″ +-- buildExpr t@(con n xs) = +-- if (is-∘ n) +-- then ″∘″ xs +-- else if (is-id n) +-- then ″id″ +-- else +-- ″[ t ↑]″ +-- buildExpr t = ″[ t ↑]″ + +-- -------------------------------------------------------------------------------- +-- -- Constructing the Solution +-- -------------------------------------------------------------------------------- + +-- constructSoln : Term → CategoryNames → Term → Term → Term +-- constructSoln cat names lhs rhs = +-- quote Category.Equiv.trans ⟨ def ⟩ 3 ⋯⟅∷⟆ cat ⟨∷⟩ +-- (quote Category.Equiv.sym ⟨ def ⟩ 3 ⋯⟅∷⟆ cat ⟨∷⟩ +-- (quote preserves-≈ ⟨ def ⟩ 3 ⋯⟅∷⟆ cat ⟨∷⟩ buildExpr names lhs ⟨∷⟩ []) ⟨∷⟩ []) +-- ⟨∷⟩ +-- (quote preserves-≈ ⟨ def ⟩ 3 ⋯⟅∷⟆ cat ⟨∷⟩ buildExpr names rhs ⟨∷⟩ []) +-- ⟨∷⟩ [] + +-- solve-macro : Term → Term → TC _ +-- solve-macro mon hole = do +-- hole′ ← inferType hole >>= normalise +-- names ← findCategoryNames mon +-- just (lhs , rhs) ← returnTC (getArgs hole′) +-- where nothing → typeError (termErr hole′ ∷ []) +-- let soln = constructSoln mon names lhs rhs +-- unify hole soln + +-- macro +-- solve : Term → Term → TC _ +-- solve = solve-macro From 30aad51ccc4fdb3398fc562a377f8130518b8f0f Mon Sep 17 00:00:00 2001 From: iwilare Date: Tue, 3 Jan 2023 16:27:41 +0100 Subject: [PATCH 22/46] =?UTF-8?q?WIP=20with=20=CE=BC-comp?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../Adjoint/Construction/Adjunctions.agda | 50 ++++++---- .../Construction/Adjunctions/Properties.agda | 95 ++++++++++--------- 2 files changed, 85 insertions(+), 60 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 6468c06a6..ce938fdb4 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --without-K --safe #-} +{-# OPTIONS --without-K --allow-unsolved-metas #-} open import Level open import Categories.Category.Core using (Category) @@ -14,10 +14,11 @@ private open import Categories.Adjoint open import Categories.Functor -open import Categories.Morphism +open import Categories.Morphism hiding (_≅_) open import Categories.Functor.Properties open import Categories.NaturalTransformation.Core open import Categories.NaturalTransformation.NaturalIsomorphism +open import Categories.NaturalTransformation.Equivalence renaming (_≃_ to _≅_) open import Categories.Morphism.Reasoning as MR open import Categories.Tactic.Category @@ -62,19 +63,10 @@ record SplitObj : Set (suc o ⊔ suc ℓ ⊔ suc e) where GF≃M.⇐.η A -} - -record Split⇒ (X Y : SplitObj) : Set (suc o ⊔ suc ℓ ⊔ suc e) where - constructor split⇒ +module _ {X Y : SplitObj} (H : Functor (SplitObj.D X) (SplitObj.D Y)) (Γ : H ∘F (SplitObj.F X) ≃ (SplitObj.F Y)) where private module X = SplitObj X module Y = SplitObj Y - open X - open Y - field - H : Functor X.D Y.D - Γ : H ∘F X.F ≃ Y.F - - private module Γ = NaturalIsomorphism Γ Γ' : NaturalTransformation X.G (Y.G ∘F H) @@ -89,6 +81,26 @@ record Split⇒ (X Y : SplitObj) : Set (suc o ⊔ suc ℓ ⊔ suc e) where zop = F⇒G (associator X.G X.F H) zap = F⇒G (associator X.G Y.F Y.G) +record Split⇒ (X Y : SplitObj) : Set (suc o ⊔ suc ℓ ⊔ suc e) where + constructor split⇒ + private + module X = SplitObj X + module Y = SplitObj Y + + field + H : Functor X.D Y.D + Γ : H ∘F X.F ≃ Y.F + + private + module Γ = NaturalIsomorphism Γ + + field + μ-comp : Y.GF≃M.F⇐G -- Y.GF≃M.F⇒G + ≅ (Y.G ∘ˡ Γ.F⇒G) + ∘ᵥ NaturalIsomorphism.F⇒G (associator X.F H Y.G) + ∘ᵥ (Γ' {X = X} {Y = Y} H Γ ∘ʳ X.F) + ∘ᵥ X.GF≃M.F⇐G + Split : Monad C → Category _ _ _ Split M = record { Obj = SplitObj @@ -107,14 +119,18 @@ Split M = record where open NaturalTransformation split-id : {A : SplitObj} → Split⇒ A A - split-id = record + split-id {A} = record { H = Categories.Functor.id ; Γ = unitorˡ - -- ; G'H≃G = unitorʳ - } + ; μ-comp = {! !} + } where module A = SplitObj A comp : {A B X : SplitObj} → Split⇒ B X → Split⇒ A B → Split⇒ A X - comp {A = A} {B = B} {X = X} (split⇒ Hᵤ HF≃F'ᵤ) (split⇒ Hᵥ HF≃F'ᵥ) = record + comp {A = A} {B = B} {X = X} (split⇒ Hᵤ HF≃F'ᵤ μ-comp) (split⇒ Hᵥ HF≃F'ᵥ μ-comp′) = record { H = Hᵤ ∘F Hᵥ ; Γ = HF≃F'ᵤ ⓘᵥ (Hᵤ ⓘˡ HF≃F'ᵥ) ⓘᵥ associator (SplitObj.F A) Hᵥ Hᵤ + ; μ-comp = {! !} -- ; G'H≃G = G'H≃Gᵥ ⓘᵥ (G'H≃Gᵤ ⓘʳ Hᵥ) ⓘᵥ sym-associator Hᵥ Hᵤ (SplitObj.G X) - } + } where + module A = SplitObj A + module B = SplitObj B + module X = SplitObj X diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda index 0f807e92a..e2bc08048 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --without-K --safe #-} +{-# OPTIONS --without-K --allow-unsolved-metas #-} open import Level open import Categories.Category.Core using (Category) @@ -33,11 +33,11 @@ EM-object = record ; G = EM.Forgetful ; adj = EM.Free⊣Forgetful ; GF≃M = EM.FF≃F - ; η-eq = begin + ; η-eq = {! !} {- begin M.F.₁ id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ id ∘ M.η.η _ ≈⟨ identityˡ ⟩ - M.η.η _ ∎ - ; μ-eq = begin + M.η.η _ ∎ -} + ; μ-eq = {! !} {- begin M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ @@ -46,7 +46,7 @@ EM-object = record (M.μ.η _ ∘ M.F.₁ id) ∘ id ≈⟨ assoc ⟩ M.μ.η _ ∘ M.F.₁ id ∘ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ M.μ.η _ ∘ M.F.₁ id ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ - M.μ.η _ ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ + M.μ.η _ ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ -} } where open Category C open HomReasoning @@ -57,11 +57,11 @@ Kl-object = record ; G = KL.Forgetful ; adj = KL.Free⊣Forgetful ; GF≃M = KL.FF≃F - ; η-eq = begin + ; η-eq = {! !} {- begin M.F.₁ id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ id ∘ M.η.η _ ≈⟨ identityˡ ⟩ - M.η.η _ ∎ - ; μ-eq = begin M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ + M.η.η _ ∎ -} + ; μ-eq = {! !} {- begin M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ M.μ.η _ ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ @@ -72,7 +72,7 @@ Kl-object = record ((M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id) ∘ id ≈⟨ assoc ⟩ (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ - (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ + (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ -} } where open Category C open HomReasoning @@ -95,15 +95,45 @@ Kl-initial = record let module A = SplitObj A module K = SplitObj Kl-object module H = Split⇒ H - module G'H≃G = NaturalIsomorphism H.G'H≃G - module HF≃F' = NaturalIsomorphism H.HF≃F' - open Category A.D - open HomReasoning in + module ! = Split⇒ ! + module Γ = NaturalIsomorphism H.Γ + module CH = C.HomReasoning + open Category A.D in record - { η = HF≃F'.⇐.η - ; η⁻¹ = HF≃F'.⇒.η - ; commute = λ f → let open A.D.HomReasoning in {! HomReasoning._∎ !} - ; iso = NaturalIsomorphism.iso (sym H.HF≃F') + { η = Γ.⇐.η + ; η⁻¹ = Γ.⇒.η + ; commute = λ f → + let useful : ∀ {x} + → A.GF≃M.⇐.η x C.≈ A.G.F₁ (Γ.⇒.η x) + C.∘ A.G.F₁ (Functor.F₁ H.H C.id) + C.∘ A.G.F₁ (Γ.⇐.η (M.F.F₀ x)) + C.∘ A.adj.unit.η (M.F.F₀ x) + useful = {! !} {-let open C.HomReasoning in + begin _ ≈⟨ H.μ-comp ⟩ + _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ elimʳ C (M.F.identity)) ⟩ + _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ C.identityʳ) ⟩ + _ ≈⟨ (refl⟩∘⟨ C.identityˡ) ⟩ + _ ≈⟨ refl⟩∘⟨ C.identityˡ ⟩ + _ ≈⟨ (refl⟩∘⟨ (A.G.F-resp-≈ (Functor.F-resp-≈ H.H M.F.identity) ⟩∘⟨refl) ⟩∘⟨refl) ⟩ + _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ C.identityˡ) ⟩ + _ ≈⟨ (refl⟩∘⟨ (elimʳ C (elimʳ C A.G.identity)) ⟩∘⟨refl) ⟩ + _ ∎-} + in let open A.D.HomReasoning + + in begin (Γ.⇐.η _ ∘ A.adj.counit.η (A.F.F₀ _) ∘ A.F.F₁ (A.GF≃M.⇐.η _ C.∘ f)) + ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.F.F-resp-≈ ((useful CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc ) CH.○ CH.refl⟩∘⟨ CH.refl⟩∘⟨ C.assoc ))) ⟩ + {! !} ≈⟨ ((refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _)) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ assoc) ⟩ + {! !} ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ (A.F.homomorphism ○ refl⟩∘⟨ A.F.homomorphism ○ refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ + {! !} ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ + {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ {! !} ⟩ + Functor.F₁ H.H f ∘ Γ.⇐.η _ ∎ + ; iso = NaturalIsomorphism.iso (sym H.Γ) }) } } @@ -121,7 +151,7 @@ Kl-initial = record adj.counit.η (F.₀ A) ∘ F.₁ (GF≃M.⇐.η A C.∘ M.η.η A) ≈⟨ refl⟩∘⟨ F.F-resp-≈ η-eq ⟩ adj.counit.η (F.₀ A) ∘ F.₁ (adj.unit.η A) ≈⟨ adj.zig ⟩ D.id ∎ - ; homomorphism = λ {X} {Y} {Z} {f} {g} → + ; homomorphism = {! !} {-λ {X} {Y} {Z} {f} {g} → let ε x = adj.counit.η x in begin ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ (M.μ.η _ C.∘ M.F.₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (assoc²'' C) ⟩ @@ -134,10 +164,10 @@ Kl-initial = record ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.homomorphism ⟩ ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g))) ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ DMR.extendʳ (adj.counit.commute _) ⟩ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ sym-assoc ⟩ - (ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g)) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ∎ + (ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g)) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ∎ -} ; F-resp-≈ = λ x → D.∘-resp-≈ʳ (F.F-resp-≈ (C.∘-resp-≈ʳ x)) } - ; HF≃F' = + ; Γ = {! !} {- let open D open D.HomReasoning in niHelper (record @@ -153,29 +183,8 @@ Kl-initial = record F.₁ f ≈⟨ D.Equiv.sym identityʳ ⟩ (F.₁ f ∘ D.id) ∎ ; iso = λ X → record { isoˡ = identityˡ ; isoʳ = identityˡ } - }) - ; G'H≃G = - let open C - open C.HomReasoning in - niHelper (record - { η = GF≃M.⇒.η - ; η⁻¹ = GF≃M.⇐.η - ; commute = λ f → - begin - GF≃M.⇒.η _ ∘ G.₁ (adj.counit.η (F.₀ _) D.∘ F.₁ (GF≃M.⇐.η _ ∘ f)) ≈⟨ refl⟩∘⟨ G.F-resp-≈ (DHR.refl⟩∘⟨ F.homomorphism) ⟩ - GF≃M.⇒.η _ ∘ G.₁ (adj.counit.η (F.₀ _) D.∘ F.₁ (GF≃M.⇐.η _) D.∘ F.₁ f) ≈⟨ refl⟩∘⟨ G.homomorphism ⟩ - GF≃M.⇒.η _ ∘ G.₁ (adj.counit.η (F.₀ _)) ∘ G.₁ (F.₁ (GF≃M.⇐.η _) D.∘ F.₁ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ G.homomorphism ⟩ - GF≃M.⇒.η _ ∘ G.₁ (adj.counit.η (F.₀ _)) ∘ (G.₁ (F.₁ (GF≃M.⇐.η _))) ∘ G.₁ (F.₁ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ introʳ C (GF≃M.iso.isoˡ _) ⟩∘⟨refl ⟩ - GF≃M.⇒.η _ ∘ G.₁ (adj.counit.η (F.₀ _)) ∘ (G.₁ (F.₁ (GF≃M.⇐.η _)) ∘ GF≃M.⇐.η _ ∘ GF≃M.⇒.η _) ∘ G.₁ (F.₁ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ extendʳ C (C.Equiv.sym (GF≃M.⇐.commute _)) ⟩∘⟨refl ⟩ - GF≃M.⇒.η _ ∘ G.₁ (adj.counit.η (F.₀ _)) ∘ (GF≃M.⇐.η _ ∘ M.F.₁ (GF≃M.⇐.η _) ∘ GF≃M.⇒.η _) ∘ G.₁ (F.₁ f) ≈⟨ solve C ⟩ -- TODO: remove this - GF≃M.⇒.η _ ∘ (G.₁ (adj.counit.η (F.₀ _)) ∘ GF≃M.⇐.η _ ∘ M.F.₁ (GF≃M.⇐.η _)) ∘ GF≃M.⇒.η _ ∘ G.₁ (F.₁ f) ≈⟨ (refl⟩∘⟨ C.Equiv.sym μ-eq ⟩∘⟨refl) ⟩ - GF≃M.⇒.η _ ∘ (GF≃M.⇐.η _ ∘ M.μ.η _) ∘ GF≃M.⇒.η _ ∘ G.₁ (F.₁ f) ≈⟨ solve C ⟩ -- TODO: remove this - GF≃M.⇒.η _ ∘ GF≃M.⇐.η _ ∘ M.μ.η _ ∘ GF≃M.⇒.η _ ∘ G.₁ (F.₁ f) ≈⟨ pullˡ C (GF≃M.iso.isoʳ _) ⟩ - C.id ∘ M.μ.η _ ∘ GF≃M.⇒.η _ ∘ G.₁ (F.₁ f) ≈⟨ identityˡ ⟩ - M.μ.η _ ∘ GF≃M.⇒.η _ ∘ G.₁ (F.₁ f) ≈⟨ pushʳ C (GF≃M.⇒.commute _) ⟩ - (M.μ.η _ ∘ M.F.₁ f) ∘ GF≃M.⇒.η _ ∎ - ; iso = GF≃M.iso - }) + }) -} + ; μ-comp = {! !} } where module adj = Adjoint adj module F = Functor F From 4b25a8bb9758a2f4957132c57ddb95cc70524818 Mon Sep 17 00:00:00 2001 From: iwilare Date: Mon, 9 Jan 2023 15:01:32 +0200 Subject: [PATCH 23/46] Minima --- .../Construction/Adjunctions/Properties.agda | 18 +++++++++++++----- 1 file changed, 13 insertions(+), 5 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda index e2bc08048..ae9a96e54 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda @@ -13,7 +13,9 @@ module M = Monad M open import Categories.Adjoint using (Adjoint) open import Categories.Functor using (Functor) open import Categories.Morphism.Reasoning -open import Categories.NaturalTransformation.NaturalIsomorphism using (NaturalIsomorphism; niHelper; sym) +open import Categories.NaturalTransformation +open NaturalTransformation using (η) +open import Categories.NaturalTransformation.NaturalIsomorphism using (NaturalIsomorphism; niHelper; sym; associator) open import Categories.Adjoint.Construction.Adjunctions M @@ -108,7 +110,7 @@ Kl-initial = record C.∘ A.G.F₁ (Functor.F₁ H.H C.id) C.∘ A.G.F₁ (Γ.⇐.η (M.F.F₀ x)) C.∘ A.adj.unit.η (M.F.F₀ x) - useful = {! !} {-let open C.HomReasoning in + useful = {! let open H in NaturalIsomorphism.F⇒G (associator ? ? ?) !} {-let open C.HomReasoning in begin _ ≈⟨ H.μ-comp ⟩ _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ elimʳ C (M.F.identity)) ⟩ _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ C.identityʳ) ⟩ @@ -121,15 +123,21 @@ Kl-initial = record in let open A.D.HomReasoning in begin (Γ.⇐.η _ ∘ A.adj.counit.η (A.F.F₀ _) ∘ A.F.F₁ (A.GF≃M.⇐.η _ C.∘ f)) - ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.F.F-resp-≈ ((useful CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc ) CH.○ CH.refl⟩∘⟨ CH.refl⟩∘⟨ C.assoc ))) ⟩ + ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.F.F-resp-≈ ((useful CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc ) CH.○ CH.refl⟩∘⟨ CH.refl⟩∘⟨ C.assoc ))) ⟩ {! !} ≈⟨ ((refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ {! !} ≈⟨ (refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _)) ⟩ {! !} ≈⟨ (refl⟩∘⟨ assoc) ⟩ {! !} ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ {! !} ≈⟨ (refl⟩∘⟨ (A.F.homomorphism ○ refl⟩∘⟨ A.F.homomorphism ○ refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ {! !} ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ - {! !} ≈⟨ {! !} ⟩ - {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ (refl⟩∘⟨ extendʳ A.D (A.adj.counit.commute _)) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.D.identityˡ) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ Γ.⇐.commute _) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ {! Functor.homomorphism H.H !}) ⟩ + + -- {! !} ≈⟨ (refl⟩∘⟨ Functor.F-resp-≈ H.H (M.η.commute _) ⟩∘⟨refl) ⟩ + -- {! !} ≈⟨ (refl⟩∘⟨ {! Functor.homomorphism H.H {f = f} {g = M.η.η _} !} ⟩∘⟨refl) ⟩ {! !} ≈⟨ {! !} ⟩ {! !} ≈⟨ {! !} ⟩ Functor.F₁ H.H f ∘ Γ.⇐.η _ ∎ From ac220917ed55d85c2d37909d63072c0cac7bfed1 Mon Sep 17 00:00:00 2001 From: fouche Date: Mon, 9 Jan 2023 16:44:15 +0200 Subject: [PATCH 24/46] aloora --- .../Construction/Adjunctions/Properties.agda | 15 ++++++++++----- 1 file changed, 10 insertions(+), 5 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda index ae9a96e54..c2fe4b3f0 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda @@ -134,12 +134,17 @@ Kl-initial = record {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.D.identityˡ) ⟩ {! !} ≈⟨ (refl⟩∘⟨ Γ.⇐.commute _) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ {! Functor.homomorphism H.H !}) ⟩ - - -- {! !} ≈⟨ (refl⟩∘⟨ Functor.F-resp-≈ H.H (M.η.commute _) ⟩∘⟨refl) ⟩ + -- {! !} ≈⟨ (Functor.F-resp-≈ H.H (C.Equiv.sym M.identityˡ) ⟩∘⟨refl) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ Functor.F-resp-≈ H.H (M.η.commute _) ⟩∘⟨refl) ⟩ -- {! !} ≈⟨ (refl⟩∘⟨ {! Functor.homomorphism H.H {f = f} {g = M.η.η _} !} ⟩∘⟨refl) ⟩ - {! !} ≈⟨ {! !} ⟩ - {! !} ≈⟨ {! !} ⟩ + -- {! !} ≈⟨ {! Functor.homomorphism H.H !} ⟩ + {! !} ≈⟨ Equiv.sym assoc ⟩ + {! !} ≈⟨ (Equiv.sym (Functor.homomorphism H.H) ⟩∘⟨refl) ⟩ + -- {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ (Functor.F-resp-≈ H.H (elimʳ C M.F.identity CH.⟩∘⟨refl) ⟩∘⟨refl) ⟩ + {! !} ≈⟨ (Functor.F-resp-≈ H.H C.sym-assoc ⟩∘⟨refl) ⟩ + {! !} ≈⟨ pushˡ A.D (Functor.homomorphism H.H) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ elimˡ A.D (Functor.identity H.H)) ⟩ Functor.F₁ H.H f ∘ Γ.⇐.η _ ∎ ; iso = NaturalIsomorphism.iso (sym H.Γ) }) From d9f6b103c0a8d231ce8154e15c5183053e150b3b Mon Sep 17 00:00:00 2001 From: fouche Date: Mon, 9 Jan 2023 16:47:06 +0200 Subject: [PATCH 25/46] pulizia --- .../Construction/Adjunctions/Properties.agda | 32 +++++++------------ 1 file changed, 11 insertions(+), 21 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda index c2fe4b3f0..5e27071f3 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda @@ -13,7 +13,7 @@ module M = Monad M open import Categories.Adjoint using (Adjoint) open import Categories.Functor using (Functor) open import Categories.Morphism.Reasoning -open import Categories.NaturalTransformation +open import Categories.NaturalTransformation hiding (id) open NaturalTransformation using (η) open import Categories.NaturalTransformation.NaturalIsomorphism using (NaturalIsomorphism; niHelper; sym; associator) @@ -35,11 +35,11 @@ EM-object = record ; G = EM.Forgetful ; adj = EM.Free⊣Forgetful ; GF≃M = EM.FF≃F - ; η-eq = {! !} {- begin + ; η-eq = begin M.F.₁ id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ id ∘ M.η.η _ ≈⟨ identityˡ ⟩ - M.η.η _ ∎ -} - ; μ-eq = {! !} {- begin + M.η.η _ ∎ + ; μ-eq = begin M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ @@ -48,7 +48,7 @@ EM-object = record (M.μ.η _ ∘ M.F.₁ id) ∘ id ≈⟨ assoc ⟩ M.μ.η _ ∘ M.F.₁ id ∘ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ M.μ.η _ ∘ M.F.₁ id ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ - M.μ.η _ ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ -} + M.μ.η _ ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ } where open Category C open HomReasoning @@ -59,11 +59,11 @@ Kl-object = record ; G = KL.Forgetful ; adj = KL.Free⊣Forgetful ; GF≃M = KL.FF≃F - ; η-eq = {! !} {- begin + ; η-eq = begin M.F.₁ id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ id ∘ M.η.η _ ≈⟨ identityˡ ⟩ - M.η.η _ ∎ -} - ; μ-eq = {! !} {- begin M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ + M.η.η _ ∎ + ; μ-eq = begin M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ M.μ.η _ ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ @@ -74,7 +74,7 @@ Kl-object = record ((M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id) ∘ id ≈⟨ assoc ⟩ (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ - (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ -} + (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ } where open Category C open HomReasoning @@ -84,12 +84,6 @@ EM-terminal = record ; !-unique = {! !} } -{- - HF≃F'.⇐.η Y ∘ A.adj.counit.η (A.F.F₀ Y) ∘ A.F.F₁ (A.GF≃M.⇐.η Y C.∘ f) - -≈ Functor.F₁ H f ∘ Γ.⇐.η X --} - Kl-initial : IsInitial Kl-object Kl-initial = record { ! = ! @@ -110,7 +104,7 @@ Kl-initial = record C.∘ A.G.F₁ (Functor.F₁ H.H C.id) C.∘ A.G.F₁ (Γ.⇐.η (M.F.F₀ x)) C.∘ A.adj.unit.η (M.F.F₀ x) - useful = {! let open H in NaturalIsomorphism.F⇒G (associator ? ? ?) !} {-let open C.HomReasoning in + useful = let open C.HomReasoning in begin _ ≈⟨ H.μ-comp ⟩ _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ elimʳ C (M.F.identity)) ⟩ _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ C.identityʳ) ⟩ @@ -119,7 +113,7 @@ Kl-initial = record _ ≈⟨ (refl⟩∘⟨ (A.G.F-resp-≈ (Functor.F-resp-≈ H.H M.F.identity) ⟩∘⟨refl) ⟩∘⟨refl) ⟩ _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ C.identityˡ) ⟩ _ ≈⟨ (refl⟩∘⟨ (elimʳ C (elimʳ C A.G.identity)) ⟩∘⟨refl) ⟩ - _ ∎-} + _ ∎ in let open A.D.HomReasoning in begin (Γ.⇐.η _ ∘ A.adj.counit.η (A.F.F₀ _) ∘ A.F.F₁ (A.GF≃M.⇐.η _ C.∘ f)) @@ -134,13 +128,9 @@ Kl-initial = record {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.D.identityˡ) ⟩ {! !} ≈⟨ (refl⟩∘⟨ Γ.⇐.commute _) ⟩ - -- {! !} ≈⟨ (Functor.F-resp-≈ H.H (C.Equiv.sym M.identityˡ) ⟩∘⟨refl) ⟩ {! !} ≈⟨ (refl⟩∘⟨ Functor.F-resp-≈ H.H (M.η.commute _) ⟩∘⟨refl) ⟩ - -- {! !} ≈⟨ (refl⟩∘⟨ {! Functor.homomorphism H.H {f = f} {g = M.η.η _} !} ⟩∘⟨refl) ⟩ - -- {! !} ≈⟨ {! Functor.homomorphism H.H !} ⟩ {! !} ≈⟨ Equiv.sym assoc ⟩ {! !} ≈⟨ (Equiv.sym (Functor.homomorphism H.H) ⟩∘⟨refl) ⟩ - -- {! !} ≈⟨ {! !} ⟩ {! !} ≈⟨ (Functor.F-resp-≈ H.H (elimʳ C M.F.identity CH.⟩∘⟨refl) ⟩∘⟨refl) ⟩ {! !} ≈⟨ (Functor.F-resp-≈ H.H C.sym-assoc ⟩∘⟨refl) ⟩ {! !} ≈⟨ pushˡ A.D (Functor.homomorphism H.H) ⟩ From 260272d6e7fa589161f1a69fc40241830a99bb84 Mon Sep 17 00:00:00 2001 From: fouche Date: Mon, 9 Jan 2023 16:49:52 +0200 Subject: [PATCH 26/46] underscore --- .../Construction/Adjunctions/Properties.agda | 42 +++++++++---------- 1 file changed, 21 insertions(+), 21 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda index 5e27071f3..4d23350eb 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda @@ -118,23 +118,23 @@ Kl-initial = record in begin (Γ.⇐.η _ ∘ A.adj.counit.η (A.F.F₀ _) ∘ A.F.F₁ (A.GF≃M.⇐.η _ C.∘ f)) ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.F.F-resp-≈ ((useful CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc ) CH.○ CH.refl⟩∘⟨ CH.refl⟩∘⟨ C.assoc ))) ⟩ - {! !} ≈⟨ ((refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _)) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ assoc) ⟩ - {! !} ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ (A.F.homomorphism ○ refl⟩∘⟨ A.F.homomorphism ○ refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ - {! !} ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ extendʳ A.D (A.adj.counit.commute _)) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.D.identityˡ) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ Γ.⇐.commute _) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ Functor.F-resp-≈ H.H (M.η.commute _) ⟩∘⟨refl) ⟩ - {! !} ≈⟨ Equiv.sym assoc ⟩ - {! !} ≈⟨ (Equiv.sym (Functor.homomorphism H.H) ⟩∘⟨refl) ⟩ - {! !} ≈⟨ (Functor.F-resp-≈ H.H (elimʳ C M.F.identity CH.⟩∘⟨refl) ⟩∘⟨refl) ⟩ - {! !} ≈⟨ (Functor.F-resp-≈ H.H C.sym-assoc ⟩∘⟨refl) ⟩ - {! !} ≈⟨ pushˡ A.D (Functor.homomorphism H.H) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ elimˡ A.D (Functor.identity H.H)) ⟩ + _ ≈⟨ ((refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ + _ ≈⟨ (refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _)) ⟩ + _ ≈⟨ (refl⟩∘⟨ assoc) ⟩ + _ ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ + _ ≈⟨ (refl⟩∘⟨ (A.F.homomorphism ○ refl⟩∘⟨ A.F.homomorphism ○ refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ + _ ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ + _ ≈⟨ (refl⟩∘⟨ extendʳ A.D (A.adj.counit.commute _)) ⟩ + _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ + _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.D.identityˡ) ⟩ + _ ≈⟨ (refl⟩∘⟨ Γ.⇐.commute _) ⟩ + _ ≈⟨ (refl⟩∘⟨ Functor.F-resp-≈ H.H (M.η.commute _) ⟩∘⟨refl) ⟩ + _ ≈⟨ Equiv.sym assoc ⟩ + _ ≈⟨ (Equiv.sym (Functor.homomorphism H.H) ⟩∘⟨refl) ⟩ + _ ≈⟨ (Functor.F-resp-≈ H.H (elimʳ C M.F.identity CH.⟩∘⟨refl) ⟩∘⟨refl) ⟩ + _ ≈⟨ (Functor.F-resp-≈ H.H C.sym-assoc ⟩∘⟨refl) ⟩ + _ ≈⟨ pushˡ A.D (Functor.homomorphism H.H) ⟩ + _ ≈⟨ (refl⟩∘⟨ elimˡ A.D (Functor.identity H.H)) ⟩ Functor.F₁ H.H f ∘ Γ.⇐.η _ ∎ ; iso = NaturalIsomorphism.iso (sym H.Γ) }) @@ -154,7 +154,7 @@ Kl-initial = record adj.counit.η (F.₀ A) ∘ F.₁ (GF≃M.⇐.η A C.∘ M.η.η A) ≈⟨ refl⟩∘⟨ F.F-resp-≈ η-eq ⟩ adj.counit.η (F.₀ A) ∘ F.₁ (adj.unit.η A) ≈⟨ adj.zig ⟩ D.id ∎ - ; homomorphism = {! !} {-λ {X} {Y} {Z} {f} {g} → + ; homomorphism = λ {X} {Y} {Z} {f} {g} → let ε x = adj.counit.η x in begin ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ (M.μ.η _ C.∘ M.F.₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (assoc²'' C) ⟩ @@ -167,10 +167,10 @@ Kl-initial = record ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.homomorphism ⟩ ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g))) ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ DMR.extendʳ (adj.counit.commute _) ⟩ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ sym-assoc ⟩ - (ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g)) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ∎ -} + (ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g)) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ∎ ; F-resp-≈ = λ x → D.∘-resp-≈ʳ (F.F-resp-≈ (C.∘-resp-≈ʳ x)) } - ; Γ = {! !} {- + ; Γ = let open D open D.HomReasoning in niHelper (record @@ -186,7 +186,7 @@ Kl-initial = record F.₁ f ≈⟨ D.Equiv.sym identityʳ ⟩ (F.₁ f ∘ D.id) ∎ ; iso = λ X → record { isoˡ = identityˡ ; isoʳ = identityˡ } - }) -} + }) ; μ-comp = {! !} } where module adj = Adjoint adj From 0c9a2457e6e84a4b41e0dbab49d3eb36369e5998 Mon Sep 17 00:00:00 2001 From: fouche Date: Mon, 9 Jan 2023 16:57:46 +0200 Subject: [PATCH 27/46] ostia --- .../Adjoint/Construction/Adjunctions/Properties.agda | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda index 4d23350eb..699b2b997 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda @@ -91,7 +91,7 @@ Kl-initial = record let module A = SplitObj A module K = SplitObj Kl-object module H = Split⇒ H - module ! = Split⇒ ! + -- module ! = Split⇒ ! module Γ = NaturalIsomorphism H.Γ module CH = C.HomReasoning open Category A.D in @@ -187,7 +187,7 @@ Kl-initial = record (F.₁ f ∘ D.id) ∎ ; iso = λ X → record { isoˡ = identityˡ ; isoʳ = identityˡ } }) - ; μ-comp = {! !} + ; μ-comp = λ { {x} → {! !} } } where module adj = Adjoint adj module F = Functor F From f13489a4cca89f5f645f5b71207244a56999cb10 Mon Sep 17 00:00:00 2001 From: fouche Date: Mon, 9 Jan 2023 17:06:16 +0200 Subject: [PATCH 28/46] nomi di mucomp --- src/Categories/Adjoint/Construction/Adjunctions.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index ce938fdb4..3711df274 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -125,7 +125,7 @@ Split M = record ; μ-comp = {! !} } where module A = SplitObj A comp : {A B X : SplitObj} → Split⇒ B X → Split⇒ A B → Split⇒ A X - comp {A = A} {B = B} {X = X} (split⇒ Hᵤ HF≃F'ᵤ μ-comp) (split⇒ Hᵥ HF≃F'ᵥ μ-comp′) = record + comp {A = A} {B = B} {X = X} (split⇒ Hᵤ HF≃F'ᵤ Aμ-comp) (split⇒ Hᵥ HF≃F'ᵥ Bμ-comp) = record { H = Hᵤ ∘F Hᵥ ; Γ = HF≃F'ᵤ ⓘᵥ (Hᵤ ⓘˡ HF≃F'ᵥ) ⓘᵥ associator (SplitObj.F A) Hᵥ Hᵤ ; μ-comp = {! !} From 399992cef9406e66acac7f430fb3891eb54743f3 Mon Sep 17 00:00:00 2001 From: fouche Date: Mon, 9 Jan 2023 19:41:18 +0200 Subject: [PATCH 29/46] un mucomp fatto --- .../Adjoint/Construction/Adjunctions.agda | 16 +++++++++++++++- 1 file changed, 15 insertions(+), 1 deletion(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 3711df274..8c9092f84 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -122,8 +122,22 @@ Split M = record split-id {A} = record { H = Categories.Functor.id ; Γ = unitorˡ - ; μ-comp = {! !} + ; μ-comp = λ { {x} → + Equiv.sym (begin {! !} ≈⟨ elimˡ C (Functor.identity A.G) ⟩ + {! !} ≈⟨ identityˡ ⟩ + {! !} ≈⟨ assoc ⟩ + {! !} ≈⟨ identityˡ ⟩ + {! !} ≈⟨ ((refl⟩∘⟨ (refl⟩∘⟨ identityʳ)) ⟩∘⟨refl) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ elimˡ C (Functor.identity A.G)) ⟩∘⟨refl ⟩ + {! !} ≈⟨ ((refl⟩∘⟨ identityˡ) ⟩∘⟨refl) ⟩ + {! !} ≈⟨ (((refl⟩∘⟨ identityˡ) ⟩∘⟨refl) ⟩∘⟨refl) ⟩ + {! !} ≈⟨ ((elimʳ C (Functor.identity A.G) ⟩∘⟨refl) ⟩∘⟨refl) ⟩ + -- {! !} ≈⟨ ({! !} ⟩∘⟨refl) ⟩ + {! !} ≈⟨ elimˡ C A.adj.zag ⟩ + {! !} ∎)} } where module A = SplitObj A + open C + open C.HomReasoning comp : {A B X : SplitObj} → Split⇒ B X → Split⇒ A B → Split⇒ A X comp {A = A} {B = B} {X = X} (split⇒ Hᵤ HF≃F'ᵤ Aμ-comp) (split⇒ Hᵥ HF≃F'ᵥ Bμ-comp) = record { H = Hᵤ ∘F Hᵥ From 69260962a660ffc4fd2f67fece913791dbcedce7 Mon Sep 17 00:00:00 2001 From: fouche Date: Mon, 9 Jan 2023 19:42:14 +0200 Subject: [PATCH 30/46] underscores --- .../Adjoint/Construction/Adjunctions.agda | 23 +++++++++---------- 1 file changed, 11 insertions(+), 12 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 8c9092f84..53aa1ab75 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -123,18 +123,17 @@ Split M = record { H = Categories.Functor.id ; Γ = unitorˡ ; μ-comp = λ { {x} → - Equiv.sym (begin {! !} ≈⟨ elimˡ C (Functor.identity A.G) ⟩ - {! !} ≈⟨ identityˡ ⟩ - {! !} ≈⟨ assoc ⟩ - {! !} ≈⟨ identityˡ ⟩ - {! !} ≈⟨ ((refl⟩∘⟨ (refl⟩∘⟨ identityʳ)) ⟩∘⟨refl) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ elimˡ C (Functor.identity A.G)) ⟩∘⟨refl ⟩ - {! !} ≈⟨ ((refl⟩∘⟨ identityˡ) ⟩∘⟨refl) ⟩ - {! !} ≈⟨ (((refl⟩∘⟨ identityˡ) ⟩∘⟨refl) ⟩∘⟨refl) ⟩ - {! !} ≈⟨ ((elimʳ C (Functor.identity A.G) ⟩∘⟨refl) ⟩∘⟨refl) ⟩ - -- {! !} ≈⟨ ({! !} ⟩∘⟨refl) ⟩ - {! !} ≈⟨ elimˡ C A.adj.zag ⟩ - {! !} ∎)} + Equiv.sym (begin _ ≈⟨ elimˡ C (Functor.identity A.G) ⟩ + _ ≈⟨ identityˡ ⟩ + _ ≈⟨ assoc ⟩ + _ ≈⟨ identityˡ ⟩ + _ ≈⟨ ((refl⟩∘⟨ (refl⟩∘⟨ identityʳ)) ⟩∘⟨refl) ⟩ + _ ≈⟨ (refl⟩∘⟨ elimˡ C (Functor.identity A.G)) ⟩∘⟨refl ⟩ + _ ≈⟨ ((refl⟩∘⟨ identityˡ) ⟩∘⟨refl) ⟩ + _ ≈⟨ (((refl⟩∘⟨ identityˡ) ⟩∘⟨refl) ⟩∘⟨refl) ⟩ + _ ≈⟨ ((elimʳ C (Functor.identity A.G) ⟩∘⟨refl) ⟩∘⟨refl) ⟩ + _ ≈⟨ elimˡ C A.adj.zag ⟩ + _ ∎)} } where module A = SplitObj A open C open C.HomReasoning From c25e87b2913448da0da4371b15c1318232495097 Mon Sep 17 00:00:00 2001 From: fouche Date: Mon, 9 Jan 2023 20:31:45 +0200 Subject: [PATCH 31/46] qualcosa fatto --- .../Adjoint/Construction/Adjunctions.agda | 14 +++++++++++--- 1 file changed, 11 insertions(+), 3 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 53aa1ab75..f6b36108a 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -138,12 +138,20 @@ Split M = record open C open C.HomReasoning comp : {A B X : SplitObj} → Split⇒ B X → Split⇒ A B → Split⇒ A X - comp {A = A} {B = B} {X = X} (split⇒ Hᵤ HF≃F'ᵤ Aμ-comp) (split⇒ Hᵥ HF≃F'ᵥ Bμ-comp) = record + comp {A = A} {B = B} {X = X} (split⇒ Hᵤ HF≃Fᵤ Aμ-comp) (split⇒ Hᵥ HF≃Fᵥ Bμ-comp) = record { H = Hᵤ ∘F Hᵥ - ; Γ = HF≃F'ᵤ ⓘᵥ (Hᵤ ⓘˡ HF≃F'ᵥ) ⓘᵥ associator (SplitObj.F A) Hᵥ Hᵤ - ; μ-comp = {! !} + ; Γ = HF≃Fᵤ ⓘᵥ (Hᵤ ⓘˡ HF≃Fᵥ) ⓘᵥ associator (SplitObj.F A) Hᵥ Hᵤ + ; μ-comp = λ { {x} → + Equiv.sym (begin {! !} ≈⟨ ( Functor.homomorphism X.G ⟩∘⟨refl) ⟩ + {! !} ≈⟨ ((refl⟩∘⟨ Functor.F-resp-≈ X.G X.D.identityʳ) ⟩∘⟨refl) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ (identityˡ ○ (identityˡ ⟩∘⟨refl))) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ (refl⟩∘⟨ (identityʳ ○ identityˡ))) ⟩∘⟨refl) ⟩ + {! !} ≈⟨ {! !} ⟩ + {! !} ∎)} -- ; G'H≃G = G'H≃Gᵥ ⓘᵥ (G'H≃Gᵤ ⓘʳ Hᵥ) ⓘᵥ sym-associator Hᵥ Hᵤ (SplitObj.G X) } where module A = SplitObj A module B = SplitObj B module X = SplitObj X + open C + open C.HomReasoning From baec081af366a031ff6805ee76417fe4cf460fcb Mon Sep 17 00:00:00 2001 From: fouche Date: Tue, 10 Jan 2023 00:16:36 +0200 Subject: [PATCH 32/46] removed ids --- src/Categories/Adjoint/Construction/Adjunctions.agda | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index f6b36108a..70146ad78 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -146,6 +146,12 @@ Split M = record {! !} ≈⟨ ((refl⟩∘⟨ Functor.F-resp-≈ X.G X.D.identityʳ) ⟩∘⟨refl) ⟩ {! !} ≈⟨ (refl⟩∘⟨ (identityˡ ○ (identityˡ ⟩∘⟨refl))) ⟩ {! !} ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ (refl⟩∘⟨ (identityʳ ○ identityˡ))) ⟩∘⟨refl) ⟩ + {! !} ≈⟨ refl⟩∘⟨ ((refl⟩∘⟨ (Functor.homomorphism X.G ⟩∘⟨refl)) ⟩∘⟨refl) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ (((refl⟩∘⟨ (identityˡ ○ Functor.identity X.G)) ⟩∘⟨refl) ⟩∘⟨refl)) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ (identityʳ ⟩∘⟨refl) ⟩∘⟨refl) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ ((refl⟩∘⟨ ((Functor.homomorphism X.G ⟩∘⟨refl) ⟩∘⟨refl)) ⟩∘⟨refl)) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ ((refl⟩∘⟨ ((elimˡ C (Functor.identity X.G) ⟩∘⟨refl) ⟩∘⟨refl)) ⟩∘⟨refl)) ⟩ + {! !} ≈⟨ {! !} ⟩ {! !} ≈⟨ {! !} ⟩ {! !} ∎)} -- ; G'H≃G = G'H≃Gᵥ ⓘᵥ (G'H≃Gᵤ ⓘʳ Hᵥ) ⓘᵥ sym-associator Hᵥ Hᵤ (SplitObj.G X) From 4543667dc0fee2a38db71255e868fc3827245333 Mon Sep 17 00:00:00 2001 From: fouche Date: Tue, 10 Jan 2023 00:30:52 +0200 Subject: [PATCH 33/46] swoosh --- src/Categories/Adjoint/Construction/Adjunctions.agda | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 70146ad78..7de711ea9 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -138,9 +138,9 @@ Split M = record open C open C.HomReasoning comp : {A B X : SplitObj} → Split⇒ B X → Split⇒ A B → Split⇒ A X - comp {A = A} {B = B} {X = X} (split⇒ Hᵤ HF≃Fᵤ Aμ-comp) (split⇒ Hᵥ HF≃Fᵥ Bμ-comp) = record + comp {A = A} {B = B} {X = X} (split⇒ Hᵤ Γᵤ Aμ-comp) (split⇒ Hᵥ Γᵥ Bμ-comp) = record { H = Hᵤ ∘F Hᵥ - ; Γ = HF≃Fᵤ ⓘᵥ (Hᵤ ⓘˡ HF≃Fᵥ) ⓘᵥ associator (SplitObj.F A) Hᵥ Hᵤ + ; Γ = Γᵤ ⓘᵥ (Hᵤ ⓘˡ Γᵥ) ⓘᵥ associator (SplitObj.F A) Hᵥ Hᵤ ; μ-comp = λ { {x} → Equiv.sym (begin {! !} ≈⟨ ( Functor.homomorphism X.G ⟩∘⟨refl) ⟩ {! !} ≈⟨ ((refl⟩∘⟨ Functor.F-resp-≈ X.G X.D.identityʳ) ⟩∘⟨refl) ⟩ From b4972e33f4644857e449047775cfb5c726862448 Mon Sep 17 00:00:00 2001 From: fouche Date: Tue, 10 Jan 2023 12:21:54 +0200 Subject: [PATCH 34/46] simplify mucomp ids --- .../Adjoint/Construction/Adjunctions.agda | 71 +++++++++++-------- 1 file changed, 41 insertions(+), 30 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 7de711ea9..d4b6dc984 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -95,11 +95,12 @@ record Split⇒ (X Y : SplitObj) : Set (suc o ⊔ suc ℓ ⊔ suc e) where module Γ = NaturalIsomorphism Γ field - μ-comp : Y.GF≃M.F⇐G -- Y.GF≃M.F⇒G - ≅ (Y.G ∘ˡ Γ.F⇒G) - ∘ᵥ NaturalIsomorphism.F⇒G (associator X.F H Y.G) - ∘ᵥ (Γ' {X = X} {Y = Y} H Γ ∘ʳ X.F) - ∘ᵥ X.GF≃M.F⇐G + -- μ-comp : Y.GF≃M.F⇐G -- Y.GF≃M.F⇒G + -- ≅ (Y.G ∘ˡ Γ.F⇒G) + -- ∘ᵥ NaturalIsomorphism.F⇒G (associator X.F H Y.G) + -- ∘ᵥ (Γ' {X = X} {Y = Y} H Γ ∘ʳ X.F) + -- ∘ᵥ X.GF≃M.F⇐G + μ-comp : ∀ {x : Obj} → Y.GF≃M.⇐.η x ≈ Y.G.F₁ (Γ.⇒.η x) ∘ ((Y.G.F₁ (Functor.F₁ H (X.adj.counit.η (X.F.F₀ x)))) ∘ Y.G.F₁ (Γ.⇐.η (X.G.F₀ (X.F.F₀ x))) ∘ (Y.adj.unit.η (X.G.F₀ (X.F.F₀ x)))) ∘ X.GF≃M.⇐.η x Split : Monad C → Category _ _ _ Split M = record @@ -122,18 +123,22 @@ Split M = record split-id {A} = record { H = Categories.Functor.id ; Γ = unitorˡ - ; μ-comp = λ { {x} → - Equiv.sym (begin _ ≈⟨ elimˡ C (Functor.identity A.G) ⟩ - _ ≈⟨ identityˡ ⟩ - _ ≈⟨ assoc ⟩ - _ ≈⟨ identityˡ ⟩ - _ ≈⟨ ((refl⟩∘⟨ (refl⟩∘⟨ identityʳ)) ⟩∘⟨refl) ⟩ - _ ≈⟨ (refl⟩∘⟨ elimˡ C (Functor.identity A.G)) ⟩∘⟨refl ⟩ - _ ≈⟨ ((refl⟩∘⟨ identityˡ) ⟩∘⟨refl) ⟩ - _ ≈⟨ (((refl⟩∘⟨ identityˡ) ⟩∘⟨refl) ⟩∘⟨refl) ⟩ - _ ≈⟨ ((elimʳ C (Functor.identity A.G) ⟩∘⟨refl) ⟩∘⟨refl) ⟩ - _ ≈⟨ elimˡ C A.adj.zag ⟩ - _ ∎)} + ; μ-comp = λ { {x} → + Equiv.sym(begin {! !} ≈⟨ elimˡ C (Functor.identity A.G) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ elimˡ C (Functor.identity A.G)) ⟩∘⟨refl ⟩ + {! !} ≈⟨ elimˡ C A.adj.zag ⟩ + {! !} ∎)} + -- Equiv.sym (begin _ ≈⟨ elimˡ C (Functor.identity A.G) ⟩ + -- _ ≈⟨ identityˡ ⟩ + -- _ ≈⟨ assoc ⟩ + -- _ ≈⟨ identityˡ ⟩ + -- _ ≈⟨ ((refl⟩∘⟨ (refl⟩∘⟨ identityʳ)) ⟩∘⟨refl) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ elimˡ C (Functor.identity A.G)) ⟩∘⟨refl ⟩ + -- _ ≈⟨ ((refl⟩∘⟨ identityˡ) ⟩∘⟨refl) ⟩ + -- _ ≈⟨ (((refl⟩∘⟨ identityˡ) ⟩∘⟨refl) ⟩∘⟨refl) ⟩ + -- _ ≈⟨ ((elimʳ C (Functor.identity A.G) ⟩∘⟨refl) ⟩∘⟨refl) ⟩ + -- _ ≈⟨ elimˡ C A.adj.zag ⟩ + -- _ ∎)} } where module A = SplitObj A open C open C.HomReasoning @@ -141,19 +146,25 @@ Split M = record comp {A = A} {B = B} {X = X} (split⇒ Hᵤ Γᵤ Aμ-comp) (split⇒ Hᵥ Γᵥ Bμ-comp) = record { H = Hᵤ ∘F Hᵥ ; Γ = Γᵤ ⓘᵥ (Hᵤ ⓘˡ Γᵥ) ⓘᵥ associator (SplitObj.F A) Hᵥ Hᵤ - ; μ-comp = λ { {x} → - Equiv.sym (begin {! !} ≈⟨ ( Functor.homomorphism X.G ⟩∘⟨refl) ⟩ - {! !} ≈⟨ ((refl⟩∘⟨ Functor.F-resp-≈ X.G X.D.identityʳ) ⟩∘⟨refl) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ (identityˡ ○ (identityˡ ⟩∘⟨refl))) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ (refl⟩∘⟨ (identityʳ ○ identityˡ))) ⟩∘⟨refl) ⟩ - {! !} ≈⟨ refl⟩∘⟨ ((refl⟩∘⟨ (Functor.homomorphism X.G ⟩∘⟨refl)) ⟩∘⟨refl) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ (((refl⟩∘⟨ (identityˡ ○ Functor.identity X.G)) ⟩∘⟨refl) ⟩∘⟨refl)) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ (identityʳ ⟩∘⟨refl) ⟩∘⟨refl) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ ((refl⟩∘⟨ ((Functor.homomorphism X.G ⟩∘⟨refl) ⟩∘⟨refl)) ⟩∘⟨refl)) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ ((refl⟩∘⟨ ((elimˡ C (Functor.identity X.G) ⟩∘⟨refl) ⟩∘⟨refl)) ⟩∘⟨refl)) ⟩ - {! !} ≈⟨ {! !} ⟩ - {! !} ≈⟨ {! !} ⟩ - {! !} ∎)} + ; μ-comp = λ { {x} → + Equiv.sym (begin {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ {! !} ⟩ + {! !} ∎) } + -- Equiv.sym (begin {! !} ≈⟨ ( Functor.homomorphism X.G ⟩∘⟨refl) ⟩ + -- {! !} ≈⟨ ((refl⟩∘⟨ Functor.F-resp-≈ X.G X.D.identityʳ) ⟩∘⟨refl) ⟩ + -- {! !} ≈⟨ (refl⟩∘⟨ (identityˡ ○ (identityˡ ⟩∘⟨refl))) ⟩ + -- {! !} ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ (refl⟩∘⟨ (identityʳ ○ identityˡ))) ⟩∘⟨refl) ⟩ + -- {! !} ≈⟨ refl⟩∘⟨ ((refl⟩∘⟨ (Functor.homomorphism X.G ⟩∘⟨refl)) ⟩∘⟨refl) ⟩ + -- {! !} ≈⟨ (refl⟩∘⟨ (((refl⟩∘⟨ (identityˡ ○ Functor.identity X.G)) ⟩∘⟨refl) ⟩∘⟨refl)) ⟩ + -- {! !} ≈⟨ (refl⟩∘⟨ (identityʳ ⟩∘⟨refl) ⟩∘⟨refl) ⟩ + -- {! !} ≈⟨ (refl⟩∘⟨ ((refl⟩∘⟨ ((Functor.homomorphism X.G ⟩∘⟨refl) ⟩∘⟨refl)) ⟩∘⟨refl)) ⟩ + -- {! !} ≈⟨ (refl⟩∘⟨ ((refl⟩∘⟨ ((elimˡ C (Functor.identity X.G) ⟩∘⟨refl) ⟩∘⟨refl)) ⟩∘⟨refl)) ⟩ + -- {! !} ≈⟨ {! !} ⟩ + -- {! !} ≈⟨ {! !} ⟩ + -- {! !} ∎)} -- ; G'H≃G = G'H≃Gᵥ ⓘᵥ (G'H≃Gᵤ ⓘʳ Hᵥ) ⓘᵥ sym-associator Hᵥ Hᵤ (SplitObj.G X) } where module A = SplitObj A From 3902c5bd57b723221f57f11d9f7657411abe0599 Mon Sep 17 00:00:00 2001 From: fouche Date: Tue, 10 Jan 2023 12:42:16 +0200 Subject: [PATCH 35/46] smol steps --- src/Categories/Adjoint/Construction/Adjunctions.agda | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index d4b6dc984..8891cc559 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -147,9 +147,9 @@ Split M = record { H = Hᵤ ∘F Hᵥ ; Γ = Γᵤ ⓘᵥ (Hᵤ ⓘˡ Γᵥ) ⓘᵥ associator (SplitObj.F A) Hᵥ Hᵤ ; μ-comp = λ { {x} → - Equiv.sym (begin {! !} ≈⟨ {! !} ⟩ - {! !} ≈⟨ {! !} ⟩ - {! !} ≈⟨ {! !} ⟩ + Equiv.sym (begin {! !} ≈⟨ (X.G.F-resp-≈ (X.D.HomReasoning.refl⟩∘⟨ X.D.identityʳ) ⟩∘⟨refl) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ ((refl⟩∘⟨ (X.G.F-resp-≈ (X.D.identityˡ X.D.HomReasoning.⟩∘⟨refl) ⟩∘⟨refl)) ⟩∘⟨refl)) ⟩ + {! !} ≈⟨ pushˡ C X.G.homomorphism ⟩ {! !} ≈⟨ {! !} ⟩ {! !} ≈⟨ {! !} ⟩ {! !} ∎) } From 42cebd6b012affa9b5a291a7bad46a1e8bd34c80 Mon Sep 17 00:00:00 2001 From: fouche Date: Tue, 10 Jan 2023 13:58:28 +0200 Subject: [PATCH 36/46] mucomp fatto --- .../Adjoint/Construction/Adjunctions.agda | 43 +--- .../Construction/Adjunctions/Properties.agda | 207 +++++++++--------- 2 files changed, 110 insertions(+), 140 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 8891cc559..518bce204 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -63,24 +63,6 @@ record SplitObj : Set (suc o ⊔ suc ℓ ⊔ suc e) where GF≃M.⇐.η A -} -module _ {X Y : SplitObj} (H : Functor (SplitObj.D X) (SplitObj.D Y)) (Γ : H ∘F (SplitObj.F X) ≃ (SplitObj.F Y)) where - private - module X = SplitObj X - module Y = SplitObj Y - module Γ = NaturalIsomorphism Γ - - Γ' : NaturalTransformation X.G (Y.G ∘F H) - Γ' = F∘id⇒F ∘ᵥ - (((Y.G ∘F H) ∘ˡ X.adj.counit) ∘ᵥ zip ∘ᵥ Y.G ∘ˡ zop) - ∘ᵥ (Y.G ∘ˡ Γ.F⇐G ∘ʳ X.G) - ∘ᵥ (zap ∘ᵥ Y.adj.unit ∘ʳ X.G) - ∘ᵥ F⇒id∘F - where - open NaturalIsomorphism - zip = F⇐G (associator (X.F ∘F X.G) H Y.G) - zop = F⇒G (associator X.G X.F H) - zap = F⇒G (associator X.G Y.F Y.G) - record Split⇒ (X Y : SplitObj) : Set (suc o ⊔ suc ℓ ⊔ suc e) where constructor split⇒ private @@ -95,11 +77,6 @@ record Split⇒ (X Y : SplitObj) : Set (suc o ⊔ suc ℓ ⊔ suc e) where module Γ = NaturalIsomorphism Γ field - -- μ-comp : Y.GF≃M.F⇐G -- Y.GF≃M.F⇒G - -- ≅ (Y.G ∘ˡ Γ.F⇒G) - -- ∘ᵥ NaturalIsomorphism.F⇒G (associator X.F H Y.G) - -- ∘ᵥ (Γ' {X = X} {Y = Y} H Γ ∘ʳ X.F) - -- ∘ᵥ X.GF≃M.F⇐G μ-comp : ∀ {x : Obj} → Y.GF≃M.⇐.η x ≈ Y.G.F₁ (Γ.⇒.η x) ∘ ((Y.G.F₁ (Functor.F₁ H (X.adj.counit.η (X.F.F₀ x)))) ∘ Y.G.F₁ (Γ.⇐.η (X.G.F₀ (X.F.F₀ x))) ∘ (Y.adj.unit.η (X.G.F₀ (X.F.F₀ x)))) ∘ X.GF≃M.⇐.η x Split : Monad C → Category _ _ _ @@ -124,21 +101,10 @@ Split M = record { H = Categories.Functor.id ; Γ = unitorˡ ; μ-comp = λ { {x} → - Equiv.sym(begin {! !} ≈⟨ elimˡ C (Functor.identity A.G) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ elimˡ C (Functor.identity A.G)) ⟩∘⟨refl ⟩ - {! !} ≈⟨ elimˡ C A.adj.zag ⟩ - {! !} ∎)} - -- Equiv.sym (begin _ ≈⟨ elimˡ C (Functor.identity A.G) ⟩ - -- _ ≈⟨ identityˡ ⟩ - -- _ ≈⟨ assoc ⟩ - -- _ ≈⟨ identityˡ ⟩ - -- _ ≈⟨ ((refl⟩∘⟨ (refl⟩∘⟨ identityʳ)) ⟩∘⟨refl) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ elimˡ C (Functor.identity A.G)) ⟩∘⟨refl ⟩ - -- _ ≈⟨ ((refl⟩∘⟨ identityˡ) ⟩∘⟨refl) ⟩ - -- _ ≈⟨ (((refl⟩∘⟨ identityˡ) ⟩∘⟨refl) ⟩∘⟨refl) ⟩ - -- _ ≈⟨ ((elimʳ C (Functor.identity A.G) ⟩∘⟨refl) ⟩∘⟨refl) ⟩ - -- _ ≈⟨ elimˡ C A.adj.zag ⟩ - -- _ ∎)} + Equiv.sym(begin _ ≈⟨ elimˡ C (Functor.identity A.G) ⟩ + _ ≈⟨ (refl⟩∘⟨ elimˡ C (Functor.identity A.G)) ⟩∘⟨refl ⟩ + _ ≈⟨ elimˡ C A.adj.zag ⟩ + _ ∎)} } where module A = SplitObj A open C open C.HomReasoning @@ -165,7 +131,6 @@ Split M = record -- {! !} ≈⟨ {! !} ⟩ -- {! !} ≈⟨ {! !} ⟩ -- {! !} ∎)} - -- ; G'H≃G = G'H≃Gᵥ ⓘᵥ (G'H≃Gᵤ ⓘʳ Hᵥ) ⓘᵥ sym-associator Hᵥ Hᵤ (SplitObj.G X) } where module A = SplitObj A module B = SplitObj B diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda index 699b2b997..0fdbc6aab 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda @@ -28,29 +28,29 @@ open import Categories.Adjoint.Construction.EilenbergMoore M as EM open import Categories.Tactic.Category -EM-object : SplitObj -EM-object = record - { D = EilenbergMoore M - ; F = EM.Free - ; G = EM.Forgetful - ; adj = EM.Free⊣Forgetful - ; GF≃M = EM.FF≃F - ; η-eq = begin - M.F.₁ id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ - id ∘ M.η.η _ ≈⟨ identityˡ ⟩ - M.η.η _ ∎ - ; μ-eq = begin - M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ - id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ - M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ - M.μ.η _ ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ - M.μ.η _ ∘ M.F.₁ id ≈⟨ Equiv.sym identityʳ ⟩ - (M.μ.η _ ∘ M.F.₁ id) ∘ id ≈⟨ assoc ⟩ - M.μ.η _ ∘ M.F.₁ id ∘ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ - M.μ.η _ ∘ M.F.₁ id ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ - M.μ.η _ ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ - } where open Category C - open HomReasoning +-- EM-object : SplitObj +-- EM-object = record +-- { D = EilenbergMoore M +-- ; F = EM.Free +-- ; G = EM.Forgetful +-- ; adj = EM.Free⊣Forgetful +-- ; GF≃M = EM.FF≃F +-- ; η-eq = begin +-- M.F.₁ id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ +-- id ∘ M.η.η _ ≈⟨ identityˡ ⟩ +-- M.η.η _ ∎ +-- ; μ-eq = begin +-- M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ +-- id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ +-- M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ +-- M.μ.η _ ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ +-- M.μ.η _ ∘ M.F.₁ id ≈⟨ Equiv.sym identityʳ ⟩ +-- (M.μ.η _ ∘ M.F.₁ id) ∘ id ≈⟨ assoc ⟩ +-- M.μ.η _ ∘ M.F.₁ id ∘ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ +-- M.μ.η _ ∘ M.F.₁ id ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ +-- M.μ.η _ ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ +-- } where open Category C +-- open HomReasoning Kl-object : SplitObj Kl-object = record @@ -63,27 +63,21 @@ Kl-object = record M.F.₁ id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ id ∘ M.η.η _ ≈⟨ identityˡ ⟩ M.η.η _ ∎ - ; μ-eq = begin M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ - id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ - M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ - M.μ.η _ ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ - M.μ.η _ ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ - M.μ.η _ ∘ M.F.₁ (M.F.₁ id) ≈⟨ Equiv.sym identityʳ ⟩ - (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ - (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ≈⟨ Equiv.sym identityʳ ⟩ - ((M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id) ∘ id ≈⟨ assoc ⟩ - (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ - (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ - (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ + ; μ-eq = {! !} -- begin M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ + -- id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ + -- M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ + -- M.μ.η _ ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + -- M.μ.η _ ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ + -- M.μ.η _ ∘ M.F.₁ (M.F.₁ id) ≈⟨ Equiv.sym identityʳ ⟩ + -- (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + -- (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ≈⟨ Equiv.sym identityʳ ⟩ + -- ((M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id) ∘ id ≈⟨ assoc ⟩ + -- (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + -- (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ + -- (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ } where open Category C open HomReasoning -EM-terminal : IsTerminal EM-object -EM-terminal = record - { ! = {! !} - ; !-unique = {! !} - } - Kl-initial : IsInitial Kl-object Kl-initial = record { ! = ! @@ -98,44 +92,44 @@ Kl-initial = record record { η = Γ.⇐.η ; η⁻¹ = Γ.⇒.η - ; commute = λ f → - let useful : ∀ {x} - → A.GF≃M.⇐.η x C.≈ A.G.F₁ (Γ.⇒.η x) - C.∘ A.G.F₁ (Functor.F₁ H.H C.id) - C.∘ A.G.F₁ (Γ.⇐.η (M.F.F₀ x)) - C.∘ A.adj.unit.η (M.F.F₀ x) - useful = let open C.HomReasoning in - begin _ ≈⟨ H.μ-comp ⟩ - _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ elimʳ C (M.F.identity)) ⟩ - _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ C.identityʳ) ⟩ - _ ≈⟨ (refl⟩∘⟨ C.identityˡ) ⟩ - _ ≈⟨ refl⟩∘⟨ C.identityˡ ⟩ - _ ≈⟨ (refl⟩∘⟨ (A.G.F-resp-≈ (Functor.F-resp-≈ H.H M.F.identity) ⟩∘⟨refl) ⟩∘⟨refl) ⟩ - _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ C.identityˡ) ⟩ - _ ≈⟨ (refl⟩∘⟨ (elimʳ C (elimʳ C A.G.identity)) ⟩∘⟨refl) ⟩ - _ ∎ - in let open A.D.HomReasoning + ; commute = λ f → {! !} + -- let useful : ∀ {x} + -- → A.GF≃M.⇐.η x C.≈ A.G.F₁ (Γ.⇒.η x) + -- C.∘ A.G.F₁ (Functor.F₁ H.H C.id) + -- C.∘ A.G.F₁ (Γ.⇐.η (M.F.F₀ x)) + -- C.∘ A.adj.unit.η (M.F.F₀ x) + -- useful = let open C.HomReasoning in + -- begin _ ≈⟨ H.μ-comp ⟩ + -- _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ elimʳ C (M.F.identity)) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ C.identityʳ) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ C.identityˡ) ⟩ + -- _ ≈⟨ refl⟩∘⟨ C.identityˡ ⟩ + -- _ ≈⟨ (refl⟩∘⟨ (A.G.F-resp-≈ (Functor.F-resp-≈ H.H M.F.identity) ⟩∘⟨refl) ⟩∘⟨refl) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ C.identityˡ) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ (elimʳ C (elimʳ C A.G.identity)) ⟩∘⟨refl) ⟩ + -- _ ∎ + -- in let open A.D.HomReasoning - in begin (Γ.⇐.η _ ∘ A.adj.counit.η (A.F.F₀ _) ∘ A.F.F₁ (A.GF≃M.⇐.η _ C.∘ f)) - ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.F.F-resp-≈ ((useful CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc ) CH.○ CH.refl⟩∘⟨ CH.refl⟩∘⟨ C.assoc ))) ⟩ - _ ≈⟨ ((refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ - _ ≈⟨ (refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _)) ⟩ - _ ≈⟨ (refl⟩∘⟨ assoc) ⟩ - _ ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ - _ ≈⟨ (refl⟩∘⟨ (A.F.homomorphism ○ refl⟩∘⟨ A.F.homomorphism ○ refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ - _ ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ - _ ≈⟨ (refl⟩∘⟨ extendʳ A.D (A.adj.counit.commute _)) ⟩ - _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ - _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.D.identityˡ) ⟩ - _ ≈⟨ (refl⟩∘⟨ Γ.⇐.commute _) ⟩ - _ ≈⟨ (refl⟩∘⟨ Functor.F-resp-≈ H.H (M.η.commute _) ⟩∘⟨refl) ⟩ - _ ≈⟨ Equiv.sym assoc ⟩ - _ ≈⟨ (Equiv.sym (Functor.homomorphism H.H) ⟩∘⟨refl) ⟩ - _ ≈⟨ (Functor.F-resp-≈ H.H (elimʳ C M.F.identity CH.⟩∘⟨refl) ⟩∘⟨refl) ⟩ - _ ≈⟨ (Functor.F-resp-≈ H.H C.sym-assoc ⟩∘⟨refl) ⟩ - _ ≈⟨ pushˡ A.D (Functor.homomorphism H.H) ⟩ - _ ≈⟨ (refl⟩∘⟨ elimˡ A.D (Functor.identity H.H)) ⟩ - Functor.F₁ H.H f ∘ Γ.⇐.η _ ∎ + -- in begin (Γ.⇐.η _ ∘ A.adj.counit.η (A.F.F₀ _) ∘ A.F.F₁ (A.GF≃M.⇐.η _ C.∘ f)) + -- ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.F.F-resp-≈ ((useful CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc ) CH.○ CH.refl⟩∘⟨ CH.refl⟩∘⟨ C.assoc ))) ⟩ + -- _ ≈⟨ ((refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _)) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ assoc) ⟩ + -- _ ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ (A.F.homomorphism ○ refl⟩∘⟨ A.F.homomorphism ○ refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ + -- _ ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ extendʳ A.D (A.adj.counit.commute _)) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.D.identityˡ) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ Γ.⇐.commute _) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ Functor.F-resp-≈ H.H (M.η.commute _) ⟩∘⟨refl) ⟩ + -- _ ≈⟨ Equiv.sym assoc ⟩ + -- _ ≈⟨ (Equiv.sym (Functor.homomorphism H.H) ⟩∘⟨refl) ⟩ + -- _ ≈⟨ (Functor.F-resp-≈ H.H (elimʳ C M.F.identity CH.⟩∘⟨refl) ⟩∘⟨refl) ⟩ + -- _ ≈⟨ (Functor.F-resp-≈ H.H C.sym-assoc ⟩∘⟨refl) ⟩ + -- _ ≈⟨ pushˡ A.D (Functor.homomorphism H.H) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ elimˡ A.D (Functor.identity H.H)) ⟩ + -- Functor.F₁ H.H f ∘ Γ.⇐.η _ ∎ ; iso = NaturalIsomorphism.iso (sym H.Γ) }) } @@ -155,19 +149,19 @@ Kl-initial = record adj.counit.η (F.₀ A) ∘ F.₁ (adj.unit.η A) ≈⟨ adj.zig ⟩ D.id ∎ ; homomorphism = λ {X} {Y} {Z} {f} {g} → - let ε x = adj.counit.η x in - begin - ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ (M.μ.η _ C.∘ M.F.₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (assoc²'' C) ⟩ - ε _ ∘ F.₁ ((GF≃M.⇐.η _ C.∘ M.μ.η _) C.∘ M.F.₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (μ-eq CHR.⟩∘⟨refl) ⟩ - ε _ ∘ F.₁ ((G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.F.₁ (GF≃M.⇐.η _)) C.∘ M.F.₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (assoc²' C) ⟩ - ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.F.₁ (GF≃M.⇐.η _) C.∘ M.F.₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ CHR.refl⟩∘⟨ pullˡ C (C.Equiv.sym M.F.homomorphism)) ⟩ - ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.F.₁ (GF≃M.⇐.η _ C.∘ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ extendʳ C (GF≃M.⇐.commute _)) ⟩ - ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ F.homomorphism ⟩ - ε _ ∘ F.₁ (G.₁ (ε (F.₀ _))) ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ DMR.extendʳ (adj.counit.commute _) ⟩ - ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.homomorphism ⟩ - ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g))) ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ DMR.extendʳ (adj.counit.commute _) ⟩ - ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ sym-assoc ⟩ - (ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g)) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ∎ + let ε x = adj.counit.η x in {! !} + -- begin + -- ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ (M.μ.η _ C.∘ M.F.₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (assoc²'' C) ⟩ + -- ε _ ∘ F.₁ ((GF≃M.⇐.η _ C.∘ M.μ.η _) C.∘ M.F.₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (μ-eq CHR.⟩∘⟨refl) ⟩ + -- ε _ ∘ F.₁ ((G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.F.₁ (GF≃M.⇐.η _)) C.∘ M.F.₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (assoc²' C) ⟩ + -- ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.F.₁ (GF≃M.⇐.η _) C.∘ M.F.₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ CHR.refl⟩∘⟨ pullˡ C (C.Equiv.sym M.F.homomorphism)) ⟩ + -- ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.F.₁ (GF≃M.⇐.η _ C.∘ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ extendʳ C (GF≃M.⇐.commute _)) ⟩ + -- ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ F.homomorphism ⟩ + -- ε _ ∘ F.₁ (G.₁ (ε (F.₀ _))) ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ DMR.extendʳ (adj.counit.commute _) ⟩ + -- ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.homomorphism ⟩ + -- ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g))) ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ DMR.extendʳ (adj.counit.commute _) ⟩ + -- ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ sym-assoc ⟩ + -- (ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g)) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ∎ ; F-resp-≈ = λ x → D.∘-resp-≈ʳ (F.F-resp-≈ (C.∘-resp-≈ʳ x)) } ; Γ = @@ -176,18 +170,29 @@ Kl-initial = record niHelper (record { η = λ _ → D.id ; η⁻¹ = λ _ → D.id - ; commute = λ f → - begin - D.id ∘ adj.counit.η _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ M.η.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C η-eq) ⟩ - D.id ∘ adj.counit.η _ ∘ F.₁ (adj.unit.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.homomorphism ⟩ - D.id ∘ adj.counit.η _ ∘ F.₁ (adj.unit.η _) ∘ F.₁ f ≈⟨ refl⟩∘⟨ pullˡ D adj.zig ⟩ - D.id ∘ D.id ∘ F.₁ f ≈⟨ identityˡ ⟩ - D.id ∘ F.₁ f ≈⟨ identityˡ ⟩ - F.₁ f ≈⟨ D.Equiv.sym identityʳ ⟩ - (F.₁ f ∘ D.id) ∎ + ; commute = λ f → {! !} + -- begin + -- D.id ∘ adj.counit.η _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ M.η.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C η-eq) ⟩ + -- D.id ∘ adj.counit.η _ ∘ F.₁ (adj.unit.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.homomorphism ⟩ + -- D.id ∘ adj.counit.η _ ∘ F.₁ (adj.unit.η _) ∘ F.₁ f ≈⟨ refl⟩∘⟨ pullˡ D adj.zig ⟩ + -- D.id ∘ D.id ∘ F.₁ f ≈⟨ identityˡ ⟩ + -- D.id ∘ F.₁ f ≈⟨ identityˡ ⟩ + -- F.₁ f ≈⟨ D.Equiv.sym identityʳ ⟩ + -- (F.₁ f ∘ D.id) ∎ ; iso = λ X → record { isoˡ = identityˡ ; isoʳ = identityˡ } }) - ; μ-comp = λ { {x} → {! !} } + ; μ-comp = λ { {x} → let open C + open C.HomReasoning in + Equiv.sym + (begin {! !} ≈⟨ elimˡ C G.identity ⟩ + {! !} ≈⟨ elimʳ C M.F.identity ⟩ + {! !} ≈⟨ (refl⟩∘⟨ elimˡ C G.identity) ⟩ + {! !} ≈⟨ (G.homomorphism ⟩∘⟨refl) ⟩ + {! !} ≈⟨ pullʳ C (adj.unit.sym-commute _) ⟩ + {! !} ≈⟨ sym-assoc ⟩ + {! !} ≈⟨ elimˡ C adj.zag ⟩ + {! !} ≈⟨ elimʳ C M.F.identity ⟩ + {! !} ∎) } } where module adj = Adjoint adj module F = Functor F From 09b1dda55a753caae7358198dffeab977bb3b0f1 Mon Sep 17 00:00:00 2001 From: fouche Date: Tue, 10 Jan 2023 16:17:40 +0200 Subject: [PATCH 37/46] typechecka fin qui --- .../Construction/Adjunctions/Pippo.agda | 84 +++++++ .../Construction/Adjunctions/Properties.agda | 220 +++++++----------- 2 files changed, 167 insertions(+), 137 deletions(-) create mode 100644 src/Categories/Adjoint/Construction/Adjunctions/Pippo.agda diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Pippo.agda b/src/Categories/Adjoint/Construction/Adjunctions/Pippo.agda new file mode 100644 index 000000000..7d02e8f1c --- /dev/null +++ b/src/Categories/Adjoint/Construction/Adjunctions/Pippo.agda @@ -0,0 +1,84 @@ +{-# OPTIONS --without-K --allow-unsolved-metas #-} + +open import Level +open import Categories.Category.Core using (Category) +open import Categories.Category +open import Categories.Monad + +module Categories.Adjoint.Construction.Adjunctions.Pippo {o ℓ e} {C : Category (o ⊔ ℓ ⊔ e) (ℓ ⊔ e) e} (M : Monad C) where + +module C = Category C +module M = Monad M + +open import Categories.Adjoint using (Adjoint) +open import Categories.Functor using (Functor) +open import Categories.Morphism.Reasoning +open import Categories.NaturalTransformation hiding (id) +open NaturalTransformation using (η) +open import Categories.NaturalTransformation.NaturalIsomorphism using (NaturalIsomorphism; niHelper; sym; associator) + +open import Categories.Adjoint.Construction.Adjunctions M + +open import Categories.Object.Terminal (Split M) +open import Categories.Object.Initial (Split M) +open import Categories.Category.Construction.EilenbergMoore +open import Categories.Category.Construction.Kleisli +open import Categories.Adjoint.Construction.Kleisli M as KL +open import Categories.Adjoint.Construction.EilenbergMoore M as EM + + +open import Categories.Adjoint.Construction.Adjunctions.Properties M + +Kl-initial : IsInitial Kl-object +Kl-initial = record + { ! = ! + ; !-unique = λ { {A} H → niHelper ( + let module A = SplitObj A + module K = SplitObj Kl-object + module H = Split⇒ H + module Γ = NaturalIsomorphism H.Γ + module CH = C.HomReasoning + open Category A.D in + record + { η = Γ.⇐.η + ; η⁻¹ = Γ.⇒.η + ; commute = λ f → let open A.D.HomReasoning in + begin {! !} ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ A.F.F-resp-≈ ((H.μ-comp CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc))))) ⟩ + _ ≈⟨ ((refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ + _ ≈⟨ (refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _)) ⟩ + _ ≈⟨ (refl⟩∘⟨ assoc) ⟩ + _ ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ A.F.homomorphism) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ pushˡ A.D A.F.homomorphism) ⟩ + {! !} ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pushˡ A.D A.F.homomorphism) ⟩ -- (refl⟩∘⟨ (refl⟩∘⟨ A.F.homomorphism)) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ extendʳ A.D (A.adj.counit.commute _)) ⟩ -- (extendʳ A.D (A.adj.counit.commute _)) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ -- (refl⟩∘⟨ refl⟩∘⟨ elimˡ A.D A.adj.zig) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ (identityˡ ○ A.F.F-resp-≈ (elimˡ C M.F.identity))) ⟩ -- (refl⟩∘⟨ refl⟩∘⟨ A.F.F-resp-≈ {! (elimˡ C M.F.identity) !}) ⟩ + {! !} ≈⟨ {! !} ⟩ + {! !} ∎ + + -- in begin (Γ.⇐.η _ ∘ A.adj.counit.η (A.F.F₀ _) ∘ A.F.F₁ (A.GF≃M.⇐.η _ C.∘ f)) + -- ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.F.F-resp-≈ ((useful CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc ) CH.○ CH.refl⟩∘⟨ CH.refl⟩∘⟨ C.assoc ))) ⟩ + -- _ ≈⟨ ((refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _)) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ assoc) ⟩ + -- _ ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ (A.F.homomorphism ○ refl⟩∘⟨ A.F.homomorphism ○ refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ + -- _ ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ extendʳ A.D (A.adj.counit.commute _)) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.D.identityˡ) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ Γ.⇐.commute _) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ Functor.F-resp-≈ H.H (M.η.commute _) ⟩∘⟨refl) ⟩ + -- _ ≈⟨ Equiv.sym assoc ⟩ + -- _ ≈⟨ (Equiv.sym (Functor.homomorphism H.H) ⟩∘⟨refl) ⟩ + -- _ ≈⟨ (Functor.F-resp-≈ H.H (elimʳ C M.F.identity CH.⟩∘⟨refl) ⟩∘⟨refl) ⟩ + -- _ ≈⟨ (Functor.F-resp-≈ H.H C.sym-assoc ⟩∘⟨refl) ⟩ + -- _ ≈⟨ pushˡ A.D (Functor.homomorphism H.H) ⟩ + -- _ ≈⟨ (refl⟩∘⟨ elimˡ A.D (Functor.identity H.H)) ⟩ + -- Functor.F₁ H.H f ∘ Γ.⇐.η _ ∎ + ; iso = NaturalIsomorphism.iso (sym H.Γ) + }) + } + } diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda index 0fdbc6aab..1d503ab82 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda @@ -7,8 +7,9 @@ open import Categories.Monad module Categories.Adjoint.Construction.Adjunctions.Properties {o ℓ e} {C : Category (o ⊔ ℓ ⊔ e) (ℓ ⊔ e) e} (M : Monad C) where -module C = Category C -module M = Monad M +private + module C = Category C + module M = Monad M open import Categories.Adjoint using (Adjoint) open import Categories.Functor using (Functor) @@ -63,143 +64,88 @@ Kl-object = record M.F.₁ id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ id ∘ M.η.η _ ≈⟨ identityˡ ⟩ M.η.η _ ∎ - ; μ-eq = {! !} -- begin M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ - -- id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ - -- M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ - -- M.μ.η _ ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ - -- M.μ.η _ ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ - -- M.μ.η _ ∘ M.F.₁ (M.F.₁ id) ≈⟨ Equiv.sym identityʳ ⟩ - -- (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ - -- (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ≈⟨ Equiv.sym identityʳ ⟩ - -- ((M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id) ∘ id ≈⟨ assoc ⟩ - -- (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ - -- (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ - -- (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ + ; μ-eq = begin M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ + id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ + M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ + M.μ.η _ ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + M.μ.η _ ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ + M.μ.η _ ∘ M.F.₁ (M.F.₁ id) ≈⟨ Equiv.sym identityʳ ⟩ + (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ≈⟨ Equiv.sym identityʳ ⟩ + ((M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id) ∘ id ≈⟨ assoc ⟩ + (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ + (M.μ.η _ ∘ M.F.₁ (M.F.₁ id)) ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ } where open Category C open HomReasoning -Kl-initial : IsInitial Kl-object -Kl-initial = record - { ! = ! - ; !-unique = λ { {A} H → niHelper ( - let module A = SplitObj A - module K = SplitObj Kl-object - module H = Split⇒ H - -- module ! = Split⇒ ! - module Γ = NaturalIsomorphism H.Γ - module CH = C.HomReasoning - open Category A.D in - record - { η = Γ.⇐.η - ; η⁻¹ = Γ.⇒.η - ; commute = λ f → {! !} - -- let useful : ∀ {x} - -- → A.GF≃M.⇐.η x C.≈ A.G.F₁ (Γ.⇒.η x) - -- C.∘ A.G.F₁ (Functor.F₁ H.H C.id) - -- C.∘ A.G.F₁ (Γ.⇐.η (M.F.F₀ x)) - -- C.∘ A.adj.unit.η (M.F.F₀ x) - -- useful = let open C.HomReasoning in - -- begin _ ≈⟨ H.μ-comp ⟩ - -- _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ elimʳ C (M.F.identity)) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ C.identityʳ) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ C.identityˡ) ⟩ - -- _ ≈⟨ refl⟩∘⟨ C.identityˡ ⟩ - -- _ ≈⟨ (refl⟩∘⟨ (A.G.F-resp-≈ (Functor.F-resp-≈ H.H M.F.identity) ⟩∘⟨refl) ⟩∘⟨refl) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ C.identityˡ) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ (elimʳ C (elimʳ C A.G.identity)) ⟩∘⟨refl) ⟩ - -- _ ∎ - -- in let open A.D.HomReasoning - -- in begin (Γ.⇐.η _ ∘ A.adj.counit.η (A.F.F₀ _) ∘ A.F.F₁ (A.GF≃M.⇐.η _ C.∘ f)) - -- ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.F.F-resp-≈ ((useful CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc ) CH.○ CH.refl⟩∘⟨ CH.refl⟩∘⟨ C.assoc ))) ⟩ - -- _ ≈⟨ ((refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _)) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ assoc) ⟩ - -- _ ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ (A.F.homomorphism ○ refl⟩∘⟨ A.F.homomorphism ○ refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ - -- _ ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ extendʳ A.D (A.adj.counit.commute _)) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.D.identityˡ) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ Γ.⇐.commute _) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ Functor.F-resp-≈ H.H (M.η.commute _) ⟩∘⟨refl) ⟩ - -- _ ≈⟨ Equiv.sym assoc ⟩ - -- _ ≈⟨ (Equiv.sym (Functor.homomorphism H.H) ⟩∘⟨refl) ⟩ - -- _ ≈⟨ (Functor.F-resp-≈ H.H (elimʳ C M.F.identity CH.⟩∘⟨refl) ⟩∘⟨refl) ⟩ - -- _ ≈⟨ (Functor.F-resp-≈ H.H C.sym-assoc ⟩∘⟨refl) ⟩ - -- _ ≈⟨ pushˡ A.D (Functor.homomorphism H.H) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ elimˡ A.D (Functor.identity H.H)) ⟩ - -- Functor.F₁ H.H f ∘ Γ.⇐.η _ ∎ - ; iso = NaturalIsomorphism.iso (sym H.Γ) +! : {A : SplitObj} → Split⇒ Kl-object A +! {splitobj D F G adj GF≃M η-eq μ-eq} = record + { H = + let open D + open D.HomReasoning in + record + { F₀ = F.₀ + ; F₁ = λ f → adj.counit.η _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) + ; identity = λ {A} → + begin + adj.counit.η (F.₀ A) ∘ F.₁ (GF≃M.⇐.η A C.∘ M.η.η A) ≈⟨ refl⟩∘⟨ F.F-resp-≈ η-eq ⟩ + adj.counit.η (F.₀ A) ∘ F.₁ (adj.unit.η A) ≈⟨ adj.zig ⟩ + D.id ∎ + ; homomorphism = λ {X} {Y} {Z} {f} {g} → + let ε x = adj.counit.η x in + begin + ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ (M.μ.η _ C.∘ M.F.₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (assoc²'' C) ⟩ + ε _ ∘ F.₁ ((GF≃M.⇐.η _ C.∘ M.μ.η _) C.∘ M.F.₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (μ-eq CHR.⟩∘⟨refl) ⟩ + ε _ ∘ F.₁ ((G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.F.₁ (GF≃M.⇐.η _)) C.∘ M.F.₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (assoc²' C) ⟩ + ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.F.₁ (GF≃M.⇐.η _) C.∘ M.F.₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ CHR.refl⟩∘⟨ pullˡ C (C.Equiv.sym M.F.homomorphism)) ⟩ + ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.F.₁ (GF≃M.⇐.η _ C.∘ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ extendʳ C (GF≃M.⇐.commute _)) ⟩ + ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ F.homomorphism ⟩ + ε _ ∘ F.₁ (G.₁ (ε (F.₀ _))) ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ DMR.extendʳ (adj.counit.commute _) ⟩ + ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.homomorphism ⟩ + ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g))) ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ DMR.extendʳ (adj.counit.commute _) ⟩ + ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ sym-assoc ⟩ + (ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g)) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ∎ + ; F-resp-≈ = λ x → D.∘-resp-≈ʳ (F.F-resp-≈ (C.∘-resp-≈ʳ x)) + } + ; Γ = + let open D + open D.HomReasoning in + niHelper (record + { η = λ _ → D.id + ; η⁻¹ = λ _ → D.id + ; commute = λ f → + begin + D.id ∘ adj.counit.η _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ M.η.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C η-eq) ⟩ + D.id ∘ adj.counit.η _ ∘ F.₁ (adj.unit.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.homomorphism ⟩ + D.id ∘ adj.counit.η _ ∘ F.₁ (adj.unit.η _) ∘ F.₁ f ≈⟨ refl⟩∘⟨ pullˡ D adj.zig ⟩ + D.id ∘ D.id ∘ F.₁ f ≈⟨ identityˡ ⟩ + D.id ∘ F.₁ f ≈⟨ identityˡ ⟩ + F.₁ f ≈⟨ D.Equiv.sym identityʳ ⟩ + (F.₁ f ∘ D.id) ∎ + ; iso = λ X → record { isoˡ = identityˡ ; isoʳ = identityˡ } }) - } - } - where - ! : {A : SplitObj} → Split⇒ Kl-object A - ! {splitobj D F G adj GF≃M η-eq μ-eq} = record - { H = - let open D - open D.HomReasoning in - record - { F₀ = F.₀ - ; F₁ = λ f → adj.counit.η _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) - ; identity = λ {A} → - begin - adj.counit.η (F.₀ A) ∘ F.₁ (GF≃M.⇐.η A C.∘ M.η.η A) ≈⟨ refl⟩∘⟨ F.F-resp-≈ η-eq ⟩ - adj.counit.η (F.₀ A) ∘ F.₁ (adj.unit.η A) ≈⟨ adj.zig ⟩ - D.id ∎ - ; homomorphism = λ {X} {Y} {Z} {f} {g} → - let ε x = adj.counit.η x in {! !} - -- begin - -- ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ (M.μ.η _ C.∘ M.F.₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (assoc²'' C) ⟩ - -- ε _ ∘ F.₁ ((GF≃M.⇐.η _ C.∘ M.μ.η _) C.∘ M.F.₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (μ-eq CHR.⟩∘⟨refl) ⟩ - -- ε _ ∘ F.₁ ((G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.F.₁ (GF≃M.⇐.η _)) C.∘ M.F.₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (assoc²' C) ⟩ - -- ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.F.₁ (GF≃M.⇐.η _) C.∘ M.F.₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ CHR.refl⟩∘⟨ pullˡ C (C.Equiv.sym M.F.homomorphism)) ⟩ - -- ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ GF≃M.⇐.η _ C.∘ M.F.₁ (GF≃M.⇐.η _ C.∘ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (CHR.refl⟩∘⟨ extendʳ C (GF≃M.⇐.commute _)) ⟩ - -- ε _ ∘ F.₁ (G.₁ (ε (F.₀ _)) C.∘ G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ F.homomorphism ⟩ - -- ε _ ∘ F.₁ (G.₁ (ε (F.₀ _))) ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ DMR.extendʳ (adj.counit.commute _) ⟩ - -- ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g)) C.∘ GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.homomorphism ⟩ - -- ε _ ∘ ε _ ∘ F.₁ (G.₁ (F.₁ (GF≃M.⇐.η _ C.∘ g))) ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ refl⟩∘⟨ DMR.extendʳ (adj.counit.commute _) ⟩ - -- ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ≈⟨ sym-assoc ⟩ - -- (ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ g)) ∘ ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ f) ∎ - ; F-resp-≈ = λ x → D.∘-resp-≈ʳ (F.F-resp-≈ (C.∘-resp-≈ʳ x)) - } - ; Γ = - let open D - open D.HomReasoning in - niHelper (record - { η = λ _ → D.id - ; η⁻¹ = λ _ → D.id - ; commute = λ f → {! !} - -- begin - -- D.id ∘ adj.counit.η _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ M.η.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C η-eq) ⟩ - -- D.id ∘ adj.counit.η _ ∘ F.₁ (adj.unit.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.homomorphism ⟩ - -- D.id ∘ adj.counit.η _ ∘ F.₁ (adj.unit.η _) ∘ F.₁ f ≈⟨ refl⟩∘⟨ pullˡ D adj.zig ⟩ - -- D.id ∘ D.id ∘ F.₁ f ≈⟨ identityˡ ⟩ - -- D.id ∘ F.₁ f ≈⟨ identityˡ ⟩ - -- F.₁ f ≈⟨ D.Equiv.sym identityʳ ⟩ - -- (F.₁ f ∘ D.id) ∎ - ; iso = λ X → record { isoˡ = identityˡ ; isoʳ = identityˡ } - }) - ; μ-comp = λ { {x} → let open C - open C.HomReasoning in - Equiv.sym - (begin {! !} ≈⟨ elimˡ C G.identity ⟩ - {! !} ≈⟨ elimʳ C M.F.identity ⟩ - {! !} ≈⟨ (refl⟩∘⟨ elimˡ C G.identity) ⟩ - {! !} ≈⟨ (G.homomorphism ⟩∘⟨refl) ⟩ - {! !} ≈⟨ pullʳ C (adj.unit.sym-commute _) ⟩ - {! !} ≈⟨ sym-assoc ⟩ - {! !} ≈⟨ elimˡ C adj.zag ⟩ - {! !} ≈⟨ elimʳ C M.F.identity ⟩ - {! !} ∎) } - } where - module adj = Adjoint adj - module F = Functor F - module G = Functor G - module GF≃M = NaturalIsomorphism GF≃M - module D = Category D - module CHR = C.HomReasoning - module DHR = D.HomReasoning - module DMR = Categories.Morphism.Reasoning D - module CMR = Categories.Morphism.Reasoning C + ; μ-comp = λ { {x} → let open C + open C.HomReasoning in + Equiv.sym + (begin _ ≈⟨ elimˡ C G.identity ⟩ + _ ≈⟨ elimʳ C M.F.identity ⟩ + _ ≈⟨ (refl⟩∘⟨ elimˡ C G.identity) ⟩ + _ ≈⟨ (G.homomorphism ⟩∘⟨refl) ⟩ + _ ≈⟨ pullʳ C (adj.unit.sym-commute _) ⟩ + _ ≈⟨ sym-assoc ⟩ + _ ≈⟨ elimˡ C adj.zag ⟩ + _ ≈⟨ elimʳ C M.F.identity ⟩ + _ ∎) } + } where + module adj = Adjoint adj + module F = Functor F + module G = Functor G + module GF≃M = NaturalIsomorphism GF≃M + module D = Category D + module CHR = C.HomReasoning + module DHR = D.HomReasoning + module DMR = Categories.Morphism.Reasoning D + module CMR = Categories.Morphism.Reasoning C + From 3c1a4162d8e3876d077611c929caa1539da0a973 Mon Sep 17 00:00:00 2001 From: fouche Date: Tue, 10 Jan 2023 16:31:38 +0200 Subject: [PATCH 38/46] KL-initial --- .../Adjoint/Construction/Adjunctions.agda | 4 +- .../Construction/Adjunctions/Pippo.agda | 52 ++++++------------- .../Construction/Adjunctions/Properties.agda | 35 +++++++++++++ 3 files changed, 54 insertions(+), 37 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 518bce204..5b2f81e28 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -73,8 +73,8 @@ record Split⇒ (X Y : SplitObj) : Set (suc o ⊔ suc ℓ ⊔ suc e) where H : Functor X.D Y.D Γ : H ∘F X.F ≃ Y.F - private - module Γ = NaturalIsomorphism Γ + module Γ = NaturalIsomorphism Γ + module H = Functor H field μ-comp : ∀ {x : Obj} → Y.GF≃M.⇐.η x ≈ Y.G.F₁ (Γ.⇒.η x) ∘ ((Y.G.F₁ (Functor.F₁ H (X.adj.counit.η (X.F.F₀ x)))) ∘ Y.G.F₁ (Γ.⇐.η (X.G.F₀ (X.F.F₀ x))) ∘ (Y.adj.unit.η (X.G.F₀ (X.F.F₀ x)))) ∘ X.GF≃M.⇐.η x diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Pippo.agda b/src/Categories/Adjoint/Construction/Adjunctions/Pippo.agda index 7d02e8f1c..dfeeab8ee 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions/Pippo.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions/Pippo.agda @@ -43,41 +43,23 @@ Kl-initial = record { η = Γ.⇐.η ; η⁻¹ = Γ.⇒.η ; commute = λ f → let open A.D.HomReasoning in - begin {! !} ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ A.F.F-resp-≈ ((H.μ-comp CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc))))) ⟩ - _ ≈⟨ ((refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ - _ ≈⟨ (refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _)) ⟩ - _ ≈⟨ (refl⟩∘⟨ assoc) ⟩ - _ ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ A.F.homomorphism) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ pushˡ A.D A.F.homomorphism) ⟩ - {! !} ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pushˡ A.D A.F.homomorphism) ⟩ -- (refl⟩∘⟨ (refl⟩∘⟨ A.F.homomorphism)) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ extendʳ A.D (A.adj.counit.commute _)) ⟩ -- (extendʳ A.D (A.adj.counit.commute _)) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ -- (refl⟩∘⟨ refl⟩∘⟨ elimˡ A.D A.adj.zig) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ (identityˡ ○ A.F.F-resp-≈ (elimˡ C M.F.identity))) ⟩ -- (refl⟩∘⟨ refl⟩∘⟨ A.F.F-resp-≈ {! (elimˡ C M.F.identity) !}) ⟩ - {! !} ≈⟨ {! !} ⟩ - {! !} ∎ - - -- in begin (Γ.⇐.η _ ∘ A.adj.counit.η (A.F.F₀ _) ∘ A.F.F₁ (A.GF≃M.⇐.η _ C.∘ f)) - -- ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.F.F-resp-≈ ((useful CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc ) CH.○ CH.refl⟩∘⟨ CH.refl⟩∘⟨ C.assoc ))) ⟩ - -- _ ≈⟨ ((refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _)) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ assoc) ⟩ - -- _ ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ (A.F.homomorphism ○ refl⟩∘⟨ A.F.homomorphism ○ refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ - -- _ ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ extendʳ A.D (A.adj.counit.commute _)) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ A.D.identityˡ) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ Γ.⇐.commute _) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ Functor.F-resp-≈ H.H (M.η.commute _) ⟩∘⟨refl) ⟩ - -- _ ≈⟨ Equiv.sym assoc ⟩ - -- _ ≈⟨ (Equiv.sym (Functor.homomorphism H.H) ⟩∘⟨refl) ⟩ - -- _ ≈⟨ (Functor.F-resp-≈ H.H (elimʳ C M.F.identity CH.⟩∘⟨refl) ⟩∘⟨refl) ⟩ - -- _ ≈⟨ (Functor.F-resp-≈ H.H C.sym-assoc ⟩∘⟨refl) ⟩ - -- _ ≈⟨ pushˡ A.D (Functor.homomorphism H.H) ⟩ - -- _ ≈⟨ (refl⟩∘⟨ elimˡ A.D (Functor.identity H.H)) ⟩ - -- Functor.F₁ H.H f ∘ Γ.⇐.η _ ∎ + begin _ ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ A.F.F-resp-≈ ((H.μ-comp CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc))))) ⟩ + _ ≈⟨ ((refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ + _ ≈⟨ (refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _)) ⟩ + _ ≈⟨ (refl⟩∘⟨ assoc) ⟩ + _ ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ + _ ≈⟨ (refl⟩∘⟨ A.F.homomorphism) ⟩ + _ ≈⟨ (refl⟩∘⟨ pushˡ A.D A.F.homomorphism) ⟩ + _ ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ + _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pushˡ A.D A.F.homomorphism) ⟩ -- (refl⟩∘⟨ (refl⟩∘⟨ A.F.homomorphism)) ⟩ + _ ≈⟨ (refl⟩∘⟨ extendʳ A.D (A.adj.counit.commute _)) ⟩ -- (extendʳ A.D (A.adj.counit.commute _)) ⟩ + _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ -- (refl⟩∘⟨ refl⟩∘⟨ elimˡ A.D A.adj.zig) ⟩ + _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ (identityˡ ○ A.F.F-resp-≈ (elimˡ C M.F.identity))) ⟩ -- (refl⟩∘⟨ refl⟩∘⟨ A.F.F-resp-≈ {! (elimˡ C M.F.identity) !}) ⟩ + _ ≈⟨ (refl⟩∘⟨ Γ.⇐.commute f) ⟩ + _ ≈⟨ pullˡ A.D (Equiv.sym H.H.homomorphism) ⟩ + _ ≈⟨ (H.H.F-resp-≈ (elimʳ C (M.F.F-resp-≈ M.F.identity CH.○ M.F.identity) CH.⟩∘⟨refl) ⟩∘⟨refl) ⟩ + _ ≈⟨ (H.H.F-resp-≈ (pullˡ C M.identityʳ CH.○ C.identityˡ) ⟩∘⟨refl) ⟩ + _ ∎ ; iso = NaturalIsomorphism.iso (sym H.Γ) }) } diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda index 1d503ab82..90c47bb70 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda @@ -149,3 +149,38 @@ Kl-object = record module DMR = Categories.Morphism.Reasoning D module CMR = Categories.Morphism.Reasoning C +Kl-initial : IsInitial Kl-object +Kl-initial = record + { ! = ! + ; !-unique = λ { {A} H → niHelper ( + let module A = SplitObj A + module K = SplitObj Kl-object + module H = Split⇒ H + module Γ = NaturalIsomorphism H.Γ + module CH = C.HomReasoning + open Category A.D in + record + { η = Γ.⇐.η + ; η⁻¹ = Γ.⇒.η + ; commute = λ f → let open A.D.HomReasoning in + begin _ ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ A.F.F-resp-≈ ((H.μ-comp CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc))))) ⟩ + _ ≈⟨ ((refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ + _ ≈⟨ (refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _)) ⟩ + _ ≈⟨ (refl⟩∘⟨ assoc) ⟩ + _ ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ + _ ≈⟨ (refl⟩∘⟨ A.F.homomorphism) ⟩ + _ ≈⟨ (refl⟩∘⟨ pushˡ A.D A.F.homomorphism) ⟩ + _ ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ + _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pushˡ A.D A.F.homomorphism) ⟩ -- (refl⟩∘⟨ (refl⟩∘⟨ A.F.homomorphism)) ⟩ + _ ≈⟨ (refl⟩∘⟨ extendʳ A.D (A.adj.counit.commute _)) ⟩ -- (extendʳ A.D (A.adj.counit.commute _)) ⟩ + _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ -- (refl⟩∘⟨ refl⟩∘⟨ elimˡ A.D A.adj.zig) ⟩ + _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ (identityˡ ○ A.F.F-resp-≈ (elimˡ C M.F.identity))) ⟩ -- (refl⟩∘⟨ refl⟩∘⟨ A.F.F-resp-≈ {! (elimˡ C M.F.identity) !}) ⟩ + _ ≈⟨ (refl⟩∘⟨ Γ.⇐.commute f) ⟩ + _ ≈⟨ pullˡ A.D (Equiv.sym H.H.homomorphism) ⟩ + _ ≈⟨ (H.H.F-resp-≈ (elimʳ C (M.F.F-resp-≈ M.F.identity CH.○ M.F.identity) CH.⟩∘⟨refl) ⟩∘⟨refl) ⟩ + _ ≈⟨ (H.H.F-resp-≈ (pullˡ C M.identityʳ CH.○ C.identityˡ) ⟩∘⟨refl) ⟩ + _ ∎ + ; iso = NaturalIsomorphism.iso (sym H.Γ) + }) + } + } \ No newline at end of file From 72e3bea9662bdfc05a9fb43f7c294ee7e9bf31f3 Mon Sep 17 00:00:00 2001 From: fouche Date: Tue, 10 Jan 2023 16:32:11 +0200 Subject: [PATCH 39/46] remove push of pippo --- .../Construction/Adjunctions/Pippo.agda | 66 ------------------- 1 file changed, 66 deletions(-) delete mode 100644 src/Categories/Adjoint/Construction/Adjunctions/Pippo.agda diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Pippo.agda b/src/Categories/Adjoint/Construction/Adjunctions/Pippo.agda deleted file mode 100644 index dfeeab8ee..000000000 --- a/src/Categories/Adjoint/Construction/Adjunctions/Pippo.agda +++ /dev/null @@ -1,66 +0,0 @@ -{-# OPTIONS --without-K --allow-unsolved-metas #-} - -open import Level -open import Categories.Category.Core using (Category) -open import Categories.Category -open import Categories.Monad - -module Categories.Adjoint.Construction.Adjunctions.Pippo {o ℓ e} {C : Category (o ⊔ ℓ ⊔ e) (ℓ ⊔ e) e} (M : Monad C) where - -module C = Category C -module M = Monad M - -open import Categories.Adjoint using (Adjoint) -open import Categories.Functor using (Functor) -open import Categories.Morphism.Reasoning -open import Categories.NaturalTransformation hiding (id) -open NaturalTransformation using (η) -open import Categories.NaturalTransformation.NaturalIsomorphism using (NaturalIsomorphism; niHelper; sym; associator) - -open import Categories.Adjoint.Construction.Adjunctions M - -open import Categories.Object.Terminal (Split M) -open import Categories.Object.Initial (Split M) -open import Categories.Category.Construction.EilenbergMoore -open import Categories.Category.Construction.Kleisli -open import Categories.Adjoint.Construction.Kleisli M as KL -open import Categories.Adjoint.Construction.EilenbergMoore M as EM - - -open import Categories.Adjoint.Construction.Adjunctions.Properties M - -Kl-initial : IsInitial Kl-object -Kl-initial = record - { ! = ! - ; !-unique = λ { {A} H → niHelper ( - let module A = SplitObj A - module K = SplitObj Kl-object - module H = Split⇒ H - module Γ = NaturalIsomorphism H.Γ - module CH = C.HomReasoning - open Category A.D in - record - { η = Γ.⇐.η - ; η⁻¹ = Γ.⇒.η - ; commute = λ f → let open A.D.HomReasoning in - begin _ ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ A.F.F-resp-≈ ((H.μ-comp CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc))))) ⟩ - _ ≈⟨ ((refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ - _ ≈⟨ (refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _)) ⟩ - _ ≈⟨ (refl⟩∘⟨ assoc) ⟩ - _ ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ - _ ≈⟨ (refl⟩∘⟨ A.F.homomorphism) ⟩ - _ ≈⟨ (refl⟩∘⟨ pushˡ A.D A.F.homomorphism) ⟩ - _ ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ - _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pushˡ A.D A.F.homomorphism) ⟩ -- (refl⟩∘⟨ (refl⟩∘⟨ A.F.homomorphism)) ⟩ - _ ≈⟨ (refl⟩∘⟨ extendʳ A.D (A.adj.counit.commute _)) ⟩ -- (extendʳ A.D (A.adj.counit.commute _)) ⟩ - _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ -- (refl⟩∘⟨ refl⟩∘⟨ elimˡ A.D A.adj.zig) ⟩ - _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ (identityˡ ○ A.F.F-resp-≈ (elimˡ C M.F.identity))) ⟩ -- (refl⟩∘⟨ refl⟩∘⟨ A.F.F-resp-≈ {! (elimˡ C M.F.identity) !}) ⟩ - _ ≈⟨ (refl⟩∘⟨ Γ.⇐.commute f) ⟩ - _ ≈⟨ pullˡ A.D (Equiv.sym H.H.homomorphism) ⟩ - _ ≈⟨ (H.H.F-resp-≈ (elimʳ C (M.F.F-resp-≈ M.F.identity CH.○ M.F.identity) CH.⟩∘⟨refl) ⟩∘⟨refl) ⟩ - _ ≈⟨ (H.H.F-resp-≈ (pullˡ C M.identityʳ CH.○ C.identityˡ) ⟩∘⟨refl) ⟩ - _ ∎ - ; iso = NaturalIsomorphism.iso (sym H.Γ) - }) - } - } From a576da1193001d5693579da02c395126eae57704 Mon Sep 17 00:00:00 2001 From: fouche Date: Tue, 10 Jan 2023 16:37:22 +0200 Subject: [PATCH 40/46] just a comp from the end --- src/Categories/Adjoint/Construction/Adjunctions.agda | 1 - 1 file changed, 1 deletion(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 5b2f81e28..2f6ca3585 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -117,7 +117,6 @@ Split M = record {! !} ≈⟨ (refl⟩∘⟨ ((refl⟩∘⟨ (X.G.F-resp-≈ (X.D.identityˡ X.D.HomReasoning.⟩∘⟨refl) ⟩∘⟨refl)) ⟩∘⟨refl)) ⟩ {! !} ≈⟨ pushˡ C X.G.homomorphism ⟩ {! !} ≈⟨ {! !} ⟩ - {! !} ≈⟨ {! !} ⟩ {! !} ∎) } -- Equiv.sym (begin {! !} ≈⟨ ( Functor.homomorphism X.G ⟩∘⟨refl) ⟩ -- {! !} ≈⟨ ((refl⟩∘⟨ Functor.F-resp-≈ X.G X.D.identityʳ) ⟩∘⟨refl) ⟩ From d2cf221a169ee6580afec419c4aea5b3cd4b4054 Mon Sep 17 00:00:00 2001 From: iwilare Date: Tue, 10 Jan 2023 21:14:14 +0200 Subject: [PATCH 41/46] Minima --- .../Adjoint/Construction/Adjunctions.agda | 67 +++++++++++++------ 1 file changed, 46 insertions(+), 21 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 2f6ca3585..a5cf08eb2 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -100,10 +100,10 @@ Split M = record split-id {A} = record { H = Categories.Functor.id ; Γ = unitorˡ - ; μ-comp = λ { {x} → - Equiv.sym(begin _ ≈⟨ elimˡ C (Functor.identity A.G) ⟩ - _ ≈⟨ (refl⟩∘⟨ elimˡ C (Functor.identity A.G)) ⟩∘⟨refl ⟩ - _ ≈⟨ elimˡ C A.adj.zag ⟩ + ; μ-comp = λ { {x} → + Equiv.sym(begin _ ≈⟨ elimˡ C (Functor.identity A.G) ⟩ + _ ≈⟨ (refl⟩∘⟨ elimˡ C (Functor.identity A.G)) ⟩∘⟨refl ⟩ + _ ≈⟨ elimˡ C A.adj.zag ⟩ _ ∎)} } where module A = SplitObj A open C @@ -112,27 +112,52 @@ Split M = record comp {A = A} {B = B} {X = X} (split⇒ Hᵤ Γᵤ Aμ-comp) (split⇒ Hᵥ Γᵥ Bμ-comp) = record { H = Hᵤ ∘F Hᵥ ; Γ = Γᵤ ⓘᵥ (Hᵤ ⓘˡ Γᵥ) ⓘᵥ associator (SplitObj.F A) Hᵥ Hᵤ - ; μ-comp = λ { {x} → - Equiv.sym (begin {! !} ≈⟨ (X.G.F-resp-≈ (X.D.HomReasoning.refl⟩∘⟨ X.D.identityʳ) ⟩∘⟨refl) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ ((refl⟩∘⟨ (X.G.F-resp-≈ (X.D.identityˡ X.D.HomReasoning.⟩∘⟨refl) ⟩∘⟨refl)) ⟩∘⟨refl)) ⟩ - {! !} ≈⟨ pushˡ C X.G.homomorphism ⟩ - {! !} ≈⟨ {! !} ⟩ + ; μ-comp = λ { {x} → + Equiv.sym (begin {! !} ≈⟨ (X.G.F-resp-≈ (X.D.HomReasoning.refl⟩∘⟨ X.D.identityʳ) ⟩∘⟨refl) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ ((refl⟩∘⟨ (X.G.F-resp-≈ (X.D.identityˡ X.D.HomReasoning.⟩∘⟨refl) ⟩∘⟨refl)) ⟩∘⟨refl)) ⟩ + {! !} ≈⟨ pushˡ C X.G.homomorphism ⟩ + {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ assoc) ⟩ + {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ Equiv.sym Bμ-comp) ⟩ + {! !} ≈⟨ Equiv.sym Aμ-comp ⟩ {! !} ∎) } - -- Equiv.sym (begin {! !} ≈⟨ ( Functor.homomorphism X.G ⟩∘⟨refl) ⟩ - -- {! !} ≈⟨ ((refl⟩∘⟨ Functor.F-resp-≈ X.G X.D.identityʳ) ⟩∘⟨refl) ⟩ - -- {! !} ≈⟨ (refl⟩∘⟨ (identityˡ ○ (identityˡ ⟩∘⟨refl))) ⟩ - -- {! !} ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ (refl⟩∘⟨ (identityʳ ○ identityˡ))) ⟩∘⟨refl) ⟩ - -- {! !} ≈⟨ refl⟩∘⟨ ((refl⟩∘⟨ (Functor.homomorphism X.G ⟩∘⟨refl)) ⟩∘⟨refl) ⟩ - -- {! !} ≈⟨ (refl⟩∘⟨ (((refl⟩∘⟨ (identityˡ ○ Functor.identity X.G)) ⟩∘⟨refl) ⟩∘⟨refl)) ⟩ - -- {! !} ≈⟨ (refl⟩∘⟨ (identityʳ ⟩∘⟨refl) ⟩∘⟨refl) ⟩ - -- {! !} ≈⟨ (refl⟩∘⟨ ((refl⟩∘⟨ ((Functor.homomorphism X.G ⟩∘⟨refl) ⟩∘⟨refl)) ⟩∘⟨refl)) ⟩ - -- {! !} ≈⟨ (refl⟩∘⟨ ((refl⟩∘⟨ ((elimˡ C (Functor.identity X.G) ⟩∘⟨refl) ⟩∘⟨refl)) ⟩∘⟨refl)) ⟩ - -- {! !} ≈⟨ {! !} ⟩ - -- {! !} ≈⟨ {! !} ⟩ - -- {! !} ∎)} + + +{- +Have + +X.G.F₁ (Γᵤ.⇒.η x) ∘ + X.G.F₁ (Hᵤ.F₁ (Γᵥ.⇒.η x)) ∘ + X.G.F₁ (Hᵤ.F₁ (Hᵥ.F₁ (A.adj.counit.η (A.F.F₀ x)))) ∘ + (X.G.F₁ + (Hᵤ.F₁ (Γᵥ.⇐.η (A.G.F₀ (A.F.F₀ x))) X.D.∘ + Γᵤ.⇐.η (A.G.F₀ (A.F.F₀ x))) + ∘ X.adj.unit.η (A.G.F₀ (A.F.F₀ x))) + ∘ A.GF≃M.⇐.η x + +Goal + + X.G.F₁ (Γᵤ.⇒.η x) ∘ + (X.G.F₁ (Hᵤ.F₁ (B.adj.counit.η (B.F.F₀ x))) ∘ + X.G.F₁ (Γᵤ.⇐.η (B.G.F₀ (B.F.F₀ x))) ∘ + X.adj.unit.η (B.G.F₀ (B.F.F₀ x))) + ∘ + B.G.F₁ (Γᵥ.⇒.η x) ∘ + (B.G.F₁ (Hᵥ.F₁ (A.adj.counit.η (A.F.F₀ x))) ∘ + B.G.F₁ (Γᵥ.⇐.η (A.G.F₀ (A.F.F₀ x))) ∘ + B.adj.unit.η (A.G.F₀ (A.F.F₀ x))) + ∘ A.GF≃M.⇐.η x +-} } where module A = SplitObj A module B = SplitObj B module X = SplitObj X open C open C.HomReasoning + module Hᵤ = Functor Hᵤ + module Hᵥ = Functor Hᵥ + module Γᵤ = NaturalIsomorphism Γᵤ + module Γᵥ = NaturalIsomorphism Γᵥ From 27dd32d4f19086fdcabdfcd3844b131836546b9f Mon Sep 17 00:00:00 2001 From: fouche Date: Wed, 11 Jan 2023 11:21:25 +0200 Subject: [PATCH 42/46] smol homomorphism --- src/Categories/Adjoint/Construction/Adjunctions.agda | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index a5cf08eb2..4439faf70 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -117,8 +117,8 @@ Split M = record {! !} ≈⟨ (refl⟩∘⟨ ((refl⟩∘⟨ (X.G.F-resp-≈ (X.D.identityˡ X.D.HomReasoning.⟩∘⟨refl) ⟩∘⟨refl)) ⟩∘⟨refl)) ⟩ {! !} ≈⟨ pushˡ C X.G.homomorphism ⟩ {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ assoc) ⟩ - {! !} ≈⟨ {! !} ⟩ - {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ (refl⟩∘⟨ ((X.G.homomorphism ⟩∘⟨refl) ⟩∘⟨refl))) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ {! A.GF≃M.⇐.η x !}) ⟩ {! !} ≈⟨ {! !} ⟩ {! !} ≈⟨ {! !} ⟩ {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ Equiv.sym Bμ-comp) ⟩ From 7712bb56dd6e870135c3d1a40cf5259df7f6c3c8 Mon Sep 17 00:00:00 2001 From: iwilare Date: Wed, 11 Jan 2023 14:08:57 +0200 Subject: [PATCH 43/46] Style fixes --- agda-categories.agda-lib | 2 +- .../Adjoint/Construction/Adjunctions.agda | 17 +++-- .../Construction/Adjunctions/Properties.agda | 76 ++++++++++--------- 3 files changed, 50 insertions(+), 45 deletions(-) diff --git a/agda-categories.agda-lib b/agda-categories.agda-lib index 3f89afb89..186e350b1 100644 --- a/agda-categories.agda-lib +++ b/agda-categories.agda-lib @@ -1,3 +1,3 @@ name: agda-categories -depend: standard-library-2.0 +depend: standard-library-1.7.1 include: src/ diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 4439faf70..90c449148 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -77,10 +77,15 @@ record Split⇒ (X Y : SplitObj) : Set (suc o ⊔ suc ℓ ⊔ suc e) where module H = Functor H field - μ-comp : ∀ {x : Obj} → Y.GF≃M.⇐.η x ≈ Y.G.F₁ (Γ.⇒.η x) ∘ ((Y.G.F₁ (Functor.F₁ H (X.adj.counit.η (X.F.F₀ x)))) ∘ Y.G.F₁ (Γ.⇐.η (X.G.F₀ (X.F.F₀ x))) ∘ (Y.adj.unit.η (X.G.F₀ (X.F.F₀ x)))) ∘ X.GF≃M.⇐.η x - -Split : Monad C → Category _ _ _ -Split M = record + μ-comp : ∀ {x : Obj} → + Y.GF≃M.⇐.η x ≈ Y.G.F₁ (Γ.⇒.η x) + ∘ ((Y.G.F₁ (Functor.F₁ H (X.adj.counit.η (X.F.F₀ x)))) + ∘ Y.G.F₁ (Γ.⇐.η (X.G.F₀ (X.F.F₀ x))) + ∘ Y.adj.unit.η (X.G.F₀ (X.F.F₀ x))) + ∘ X.GF≃M.⇐.η x + +Split : Category _ _ _ +Split = record { Obj = SplitObj ; _⇒_ = Split⇒ ; _≈_ = λ U V → Split⇒.H U ≃ Split⇒.H V @@ -118,14 +123,12 @@ Split M = record {! !} ≈⟨ pushˡ C X.G.homomorphism ⟩ {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ assoc) ⟩ {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ (refl⟩∘⟨ ((X.G.homomorphism ⟩∘⟨refl) ⟩∘⟨refl))) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ {! A.GF≃M.⇐.η x !}) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ {! !}) ⟩ {! !} ≈⟨ {! !} ⟩ {! !} ≈⟨ {! !} ⟩ {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ Equiv.sym Bμ-comp) ⟩ {! !} ≈⟨ Equiv.sym Aμ-comp ⟩ {! !} ∎) } - - {- Have diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda index 90c47bb70..bc8f8306e 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda @@ -20,8 +20,8 @@ open import Categories.NaturalTransformation.NaturalIsomorphism using (NaturalIs open import Categories.Adjoint.Construction.Adjunctions M -open import Categories.Object.Terminal (Split M) -open import Categories.Object.Initial (Split M) +open import Categories.Object.Terminal Split +open import Categories.Object.Initial Split open import Categories.Category.Construction.EilenbergMoore open import Categories.Category.Construction.Kleisli open import Categories.Adjoint.Construction.Kleisli M as KL @@ -64,9 +64,9 @@ Kl-object = record M.F.₁ id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ id ∘ M.η.η _ ≈⟨ identityˡ ⟩ M.η.η _ ∎ - ; μ-eq = begin M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ - id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ - M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ + ; μ-eq = begin + M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ + id ∘ M.μ.η _ ≈⟨ id-comm-sym C ⟩ M.μ.η _ ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ M.μ.η _ ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ M.μ.η _ ∘ M.F.₁ (M.F.₁ id) ≈⟨ Equiv.sym identityʳ ⟩ @@ -94,7 +94,7 @@ Kl-object = record adj.counit.η (F.₀ A) ∘ F.₁ (adj.unit.η A) ≈⟨ adj.zig ⟩ D.id ∎ ; homomorphism = λ {X} {Y} {Z} {f} {g} → - let ε x = adj.counit.η x in + let ε x = adj.counit.η x in begin ε _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ (M.μ.η _ C.∘ M.F.₁ g) C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (assoc²'' C) ⟩ ε _ ∘ F.₁ ((GF≃M.⇐.η _ C.∘ M.μ.η _) C.∘ M.F.₁ g C.∘ f) ≈⟨ refl⟩∘⟨ F.F-resp-≈ (μ-eq CHR.⟩∘⟨refl) ⟩ @@ -115,7 +115,7 @@ Kl-object = record niHelper (record { η = λ _ → D.id ; η⁻¹ = λ _ → D.id - ; commute = λ f → + ; commute = λ f → begin D.id ∘ adj.counit.η _ ∘ F.₁ (GF≃M.⇐.η _ C.∘ M.η.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.F-resp-≈ (pullˡ C η-eq) ⟩ D.id ∘ adj.counit.η _ ∘ F.₁ (adj.unit.η _ C.∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F.homomorphism ⟩ @@ -127,17 +127,18 @@ Kl-object = record ; iso = λ X → record { isoˡ = identityˡ ; isoʳ = identityˡ } }) ; μ-comp = λ { {x} → let open C - open C.HomReasoning in - Equiv.sym - (begin _ ≈⟨ elimˡ C G.identity ⟩ - _ ≈⟨ elimʳ C M.F.identity ⟩ - _ ≈⟨ (refl⟩∘⟨ elimˡ C G.identity) ⟩ - _ ≈⟨ (G.homomorphism ⟩∘⟨refl) ⟩ - _ ≈⟨ pullʳ C (adj.unit.sym-commute _) ⟩ - _ ≈⟨ sym-assoc ⟩ - _ ≈⟨ elimˡ C adj.zag ⟩ - _ ≈⟨ elimʳ C M.F.identity ⟩ - _ ∎) } + open C.HomReasoning in + Equiv.sym + (begin + _ ≈⟨ elimˡ C G.identity ⟩ + _ ≈⟨ elimʳ C M.F.identity ⟩ + _ ≈⟨ (refl⟩∘⟨ elimˡ C G.identity) ⟩ + _ ≈⟨ (G.homomorphism ⟩∘⟨refl) ⟩ + _ ≈⟨ pullʳ C (adj.unit.sym-commute _) ⟩ + _ ≈⟨ sym-assoc ⟩ + _ ≈⟨ elimˡ C adj.zag ⟩ + _ ≈⟨ elimʳ C M.F.identity ⟩ + _ ∎) } } where module adj = Adjoint adj module F = Functor F @@ -162,25 +163,26 @@ Kl-initial = record record { η = Γ.⇐.η ; η⁻¹ = Γ.⇒.η - ; commute = λ f → let open A.D.HomReasoning in - begin _ ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ A.F.F-resp-≈ ((H.μ-comp CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc))))) ⟩ - _ ≈⟨ ((refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ - _ ≈⟨ (refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _)) ⟩ - _ ≈⟨ (refl⟩∘⟨ assoc) ⟩ - _ ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ - _ ≈⟨ (refl⟩∘⟨ A.F.homomorphism) ⟩ - _ ≈⟨ (refl⟩∘⟨ pushˡ A.D A.F.homomorphism) ⟩ - _ ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ - _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pushˡ A.D A.F.homomorphism) ⟩ -- (refl⟩∘⟨ (refl⟩∘⟨ A.F.homomorphism)) ⟩ - _ ≈⟨ (refl⟩∘⟨ extendʳ A.D (A.adj.counit.commute _)) ⟩ -- (extendʳ A.D (A.adj.counit.commute _)) ⟩ - _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ -- (refl⟩∘⟨ refl⟩∘⟨ elimˡ A.D A.adj.zig) ⟩ - _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ (identityˡ ○ A.F.F-resp-≈ (elimˡ C M.F.identity))) ⟩ -- (refl⟩∘⟨ refl⟩∘⟨ A.F.F-resp-≈ {! (elimˡ C M.F.identity) !}) ⟩ - _ ≈⟨ (refl⟩∘⟨ Γ.⇐.commute f) ⟩ - _ ≈⟨ pullˡ A.D (Equiv.sym H.H.homomorphism) ⟩ - _ ≈⟨ (H.H.F-resp-≈ (elimʳ C (M.F.F-resp-≈ M.F.identity CH.○ M.F.identity) CH.⟩∘⟨refl) ⟩∘⟨refl) ⟩ - _ ≈⟨ (H.H.F-resp-≈ (pullˡ C M.identityʳ CH.○ C.identityˡ) ⟩∘⟨refl) ⟩ - _ ∎ + ; commute = λ f → let open A.D.HomReasoning in + begin + _ ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ A.F.F-resp-≈ ((H.μ-comp CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc))))) ⟩ + _ ≈⟨ ((refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ + _ ≈⟨ (refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _)) ⟩ + _ ≈⟨ (refl⟩∘⟨ assoc) ⟩ + _ ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ + _ ≈⟨ (refl⟩∘⟨ A.F.homomorphism) ⟩ + _ ≈⟨ (refl⟩∘⟨ pushˡ A.D A.F.homomorphism) ⟩ + _ ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ + _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pushˡ A.D A.F.homomorphism) ⟩ + _ ≈⟨ (refl⟩∘⟨ extendʳ A.D (A.adj.counit.commute _)) ⟩ + _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ + _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ (identityˡ ○ A.F.F-resp-≈ (elimˡ C M.F.identity))) ⟩ + _ ≈⟨ (refl⟩∘⟨ Γ.⇐.commute f) ⟩ + _ ≈⟨ pullˡ A.D (Equiv.sym H.H.homomorphism) ⟩ + _ ≈⟨ (H.H.F-resp-≈ (elimʳ C (M.F.F-resp-≈ M.F.identity CH.○ M.F.identity) CH.⟩∘⟨refl) ⟩∘⟨refl) ⟩ + _ ≈⟨ (H.H.F-resp-≈ (pullˡ C M.identityʳ CH.○ C.identityˡ) ⟩∘⟨refl) ⟩ + _ ∎ ; iso = NaturalIsomorphism.iso (sym H.Γ) }) } - } \ No newline at end of file + } From 20c26f891706d594442c5ec25823cc46fc7d4eae Mon Sep 17 00:00:00 2001 From: iwilare Date: Wed, 11 Jan 2023 15:14:06 +0200 Subject: [PATCH 44/46] Make terms explicit --- .../Construction/Adjunctions/Properties.agda | 131 ++++++++++-------- 1 file changed, 77 insertions(+), 54 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda index bc8f8306e..38e514049 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions/Properties.agda @@ -12,7 +12,7 @@ private module M = Monad M open import Categories.Adjoint using (Adjoint) -open import Categories.Functor using (Functor) +open import Categories.Functor using (Functor; _∘F_) renaming (id to idF) open import Categories.Morphism.Reasoning open import Categories.NaturalTransformation hiding (id) open NaturalTransformation using (η) @@ -29,29 +29,28 @@ open import Categories.Adjoint.Construction.EilenbergMoore M as EM open import Categories.Tactic.Category --- EM-object : SplitObj --- EM-object = record --- { D = EilenbergMoore M --- ; F = EM.Free --- ; G = EM.Forgetful --- ; adj = EM.Free⊣Forgetful --- ; GF≃M = EM.FF≃F --- ; η-eq = begin --- M.F.₁ id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ --- id ∘ M.η.η _ ≈⟨ identityˡ ⟩ --- M.η.η _ ∎ --- ; μ-eq = begin --- M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ --- id ∘ M.μ.η _ ≈⟨ identityˡ ⟩ --- M.μ.η _ ≈⟨ Equiv.sym identityʳ ⟩ --- M.μ.η _ ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ --- M.μ.η _ ∘ M.F.₁ id ≈⟨ Equiv.sym identityʳ ⟩ --- (M.μ.η _ ∘ M.F.₁ id) ∘ id ≈⟨ assoc ⟩ --- M.μ.η _ ∘ M.F.₁ id ∘ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ --- M.μ.η _ ∘ M.F.₁ id ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ --- M.μ.η _ ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ --- } where open Category C --- open HomReasoning +EM-object : SplitObj +EM-object = record + { D = EilenbergMoore M + ; F = EM.Free + ; G = EM.Forgetful + ; adj = EM.Free⊣Forgetful + ; GF≃M = EM.FF≃F + ; η-eq = begin + M.F.₁ id ∘ M.η.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ + id ∘ M.η.η _ ≈⟨ identityˡ ⟩ + M.η.η _ ∎ + ; μ-eq = begin + M.F.₁ id ∘ M.μ.η _ ≈⟨ M.F.identity ⟩∘⟨refl ⟩ + id ∘ M.μ.η _ ≈⟨ id-comm-sym C ⟩ + M.μ.η _ ∘ id ≈⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + M.μ.η _ ∘ M.F.₁ id ≈⟨ Equiv.sym identityʳ ⟩ + (M.μ.η _ ∘ M.F.₁ id) ∘ id ≈⟨ assoc ⟩ + M.μ.η _ ∘ M.F.₁ id ∘ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ Equiv.sym M.F.identity ⟩ + M.μ.η _ ∘ M.F.₁ id ∘ M.F.₁ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ M.F.F-resp-≈ (Equiv.sym M.F.identity) ⟩ + M.μ.η _ ∘ M.F.₁ id ∘ M.F.₁ (M.F.₁ id) ∎ + } where open Category C + open HomReasoning Kl-object : SplitObj Kl-object = record @@ -126,19 +125,27 @@ Kl-object = record (F.₁ f ∘ D.id) ∎ ; iso = λ X → record { isoˡ = identityˡ ; isoʳ = identityˡ } }) - ; μ-comp = λ { {x} → let open C - open C.HomReasoning in - Equiv.sym - (begin - _ ≈⟨ elimˡ C G.identity ⟩ - _ ≈⟨ elimʳ C M.F.identity ⟩ - _ ≈⟨ (refl⟩∘⟨ elimˡ C G.identity) ⟩ - _ ≈⟨ (G.homomorphism ⟩∘⟨refl) ⟩ - _ ≈⟨ pullʳ C (adj.unit.sym-commute _) ⟩ - _ ≈⟨ sym-assoc ⟩ - _ ≈⟨ elimˡ C adj.zag ⟩ - _ ≈⟨ elimʳ C M.F.identity ⟩ - _ ∎) } + ; μ-comp = λ { {x} → + let open C + open C.HomReasoning in + Equiv.sym (begin + G.F₁ D.id ∘ (G.F₁ (adj.counit.η (F.₀ x) D.∘ F.₁ (GF≃M.⇐.η x ∘ M.F.F₁ id)) ∘ G.F₁ D.id ∘ adj.unit.η (M.F.F₀ x)) ∘ M.F.F₁ id + ≈⟨ elimˡ C G.identity ⟩ + (G.F₁ (adj.counit.η (F.₀ x) D.∘ F.₁ (GF≃M.⇐.η x ∘ M.F.F₁ id)) ∘ G.F₁ D.id ∘ adj.unit.η (M.F.F₀ x)) ∘ M.F.F₁ id + ≈⟨ elimʳ C M.F.identity ⟩ + G.F₁ (adj.counit.η (F.₀ x) D.∘ F.₁ (GF≃M.⇐.η x ∘ M.F.F₁ id)) ∘ G.F₁ D.id ∘ adj.unit.η (M.F.F₀ x) + ≈⟨ refl⟩∘⟨ elimˡ C G.identity ⟩ + (G.F₁ ( adj.counit.η (F.₀ x) D.∘ F.₁ (GF≃M.⇐.η x ∘ M.F.F₁ id)) ∘ adj.unit.η (M.F.F₀ x)) + ≈⟨ (G.homomorphism ⟩∘⟨refl) ⟩ + ((G.F₁ (adj.counit.η (F.₀ x)) ∘ Functor.F₁ (G ∘F F) (GF≃M.⇐.η x ∘ M.F.F₁ id)) ∘ adj.unit.η (M.F.F₀ x)) + ≈⟨ pullʳ C (adj.unit.sym-commute _) ⟩ + (G.F₁ (adj.counit.η (F.₀ x)) ∘ adj.unit.η (G.F₀ (F.₀ x)) ∘ GF≃M.⇐.η x ∘ M.F.F₁ id) + ≈⟨ sym-assoc ⟩ + (G.F₁ (adj.counit.η (F.₀ x)) ∘ adj.unit.η (G.F₀ (F.₀ x))) ∘ GF≃M.⇐.η x ∘ M.F.F₁ id + ≈⟨ elimˡ C adj.zag ⟩ + (GF≃M.⇐.η x ∘ M.F.F₁ id) + ≈⟨ elimʳ C M.F.identity ⟩ + GF≃M.⇐.η x ∎) } } where module adj = Adjoint adj module F = Functor F @@ -165,23 +172,39 @@ Kl-initial = record ; η⁻¹ = Γ.⇒.η ; commute = λ f → let open A.D.HomReasoning in begin - _ ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ A.F.F-resp-≈ ((H.μ-comp CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc))))) ⟩ - _ ≈⟨ ((refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism)) ⟩ - _ ≈⟨ (refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _)) ⟩ - _ ≈⟨ (refl⟩∘⟨ assoc) ⟩ - _ ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ - _ ≈⟨ (refl⟩∘⟨ A.F.homomorphism) ⟩ - _ ≈⟨ (refl⟩∘⟨ pushˡ A.D A.F.homomorphism) ⟩ - _ ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ - _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pushˡ A.D A.F.homomorphism) ⟩ - _ ≈⟨ (refl⟩∘⟨ extendʳ A.D (A.adj.counit.commute _)) ⟩ - _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig) ⟩ - _ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ (identityˡ ○ A.F.F-resp-≈ (elimˡ C M.F.identity))) ⟩ - _ ≈⟨ (refl⟩∘⟨ Γ.⇐.commute f) ⟩ - _ ≈⟨ pullˡ A.D (Equiv.sym H.H.homomorphism) ⟩ - _ ≈⟨ (H.H.F-resp-≈ (elimʳ C (M.F.F-resp-≈ M.F.identity CH.○ M.F.identity) CH.⟩∘⟨refl) ⟩∘⟨refl) ⟩ - _ ≈⟨ (H.H.F-resp-≈ (pullˡ C M.identityʳ CH.○ C.identityˡ) ⟩∘⟨refl) ⟩ - _ ∎ + Γ.⇐.η _ ∘ A.adj.counit.η (A.F.F₀ _) ∘ A.F.F₁ (A.GF≃M.⇐.η _ C.∘ f) + ≈⟨ refl⟩∘⟨ refl⟩∘⟨ A.F.F-resp-≈ ((H.μ-comp CH.⟩∘⟨refl) CH.○ (C.assoc CH.○ (CH.refl⟩∘⟨ C.assoc))) ⟩ + Γ.⇐.η _ ∘ A.adj.counit.η (A.F.F₀ _) ∘ A.F.F₁ (A.G.₁ (Γ.⇒.η _) C.∘ (A.G.₁ (H.H.F₁ (M.F.F₁ C.id)) C.∘ (A.G.₁ (Γ.⇐.η (M.F.F₀ _)) C.∘ A.adj.unit.η (M.F.F₀ _))) C.∘ M.F.F₁ C.id C.∘ f) + ≈⟨ refl⟩∘⟨ refl⟩∘⟨ A.F.homomorphism ⟩ + Γ.⇐.η _ ∘ A.adj.counit.η (A.F.F₀ _) ∘ Functor.F₁ (A.F Categories.Functor.∘F A.G) (Γ.⇒.η _) ∘ A.F.F₁ ((A.G.₁ (H.H.F₁ (M.F.F₁ C.id)) C.∘ A.G.₁ (Γ.⇐.η (M.F.F₀ _)) C.∘ A.adj.unit.η (M.F.F₀ _)) C.∘ M.F.F₁ C.id C.∘ f) + ≈⟨ refl⟩∘⟨ pullˡ A.D (A.adj.counit.commute _) ⟩ + Γ.⇐.η _ ∘ (Γ.⇒.η _ ∘ A.adj.counit.η (H.H.F₀ (Functor.₀ K.F _))) ∘ A.F.F₁ ((A.G.₁ (H.H.F₁ (M.F.F₁ C.id)) C.∘ (A.G.₁ (Γ.⇐.η (M.F.F₀ _)) C.∘ A.adj.unit.η (M.F.F₀ _))) C.∘ M.F.F₁ C.id C.∘ f) + ≈⟨ refl⟩∘⟨ assoc ⟩ + Γ.⇐.η _ ∘ Γ.⇒.η _ ∘ A.adj.counit.η (H.H.F₀ (Functor.₀ K.F _)) ∘ A.F.F₁ ((A.G.₁ (H.H.F₁ (M.F.F₁ C.id)) C.∘ A.G.₁ (Γ.⇐.η (M.F.F₀ _)) C.∘ A.adj.unit.η (M.F.F₀ _)) C.∘ M.F.F₁ C.id C.∘ f) + ≈⟨ cancelˡ A.D (Γ.iso.isoˡ _) ⟩ + A.adj.counit.η (H.H.F₀ (Functor.₀ K.F _)) ∘ A.F.F₁ ((A.G.₁ (H.H.F₁ (M.F.F₁ C.id)) C.∘ (A.G.₁ (Γ.⇐.η (M.F.F₀ _)) C.∘ A.adj.unit.η (M.F.F₀ _))) C.∘ M.F.F₁ C.id C.∘ f) + ≈⟨ refl⟩∘⟨ A.F.homomorphism ⟩ + A.adj.counit.η (H.H.F₀ (Functor.₀ K.F _)) ∘ A.F.F₁ (A.G.₁ (H.H.F₁ (M.F.F₁ C.id)) C.∘ A.G.₁ (Γ.⇐.η (M.F.F₀ _)) C.∘ A.adj.unit.η (M.F.F₀ _)) ∘ A.F.F₁ (M.F.F₁ C.id C.∘ f) + ≈⟨ refl⟩∘⟨ pushˡ A.D A.F.homomorphism ⟩ + A.adj.counit.η (H.H.F₀ (Functor.₀ K.F _)) ∘ Functor.F₁ (A.F Categories.Functor.∘F A.G) (H.H.F₁ (M.F.F₁ C.id)) ∘ A.F.F₁ (A.G.₁ (Γ.⇐.η (M.F.F₀ _)) C.∘ A.adj.unit.η (M.F.F₀ _)) ∘ A.F.F₁ (M.F.F₁ C.id C.∘ f) + ≈⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ + H.H.F₁ (M.F.F₁ C.id) ∘ A.adj.counit.η (H.H.F₀ (Functor.₀ K.F (M.F.F₀ _))) ∘ A.F.F₁ (A.G.₁ (Γ.⇐.η (M.F.F₀ _)) C.∘ A.adj.unit.η (M.F.F₀ _)) ∘ A.F.F₁ (M.F.F₁ C.id C.∘ f) + ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ A.D A.F.homomorphism ⟩ + H.H.F₁ (M.F.F₁ C.id) ∘ A.adj.counit.η (H.H.F₀ (Functor.₀ K.F (M.F.F₀ _))) ∘ Functor.F₁ (A.F Categories.Functor.∘F A.G) (Γ.⇐.η (M.F.F₀ _)) ∘ A.F.F₁ (A.adj.unit.η (M.F.F₀ _)) ∘ A.F.F₁ (M.F.F₁ C.id C.∘ f) + ≈⟨ refl⟩∘⟨ extendʳ A.D (A.adj.counit.commute _) ⟩ + H.H.F₁ (M.F.F₁ C.id) ∘ Γ.⇐.η (M.F.F₀ _) ∘ A.adj.counit.η (A.F.F₀ (M.F.F₀ _)) ∘ A.F.F₁ (A.adj.unit.η (M.F.F₀ _)) ∘ A.F.F₁ (M.F.F₁ C.id C.∘ f) + ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pullˡ A.D A.adj.zig ⟩ + H.H.F₁ (M.F.F₁ C.id) ∘ Γ.⇐.η (M.F.F₀ _) ∘ id ∘ A.F.F₁ (M.F.F₁ C.id C.∘ f) + ≈⟨ refl⟩∘⟨ refl⟩∘⟨ (identityˡ ○ A.F.F-resp-≈ (elimˡ C M.F.identity)) ⟩ + H.H.F₁ (M.F.F₁ C.id) ∘ Γ.⇐.η (M.F.F₀ _) ∘ A.F.F₁ f + ≈⟨ refl⟩∘⟨ Γ.⇐.commute f ⟩ + H.H.F₁ (M.F.F₁ C.id) ∘ H.H.F₁ (M.η.η (M.F.F₀ (Functor.₀ K.F _)) C.∘ f) ∘ Γ.⇐.η _ + ≈⟨ pullˡ A.D (Equiv.sym H.H.homomorphism) ⟩ + H.H.F₁ ((M.μ.η (Functor.₀ K.F _) C.∘ M.F.F₁ (M.F.F₁ C.id)) C.∘ M.η.η (M.F.F₀ (Functor.₀ K.F _)) C.∘ f) ∘ Γ.⇐.η _ + ≈⟨ H.H.F-resp-≈ (elimʳ C (M.F.F-resp-≈ M.F.identity CH.○ M.F.identity) CH.⟩∘⟨refl) ⟩∘⟨refl ⟩ + H.H.F₁ (M.μ.η (Functor.₀ K.F _) C.∘ M.η.η (M.F.F₀ (Functor.₀ K.F _)) C.∘ f) ∘ Γ.⇐.η _ + ≈⟨ H.H.F-resp-≈ (pullˡ C M.identityʳ CH.○ C.identityˡ) ⟩∘⟨refl ⟩ + H.H.F₁ f ∘ Γ.⇐.η _ ∎ ; iso = NaturalIsomorphism.iso (sym H.Γ) }) } From e963814c83e0af69defc688d0574828dcf7c6069 Mon Sep 17 00:00:00 2001 From: iwilare Date: Wed, 11 Jan 2023 16:46:13 +0200 Subject: [PATCH 45/46] Revert Categories.Tactic.Category --- src/Categories/Tactic/Category.agda | 324 ++++++++++++++-------------- 1 file changed, 162 insertions(+), 162 deletions(-) diff --git a/src/Categories/Tactic/Category.agda b/src/Categories/Tactic/Category.agda index 5c0a2e052..5f791c268 100644 --- a/src/Categories/Tactic/Category.agda +++ b/src/Categories/Tactic/Category.agda @@ -19,165 +19,165 @@ open import Data.List as List using (List; _∷_; []) open import Data.Product as Product using (_×_; _,_) open import Agda.Builtin.Reflection --- open import Reflection.Argument --- open import Reflection.Term using (getName; _⋯⟅∷⟆_) --- open import Reflection.TypeChecking.Monad.Syntax - --- module _ {o ℓ e} (𝒞 : Category o ℓ e) where - --- open Category 𝒞 --- open HomReasoning --- open Equiv - --- private --- variable --- A B C : Obj --- f g : A ⇒ B - --- -------------------------------------------------------------------------------- --- -- An 'Expr' reifies the parentheses/identity morphisms of some series of --- -- compositions of morphisms into a data structure. In fact, this is also --- -- a category! --- -------------------------------------------------------------------------------- --- data Expr : Obj → Obj → Set (o ⊔ ℓ) where --- _∘′_ : ∀ {A B C} → Expr B C → Expr A B → Expr A C --- id′ : ∀ {A} → Expr A A --- [_↑] : ∀ {A B} → A ⇒ B → Expr A B - --- -- Embed a morphism in 'Expr' back into '𝒞' without normalizing. --- [_↓] : Expr A B → A ⇒ B --- [ f ∘′ g ↓] = [ f ↓] ∘ [ g ↓] --- [ id′ ↓] = id --- [ [ f ↑] ↓] = f - --- -- Convert an 'Expr' back into a morphism, while normalizing --- -- --- -- This actually embeds the morphism into the category of copresheaves --- -- on 𝒞, which obeys the category laws up to beta-eta equality. --- -- This lets us normalize away all the associations/identity morphisms. --- embed : Expr B C → A ⇒ B → A ⇒ C --- embed (f ∘′ g) h = embed f (embed g h) --- embed id′ h = h --- embed [ f ↑] h = f ∘ h - - --- preserves-≈′ : ∀ (f : Expr B C) → (h : A ⇒ B) → embed f id ∘ h ≈ embed f h --- preserves-≈′ id′ f = identityˡ --- preserves-≈′ [ x ↑] f = ∘-resp-≈ˡ identityʳ --- preserves-≈′ (f ∘′ g) h = begin --- embed (f ∘′ g) id ∘ h ≡⟨⟩ --- embed f (embed g id) ∘ h ≈˘⟨ preserves-≈′ f (embed g id) ⟩∘⟨refl ⟩ --- (embed f id ∘ embed g id) ∘ h ≈⟨ assoc ⟩ --- embed f id ∘ embed g id ∘ h ≈⟨ refl⟩∘⟨ preserves-≈′ g h ⟩ --- embed f id ∘ embed g h ≈⟨ preserves-≈′ f (embed g h) ⟩ --- embed (f ∘′ g) h ∎ - --- preserves-≈ : ∀ (f : Expr A B) → embed f id ≈ [ f ↓] --- preserves-≈ id′ = refl --- preserves-≈ [ x ↑] = identityʳ --- preserves-≈ (f ∘′ g) = begin --- embed (f ∘′ g) id ≈˘⟨ preserves-≈′ f (embed g id) ⟩ --- embed f id ∘ embed g id ≈⟨ preserves-≈ f ⟩∘⟨ preserves-≈ g ⟩ --- [ f ↓] ∘ [ g ↓] ≡⟨⟩ --- [ f ∘′ g ↓] ∎ - --- -------------------------------------------------------------------------------- --- -- Reflection Helpers --- -------------------------------------------------------------------------------- - --- _==_ = primQNameEquality --- {-# INLINE _==_ #-} - --- getArgs : Term → Maybe (Term × Term) --- getArgs (def _ xs) = go xs --- where --- go : List (Arg Term) → Maybe (Term × Term) --- go (vArg x ∷ vArg y ∷ []) = just (x , y) --- go (x ∷ xs) = go xs --- go _ = nothing --- getArgs _ = nothing - --- -------------------------------------------------------------------------------- --- -- Getting Category Names --- -------------------------------------------------------------------------------- - --- record CategoryNames : Set where --- field --- is-∘ : Name → Bool --- is-id : Name → Bool - --- buildMatcher : Name → Maybe Name → Name → Bool --- buildMatcher n nothing x = n == x --- buildMatcher n (just m) x = n == x ∨ m == x - --- findCategoryNames : Term → TC CategoryNames --- findCategoryNames cat = do --- ∘-altName ← normalise (def (quote Category._∘_) (3 ⋯⟅∷⟆ cat ⟨∷⟩ [])) --- id-altName ← normalise (def (quote Category.id) (3 ⋯⟅∷⟆ cat ⟨∷⟩ [])) --- returnTC record --- { is-∘ = buildMatcher (quote Category._∘_) (getName ∘-altName) --- ; is-id = buildMatcher (quote Category.id) (getName id-altName) --- } - --- -------------------------------------------------------------------------------- --- -- Constructing an Expr --- -------------------------------------------------------------------------------- - --- ″id″ : Term --- ″id″ = quote id′ ⟨ con ⟩ [] - --- ″[_↑]″ : Term → Term --- ″[ t ↑]″ = quote [_↑] ⟨ con ⟩ (t ⟨∷⟩ []) - --- module _ (names : CategoryNames) where - --- open CategoryNames names - --- mutual --- ″∘″ : List (Arg Term) → Term --- ″∘″ (x ⟨∷⟩ y ⟨∷⟩ xs) = quote _∘′_ ⟨ con ⟩ buildExpr x ⟨∷⟩ buildExpr y ⟨∷⟩ [] --- ″∘″ (x ∷ xs) = ″∘″ xs --- ″∘″ _ = unknown - --- buildExpr : Term → Term --- buildExpr t@(def n xs) = --- if (is-∘ n) --- then ″∘″ xs --- else if (is-id n) --- then ″id″ --- else --- ″[ t ↑]″ --- buildExpr t@(con n xs) = --- if (is-∘ n) --- then ″∘″ xs --- else if (is-id n) --- then ″id″ --- else --- ″[ t ↑]″ --- buildExpr t = ″[ t ↑]″ - --- -------------------------------------------------------------------------------- --- -- Constructing the Solution --- -------------------------------------------------------------------------------- - --- constructSoln : Term → CategoryNames → Term → Term → Term --- constructSoln cat names lhs rhs = --- quote Category.Equiv.trans ⟨ def ⟩ 3 ⋯⟅∷⟆ cat ⟨∷⟩ --- (quote Category.Equiv.sym ⟨ def ⟩ 3 ⋯⟅∷⟆ cat ⟨∷⟩ --- (quote preserves-≈ ⟨ def ⟩ 3 ⋯⟅∷⟆ cat ⟨∷⟩ buildExpr names lhs ⟨∷⟩ []) ⟨∷⟩ []) --- ⟨∷⟩ --- (quote preserves-≈ ⟨ def ⟩ 3 ⋯⟅∷⟆ cat ⟨∷⟩ buildExpr names rhs ⟨∷⟩ []) --- ⟨∷⟩ [] - --- solve-macro : Term → Term → TC _ --- solve-macro mon hole = do --- hole′ ← inferType hole >>= normalise --- names ← findCategoryNames mon --- just (lhs , rhs) ← returnTC (getArgs hole′) --- where nothing → typeError (termErr hole′ ∷ []) --- let soln = constructSoln mon names lhs rhs --- unify hole soln - --- macro --- solve : Term → Term → TC _ --- solve = solve-macro +open import Reflection.Argument +open import Reflection.Term using (getName; _⋯⟅∷⟆_) +open import Reflection.TypeChecking.Monad.Syntax + +module _ {o ℓ e} (𝒞 : Category o ℓ e) where + + open Category 𝒞 + open HomReasoning + open Equiv + + private + variable + A B C : Obj + f g : A ⇒ B + + -------------------------------------------------------------------------------- + -- An 'Expr' reifies the parentheses/identity morphisms of some series of + -- compositions of morphisms into a data structure. In fact, this is also + -- a category! + -------------------------------------------------------------------------------- + data Expr : Obj → Obj → Set (o ⊔ ℓ) where + _∘′_ : ∀ {A B C} → Expr B C → Expr A B → Expr A C + id′ : ∀ {A} → Expr A A + [_↑] : ∀ {A B} → A ⇒ B → Expr A B + + -- Embed a morphism in 'Expr' back into '𝒞' without normalizing. + [_↓] : Expr A B → A ⇒ B + [ f ∘′ g ↓] = [ f ↓] ∘ [ g ↓] + [ id′ ↓] = id + [ [ f ↑] ↓] = f + + -- Convert an 'Expr' back into a morphism, while normalizing + -- + -- This actually embeds the morphism into the category of copresheaves + -- on 𝒞, which obeys the category laws up to beta-eta equality. + -- This lets us normalize away all the associations/identity morphisms. + embed : Expr B C → A ⇒ B → A ⇒ C + embed (f ∘′ g) h = embed f (embed g h) + embed id′ h = h + embed [ f ↑] h = f ∘ h + + + preserves-≈′ : ∀ (f : Expr B C) → (h : A ⇒ B) → embed f id ∘ h ≈ embed f h + preserves-≈′ id′ f = identityˡ + preserves-≈′ [ x ↑] f = ∘-resp-≈ˡ identityʳ + preserves-≈′ (f ∘′ g) h = begin + embed (f ∘′ g) id ∘ h ≡⟨⟩ + embed f (embed g id) ∘ h ≈˘⟨ preserves-≈′ f (embed g id) ⟩∘⟨refl ⟩ + (embed f id ∘ embed g id) ∘ h ≈⟨ assoc ⟩ + embed f id ∘ embed g id ∘ h ≈⟨ refl⟩∘⟨ preserves-≈′ g h ⟩ + embed f id ∘ embed g h ≈⟨ preserves-≈′ f (embed g h) ⟩ + embed (f ∘′ g) h ∎ + + preserves-≈ : ∀ (f : Expr A B) → embed f id ≈ [ f ↓] + preserves-≈ id′ = refl + preserves-≈ [ x ↑] = identityʳ + preserves-≈ (f ∘′ g) = begin + embed (f ∘′ g) id ≈˘⟨ preserves-≈′ f (embed g id) ⟩ + embed f id ∘ embed g id ≈⟨ preserves-≈ f ⟩∘⟨ preserves-≈ g ⟩ + [ f ↓] ∘ [ g ↓] ≡⟨⟩ + [ f ∘′ g ↓] ∎ + +-------------------------------------------------------------------------------- +-- Reflection Helpers +-------------------------------------------------------------------------------- + +_==_ = primQNameEquality +{-# INLINE _==_ #-} + +getArgs : Term → Maybe (Term × Term) +getArgs (def _ xs) = go xs + where + go : List (Arg Term) → Maybe (Term × Term) + go (vArg x ∷ vArg y ∷ []) = just (x , y) + go (x ∷ xs) = go xs + go _ = nothing +getArgs _ = nothing + +-------------------------------------------------------------------------------- +-- Getting Category Names +-------------------------------------------------------------------------------- + +record CategoryNames : Set where + field + is-∘ : Name → Bool + is-id : Name → Bool + +buildMatcher : Name → Maybe Name → Name → Bool +buildMatcher n nothing x = n == x +buildMatcher n (just m) x = n == x ∨ m == x + +findCategoryNames : Term → TC CategoryNames +findCategoryNames cat = do + ∘-altName ← normalise (def (quote Category._∘_) (3 ⋯⟅∷⟆ cat ⟨∷⟩ [])) + id-altName ← normalise (def (quote Category.id) (3 ⋯⟅∷⟆ cat ⟨∷⟩ [])) + returnTC record + { is-∘ = buildMatcher (quote Category._∘_) (getName ∘-altName) + ; is-id = buildMatcher (quote Category.id) (getName id-altName) + } + +-------------------------------------------------------------------------------- +-- Constructing an Expr +-------------------------------------------------------------------------------- + +″id″ : Term +″id″ = quote id′ ⟨ con ⟩ [] + +″[_↑]″ : Term → Term +″[ t ↑]″ = quote [_↑] ⟨ con ⟩ (t ⟨∷⟩ []) + +module _ (names : CategoryNames) where + + open CategoryNames names + + mutual + ″∘″ : List (Arg Term) → Term + ″∘″ (x ⟨∷⟩ y ⟨∷⟩ xs) = quote _∘′_ ⟨ con ⟩ buildExpr x ⟨∷⟩ buildExpr y ⟨∷⟩ [] + ″∘″ (x ∷ xs) = ″∘″ xs + ″∘″ _ = unknown + + buildExpr : Term → Term + buildExpr t@(def n xs) = + if (is-∘ n) + then ″∘″ xs + else if (is-id n) + then ″id″ + else + ″[ t ↑]″ + buildExpr t@(con n xs) = + if (is-∘ n) + then ″∘″ xs + else if (is-id n) + then ″id″ + else + ″[ t ↑]″ + buildExpr t = ″[ t ↑]″ + +-------------------------------------------------------------------------------- +-- Constructing the Solution +-------------------------------------------------------------------------------- + +constructSoln : Term → CategoryNames → Term → Term → Term +constructSoln cat names lhs rhs = + quote Category.Equiv.trans ⟨ def ⟩ 3 ⋯⟅∷⟆ cat ⟨∷⟩ + (quote Category.Equiv.sym ⟨ def ⟩ 3 ⋯⟅∷⟆ cat ⟨∷⟩ + (quote preserves-≈ ⟨ def ⟩ 3 ⋯⟅∷⟆ cat ⟨∷⟩ buildExpr names lhs ⟨∷⟩ []) ⟨∷⟩ []) + ⟨∷⟩ + (quote preserves-≈ ⟨ def ⟩ 3 ⋯⟅∷⟆ cat ⟨∷⟩ buildExpr names rhs ⟨∷⟩ []) + ⟨∷⟩ [] + +solve-macro : Term → Term → TC _ +solve-macro mon hole = do + hole′ ← inferType hole >>= normalise + names ← findCategoryNames mon + just (lhs , rhs) ← returnTC (getArgs hole′) + where nothing → typeError (termErr hole′ ∷ []) + let soln = constructSoln mon names lhs rhs + unify hole soln + +macro + solve : Term → Term → TC _ + solve = solve-macro From 3cee6a10bbce81674709fa011e7eb04408fdeeba Mon Sep 17 00:00:00 2001 From: iwilare Date: Thu, 12 Jan 2023 11:38:00 +0200 Subject: [PATCH 46/46] Small progress on mu-comp --- .../Adjoint/Construction/Adjunctions.agda | 14 ++++++++++++-- 1 file changed, 12 insertions(+), 2 deletions(-) diff --git a/src/Categories/Adjoint/Construction/Adjunctions.agda b/src/Categories/Adjoint/Construction/Adjunctions.agda index 90c449148..8ba082a2d 100644 --- a/src/Categories/Adjoint/Construction/Adjunctions.agda +++ b/src/Categories/Adjoint/Construction/Adjunctions.agda @@ -123,12 +123,22 @@ Split = record {! !} ≈⟨ pushˡ C X.G.homomorphism ⟩ {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ assoc) ⟩ {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ (refl⟩∘⟨ ((X.G.homomorphism ⟩∘⟨refl) ⟩∘⟨refl))) ⟩ - {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ {! !}) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ (assoc ○ assoc)) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ X.adj.unit.commute _) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ (refl⟩∘⟨ extendʳ C (Equiv.sym X.G.homomorphism ○ X.G.F-resp-≈ (Γᵤ.⇐.commute _) ○ X.G.homomorphism))) ⟩ + {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ (extendʳ C (Equiv.sym (Functor.homomorphism (X.G ∘F Hᵤ)) ○ Functor.F-resp-≈ ((X.G ∘F Hᵤ)) (Γᵥ.⇐.commute _) ○ Functor.homomorphism ((X.G ∘F Hᵤ))))) ⟩ {! !} ≈⟨ {! !} ⟩ - {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ {! !} ⟩ + {! !} ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ Equiv.sym Bμ-comp) ⟩ {! !} ≈⟨ Equiv.sym Aμ-comp ⟩ {! !} ∎) } +{- +Equiv.sym X.G.homomorphism ○ X.G.F-resp-≈ (Γᵤ.⇐.commute _) ○ X.G.homomorphism +-} + + {- Have