mirror of
https://git8.cs.fau.de/theses/bsc-leon-vatthauer.git
synced 2024-05-31 07:28:34 +02:00
362 KiB
362 KiB
module Monad.Instance.K.Strong {o ℓ e} (ambient : Ambient o ℓ e) (MK : MIK.MonadK ambient) where open Ambient ambient open import Category.Construction.ElgotAlgebras cocartesian open import Algebra.Elgot cocartesian open import Algebra.Elgot.Free cocartesian using (FreeElgotAlgebra; elgotForgetfulF) open import Algebra.Elgot.Stable distributive using (IsStableFreeElgotAlgebra) open MIK ambient open MonadK MK open Equiv open MR C open M C open HomReasoning
The monad K is strong
K is a strong monad with the strength defined as η ♯
, where ♯ is the operator we get from stability.
Verifying the axioms of strength is straightforward once you know the procedure, since the proofs are all very similar.
For example the proof of identityˡ
i.e. K₁ π₂ ∘ τ ≈ π₂
goes as follows:
- find a morphism
f
such thatK₁ π₂ ∘ τ ≈ f ♯ ≈ π₂
- show that
K₁ π₂ ∘ τ
is iteration preserving and satisfies the stabiltiy law - show that
π₂
is iteration preserving and satisfies the stabiltiy law
=> by uniqueness of f ♯
we are done
The following diagram demonstrates this:
-- we use properties of the kleisli representation as well as the 'normal' monad representation open kleisliK using (extend) open monadK using (μ) -- defining τ private -- some helper definitions to make our life easier η = λ Z → FreeObject.η (freealgebras Z) _♯ = λ {A X Y} f → IsStableFreeElgotAlgebra.[_,_]♯ {Y = X} (stable X) {X = A} (algebras Y) f _# = λ {A} {X} f → Elgot-Algebra._# (algebras A) {X = X} f open IsStableFreeElgotAlgebra using (♯-law; ♯-preserving; ♯-unique) open Elgot-Algebra using (#-Uniformity; #-Fixpoint; #-resp-≈) module _ (P : Category.Obj (CProduct C C)) where private X = proj₁ P Y = proj₂ P τ : X × K.₀ Y ⇒ K.₀ (X × Y) τ = η (X × Y) ♯ τ-η : τ ∘ (idC ⁂ η Y) ≈ η (X × Y) τ-η = sym (♯-law (stable Y) (η (X × Y))) -- for K not only strengthening with 1 is irrelevant τ-π₂ : K.₁ π₂ ∘ τ ≈ π₂ τ-π₂ = begin K.₁ π₂ ∘ τ ≈⟨ ♯-unique (stable Y) (η _ ∘ π₂) (K.₁ π₂ ∘ τ) comm₁ comm₂ ⟩ (η _ ∘ π₂) ♯ ≈⟨ sym (♯-unique (stable Y) (η _ ∘ π₂) π₂ (sym π₂∘⁂) comm₃) ⟩ π₂ ∎ where comm₁ : η _ ∘ π₂ ≈ (K.₁ π₂ ∘ τ) ∘ (idC ⁂ η _) comm₁ = sym (begin (K.₁ π₂ ∘ τ) ∘ (idC ⁂ η _) ≈⟨ pullʳ τ-η ⟩ K.₁ π₂ ∘ η _ ≈⟨ (sym (F₁⇒extend monadK π₂)) ⟩∘⟨refl ⟩ extend (η _ ∘ π₂) ∘ η _ ≈⟨ kleisliK.identityʳ ⟩ η _ ∘ π₂ ∎) comm₂ : ∀ {Z : Obj} (h : Z ⇒ K.₀ Y + Z) → (K.₁ π₂ ∘ τ) ∘ (idC ⁂ h # ) ≈ ((K.₁ π₂ ∘ τ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))# comm₂ {Z} h = begin (K.₁ π₂ ∘ τ) ∘ (idC ⁂ h #) ≈⟨ pullʳ (♯-preserving (stable Y) (η _) h) ⟩ K.₁ π₂ ∘ ((τ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # ≈⟨ Elgot-Algebra-Morphism.preserves ((freealgebras (X × Y) FreeObject.*) (η _ ∘ π₂)) ⟩ ((K.₁ π₂ +₁ idC) ∘ (τ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # ≈⟨ #-resp-≈ (algebras Y) (pullˡ (+₁∘+₁ ○ +₁-cong₂ refl identity²)) ⟩ ((K.₁ π₂ ∘ τ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))# ∎ comm₃ : ∀ {Z : Obj} (h : Z ⇒ K.₀ Y + Z) → π₂ ∘ (idC ⁂ h #) ≈ ((π₂ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # comm₃ {Z} h = begin π₂ ∘ (idC ⁂ h #) ≈⟨ π₂∘⁂ ⟩ h # ∘ π₂ ≈⟨ sym (#-Uniformity (algebras Y) uni-helper) ⟩ ((π₂ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # ∎ where uni-helper = begin (idC +₁ π₂) ∘ (π₂ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h) ≈⟨ pullˡ +₁∘+₁ ⟩ (idC ∘ π₂ +₁ π₂ ∘ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h) ≈⟨ (+₁-cong₂ identityˡ identityʳ) ⟩∘⟨refl ⟩ (π₂ +₁ π₂) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h) ≈⟨ pullˡ distributeˡ⁻¹-π₂ ⟩ π₂ ∘ (idC ⁂ h) ≈⟨ project₂ ⟩ h ∘ π₂ ∎ τ-comm : ∀ {X Y Z : Obj} (h : Z ⇒ K.₀ Y + Z) → τ (X , Y) ∘ (idC ⁂ h #) ≈ ((τ (X , Y) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))# τ-comm {X} {Y} {Z} h = ♯-preserving (stable Y) (η (X × Y)) h K₁η : ∀ {X Y} (f : X ⇒ Y) → K.₁ f ∘ η X ≈ η Y ∘ f K₁η {X} {Y} f = begin K.₁ f ∘ η X ≈⟨ (sym (F₁⇒extend monadK f)) ⟩∘⟨refl ⟩ extend (η Y ∘ f) ∘ η X ≈⟨ kleisliK.identityʳ ⟩ η Y ∘ f ∎ KStrength : Strength monoidal monadK KStrength = record { strengthen = ntHelper (record { η = τ ; commute = commute' }) ; identityˡ = λ {X} → τ-π₂ (Terminal.⊤ terminal , X) ; η-comm = λ {A} {B} → τ-η (A , B) ; μ-η-comm = μ-η-comm' ; strength-assoc = strength-assoc' } where commute' : ∀ {P₁ : Category.Obj (CProduct C C)} {P₂ : Category.Obj (CProduct C C)} (fg : _[_,_] (CProduct C C) P₁ P₂) → τ P₂ ∘ ((proj₁ fg) ⁂ K.₁ (proj₂ fg)) ≈ K.₁ ((proj₁ fg) ⁂ (proj₂ fg)) ∘ τ P₁ commute' {(U , V)} {(W , X)} (f , g) = begin τ _ ∘ (f ⁂ K.₁ g) ≈⟨ ♯-unique (stable V) (η (W × X) ∘ (f ⁂ g)) (τ _ ∘ (f ⁂ K.₁ g)) comm₁ comm₂ ⟩ (η _ ∘ (f ⁂ g)) ♯ ≈⟨ sym (♯-unique (stable V) (η (W × X) ∘ (f ⁂ g)) (K.₁ (f ⁂ g) ∘ τ _) comm₃ comm₄) ⟩ K.₁ (f ⁂ g) ∘ τ _ ∎ where comm₁ : η (W × X) ∘ (f ⁂ g) ≈ (τ (W , X) ∘ (f ⁂ K.₁ g)) ∘ (idC ⁂ η V) comm₁ = sym (begin (τ (W , X) ∘ (f ⁂ K.₁ g)) ∘ (idC ⁂ η V) ≈⟨ pullʳ ⁂∘⁂ ⟩ τ (W , X) ∘ (f ∘ idC ⁂ K.₁ g ∘ η V) ≈⟨ refl⟩∘⟨ (⁂-cong₂ id-comm (K₁η g)) ⟩ τ (W , X) ∘ (idC ∘ f ⁂ η X ∘ g) ≈⟨ refl⟩∘⟨ (sym ⁂∘⁂) ⟩ τ (W , X) ∘ (idC ⁂ η X) ∘ (f ⁂ g) ≈⟨ pullˡ (τ-η (W , X)) ⟩ η (W × X) ∘ (f ⁂ g) ∎) comm₃ : η (W × X) ∘ (f ⁂ g) ≈ (K.₁ (f ⁂ g) ∘ τ (U , V)) ∘ (idC ⁂ η V) comm₃ = sym (begin (K.₁ (f ⁂ g) ∘ τ (U , V)) ∘ (idC ⁂ η V) ≈⟨ pullʳ (τ-η (U , V)) ⟩ K.₁ (f ⁂ g) ∘ η (U × V) ≈⟨ K₁η (f ⁂ g) ⟩ η (W × X) ∘ (f ⁂ g) ∎) comm₂ : ∀ {Z : Obj} (h : Z ⇒ K.₀ V + Z) → (τ (W , X) ∘ (f ⁂ K.₁ g)) ∘ (idC ⁂ h #) ≈ ((τ (W , X) ∘ (f ⁂ K.₁ g) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))# comm₂ {Z} h = begin (τ (W , X) ∘ (f ⁂ K.₁ g)) ∘ (idC ⁂ h #) ≈⟨ pullʳ ⁂∘⁂ ⟩ τ (W , X) ∘ (f ∘ idC ⁂ K.₁ g ∘ (h #)) ≈⟨ refl⟩∘⟨ (⁂-cong₂ id-comm ((Elgot-Algebra-Morphism.preserves (((freealgebras _) FreeObject.*) (η X ∘ g))) ○ sym identityʳ)) ⟩ τ (W , X) ∘ (idC ∘ f ⁂ ((K.₁ g +₁ idC) ∘ h) # ∘ idC) ≈⟨ refl⟩∘⟨ (sym ⁂∘⁂) ⟩ τ (W , X) ∘ (idC ⁂ ((K.₁ g +₁ idC) ∘ h) #) ∘ (f ⁂ idC) ≈⟨ pullˡ (♯-preserving (stable _) (η _) ((K.₁ g +₁ idC) ∘ h)) ⟩ ((τ (W , X) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ (K.₁ g +₁ idC) ∘ h)) # ∘ (f ⁂ idC) ≈⟨ sym (#-Uniformity (algebras _) uni-helper) ⟩ ((τ (W , X) ∘ (f ⁂ K.₁ g) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))# ∎ where uni-helper = begin (idC +₁ f ⁂ idC) ∘ (τ (W , X) ∘ (f ⁂ K.₁ g) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h) ≈⟨ pullˡ +₁∘+₁ ⟩ (idC ∘ τ (W , X) ∘ (f ⁂ K.₁ g) +₁ (f ⁂ idC) ∘ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h) ≈⟨ (+₁-cong₂ identityˡ id-comm) ⟩∘⟨refl ⟩ (τ (W , X) ∘ (f ⁂ K.₁ g) +₁ idC ∘ (f ⁂ idC)) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h) ≈⟨ (sym +₁∘+₁) ⟩∘⟨refl ⟩ ((τ (W , X) +₁ idC) ∘ ((f ⁂ K.₁ g) +₁ (f ⁂ idC))) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h) ≈⟨ pullʳ (pullˡ (distributeˡ⁻¹-natural f (K.₁ g) idC)) ⟩ (τ (W , X) +₁ idC) ∘ (distributeˡ⁻¹ ∘ (f ⁂ (K.₁ g +₁ idC))) ∘ (idC ⁂ h) ≈⟨ refl⟩∘⟨ (pullʳ (⁂∘⁂ ○ ⁂-cong₂ identityʳ refl)) ⟩ (τ (W , X) +₁ idC) ∘ distributeˡ⁻¹ ∘ (f ⁂ (K.₁ g +₁ idC) ∘ h) ≈˘⟨ pullʳ (pullʳ (⁂∘⁂ ○ ⁂-cong₂ identityˡ identityʳ)) ⟩ ((τ (W , X) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ (K.₁ g +₁ idC) ∘ h)) ∘ (f ⁂ idC) ∎ comm₄ : ∀ {Z : Obj} (h : Z ⇒ K.₀ V + Z) → (K.₁ (f ⁂ g) ∘ τ (U , V)) ∘ (idC ⁂ h #) ≈ ((K.₁ (f ⁂ g) ∘ τ (U , V) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # comm₄ {Z} h = begin (K.₁ (f ⁂ g) ∘ τ (U , V)) ∘ (idC ⁂ (h #)) ≈⟨ pullʳ (τ-comm h) ⟩ K.₁ (f ⁂ g) ∘ ((τ (U , V) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # ≈⟨ Elgot-Algebra-Morphism.preserves (((freealgebras _) FreeObject.*) (η (W × X) ∘ (f ⁂ g))) ⟩ ((K.₁ (f ⁂ g) +₁ idC) ∘ (τ (U , V) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # ≈⟨ #-resp-≈ (algebras (W × X)) (pullˡ (+₁∘+₁ ○ +₁-cong₂ refl identity²)) ⟩ ((K.₁ (f ⁂ g) ∘ τ (U , V) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # ∎ μ-η-comm' : ∀ {A B} → μ.η _ ∘ K.₁ (τ _) ∘ τ (A , K.₀ B) ≈ τ _ ∘ (idC ⁂ μ.η _) μ-η-comm' {A} {B} = begin μ.η _ ∘ K.₁ (τ _) ∘ τ _ ≈⟨ ♯-unique (stable (K.₀ B)) (τ (A , B)) (μ.η _ ∘ K.₁ (τ _) ∘ τ _) comm₁ comm₂ ⟩ (τ _ ♯) ≈⟨ sym (♯-unique (stable (K.₀ B)) (τ (A , B)) (τ _ ∘ (idC ⁂ μ.η _)) (sym (cancelʳ (⁂∘⁂ ○ ⁂-cong₂ identity² monadK.identityʳ ○ ⟨⟩-unique id-comm id-comm))) comm₃) ⟩ τ _ ∘ (idC ⁂ μ.η _) ∎ where comm₁ : τ (A , B) ≈ (μ.η _ ∘ K.₁ (τ _) ∘ τ _) ∘ (idC ⁂ η _) comm₁ = sym (begin (μ.η _ ∘ K.₁ (τ _) ∘ τ _) ∘ (idC ⁂ η _) ≈⟨ pullʳ (pullʳ (τ-η _)) ⟩ μ.η _ ∘ K.₁ (τ _) ∘ η _ ≈⟨ refl⟩∘⟨ (K₁η (τ (A , B))) ⟩ μ.η _ ∘ η _ ∘ τ _ ≈⟨ cancelˡ monadK.identityʳ ⟩ τ _ ∎) comm₂ : ∀ {Z : Obj} (h : Z ⇒ K.₀ (K.₀ B) + Z) → (μ.η _ ∘ K.₁ (τ _) ∘ τ _) ∘ (idC ⁂ h #) ≈ ((μ.η _ ∘ K.₁ (τ (A , B)) ∘ τ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # comm₂ {Z} h = begin (μ.η _ ∘ K.₁ (τ _) ∘ τ _) ∘ (idC ⁂ h #) ≈⟨ pullʳ (pullʳ (τ-comm h)) ⟩ μ.η _ ∘ K.₁ (τ _) ∘ (((τ (A , K.₀ B) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) #) ≈⟨ refl⟩∘⟨ (Elgot-Algebra-Morphism.preserves (((freealgebras _) FreeObject.*) (η _ ∘ τ _))) ⟩ μ.η _ ∘ ((K.₁ (τ _) +₁ idC) ∘ (τ (A , K.₀ B) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # ≈⟨ Elgot-Algebra-Morphism.preserves (((freealgebras _) FreeObject.*) idC) ⟩ ((μ.η _ +₁ idC) ∘ (K.₁ (τ _) +₁ idC) ∘ (τ (A , K.₀ B) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # ≈⟨ #-resp-≈ (algebras _) (pullˡ +₁∘+₁) ⟩ ((μ.η _ ∘ K.₁ (τ _) +₁ idC ∘ idC) ∘ (τ (A , K.₀ B) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # ≈⟨ #-resp-≈ (algebras _) (pullˡ +₁∘+₁) ⟩ (((μ.η _ ∘ K.₁ (τ _)) ∘ τ _ +₁ (idC ∘ idC) ∘ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # ≈⟨ #-resp-≈ (algebras _) ((+₁-cong₂ assoc (cancelʳ identity²)) ⟩∘⟨refl) ⟩ ((μ.η _ ∘ K.₁ (τ (A , B)) ∘ τ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # ∎ comm₃ : ∀ {Z : Obj} (h : Z ⇒ K.₀ (K.₀ B) + Z) → (τ _ ∘ (idC ⁂ μ.η _)) ∘ (idC ⁂ h #) ≈ ((τ _ ∘ (idC ⁂ μ.η _) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # comm₃ {Z} h = begin (τ _ ∘ (idC ⁂ μ.η _)) ∘ (idC ⁂ h #) ≈⟨ pullʳ ⁂∘⁂ ⟩ τ _ ∘ (idC ∘ idC ⁂ μ.η _ ∘ h #) ≈⟨ refl⟩∘⟨ (⁂-cong₂ identity² (Elgot-Algebra-Morphism.preserves (((freealgebras _) FreeObject.*) idC))) ⟩ τ _ ∘ (idC ⁂ ((μ.η _ +₁ idC) ∘ h) #) ≈⟨ τ-comm ((μ.η B +₁ idC) ∘ h) ⟩ ((τ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ (μ.η B +₁ idC) ∘ h)) # ≈⟨ #-resp-≈ (algebras _) (refl⟩∘⟨ (refl⟩∘⟨ (⁂-cong₂ (sym identity²) refl ○ sym ⁂∘⁂))) ⟩ ((τ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ (μ.η B +₁ idC)) ∘ (idC ⁂ h)) # ≈⟨ #-resp-≈ (algebras _) (refl⟩∘⟨ (pullˡ (sym (distributeˡ⁻¹-natural idC (μ.η B) idC)))) ⟩ ((τ _ +₁ idC) ∘ ((idC ⁂ μ.η B +₁ idC ⁂ idC) ∘ distributeˡ⁻¹) ∘ (idC ⁂ h)) # ≈⟨ #-resp-≈ (algebras _) (pullˡ (pullˡ (+₁∘+₁ ○ +₁-cong₂ refl (elimʳ (⟨⟩-unique id-comm id-comm))))) ⟩ (((τ _ ∘ (idC ⁂ μ.η B) +₁ idC) ∘ distributeˡ⁻¹) ∘ (idC ⁂ h)) # ≈⟨ #-resp-≈ (algebras _) assoc ⟩ ((τ _ ∘ (idC ⁂ μ.η _) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # ∎ strength-assoc' : ∀ {X Y Z} → K.₁ assocˡ ∘ τ (X × Y , Z) ≈ τ (X , Y × Z) ∘ (idC ⁂ τ (Y , Z)) ∘ assocˡ strength-assoc' {X} {Y} {Z} = begin K.₁ assocˡ ∘ τ _ ≈⟨ ♯-unique (stable _) (η (X × Y × Z) ∘ assocˡ) (K.₁ assocˡ ∘ τ _) (sym (pullʳ (τ-η _) ○ K₁η _)) comm₁ ⟩ ((η (X × Y × Z) ∘ assocˡ) ♯) ≈⟨ sym (♯-unique (stable _) (η (X × Y × Z) ∘ assocˡ) (τ _ ∘ (idC ⁂ τ _) ∘ assocˡ) comm₂ comm₃) ⟩ τ _ ∘ (idC ⁂ τ _) ∘ assocˡ ∎ where comm₁ : ∀ {A : Obj} (h : A ⇒ K.₀ Z + A) → (K.₁ assocˡ ∘ τ _) ∘ (idC ⁂ h #) ≈ ((K.₁ assocˡ ∘ τ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # comm₁ {A} h = begin (K.₁ assocˡ ∘ τ _) ∘ (idC ⁂ h #) ≈⟨ pullʳ (τ-comm h) ⟩ K.₁ assocˡ ∘ ((τ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))# ≈⟨ Elgot-Algebra-Morphism.preserves (((freealgebras _) FreeObject.*) _) ⟩ ((K.₁ assocˡ +₁ idC) ∘ (τ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))# ≈⟨ #-resp-≈ (algebras _) (pullˡ (+₁∘+₁ ○ +₁-cong₂ refl identity²)) ⟩ ((K.₁ assocˡ ∘ τ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # ∎ comm₂ : η (X × Y × Z) ∘ assocˡ ≈ (τ _ ∘ (idC ⁂ τ _) ∘ assocˡ) ∘ (idC ⁂ η _) comm₂ = sym (begin (τ _ ∘ (idC ⁂ τ _) ∘ assocˡ) ∘ (idC ⁂ η _) ≈⟨ (refl⟩∘⟨ ⁂∘⟨⟩) ⟩∘⟨refl ⟩ (τ _ ∘ ⟨ idC ∘ π₁ ∘ π₁ , τ _ ∘ ⟨ π₂ ∘ π₁ , π₂ ⟩ ⟩) ∘ (idC ⁂ η _) ≈⟨ pullʳ ⟨⟩∘ ⟩ τ _ ∘ ⟨ (idC ∘ π₁ ∘ π₁) ∘ (idC ⁂ η _) , (τ _ ∘ ⟨ π₂ ∘ π₁ , π₂ ⟩) ∘ (idC ⁂ η _) ⟩ ≈⟨ refl⟩∘⟨ (⟨⟩-cong₂ (identityˡ ⟩∘⟨refl ○ pullʳ π₁∘⁂) (pullʳ ⟨⟩∘)) ⟩ τ _ ∘ ⟨ π₁ ∘ idC ∘ π₁ , τ _ ∘ ⟨ (π₂ ∘ π₁) ∘ (idC ⁂ η _) , π₂ ∘ (idC ⁂ η _) ⟩ ⟩ ≈⟨ refl⟩∘⟨ (⟨⟩-cong₂ (refl⟩∘⟨ identityˡ) (refl⟩∘⟨ (⟨⟩-cong₂ (pullʳ π₁∘⁂) π₂∘⁂))) ⟩ τ _ ∘ ⟨ π₁ ∘ π₁ , τ _ ∘ ⟨ π₂ ∘ idC ∘ π₁ , η _ ∘ π₂ ⟩ ⟩ ≈⟨ refl⟩∘⟨ (⟨⟩-cong₂ (sym identityˡ) (refl⟩∘⟨ ((⟨⟩-cong₂ (sym identityˡ) refl) ○ sym ⁂∘⟨⟩))) ⟩ τ _ ∘ ⟨ idC ∘ π₁ ∘ π₁ , τ _ ∘ (idC ⁂ η _) ∘ ⟨ π₂ ∘ idC ∘ π₁ , π₂ ⟩ ⟩ ≈⟨ refl⟩∘⟨ (⟨⟩-cong₂ refl (pullˡ (τ-η (Y , Z)))) ⟩ τ _ ∘ ⟨ idC ∘ π₁ ∘ π₁ , η _ ∘ ⟨ π₂ ∘ idC ∘ π₁ , π₂ ⟩ ⟩ ≈⟨ refl⟩∘⟨ (sym ⁂∘⟨⟩) ⟩ τ _ ∘ (idC ⁂ η _) ∘ ⟨ π₁ ∘ π₁ , ⟨ π₂ ∘ idC ∘ π₁ , π₂ ⟩ ⟩ ≈⟨ pullˡ (τ-η _) ⟩ η _ ∘ ⟨ π₁ ∘ π₁ , ⟨ π₂ ∘ idC ∘ π₁ , π₂ ⟩ ⟩ ≈⟨ refl⟩∘⟨ ⟨⟩-cong₂ refl (⟨⟩-cong₂ (refl⟩∘⟨ identityˡ) refl) ⟩ η (X × Y × Z) ∘ assocˡ ∎) comm₃ : ∀ {A : Obj} (h : A ⇒ K.₀ Z + A) → (τ _ ∘ (idC ⁂ τ _) ∘ assocˡ) ∘ (idC ⁂ h #) ≈ ((τ _ ∘ (idC ⁂ τ _) ∘ assocˡ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # comm₃ {A} h = begin (τ _ ∘ (idC ⁂ τ _) ∘ assocˡ) ∘ (idC ⁂ h #) ≈⟨ (refl⟩∘⟨ ⁂∘⟨⟩) ⟩∘⟨refl ⟩ (τ _ ∘ ⟨ idC ∘ π₁ ∘ π₁ , τ _ ∘ ⟨ π₂ ∘ π₁ , π₂ ⟩ ⟩) ∘ (idC ⁂ h #) ≈⟨ pullʳ ⟨⟩∘ ⟩ τ _ ∘ ⟨ (idC ∘ π₁ ∘ π₁) ∘ (idC ⁂ h #) , (τ _ ∘ ⟨ π₂ ∘ π₁ , π₂ ⟩) ∘ (idC ⁂ h #) ⟩ ≈⟨ refl⟩∘⟨ (⟨⟩-cong₂ (identityˡ ⟩∘⟨refl ○ pullʳ π₁∘⁂) (pullʳ ⟨⟩∘)) ⟩ τ _ ∘ ⟨ π₁ ∘ idC ∘ π₁ , τ _ ∘ ⟨ (π₂ ∘ π₁) ∘ (idC ⁂ h #) , π₂ ∘ (idC ⁂ h #) ⟩ ⟩ ≈⟨ refl⟩∘⟨ ⟨⟩-cong₂ (refl⟩∘⟨ identityˡ) (refl⟩∘⟨ (⟨⟩-cong₂ (pullʳ π₁∘⁂) π₂∘⁂)) ⟩ τ _ ∘ ⟨ π₁ ∘ π₁ , τ _ ∘ ⟨ π₂ ∘ idC ∘ π₁ , h # ∘ π₂ ⟩ ⟩ ≈⟨ refl⟩∘⟨ (⟨⟩-cong₂ refl (refl⟩∘⟨ (⟨⟩-cong₂ ((refl⟩∘⟨ identityˡ) ○ sym identityˡ) refl))) ⟩ τ _ ∘ ⟨ π₁ ∘ π₁ , τ _ ∘ ⟨ idC ∘ π₂ ∘ π₁ , h # ∘ π₂ ⟩ ⟩ ≈⟨ refl⟩∘⟨ ⟨⟩-cong₂ refl (refl⟩∘⟨ (sym ⁂∘⟨⟩)) ⟩ τ _ ∘ ⟨ π₁ ∘ π₁ , τ _ ∘ (idC ⁂ h #) ∘ ⟨ π₂ ∘ π₁ , π₂ ⟩ ⟩ ≈⟨ refl⟩∘⟨ (⟨⟩-cong₂ (sym identityˡ) (pullˡ (τ-comm h))) ⟩ τ _ ∘ ⟨ idC ∘ π₁ ∘ π₁ , (((τ (Y , Z) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) #) ∘ ⟨ π₂ ∘ π₁ , π₂ ⟩ ⟩ ≈⟨ refl⟩∘⟨ (sym ⁂∘⟨⟩) ⟩ τ _ ∘ (idC ⁂ ((τ (Y , Z) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) #) ∘ assocˡ ≈⟨ pullˡ (τ-comm _) ⟩ ((τ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ (τ (Y , Z) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))) # ∘ assocˡ ≈⟨ sym (#-Uniformity (algebras _) uni-helper) ⟩ ((τ _ ∘ (idC ⁂ τ _) ∘ assocˡ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) # ∎ where uni-helper : (idC +₁ assocˡ) ∘ (τ _ ∘ (idC ⁂ τ (Y , Z)) ∘ assocˡ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h) ≈ ((τ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ (τ (Y , Z) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))) ∘ assocˡ uni-helper = begin (idC +₁ assocˡ) ∘ (τ _ ∘ (idC ⁂ τ (Y , Z)) ∘ assocˡ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h) ≈⟨ pullˡ +₁∘+₁ ⟩ (idC ∘ τ _ ∘ (idC ⁂ τ (Y , Z)) ∘ assocˡ +₁ assocˡ ∘ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h) ≈⟨ (+₁-cong₂ identityˡ id-comm) ⟩∘⟨refl ⟩ (τ _ ∘ (idC ⁂ τ (Y , Z)) ∘ assocˡ +₁ idC ∘ assocˡ) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h) ≈˘⟨ (+₁∘+₁ ○ +₁-cong₂ assoc refl) ⟩∘⟨refl ⟩ ((τ _ ∘ (idC ⁂ τ (Y , Z)) +₁ idC) ∘ (assocˡ +₁ assocˡ)) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h) ≈⟨ pullʳ (pullˡ (sym distributeˡ⁻¹-assoc)) ⟩ (τ _ ∘ (idC ⁂ τ (Y , Z)) +₁ idC) ∘ (distributeˡ⁻¹ ∘ (idC ⁂ distributeˡ⁻¹) ∘ assocˡ) ∘ (idC ⁂ h) ≈⟨ refl⟩∘⟨ assoc²' ⟩ (τ _ ∘ (idC ⁂ τ _) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ distributeˡ⁻¹) ∘ assocˡ ∘ (idC ⁂ h) ≈˘⟨ (+₁-cong₂ refl (elimʳ (⟨⟩-unique id-comm id-comm))) ⟩∘⟨refl ⟩ (τ _ ∘ (idC ⁂ τ _) +₁ idC ∘ (idC ⁂ idC)) ∘ distributeˡ⁻¹ ∘ (idC ⁂ distributeˡ⁻¹) ∘ assocˡ ∘ (idC ⁂ h) ≈˘⟨ assoc ○ assoc ⟩ (((τ _ ∘ (idC ⁂ τ _) +₁ idC ∘ (idC ⁂ idC)) ∘ distributeˡ⁻¹) ∘ (idC ⁂ distributeˡ⁻¹)) ∘ _≅_.to ×-assoc ∘ (idC ⁂ h) ≈˘⟨ pullˡ (pullˡ (pullˡ +₁∘+₁)) ⟩ (τ _ +₁ idC) ∘ ((((idC ⁂ τ _) +₁ (idC ⁂ idC)) ∘ distributeˡ⁻¹) ∘ (idC ⁂ distributeˡ⁻¹)) ∘ assocˡ ∘ (idC ⁂ h) ≈⟨ refl⟩∘⟨ ((distributeˡ⁻¹-natural idC (τ (Y , Z)) idC) ⟩∘⟨refl) ⟩∘⟨refl ⟩ (τ _ +₁ idC) ∘ ((distributeˡ⁻¹ ∘ (idC ⁂ (τ (Y , Z) +₁ idC))) ∘ (idC ⁂ distributeˡ⁻¹)) ∘ assocˡ ∘ (idC ⁂ h) ≈⟨ refl⟩∘⟨ (assoc ○ assoc ○ refl⟩∘⟨ sym-assoc) ⟩ (τ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ ((idC ⁂ (τ (Y , Z) +₁ idC)) ∘ (idC ⁂ distributeˡ⁻¹)) ∘ assocˡ ∘ (idC ⁂ h) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ (⁂∘⁂ ○ ⁂-cong₂ identity² refl) ⟩∘⟨refl ⟩ (τ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ (τ (Y , Z) +₁ idC) ∘ distributeˡ⁻¹) ∘ assocˡ ∘ (idC ⁂ h) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ ⁂-cong₂ (sym (⟨⟩-unique id-comm id-comm)) refl ⟩ (τ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ (τ (Y , Z) +₁ idC) ∘ distributeˡ⁻¹) ∘ assocˡ ∘ ((idC ⁂ idC) ⁂ h) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ assocˡ∘⁂ ⟩ (τ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ (τ (Y , Z) +₁ idC) ∘ distributeˡ⁻¹) ∘ (idC ⁂ (idC ⁂ h)) ∘ assocˡ ≈˘⟨ refl⟩∘⟨ refl⟩∘⟨ assoc ⟩ (τ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ ((idC ⁂ (τ (Y , Z) +₁ idC) ∘ distributeˡ⁻¹) ∘ (idC ⁂ (idC ⁂ h))) ∘ assocˡ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ ⁂∘⁂ ⟩∘⟨refl ⟩ (τ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ∘ idC ⁂ ((τ (Y , Z) +₁ idC) ∘ distributeˡ⁻¹) ∘ (idC ⁂ h)) ∘ assocˡ ≈⟨ refl⟩∘⟨ (refl⟩∘⟨ ((⁂-cong₂ identity² assoc) ⟩∘⟨refl) ○ sym-assoc) ○ sym-assoc ⟩ ((τ _ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ (τ (Y , Z) +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))) ∘ assocˡ ∎ KStrong : StrongMonad {C = C} monoidal KStrong = record { M = monadK ; strength = KStrength } τ-comm-id : ∀ {X Y Z} (f : X ⇒ Y) → τ (Y , Z) ∘ (f ⁂ idC) ≈ K.₁ (f ⁂ idC) ∘ τ (X , Z) τ-comm-id {X} {Y} {Z} f = begin τ (Y , Z) ∘ (f ⁂ idC) ≈⟨ refl⟩∘⟨ (⁂-cong₂ refl (sym K.identity)) ⟩ τ (Y , Z) ∘ (f ⁂ K.₁ idC) ≈⟨ strengthen.commute (f , idC) ⟩ K.₁ (f ⁂ idC) ∘ τ (X , Z) ∎ where open Strength KStrength using (strengthen) module strongK = StrongMonad KStrong