work on new proof principle

This commit is contained in:
Leon Vatthauer 2023-11-03 17:33:44 +01:00
parent 446d55f1f5
commit 7f336282ed
Signed by: leonv
SSH key fingerprint: SHA256:G4+ddwoZmhLPRB1agvXzZMXIzkVJ36dUYZXf5NxT+u8
3 changed files with 154 additions and 106 deletions

View file

@ -66,7 +66,57 @@ This file contains some typedefs and records concerning different algebras.
→ f ≈ g ∘ (idC ⁂ η) → f ≈ g ∘ (idC ⁂ η)
→ (∀ {Z : Obj} (h : Z ⇒ Uniform-Iteration-Algebra.A FY + Z) → g ∘ (idC ⁂ h #) ≈ Uniform-Iteration-Algebra._# B ((g +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))) → (∀ {Z : Obj} (h : Z ⇒ Uniform-Iteration-Algebra.A FY + Z) → g ∘ (idC ⁂ h #) ≈ Uniform-Iteration-Algebra._# B ((g +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)))
→ g ≈ [ B , f ]♯ → g ≈ [ B , f ]♯
[_,_]♯ˡ : ∀ {X : Obj} (A : Uniform-Iteration-Algebra) (f : Y × X ⇒ Uniform-Iteration-Algebra.A A) → Uniform-Iteration-Algebra.A FY × X ⇒ Uniform-Iteration-Algebra.A A
[_,_]♯ˡ {X} A f = [ A , (f ∘ swap) ]♯ ∘ swap
♯ˡ-law : ∀ {X : Obj} {A : Uniform-Iteration-Algebra} (f : Y × X ⇒ Uniform-Iteration-Algebra.A A) → f ≈ [ A , f ]♯ˡ ∘ (η ⁂ idC)
♯ˡ-law {X} {A} f = begin
f ≈⟨ introʳ swap∘swap ⟩
f ∘ swap ∘ swap ≈⟨ pullˡ (♯-law (f ∘ swap)) ⟩
([ A , f ∘ swap ]♯ ∘ (idC ⁂ η)) ∘ swap ≈⟨ pullʳ (sym swap∘⁂) ⟩
[ A , (f ∘ swap) ]♯ ∘ swap ∘ (η ⁂ idC) ≈⟨ sym-assoc ⟩
([ A , (f ∘ swap) ]♯ ∘ swap) ∘ (η ⁂ idC) ∎
♯ˡ-preserving : ∀ {X : Obj} {B : Uniform-Iteration-Algebra} (f : Y × X ⇒ Uniform-Iteration-Algebra.A B) {Z : Obj} (h : Z ⇒ Uniform-Iteration-Algebra.A FY + Z) → [ B , f ]♯ˡ ∘ (h # ⁂ idC) ≈ Uniform-Iteration-Algebra._# B (([ B , f ]♯ˡ +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC))
♯ˡ-preserving {X} {B} f {Z} h = begin
([ B , (f ∘ swap) ]♯ ∘ swap) ∘ ((h #) ⁂ idC) ≈⟨ pullʳ swap∘⁂ ⟩
[ B , (f ∘ swap) ]♯ ∘ (idC ⁂ h #) ∘ swap ≈⟨ pullˡ (♯-preserving (f ∘ swap) h) ⟩
(([ B , (f ∘ swap) ]♯ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) #ᵇ ∘ swap ≈⟨ sym (#ᵇ-Uniformity uni-helper) ⟩
((([ B , (f ∘ swap) ]♯ ∘ swap) +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC)) #ᵇ
where
open Uniform-Iteration-Algebra B using () renaming (_# to _#ᵇ; #-Uniformity to #ᵇ-Uniformity)
uni-helper : (idC +₁ swap) ∘ ([ B , f ∘ swap ]♯ ∘ swap +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC) ≈ (([ B , f ∘ swap ]♯ coproducts.+₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) ∘ swap
uni-helper = begin
(idC +₁ swap) ∘ ([ B , f ∘ swap ]♯ ∘ swap +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC) ≈⟨ pullˡ +₁∘+₁ ⟩
(idC ∘ [ B , f ∘ swap ]♯ ∘ swap +₁ swap ∘ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC) ≈⟨ (+₁-cong₂ identityˡ id-comm) ⟩∘⟨refl ⟩
([ B , f ∘ swap ]♯ ∘ swap +₁ idC ∘ swap) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC) ≈⟨ (sym +₁∘+₁) ⟩∘⟨refl ⟩
(([ B , f ∘ swap ]♯ +₁ idC) ∘ (swap +₁ swap)) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC) ≈⟨ pullʳ (pullˡ (sym distributeˡ⁻¹∘swap)) ⟩
([ B , f ∘ swap ]♯ +₁ idC) ∘ (distributeˡ⁻¹ ∘ swap) ∘ (h ⁂ idC) ≈⟨ (refl⟩∘⟨ (pullʳ swap∘⁂ ○ sym-assoc)) ○ sym-assoc ⟩
(([ B , f ∘ swap ]♯ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) ∘ swap ∎
♯ˡ-unique : ∀ {X : Obj} {B : Uniform-Iteration-Algebra} (f : Y × X ⇒ Uniform-Iteration-Algebra.A B) (g : Uniform-Iteration-Algebra.A FY × X ⇒ Uniform-Iteration-Algebra.A B)
→ f ≈ g ∘ (η ⁂ idC)
→ ({Z : Obj} (h : Z ⇒ Uniform-Iteration-Algebra.A FY + Z) → g ∘ (h # ⁂ idC) ≈ Uniform-Iteration-Algebra._# B ((g +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC)))
→ g ≈ [ B , f ]♯ˡ
♯ˡ-unique {X} {B} f g g-law g-preserving = begin
g ≈⟨ introʳ swap∘swap ⟩
g ∘ swap ∘ swap ≈⟨ sym-assoc ⟩
(g ∘ swap) ∘ swap ≈⟨ (♯-unique (f ∘ swap) (g ∘ swap) helper₁ helper₂) ⟩∘⟨refl ⟩
[ B , (f ∘ swap) ]♯ ∘ swap ∎
where
open Uniform-Iteration-Algebra B using () renaming (_# to _#ᵇ; #-Uniformity to #ᵇ-Uniformity)
helper₁ : f ∘ swap ≈ (g ∘ swap) ∘ (idC ⁂ η)
helper₁ = begin
f ∘ swap ≈⟨ g-law ⟩∘⟨refl ⟩
(g ∘ (η ⁂ idC)) ∘ swap ≈⟨ pullʳ (sym swap∘⁂) ⟩
g ∘ swap ∘ (idC ⁂ η) ≈⟨ sym-assoc ⟩
(g ∘ swap) ∘ (idC ⁂ η) ∎
helper₂ : ∀ {Z : Obj} (h : Z ⇒ Uniform-Iteration-Algebra.A FY + Z) → (g ∘ swap) ∘ (idC ⁂ h #) ≈ ((g ∘ swap +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) #ᵇ
helper₂ {Z} h = begin
(g ∘ swap) ∘ (idC ⁂ h #) ≈⟨ pullʳ swap∘⁂ ⟩
g ∘ (h # ⁂ idC) ∘ swap ≈⟨ pullˡ (g-preserving h) ⟩
((g +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC)) #ᵇ ∘ swap ≈⟨ sym (#ᵇ-Uniformity uni-helper) ⟩
((g ∘ swap +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) #ᵇ
where
uni-helper : (idC +₁ swap) ∘ (g ∘ swap +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h) ≈ ((g +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC)) ∘ swap
uni-helper = pullˡ +₁∘+₁ ○ (+₁-cong₂ identityˡ id-comm) ⟩∘⟨refl ○ (sym +₁∘+₁) ⟩∘⟨refl ○ pullʳ (pullˡ (sym distributeʳ⁻¹∘swap)) ○ (refl⟩∘⟨ (pullʳ swap∘⁂ ○ sym-assoc)) ○ sym-assoc
record StableFreeUniformIterationAlgebra : Set (suc o ⊔ suc ⊔ suc e) where record StableFreeUniformIterationAlgebra : Set (suc o ⊔ suc ⊔ suc e) where
field field
Y : Obj Y : Obj

View file

@ -88,6 +88,15 @@ module Category.Instance.AmbientCategory where
[ (idC ⁂ i₁) ∘ swap , (idC ⁂ i₂) ∘ swap ] ∘ distributeʳ⁻¹ ≈⟨ sym (pullˡ []∘+₁) ⟩ [ (idC ⁂ i₁) ∘ swap , (idC ⁂ i₂) ∘ swap ] ∘ distributeʳ⁻¹ ≈⟨ sym (pullˡ []∘+₁) ⟩
distributeˡ ∘ (swap +₁ swap) ∘ distributeʳ⁻¹ ∎) distributeˡ ∘ (swap +₁ swap) ∘ distributeʳ⁻¹ ∎)
distributeʳ⁻¹∘swap : ∀ {A B C : Obj} → distributeʳ⁻¹ ∘ swap ≈ (swap +₁ swap) ∘ distributeˡ⁻¹ {A} {B} {C}
distributeʳ⁻¹∘swap = Iso⇒Mono C (IsIso.iso isIsoʳ) (distributeʳ⁻¹ ∘ swap) ((swap +₁ swap) ∘ distributeˡ⁻¹) (begin
(distributeʳ ∘ distributeʳ⁻¹ ∘ swap) ≈⟨ cancelˡ (IsIso.isoʳ isIsoʳ) ⟩
swap ≈⟨ sym (cancelʳ (IsIso.isoʳ isIsoˡ)) ⟩
((swap ∘ distributeˡ) ∘ distributeˡ⁻¹) ≈⟨ (∘[] ⟩∘⟨refl) ⟩
[ swap ∘ (idC ⁂ i₁) , swap ∘ (idC ⁂ i₂) ] ∘ distributeˡ⁻¹ ≈⟨ sym (([]-cong₂ (sym swap∘⁂) (sym swap∘⁂)) ⟩∘⟨refl) ⟩
[ (i₁ ⁂ idC) ∘ swap , (i₂ ⁂ idC) ∘ swap ] ∘ distributeˡ⁻¹ ≈⟨ sym (pullˡ []∘+₁) ⟩
(distributeʳ ∘ (swap +₁ swap) ∘ distributeˡ⁻¹) ∎)
dstr-law₁ : ∀ {A B C} → distributeˡ⁻¹ {A} {B} {C} ∘ (idC ⁂ i₁) ≈ i₁ dstr-law₁ : ∀ {A B C} → distributeˡ⁻¹ {A} {B} {C} ∘ (idC ⁂ i₁) ≈ i₁
dstr-law₁ = (refl⟩∘⟨ (sym inject₁)) ○ (cancelˡ (IsIso.isoˡ isIsoˡ)) dstr-law₁ = (refl⟩∘⟨ (sym inject₁)) ○ (cancelˡ (IsIso.isoˡ isIsoˡ))
dstr-law₂ : ∀ {A B C} → distributeˡ⁻¹ {A} {B} {C} ∘ (idC ⁂ i₂) ≈ i₂ dstr-law₂ : ∀ {A B C} → distributeˡ⁻¹ {A} {B} {C} ∘ (idC ⁂ i₂) ≈ i₂

View file

@ -13,17 +13,17 @@ import Monad.Instance.K as MIK
```agda ```agda
module Monad.Instance.K.Commutative {o e} (ambient : Ambient o e) (MK : MIK.MonadK ambient) where module Monad.Instance.K.Commutative {o e} (ambient : Ambient o e) (MK : MIK.MonadK ambient) where
open Ambient ambient open Ambient ambient
open MIK ambient open MIK ambient
open MonadK MK open MonadK MK
open import Monad.Instance.K.Strong ambient MK open import Monad.Instance.K.Strong ambient MK
open import Category.Construction.UniformIterationAlgebras ambient open import Category.Construction.UniformIterationAlgebras ambient
open import Algebra.UniformIterationAlgebra ambient open import Algebra.UniformIterationAlgebra ambient
open import Algebra.Properties ambient using (FreeUniformIterationAlgebra; uniformForgetfulF; IsStableFreeUniformIterationAlgebra) open import Algebra.Properties ambient using (FreeUniformIterationAlgebra; uniformForgetfulF; IsStableFreeUniformIterationAlgebra)
open Equiv open Equiv
open HomReasoning open HomReasoning
open MR C open MR C
-- open M C -- open M C
``` ```
@ -34,32 +34,49 @@ The proof is analogous to the ones for strength, the relevant diagram is:
<iframe class="quiver-embed" src="https://q.uiver.app/#q=WzAsNyxbMCwxLCJLWCBcXHRpbWVzIEtZIl0sWzEsMCwiSyhLWCBcXHRpbWVzIFkpIl0sWzIsMCwiSyhLKFggXFx0aW1lcyBZKSkiXSxbMywxLCJLKFggXFx0aW1lcyBZKSJdLFsxLDIsIksoWCBcXHRpbWVzIEtZKSJdLFsyLDIsIksoSyhYIFxcdGltZXMgWSkpIl0sWzAsNCwiS1ggXFx0aW1lcyBZIl0sWzAsMSwiXFx0YXUiXSxbMSwyLCJLXFxoYXR7XFx0YXV9Il0sWzIsMywiXFxtdSJdLFswLDQsIlxcaGF0e1xcdGF1fSIsMl0sWzQsNSwiS1xcdGF1IiwyXSxbNSwzLCJcXG11IiwyXSxbNiwwLCJpZCBcXHRpbWVzIFxcZXRhIl0sWzYsMywiXFxoYXR7XFx0YXV9IiwwLHsiY3VydmUiOjV9XSxbMCwzLCJcXGhhdHtcXHRhdX1eXFwjIl1d&embed" width="974" height="688" style="border-radius: 8px; border: none;"></iframe> <iframe class="quiver-embed" src="https://q.uiver.app/#q=WzAsNyxbMCwxLCJLWCBcXHRpbWVzIEtZIl0sWzEsMCwiSyhLWCBcXHRpbWVzIFkpIl0sWzIsMCwiSyhLKFggXFx0aW1lcyBZKSkiXSxbMywxLCJLKFggXFx0aW1lcyBZKSJdLFsxLDIsIksoWCBcXHRpbWVzIEtZKSJdLFsyLDIsIksoSyhYIFxcdGltZXMgWSkpIl0sWzAsNCwiS1ggXFx0aW1lcyBZIl0sWzAsMSwiXFx0YXUiXSxbMSwyLCJLXFxoYXR7XFx0YXV9Il0sWzIsMywiXFxtdSJdLFswLDQsIlxcaGF0e1xcdGF1fSIsMl0sWzQsNSwiS1xcdGF1IiwyXSxbNSwzLCJcXG11IiwyXSxbNiwwLCJpZCBcXHRpbWVzIFxcZXRhIl0sWzYsMywiXFxoYXR7XFx0YXV9IiwwLHsiY3VydmUiOjV9XSxbMCwzLCJcXGhhdHtcXHRhdX1eXFwjIl1d&embed" width="974" height="688" style="border-radius: 8px; border: none;"></iframe>
```agda ```agda
KCommutative : CommutativeMonad {C = C} {V = monoidal} symmetric KStrong
KCommutative = record { commutes = commutes' }
where
open monadK using (μ)
open StrongMonad KStrong using (strengthen)
open IsStableFreeUniformIterationAlgebra using (♯-law; ♯-preserving; ♯-unique)
open Uniform-Iteration-Algebra using (#-Uniformity; #-Fixpoint; #-resp-≈)
-- some helper definitions to make our life easier open monadK using (μ)
open StrongMonad KStrong using (strengthen)
open IsStableFreeUniformIterationAlgebra using (♯-law; ♯-preserving; ♯-unique; ♯ˡ-unique; ♯ˡ-preserving; ♯ˡ-law)
open Uniform-Iteration-Algebra using (#-Uniformity; #-Fixpoint; #-resp-≈)
-- some helper definitions to make our life easier
private
η = λ Z → FreeObject.η (freealgebras Z) η = λ Z → FreeObject.η (freealgebras Z)
_♯ = λ {A X Y} f → IsStableFreeUniformIterationAlgebra.[_,_]♯ {Y = X} (stable X) {X = A} (algebras Y) f _♯ = λ {A X Y} f → IsStableFreeUniformIterationAlgebra.[_,_]♯ {Y = X} (stable X) {X = A} (algebras Y) f
_♯ˡ = λ {A X Y} f → IsStableFreeUniformIterationAlgebra.[_,_]♯ˡ {Y = X} (stable X) {X = A} (algebras Y) f
_# = λ {A} {X} f → Uniform-Iteration-Algebra._# (algebras A) {X = X} f _# = λ {A} {X} f → Uniform-Iteration-Algebra._# (algebras A) {X = X} f
σ : ∀ ((X , Y) : Obj ×f Obj) → K.₀ X × Y ⇒ K.₀ (X × Y) σ : ∀ ((X , Y) : Obj ×f Obj) → K.₀ X × Y ⇒ K.₀ (X × Y)
σ _ = K.₁ swap ∘ (τ _) ∘ swap σ _ = K.₁ swap ∘ (τ _) ∘ swap
σ-preserve : ∀ {X Y Z : Obj} (h : Z ⇒ K.₀ Y + Z) → σ (Y , X) ∘ (h # ⁂ idC) ≈ ((σ _ +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC))# proof-principle : ∀ {X Y} (f g : K.₀ X × K.₀ Y ⇒ K.₀ (X × Y)) → f ∘ (η _ ⁂ η _) ≈ g ∘ (η _ ⁂ η _) → (∀ {A} (h : A ⇒ K.₀ Y + A) → f ∘ (idC ⁂ h #) ≈ ((f +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))#) → (∀ {A} (h : A ⇒ K.₀ X + A) → f ∘ (h # ⁂ idC) ≈ ((f +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC)) #) → (∀ {A} (h : A ⇒ K.₀ Y + A) → g ∘ (idC ⁂ h #) ≈ ((g +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))#) → (∀ {A} (h : A ⇒ K.₀ X + A) → g ∘ (h # ⁂ idC) ≈ ((g +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC)) #) → f ≈ g
{- proof-principle {X} {Y} f g η-eq f-iter₁ f-iter₂ g-iter₁ g-iter₂ = begin
K.₁ swap ∘ τ ∘ swap ∘ (h # ⁂ idC) f ≈⟨ ♯-unique (stable _) (f ∘ (idC ⁂ η Y)) f refl (λ h → f-iter₁ h) ⟩
≈ K.₁ swap ∘ τ ∘ (idC ⁂ h #) ∘ swap (f ∘ (idC ⁂ η _)) ♯ ≈⟨ sym (♯-unique (stable _) (f ∘ (idC ⁂ η Y)) g helper₁ {! !}) ⟩
≈ K.₁ swap ∘ ()# ∘ swap g ∎
≈ ((K.₁ swap +₁ τ) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))# ∘ swap where
-} helper₁ : f ∘ (idC ⁂ η Y) ≈ g ∘ (idC ⁂ η Y)
σ-preserve {Z} h = {! !} helper₁ = begin
σ-preserve' : ∀ {X Y Z : Obj} (h : Z ⇒ K.₀ Y + Z) → σ (X , K.₀ Y) ∘ (idC ⁂ h #) ≈ ((σ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))# f ∘ (idC ⁂ η Y) ≈⟨ ♯ˡ-unique (stable _) (f ∘ (η X ⁂ η Y)) (f ∘ (idC ⁂ η Y)) (sym (pullʳ (⁂∘⁂ ○ (⁂-cong₂ identityˡ identityʳ)))) {! !} ⟩
σ-preserve' {Z} h = {! !} (f ∘ (η X ⁂ η Y)) ♯ˡ ≈⟨ {! !} ⟩
g ∘ (idC ⁂ η Y) ∎
where
comm₁ : ∀ {Z : Obj} (h : Z ⇒ K.₀ X + Z) → (f ∘ (idC ⁂ η Y)) ∘ (h # ⁂ idC) ≈ ((f ∘ (idC ⁂ η Y) +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC)) #
comm₁ {Z} h = begin
(f ∘ (idC ⁂ η Y)) ∘ (h # ⁂ idC) ≈⟨ pullʳ ⁂∘⁂ ⟩
f ∘ (idC ∘ h # ⁂ η Y ∘ idC) ≈⟨ refl⟩∘⟨ (⁂-cong₂ id-comm-sym id-comm) ⟩
f ∘ (h # ∘ idC ⁂ idC ∘ η Y) ≈⟨ refl⟩∘⟨ sym ⁂∘⁂ ⟩
f ∘ (h # ⁂ idC) ∘ (idC ⁂ η Y) ≈⟨ pullˡ (f-iter₂ h) ⟩
(((f +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC)) #) ∘ (idC ⁂ η Y) ≈⟨ sym (#-Uniformity (algebras _) uni-helper) ⟩
((f ∘ (idC ⁂ η Y) +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC)) # ∎
where
uni-helper : (idC +₁ idC ⁂ η Y) ∘ (f ∘ (idC ⁂ η Y) +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC) ≈ ((f +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC)) ∘ (idC ⁂ η Y)
uni-helper = {! !}
KCommutative : CommutativeMonad {C = C} {V = monoidal} symmetric KStrong
KCommutative = record { commutes = commutes' }
where
commutes' : ∀ {X Y : Obj} → μ.η _ ∘ K.₁ (σ _) ∘ τ (K.₀ X , Y) ≈ μ.η _ ∘ K.₁ (τ _) ∘ σ _ commutes' : ∀ {X Y : Obj} → μ.η _ ∘ K.₁ (σ _) ∘ τ (K.₀ X , Y) ≈ μ.η _ ∘ K.₁ (τ _) ∘ σ _
commutes' {X} {Y} = begin commutes' {X} {Y} = begin
μ.η _ ∘ K.₁ (σ _) ∘ τ _ ≈⟨ ♯-unique (stable _) (σ _) (μ.η (X × Y) ∘ K.₁ (σ _) ∘ τ _) comm₁ comm₂ ⟩ μ.η _ ∘ K.₁ (σ _) ∘ τ _ ≈⟨ ♯-unique (stable _) (σ _) (μ.η (X × Y) ∘ K.₁ (σ _) ∘ τ _) comm₁ comm₂ ⟩
@ -99,36 +116,8 @@ The proof is analogous to the ones for strength, the relevant diagram is:
comm₄ {Z} h = begin comm₄ {Z} h = begin
(μ.η _ ∘ K.₁ (τ _) ∘ σ _) ∘ (idC ⁂ h #) ≈⟨ {! !} ⟩ (μ.η _ ∘ K.₁ (τ _) ∘ σ _) ∘ (idC ⁂ h #) ≈⟨ {! !} ⟩
(μ.η _ ∘ K.₁ (τ _) ∘ σ _) ∘ ((i₁ # ∘ idC) ⁂ h #) ≈˘⟨ {! !} ⟩ (μ.η _ ∘ K.₁ (τ _) ∘ σ _) ∘ ((i₁ # ∘ idC) ⁂ h #) ≈˘⟨ {! !} ⟩
(μ.η _ ∘ K.₁ (τ _) ∘ σ _) ∘ (((i₁ #) ⁂ h #)) ≈˘⟨ refl⟩∘⟨ ⟨⟩-cong₂ (#-Uniformity (algebras _) helper₁) {! !} ⟩
(μ.η _ ∘ K.₁ (τ _) ∘ σ _) ∘ ⟨ ((π₁ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # , ((π₂ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # ⟩ ≈⟨ {! !} ⟩
{! !} ≈⟨ {! !} ⟩ {! !} ≈⟨ {! !} ⟩
{! !} ≈⟨ {! !} ⟩ {! !} ≈⟨ {! !} ⟩
{! !} ≈⟨ {! !} ⟩ {! !} ≈⟨ {! !} ⟩
((μ.η _ ∘ K.₁ (τ _) ∘ σ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))# ∎ ((μ.η _ ∘ K.₁ (τ _) ∘ σ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))# ∎
where
-- this leads nowhere
helper₁ : (idC +₁ π₁) ∘ (π₁ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h) ≈ i₁ ∘ π₁
helper₁ = begin
(idC +₁ π₁) ∘ (π₁ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h) ≈⟨ pullˡ (+₁∘+₁ ○ +₁-cong₂ identityˡ identityʳ) ⟩
(π₁ +₁ π₁) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h) ≈⟨ {! !} ⟩
i₁ ∘ π₁ ∘ (idC ⁂ h) ≈⟨ refl⟩∘⟨ π₁∘⁂ ⟩
i₁ ∘ idC ∘ π₁ ≈⟨ refl⟩∘⟨ identityˡ ⟩
i₁ ∘ π₁ ∎
test : ((τ (X , Y) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))# ∘ swap ≈ ((τ (X , Y) ∘ swap +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC)) #
test = sym (#-Uniformity (algebras _) (sym (begin
((τ (X , Y) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) ∘ swap ≈⟨ pullʳ (pullʳ (sym swap∘⁂)) ⟩
(τ (X , Y) +₁ idC) ∘ distributeˡ⁻¹ ∘ swap ∘ (h ⁂ idC) ≈⟨ refl⟩∘⟨ (pullˡ distributeˡ⁻¹∘swap) ⟩
(τ (X , Y) +₁ idC) ∘ ((swap +₁ swap) ∘ distributeʳ⁻¹) ∘ (h ⁂ idC) ≈⟨ pullˡ (pullˡ (+₁∘+₁ ○ +₁-cong₂ (sym identityˡ) id-comm-sym)) ⟩
((idC ∘ τ (X , Y) ∘ swap +₁ swap ∘ idC) ∘ distributeʳ⁻¹) ∘ (h ⁂ idC) ≈⟨ assoc ○ (sym +₁∘+₁) ⟩∘⟨refl ⟩
((idC +₁ swap) ∘ (τ (X , Y) ∘ swap +₁ idC)) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC) ≈⟨ assoc ⟩
(idC +₁ swap) ∘ (τ (X , Y) ∘ swap +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC) ∎)))
helper : τ _ ∘ (h # ⁂ idC) ∘ swap ≈ ((τ _ +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC)) # ∘ swap
helper = {! !}
τ∘swap-preserving : τ (K.₀ Y , X) ∘ (h # ⁂ idC) ≈ ((τ _ +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC)) #
τ∘swap-preserving = begin
τ (K.₀ Y , X) ∘ (h # ⁂ idC) ≈⟨ {! !} ⟩
τ (K.₀ Y , X) ∘ (h # ⁂ K.₁ idC) ≈⟨ {! !} ⟩
K.₁ (h # ⁂ idC) ∘ τ _ ≈⟨ {! !} ⟩
{! !} ≈⟨ {! !} ⟩
((τ _ +₁ idC) ∘ distributeʳ⁻¹ ∘ (h ⁂ idC)) # ∎
``` ```