Align proofs

This commit is contained in:
Leon Vatthauer 2023-10-22 15:07:39 +02:00
parent 544b604ce8
commit a71045161a
Signed by: leonv
SSH key fingerprint: SHA256:G4+ddwoZmhLPRB1agvXzZMXIzkVJ36dUYZXf5NxT+u8
2 changed files with 129 additions and 130 deletions

View file

@ -238,37 +238,37 @@ and second that `extend f` is the unique morphism satisfying the commutative dia
guarded-unique : ∀ {X Y} (g : X ⇒ D₀ (Y + X)) (f i : X ⇒ D₀ Y) → is-guarded g → f ≈ extend' ([ now , f ]) ∘ g → i ≈ extend' ([ now , i ]) ∘ g → f ≈ i guarded-unique : ∀ {X Y} (g : X ⇒ D₀ (Y + X)) (f i : X ⇒ D₀ Y) → is-guarded g → f ≈ extend' ([ now , f ]) ∘ g → i ≈ extend' ([ now , i ]) ∘ g → f ≈ i
guarded-unique {X} {Y} g f i (h , guarded) eqf eqi = begin guarded-unique {X} {Y} g f i (h , guarded) eqf eqi = begin
f ≈⟨ eqf ⟩ f ≈⟨ eqf ⟩
extend' ([ now , f ]) ∘ g ≈⟨ (sym (Terminal.!-unique (coalgebras Y) (record { f = extend' ([ now , f ]) ; commutes = begin extend' ([ now , f ]) ∘ g ≈⟨ (sym (Terminal.!-unique (coalgebras Y) (record { f = extend' ([ now , f ]) ; commutes = begin
out ∘ extend' ([ now , f ]) ≈⟨ extendlaw ([ now , f ]) ⟩ out ∘ extend' ([ now , f ]) ≈⟨ extendlaw ([ now , f ]) ⟩
[ out ∘ [ now , f ] , i₂ ∘ extend' ([ now , f ]) ] ∘ out ≈⟨ ([]-cong₂ ∘[] refl) ⟩∘⟨refl ⟩ [ out ∘ [ now , f ] , i₂ ∘ extend' ([ now , f ]) ] ∘ out ≈⟨ ([]-cong₂ ∘[] refl) ⟩∘⟨refl ⟩
[ [ out ∘ now , out ∘ f ] , i₂ ∘ extend' ([ now , f ]) ] ∘ out ≈⟨ ([]-cong₂ ([]-cong₂ unitlaw (helper f eqf)) refl) ⟩∘⟨refl ⟩ [ [ out ∘ now , out ∘ f ] , i₂ ∘ extend' ([ now , f ]) ] ∘ out ≈⟨ ([]-cong₂ ([]-cong₂ unitlaw (helper f eqf)) refl) ⟩∘⟨refl ⟩
[ [ i₁ , (idC +₁ extend' ([ now , f ])) ∘ h ] , i₂ ∘ extend' ([ now , f ]) ] ∘ out ≈˘⟨ ([]-cong₂ ([]-cong₂ identityʳ refl) refl) ⟩∘⟨refl ⟩ [ [ i₁ , (idC +₁ extend' ([ now , f ])) ∘ h ] , i₂ ∘ extend' ([ now , f ]) ] ∘ out ≈˘⟨ ([]-cong₂ ([]-cong₂ identityʳ refl) refl) ⟩∘⟨refl ⟩
[ [ i₁ ∘ idC , (idC +₁ extend' ([ now , f ])) ∘ h ] , i₂ ∘ extend' ([ now , f ]) ] ∘ out ≈˘⟨ ([]-cong₂ ([]-cong₂ +₁∘i₁ refl) refl) ⟩∘⟨refl ⟩ [ [ i₁ ∘ idC , (idC +₁ extend' ([ now , f ])) ∘ h ] , i₂ ∘ extend' ([ now , f ]) ] ∘ out ≈˘⟨ ([]-cong₂ ([]-cong₂ +₁∘i₁ refl) refl) ⟩∘⟨refl ⟩
[ [ (idC +₁ extend' ([ now , f ])) ∘ i₁ , (idC +₁ extend' ([ now , f ])) ∘ h ] , i₂ ∘ extend' ([ now , f ]) ] ∘ out ≈˘⟨ ([]-cong₂ ∘[] +₁∘i₂) ⟩∘⟨refl ⟩ [ [ (idC +₁ extend' ([ now , f ])) ∘ i₁ , (idC +₁ extend' ([ now , f ])) ∘ h ] , i₂ ∘ extend' ([ now , f ]) ] ∘ out ≈˘⟨ ([]-cong₂ ∘[] +₁∘i₂) ⟩∘⟨refl ⟩
[ (idC +₁ extend' ([ now , f ])) ∘ [ i₁ , h ] , (idC +₁ extend' ([ now , f ])) ∘ i₂ ] ∘ out ≈˘⟨ pullˡ ∘[] ⟩ [ (idC +₁ extend' ([ now , f ])) ∘ [ i₁ , h ] , (idC +₁ extend' ([ now , f ])) ∘ i₂ ] ∘ out ≈˘⟨ pullˡ ∘[] ⟩
(idC +₁ extend' ([ now , f ])) ∘ [ [ i₁ , h ] , i₂ ] ∘ out ∎ }))) ⟩∘⟨refl ⟩ (idC +₁ extend' ([ now , f ])) ∘ [ [ i₁ , h ] , i₂ ] ∘ out ∎ }))) ⟩∘⟨refl ⟩
u (Terminal.! (coalgebras Y)) ∘ g ≈⟨ (Terminal.!-unique (coalgebras Y) (record { f = extend' ([ now , i ]) ; commutes = begin u (Terminal.! (coalgebras Y)) ∘ g ≈⟨ (Terminal.!-unique (coalgebras Y) (record { f = extend' ([ now , i ]) ; commutes = begin
out ∘ extend' ([ now , i ]) ≈⟨ extendlaw ([ now , i ]) ⟩ out ∘ extend' ([ now , i ]) ≈⟨ extendlaw ([ now , i ]) ⟩
[ out ∘ [ now , i ] , i₂ ∘ extend' ([ now , i ]) ] ∘ out ≈⟨ ([]-cong₂ ∘[] refl) ⟩∘⟨refl ⟩ [ out ∘ [ now , i ] , i₂ ∘ extend' ([ now , i ]) ] ∘ out ≈⟨ ([]-cong₂ ∘[] refl) ⟩∘⟨refl ⟩
[ [ out ∘ now , out ∘ i ] , i₂ ∘ extend' ([ now , i ]) ] ∘ out ≈⟨ ([]-cong₂ ([]-cong₂ unitlaw (helper i eqi)) refl) ⟩∘⟨refl ⟩ [ [ out ∘ now , out ∘ i ] , i₂ ∘ extend' ([ now , i ]) ] ∘ out ≈⟨ ([]-cong₂ ([]-cong₂ unitlaw (helper i eqi)) refl) ⟩∘⟨refl ⟩
[ [ i₁ , (idC +₁ extend' ([ now , i ])) ∘ h ] , i₂ ∘ extend' ([ now , i ]) ] ∘ out ≈˘⟨ ([]-cong₂ ([]-cong₂ identityʳ refl) refl) ⟩∘⟨refl ⟩ [ [ i₁ , (idC +₁ extend' ([ now , i ])) ∘ h ] , i₂ ∘ extend' ([ now , i ]) ] ∘ out ≈˘⟨ ([]-cong₂ ([]-cong₂ identityʳ refl) refl) ⟩∘⟨refl ⟩
[ [ i₁ ∘ idC , (idC +₁ extend' ([ now , i ])) ∘ h ] , i₂ ∘ extend' ([ now , i ]) ] ∘ out ≈˘⟨ ([]-cong₂ ([]-cong₂ +₁∘i₁ refl) refl) ⟩∘⟨refl ⟩ [ [ i₁ ∘ idC , (idC +₁ extend' ([ now , i ])) ∘ h ] , i₂ ∘ extend' ([ now , i ]) ] ∘ out ≈˘⟨ ([]-cong₂ ([]-cong₂ +₁∘i₁ refl) refl) ⟩∘⟨refl ⟩
[ [ (idC +₁ extend' ([ now , i ])) ∘ i₁ , (idC +₁ extend' ([ now , i ])) ∘ h ] , i₂ ∘ extend' ([ now , i ]) ] ∘ out ≈˘⟨ ([]-cong₂ ∘[] +₁∘i₂) ⟩∘⟨refl ⟩ [ [ (idC +₁ extend' ([ now , i ])) ∘ i₁ , (idC +₁ extend' ([ now , i ])) ∘ h ] , i₂ ∘ extend' ([ now , i ]) ] ∘ out ≈˘⟨ ([]-cong₂ ∘[] +₁∘i₂) ⟩∘⟨refl ⟩
[ (idC +₁ extend' ([ now , i ])) ∘ [ i₁ , h ] , (idC +₁ extend' ([ now , i ])) ∘ i₂ ] ∘ out ≈˘⟨ pullˡ ∘[] ⟩ [ (idC +₁ extend' ([ now , i ])) ∘ [ i₁ , h ] , (idC +₁ extend' ([ now , i ])) ∘ i₂ ] ∘ out ≈˘⟨ pullˡ ∘[] ⟩
(idC +₁ extend' ([ now , i ])) ∘ [ [ i₁ , h ] , i₂ ] ∘ out ∎ })) ⟩∘⟨refl ⟩ (idC +₁ extend' ([ now , i ])) ∘ [ [ i₁ , h ] , i₂ ] ∘ out ∎ })) ⟩∘⟨refl ⟩
extend' ([ now , i ]) ∘ g ≈⟨ sym eqi ⟩ extend' ([ now , i ]) ∘ g ≈⟨ sym eqi ⟩
i ∎ i
where where
helper : ∀ (f : X ⇒ D₀ Y) → f ≈ extend' ([ now , f ]) ∘ g → out ∘ f ≈ (idC +₁ extend' ([ now , f ])) ∘ h helper : ∀ (f : X ⇒ D₀ Y) → f ≈ extend' ([ now , f ]) ∘ g → out ∘ f ≈ (idC +₁ extend' ([ now , f ])) ∘ h
helper f eqf = begin helper f eqf = begin
out ∘ f ≈⟨ refl⟩∘⟨ eqf ⟩ out ∘ f ≈⟨ refl⟩∘⟨ eqf ⟩
out ∘ extend' ([ now , f ]) ∘ g ≈⟨ pullˡ (extendlaw ([ now , f ])) ⟩ out ∘ extend' ([ now , f ]) ∘ g ≈⟨ pullˡ (extendlaw ([ now , f ])) ⟩
([ out ∘ [ now , f ] , i₂ ∘ extend' ([ now , f ]) ] ∘ out) ∘ g ≈⟨ pullʳ (sym guarded) ⟩ ([ out ∘ [ now , f ] , i₂ ∘ extend' ([ now , f ]) ] ∘ out) ∘ g ≈⟨ pullʳ (sym guarded) ⟩
[ out ∘ [ now , f ] , i₂ ∘ extend' ([ now , f ]) ] ∘ (i₁ +₁ idC) ∘ h ≈⟨ pullˡ []∘+₁ ⟩ [ out ∘ [ now , f ] , i₂ ∘ extend' ([ now , f ]) ] ∘ (i₁ +₁ idC) ∘ h ≈⟨ pullˡ []∘+₁ ⟩
[ (out ∘ [ now , f ]) ∘ i₁ , (i₂ ∘ extend' ([ now , f ])) ∘ idC ] ∘ h ≈⟨ ([]-cong₂ (pullʳ inject₁) identityʳ) ⟩∘⟨refl ⟩ [ (out ∘ [ now , f ]) ∘ i₁ , (i₂ ∘ extend' ([ now , f ])) ∘ idC ] ∘ h ≈⟨ ([]-cong₂ (pullʳ inject₁) identityʳ) ⟩∘⟨refl ⟩
[ out ∘ now , i₂ ∘ extend' ([ now , f ]) ] ∘ h ≈⟨ ([]-cong₂ (unitlaw ○ sym identityʳ) refl) ⟩∘⟨refl ⟩ [ out ∘ now , i₂ ∘ extend' ([ now , f ]) ] ∘ h ≈⟨ ([]-cong₂ (unitlaw ○ sym identityʳ) refl) ⟩∘⟨refl ⟩
(idC +₁ extend' ([ now , f ])) ∘ h ∎ (idC +₁ extend' ([ now , f ])) ∘ h
out-law : ∀ {X Y} (f : X ⇒ Y) → out {Y} ∘ extend' (now ∘ f) ≈ (f +₁ extend' (now ∘ f)) ∘ out {X} out-law : ∀ {X Y} (f : X ⇒ Y) → out {Y} ∘ extend' (now ∘ f) ≈ (f +₁ extend' (now ∘ f)) ∘ out {X}
out-law {X} {Y} f = begin out-law {X} {Y} f = begin

View file

@ -35,8 +35,8 @@ module Monad.Instance.Delay.Commutative {o e} (ambient : Ambient o e) (D
Kleisli⇒Monad⇒Kleisli K {X} {Y} f = begin Kleisli⇒Monad⇒Kleisli K {X} {Y} f = begin
extend idC ∘ extend (unit ∘ f) ≈⟨ sym k-assoc ⟩ extend idC ∘ extend (unit ∘ f) ≈⟨ sym k-assoc ⟩
extend (extend idC ∘ unit ∘ f) ≈⟨ extend-≈ (pullˡ k-identityʳ) ⟩ extend (extend idC ∘ unit ∘ f) ≈⟨ extend-≈ (pullˡ k-identityʳ) ⟩
extend (idC ∘ f) ≈⟨ extend-≈ (identityˡ) ⟩ extend (idC ∘ f) ≈⟨ extend-≈ (identityˡ) ⟩
extend f ∎ extend f
where open RMonad K using (unit; extend; extend-≈) renaming (assoc to k-assoc; identityʳ to k-identityʳ) where open RMonad K using (unit; extend; extend-≈) renaming (assoc to k-assoc; identityʳ to k-identityʳ)
open DelayM D open DelayM D
@ -44,7 +44,6 @@ module Monad.Instance.Delay.Commutative {o e} (ambient : Ambient o e) (D
open Functor functor using () renaming (F₁ to D₁; identity to D-identity; homomorphism to D-homomorphism; F-resp-≈ to D-resp-≈) open Functor functor using () renaming (F₁ to D₁; identity to D-identity; homomorphism to D-homomorphism; F-resp-≈ to D-resp-≈)
open RMonad kleisli using (extend; extend-≈) renaming (assoc to k-assoc; identityʳ to k-identityʳ; identityˡ to k-identityˡ) open RMonad kleisli using (extend; extend-≈) renaming (assoc to k-assoc; identityʳ to k-identityʳ; identityˡ to k-identityˡ)
open Monad monad using (η; μ) open Monad monad using (η; μ)
-- open StrongMonad strongMonad using (strengthen)
open NaturalTransformation (StrongMonad.strengthen strongMonad) using () renaming (commute to τ-commute) open NaturalTransformation (StrongMonad.strengthen strongMonad) using () renaming (commute to τ-commute)
``` ```
@ -61,32 +60,32 @@ module Monad.Instance.Delay.Commutative {o e} (ambient : Ambient o e) (D
σ {X} {Y} = D₁ swap ∘ τ ∘ swap σ {X} {Y} = D₁ swap ∘ τ ∘ swap
σ-coalg : ∀ (X Y : Obj) → F-Coalgebra-Morphism {F = (X × Y) +- } (record { A = D₀ X × Y ; α = distributeʳ⁻¹ {Y} {X} {D₀ X} ∘ (out {X} ⁂ idC) }) (record { A = D₀ (X × Y) ; α = out {X × Y} }) σ-coalg : ∀ (X Y : Obj) → F-Coalgebra-Morphism {F = (X × Y) +- } (record { A = D₀ X × Y ; α = distributeʳ⁻¹ {Y} {X} {D₀ X} ∘ (out {X} ⁂ idC) }) (record { A = D₀ (X × Y) ; α = out {X × Y} })
σ-coalg X Y = record { f = σ ; commutes = begin σ-coalg X Y = record { f = σ ; commutes = begin
out ∘ σ ≈⟨ pullˡ (out-law swap) ⟩ out ∘ σ ≈⟨ pullˡ (out-law swap) ⟩
((swap +₁ D₁ swap) ∘ out) ∘ τ ∘ swap ≈⟨ pullˡ (pullʳ (τ-law (Y , X))) ⟩ ((swap +₁ D₁ swap) ∘ out) ∘ τ ∘ swap ≈⟨ pullˡ (pullʳ (τ-law (Y , X))) ⟩
((swap +₁ D₁ swap) ∘ (idC +₁ τ) ∘ distributeˡ⁻¹ ∘ (idC ⁂ out)) ∘ swap ≈⟨ pullʳ (pullʳ (pullʳ (sym swap∘⁂))) ⟩ ((swap +₁ D₁ swap) ∘ (idC +₁ τ) ∘ distributeˡ⁻¹ ∘ (idC ⁂ out)) ∘ swap ≈⟨ pullʳ (pullʳ (pullʳ (sym swap∘⁂))) ⟩
(swap +₁ D₁ swap) ∘ (idC +₁ τ) ∘ distributeˡ⁻¹ ∘ swap ∘ (out ⁂ idC) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pullˡ distributeˡ⁻¹∘swap ⟩ (swap +₁ D₁ swap) ∘ (idC +₁ τ) ∘ distributeˡ⁻¹ ∘ swap ∘ (out ⁂ idC) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pullˡ distributeˡ⁻¹∘swap ⟩
(swap +₁ D₁ swap) ∘ (idC +₁ τ) ∘ ((swap +₁ swap) ∘ distributeʳ⁻¹) ∘ (out ⁂ idC) ≈⟨ pullˡ +₁∘+₁ ⟩ (swap +₁ D₁ swap) ∘ (idC +₁ τ) ∘ ((swap +₁ swap) ∘ distributeʳ⁻¹) ∘ (out ⁂ idC) ≈⟨ pullˡ +₁∘+₁ ⟩
(swap ∘ idC +₁ D₁ swap ∘ τ) ∘ ((swap +₁ swap) ∘ distributeʳ⁻¹) ∘ (out ⁂ idC) ≈⟨ pullˡ (pullˡ +₁∘+₁) ⟩ (swap ∘ idC +₁ D₁ swap ∘ τ) ∘ ((swap +₁ swap) ∘ distributeʳ⁻¹) ∘ (out ⁂ idC) ≈⟨ pullˡ (pullˡ +₁∘+₁) ⟩
(((swap ∘ idC) ∘ swap +₁ (D₁ swap ∘ τ) ∘ swap) ∘ distributeʳ⁻¹) ∘ (out ⁂ idC) ≈⟨ ((+₁-cong₂ (identityʳ ⟩∘⟨refl ○ swap∘swap) assoc) ⟩∘⟨refl) ⟩∘⟨refl ⟩ (((swap ∘ idC) ∘ swap +₁ (D₁ swap ∘ τ) ∘ swap) ∘ distributeʳ⁻¹) ∘ (out ⁂ idC) ≈⟨ ((+₁-cong₂ (identityʳ ⟩∘⟨refl ○ swap∘swap) assoc) ⟩∘⟨refl) ⟩∘⟨refl ⟩
((idC +₁ D₁ swap ∘ τ ∘ swap) ∘ distributeʳ⁻¹) ∘ (out ⁂ idC) ≈⟨ assoc ⟩ ((idC +₁ D₁ swap ∘ τ ∘ swap) ∘ distributeʳ⁻¹) ∘ (out ⁂ idC) ≈⟨ assoc ⟩
(idC +₁ σ) ∘ distributeʳ⁻¹ ∘ (out ⁂ idC) ∎ } (idC +₁ σ) ∘ distributeʳ⁻¹ ∘ (out ⁂ idC) ∎ }
σ-helper : ∀ {X Y : Obj} → σ {X} {Y} ∘ (out⁻¹ ⁂ idC) ≈ out⁻¹ ∘ (idC +₁ σ) ∘ distributeʳ⁻¹ σ-helper : ∀ {X Y : Obj} → σ {X} {Y} ∘ (out⁻¹ ⁂ idC) ≈ out⁻¹ ∘ (idC +₁ σ) ∘ distributeʳ⁻¹
σ-helper {X} {Y} = begin σ-helper {X} {Y} = begin
σ ∘ (out⁻¹ ⁂ idC) ≈⟨ introˡ (_≅_.isoˡ out-≅) ⟩ σ ∘ (out⁻¹ ⁂ idC) ≈⟨ introˡ (_≅_.isoˡ out-≅) ⟩
(out⁻¹ ∘ out) ∘ σ ∘ (out⁻¹ ⁂ idC) ≈⟨ pullʳ (pullˡ (u-commutes (σ-coalg X Y))) ⟩ (out⁻¹ ∘ out) ∘ σ ∘ (out⁻¹ ⁂ idC) ≈⟨ pullʳ (pullˡ (u-commutes (σ-coalg X Y))) ⟩
out⁻¹ ∘ ((idC +₁ σ) ∘ distributeʳ⁻¹ ∘ (out ⁂ idC)) ∘ (out⁻¹ ⁂ idC) ≈⟨ refl⟩∘⟨ (pullʳ (cancelʳ (⁂∘⁂ ○ ⁂-cong₂ (_≅_.isoʳ out-≅) identity² ○ ⟨⟩-unique id-comm id-comm))) ⟩ out⁻¹ ∘ ((idC +₁ σ) ∘ distributeʳ⁻¹ ∘ (out ⁂ idC)) ∘ (out⁻¹ ⁂ idC) ≈⟨ refl⟩∘⟨ (pullʳ (cancelʳ (⁂∘⁂ ○ ⁂-cong₂ (_≅_.isoʳ out-≅) identity² ○ ⟨⟩-unique id-comm id-comm))) ⟩
out⁻¹ ∘ (idC +₁ σ) ∘ distributeʳ⁻¹ ∎ out⁻¹ ∘ (idC +₁ σ) ∘ distributeʳ⁻¹
-- TODO this should be in commutative, it expresses that σ is natural -- TODO this should be in commutative, it expresses that σ is natural
σ-commute : ∀ {U V W X : Obj} (f : U ⇒ V) (g : W ⇒ X) → σ ∘ (extend (now ∘ f) ⁂ g) ≈ extend (now ∘ (f ⁂ g)) ∘ σ σ-commute : ∀ {U V W X : Obj} (f : U ⇒ V) (g : W ⇒ X) → σ ∘ (extend (now ∘ f) ⁂ g) ≈ extend (now ∘ (f ⁂ g)) ∘ σ
σ-commute {U} {V} {W} {X} f g = begin σ-commute {U} {V} {W} {X} f g = begin
σ ∘ (D₁ f ⁂ g) ≈⟨ pullʳ (pullʳ swap∘⁂) ⟩ σ ∘ (D₁ f ⁂ g) ≈⟨ pullʳ (pullʳ swap∘⁂) ⟩
D₁ swap ∘ τ ∘ (g ⁂ extend (now ∘ f)) ∘ swap ≈⟨ refl⟩∘⟨ (pullˡ (τ-commute (g , f))) ⟩ D₁ swap ∘ τ ∘ (g ⁂ extend (now ∘ f)) ∘ swap ≈⟨ refl⟩∘⟨ (pullˡ (τ-commute (g , f))) ⟩
D₁ swap ∘ (D₁ (g ⁂ f) ∘ τ) ∘ swap ≈⟨ pullˡ (pullˡ (sym D-homomorphism)) ⟩ D₁ swap ∘ (D₁ (g ⁂ f) ∘ τ) ∘ swap ≈⟨ pullˡ (pullˡ (sym D-homomorphism)) ⟩
(D₁ (swap ∘ (g ⁂ f)) ∘ τ) ∘ swap ≈⟨ ((D-resp-≈ swap∘⁂) ⟩∘⟨refl) ⟩∘⟨refl ⟩ (D₁ (swap ∘ (g ⁂ f)) ∘ τ) ∘ swap ≈⟨ ((D-resp-≈ swap∘⁂) ⟩∘⟨refl) ⟩∘⟨refl ⟩
(D₁ ((f ⁂ g) ∘ swap) ∘ τ) ∘ swap ≈⟨ pushˡ D-homomorphism ⟩∘⟨refl ⟩ (D₁ ((f ⁂ g) ∘ swap) ∘ τ) ∘ swap ≈⟨ pushˡ D-homomorphism ⟩∘⟨refl ⟩
(D₁ (f ⁂ g) ∘ D₁ swap ∘ τ) ∘ swap ≈⟨ assoc²' ⟩ (D₁ (f ⁂ g) ∘ D₁ swap ∘ τ) ∘ swap ≈⟨ assoc²' ⟩
D₁ (f ⁂ g) ∘ σ D₁ (f ⁂ g) ∘ σ
commutes' : ∀ {X Y} → extend σ ∘ τ {D₀ X} {Y} ≈ extend τ ∘ σ commutes' : ∀ {X Y} → extend σ ∘ τ {D₀ X} {Y} ≈ extend τ ∘ σ
commutes' {X} {Y} = guarded-unique g (extend σ ∘ τ) (extend τ ∘ σ) guarded (fixpoint-eq fixpoint₁) (fixpoint-eq fixpoint₂) commutes' {X} {Y} = guarded-unique g (extend σ ∘ τ) (extend τ ∘ σ) guarded (fixpoint-eq fixpoint₁) (fixpoint-eq fixpoint₂)
@ -95,118 +94,118 @@ module Monad.Instance.Delay.Commutative {o e} (ambient : Ambient o e) (D
g = out⁻¹ ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w g = out⁻¹ ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w
guarded : is-guarded g guarded : is-guarded g
guarded = [ idC +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w , (begin guarded = [ idC +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w , (begin
(i₁ +₁ idC) ∘ [ idC +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈⟨ pullˡ ∘[] ⟩ (i₁ +₁ idC) ∘ [ idC +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈⟨ pullˡ ∘[] ⟩
[ (i₁ +₁ idC) ∘ (idC +₁ D₁ i₁ ∘ σ) , (i₁ +₁ idC) ∘ i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈⟨ ([]-cong₂ +₁∘+₁ (pullˡ +₁∘i₂)) ⟩∘⟨refl ⟩ [ (i₁ +₁ idC) ∘ (idC +₁ D₁ i₁ ∘ σ) , (i₁ +₁ idC) ∘ i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈⟨ ([]-cong₂ +₁∘+₁ (pullˡ +₁∘i₂)) ⟩∘⟨refl ⟩
[ (i₁ ∘ idC +₁ idC ∘ D₁ i₁ ∘ σ) , (i₂ ∘ idC) ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈⟨ ([]-cong₂ (+₁-cong₂ identityʳ identityˡ) (identityʳ ⟩∘⟨refl)) ⟩∘⟨refl ⟩ [ (i₁ ∘ idC +₁ idC ∘ D₁ i₁ ∘ σ) , (i₂ ∘ idC) ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈⟨ ([]-cong₂ (+₁-cong₂ identityʳ identityˡ) (identityʳ ⟩∘⟨refl)) ⟩∘⟨refl ⟩
[ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈⟨ sym (cancelˡ (_≅_.isoʳ out-≅)) ⟩ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈⟨ sym (cancelˡ (_≅_.isoʳ out-≅)) ⟩
out ∘ out⁻¹ ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ∎) out ∘ out⁻¹ ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ∎)
helper₁ : (D₁ distributeʳ⁻¹) ∘ τ ≈ [ D₁ i₁ ∘ τ , D₁ i₂ ∘ τ ] ∘ distributeʳ⁻¹ helper₁ : (D₁ distributeʳ⁻¹) ∘ τ ≈ [ D₁ i₁ ∘ τ , D₁ i₂ ∘ τ ] ∘ distributeʳ⁻¹
helper₁ = Iso⇒Epi (IsIso.iso isIsoʳ) ((D₁ distributeʳ⁻¹) ∘ τ) ([ D₁ i₁ ∘ τ , D₁ i₂ ∘ τ ] ∘ distributeʳ⁻¹) (begin helper₁ = Iso⇒Epi (IsIso.iso isIsoʳ) ((D₁ distributeʳ⁻¹) ∘ τ) ([ D₁ i₁ ∘ τ , D₁ i₂ ∘ τ ] ∘ distributeʳ⁻¹) (begin
((D₁ distributeʳ⁻¹) ∘ τ) ∘ distributeʳ ≈⟨ ∘[] ⟩ ((D₁ distributeʳ⁻¹) ∘ τ) ∘ distributeʳ ≈⟨ ∘[] ⟩
[ ((D₁ distributeʳ⁻¹) ∘ τ) ∘ (i₁ ⁂ idC) , ((D₁ distributeʳ⁻¹) ∘ τ) ∘ (i₂ ⁂ idC) ] ≈⟨ []-cong₂ (refl⟩∘⟨ (⁂-cong₂ refl (sym D-identity))) (refl⟩∘⟨ (⁂-cong₂ refl (sym D-identity))) ⟩ [ ((D₁ distributeʳ⁻¹) ∘ τ) ∘ (i₁ ⁂ idC) , ((D₁ distributeʳ⁻¹) ∘ τ) ∘ (i₂ ⁂ idC) ] ≈⟨ []-cong₂ (refl⟩∘⟨ (⁂-cong₂ refl (sym D-identity))) (refl⟩∘⟨ (⁂-cong₂ refl (sym D-identity))) ⟩
[ ((D₁ distributeʳ⁻¹) ∘ τ) ∘ (i₁ ⁂ D₁ idC) , ((D₁ distributeʳ⁻¹) ∘ τ) ∘ (i₂ ⁂ D₁ idC) ] ≈⟨ []-cong₂ (pullʳ (τ-commute (i₁ , idC))) (pullʳ (τ-commute (i₂ , idC))) ⟩ [ ((D₁ distributeʳ⁻¹) ∘ τ) ∘ (i₁ ⁂ D₁ idC) , ((D₁ distributeʳ⁻¹) ∘ τ) ∘ (i₂ ⁂ D₁ idC) ] ≈⟨ []-cong₂ (pullʳ (τ-commute (i₁ , idC))) (pullʳ (τ-commute (i₂ , idC))) ⟩
[ (D₁ distributeʳ⁻¹) ∘ D₁ (i₁ ⁂ idC) ∘ τ , (D₁ distributeʳ⁻¹) ∘ D₁ (i₂ ⁂ idC) ∘ τ ] ≈⟨ []-cong₂ (pullˡ (sym D-homomorphism)) (pullˡ (sym D-homomorphism)) ⟩ [ (D₁ distributeʳ⁻¹) ∘ D₁ (i₁ ⁂ idC) ∘ τ , (D₁ distributeʳ⁻¹) ∘ D₁ (i₂ ⁂ idC) ∘ τ ] ≈⟨ []-cong₂ (pullˡ (sym D-homomorphism)) (pullˡ (sym D-homomorphism)) ⟩
[ D₁ (distributeʳ⁻¹ ∘ (i₁ ⁂ idC)) ∘ τ , D₁ (distributeʳ⁻¹ ∘ (i₂ ⁂ idC)) ∘ τ ] ≈⟨ []-cong₂ (D-resp-≈ dstr-law₃ ⟩∘⟨refl) ((D-resp-≈ dstr-law₄) ⟩∘⟨refl) ⟩ [ D₁ (distributeʳ⁻¹ ∘ (i₁ ⁂ idC)) ∘ τ , D₁ (distributeʳ⁻¹ ∘ (i₂ ⁂ idC)) ∘ τ ] ≈⟨ []-cong₂ (D-resp-≈ dstr-law₃ ⟩∘⟨refl) ((D-resp-≈ dstr-law₄) ⟩∘⟨refl) ⟩
[ D₁ i₁ ∘ τ , D₁ i₂ ∘ τ ] ≈˘⟨ cancelʳ (IsIso.isoˡ isIsoʳ) ⟩ [ D₁ i₁ ∘ τ , D₁ i₂ ∘ τ ] ≈˘⟨ cancelʳ (IsIso.isoˡ isIsoʳ) ⟩
([ D₁ i₁ ∘ τ , D₁ i₂ ∘ τ ] ∘ distributeʳ⁻¹) ∘ distributeʳ ∎) ([ D₁ i₁ ∘ τ , D₁ i₂ ∘ τ ] ∘ distributeʳ⁻¹) ∘ distributeʳ ∎)
fixpoint₁ = Iso⇒Mono (_≅_.iso out-≅) (extend σ ∘ τ) (out⁻¹ ∘ [ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ extend σ ∘ τ ] ] ∘ w) (begin fixpoint₁ = Iso⇒Mono (_≅_.iso out-≅) (extend σ ∘ τ) (out⁻¹ ∘ [ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ extend σ ∘ τ ] ] ∘ w) (begin
out ∘ extend σ ∘ τ ≈⟨ pullˡ (extendlaw σ) ⟩ out ∘ extend σ ∘ τ ≈⟨ pullˡ (extendlaw σ) ⟩
([ out ∘ σ , i₂ ∘ extend' σ ] ∘ out) ∘ τ ≈⟨ pullʳ (τ-law (D₀ X , Y)) ⟩ ([ out ∘ σ , i₂ ∘ extend' σ ] ∘ out) ∘ τ ≈⟨ pullʳ (τ-law (D₀ X , Y)) ⟩
[ out ∘ σ , i₂ ∘ extend' σ ] ∘ (idC +₁ τ) ∘ distributeˡ⁻¹ ∘ (idC ⁂ out) ≈⟨ pullˡ []∘+₁ ⟩ [ out ∘ σ , i₂ ∘ extend' σ ] ∘ (idC +₁ τ) ∘ distributeˡ⁻¹ ∘ (idC ⁂ out) ≈⟨ pullˡ []∘+₁ ⟩
[ (out ∘ σ) ∘ idC , (i₂ ∘ extend' σ) ∘ τ ] ∘ distributeˡ⁻¹ ∘ (idC ⁂ out) ≈⟨ ([]-cong₂ (identityʳ ○ u-commutes (σ-coalg X Y)) assoc) ⟩∘⟨refl ⟩ [ (out ∘ σ) ∘ idC , (i₂ ∘ extend' σ) ∘ τ ] ∘ distributeˡ⁻¹ ∘ (idC ⁂ out) ≈⟨ ([]-cong₂ (identityʳ ○ u-commutes (σ-coalg X Y)) assoc) ⟩∘⟨refl ⟩
[ (idC +₁ σ) ∘ distributeʳ⁻¹ ∘ (out ⁂ idC) , i₂ ∘ extend' σ ∘ τ ] ∘ distributeˡ⁻¹ ∘ (idC ⁂ out) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ ⁂-cong₂ (sym (_≅_.isoˡ out-≅)) refl ⟩ [ (idC +₁ σ) ∘ distributeʳ⁻¹ ∘ (out ⁂ idC) , i₂ ∘ extend' σ ∘ τ ] ∘ distributeˡ⁻¹ ∘ (idC ⁂ out) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ ⁂-cong₂ (sym (_≅_.isoˡ out-≅)) refl ⟩
[ (idC +₁ σ) ∘ distributeʳ⁻¹ ∘ (out ⁂ idC) , i₂ ∘ extend' σ ∘ τ ] ∘ distributeˡ⁻¹ ∘ (out⁻¹ ∘ out ⁂ out) ≈⟨ sym (refl⟩∘⟨ refl⟩∘⟨ (⁂∘⁂ ○ ⁂-cong₂ refl (elimˡ ([]-unique id-comm-sym id-comm-sym)))) ⟩ [ (idC +₁ σ) ∘ distributeʳ⁻¹ ∘ (out ⁂ idC) , i₂ ∘ extend' σ ∘ τ ] ∘ distributeˡ⁻¹ ∘ (out⁻¹ ∘ out ⁂ out) ≈⟨ sym (refl⟩∘⟨ refl⟩∘⟨ (⁂∘⁂ ○ ⁂-cong₂ refl (elimˡ ([]-unique id-comm-sym id-comm-sym)))) ⟩
[ (idC +₁ σ) ∘ distributeʳ⁻¹ ∘ (out ⁂ idC) , i₂ ∘ extend' σ ∘ τ ] ∘ distributeˡ⁻¹ ∘ (out⁻¹ ⁂ (idC +₁ idC)) ∘ (out ⁂ out) ≈⟨ refl⟩∘⟨ pullˡ (sym (distribute₁ out⁻¹ idC idC)) ⟩ [ (idC +₁ σ) ∘ distributeʳ⁻¹ ∘ (out ⁂ idC) , i₂ ∘ extend' σ ∘ τ ] ∘ distributeˡ⁻¹ ∘ (out⁻¹ ⁂ (idC +₁ idC)) ∘ (out ⁂ out) ≈⟨ refl⟩∘⟨ pullˡ (sym (distribute₁ out⁻¹ idC idC)) ⟩
[ (idC +₁ σ) ∘ distributeʳ⁻¹ ∘ (out ⁂ idC) , i₂ ∘ extend' σ ∘ τ ] ∘ (((out⁻¹ ⁂ idC) +₁ (out⁻¹ ⁂ idC)) ∘ distributeˡ⁻¹) ∘ (out ⁂ out) ≈⟨ pullˡ (pullˡ []∘+₁) ⟩ [ (idC +₁ σ) ∘ distributeʳ⁻¹ ∘ (out ⁂ idC) , i₂ ∘ extend' σ ∘ τ ] ∘ (((out⁻¹ ⁂ idC) +₁ (out⁻¹ ⁂ idC)) ∘ distributeˡ⁻¹) ∘ (out ⁂ out) ≈⟨ pullˡ (pullˡ []∘+₁) ⟩
([ ((idC +₁ σ) ∘ distributeʳ⁻¹ ∘ (out ⁂ idC)) ∘ (out⁻¹ ⁂ idC) , (i₂ ∘ extend' σ ∘ τ) ∘ (out⁻¹ ⁂ idC) ] ∘ distributeˡ⁻¹) ∘ (out ⁂ out) ≈⟨ assoc ○ ([]-cong₂ (pullʳ (cancelʳ (⁂∘⁂ ○ ⁂-cong₂ (_≅_.isoʳ out-≅) identity² ○ ⟨⟩-unique id-comm id-comm))) (refl⟩∘⟨ (⁂-cong₂ refl (sym D-identity)) ○ (pullʳ (pullʳ (τ-commute (out⁻¹ , idC)))))) ⟩∘⟨refl ⟩ ([ ((idC +₁ σ) ∘ distributeʳ⁻¹ ∘ (out ⁂ idC)) ∘ (out⁻¹ ⁂ idC) , (i₂ ∘ extend' σ ∘ τ) ∘ (out⁻¹ ⁂ idC) ] ∘ distributeˡ⁻¹) ∘ (out ⁂ out) ≈⟨ assoc ○ ([]-cong₂ (pullʳ (cancelʳ (⁂∘⁂ ○ ⁂-cong₂ (_≅_.isoʳ out-≅) identity² ○ ⟨⟩-unique id-comm id-comm))) (refl⟩∘⟨ (⁂-cong₂ refl (sym D-identity)) ○ (pullʳ (pullʳ (τ-commute (out⁻¹ , idC)))))) ⟩∘⟨refl ⟩
[ (idC +₁ σ) ∘ distributeʳ⁻¹ , i₂ ∘ extend' σ ∘ D₁ (out⁻¹ ⁂ idC) ∘ τ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ (pullˡ (sym k-assoc)) ○ refl⟩∘⟨ ((extend-≈ (pullˡ k-identityʳ)) ⟩∘⟨refl))) ⟩∘⟨refl ⟩ [ (idC +₁ σ) ∘ distributeʳ⁻¹ , i₂ ∘ extend' σ ∘ D₁ (out⁻¹ ⁂ idC) ∘ τ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ (pullˡ (sym k-assoc)) ○ refl⟩∘⟨ ((extend-≈ (pullˡ k-identityʳ)) ⟩∘⟨refl))) ⟩∘⟨refl ⟩
[ (idC +₁ σ) ∘ distributeʳ⁻¹ , i₂ ∘ extend' (σ ∘ (out⁻¹ ⁂ idC)) ∘ τ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ ((extend-≈ σ-helper) ⟩∘⟨refl))) ⟩∘⟨refl ⟩ [ (idC +₁ σ) ∘ distributeʳ⁻¹ , i₂ ∘ extend' (σ ∘ (out⁻¹ ⁂ idC)) ∘ τ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ ((extend-≈ σ-helper) ⟩∘⟨refl))) ⟩∘⟨refl ⟩
[ (idC +₁ σ) ∘ distributeʳ⁻¹ , i₂ ∘ extend' (out⁻¹ ∘ (idC +₁ σ) ∘ distributeʳ⁻¹) ∘ τ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈˘⟨ ([]-cong₂ refl (refl⟩∘⟨ ((sym k-assoc ○ extend-≈ (pullˡ k-identityʳ) ○ extend-≈ assoc) ⟩∘⟨refl))) ⟩∘⟨refl ⟩ [ (idC +₁ σ) ∘ distributeʳ⁻¹ , i₂ ∘ extend' (out⁻¹ ∘ (idC +₁ σ) ∘ distributeʳ⁻¹) ∘ τ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈˘⟨ ([]-cong₂ refl (refl⟩∘⟨ ((sym k-assoc ○ extend-≈ (pullˡ k-identityʳ) ○ extend-≈ assoc) ⟩∘⟨refl))) ⟩∘⟨refl ⟩
[ (idC +₁ σ) ∘ distributeʳ⁻¹ , i₂ ∘ (extend' (out⁻¹ ∘ (idC +₁ σ)) ∘ D₁ distributeʳ⁻¹) ∘ τ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ pullʳ helper₁)) ⟩∘⟨refl ⟩ [ (idC +₁ σ) ∘ distributeʳ⁻¹ , i₂ ∘ (extend' (out⁻¹ ∘ (idC +₁ σ)) ∘ D₁ distributeʳ⁻¹) ∘ τ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ pullʳ helper₁)) ⟩∘⟨refl ⟩
[ (idC +₁ σ) ∘ distributeʳ⁻¹ , i₂ ∘ extend' (out⁻¹ ∘ (idC +₁ σ)) ∘ [ D₁ i₁ ∘ τ , D₁ i₂ ∘ τ ] ∘ distributeʳ⁻¹ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈˘⟨ pullˡ ([]∘+₁ ○ []-cong₂ refl assoc²') ⟩ [ (idC +₁ σ) ∘ distributeʳ⁻¹ , i₂ ∘ extend' (out⁻¹ ∘ (idC +₁ σ)) ∘ [ D₁ i₁ ∘ τ , D₁ i₂ ∘ τ ] ∘ distributeʳ⁻¹ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈˘⟨ pullˡ ([]∘+₁ ○ []-cong₂ refl assoc²') ⟩
[ (idC +₁ σ) , i₂ ∘ extend' (out⁻¹ ∘ (idC +₁ σ)) ∘ [ D₁ i₁ ∘ τ , D₁ i₂ ∘ τ ] ] ∘ (distributeʳ⁻¹ +₁ distributeʳ⁻¹) ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ ∘[])) ⟩∘⟨refl ⟩ [ (idC +₁ σ) , i₂ ∘ extend' (out⁻¹ ∘ (idC +₁ σ)) ∘ [ D₁ i₁ ∘ τ , D₁ i₂ ∘ τ ] ] ∘ (distributeʳ⁻¹ +₁ distributeʳ⁻¹) ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ ∘[])) ⟩∘⟨refl ⟩
[ (idC +₁ σ) , i₂ ∘ [ extend' (out⁻¹ ∘ (idC +₁ σ)) ∘ D₁ i₁ ∘ τ , extend' (out⁻¹ ∘ (idC +₁ σ)) ∘ D₁ i₂ ∘ τ ] ] ∘ (distributeʳ⁻¹ +₁ distributeʳ⁻¹) ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ ([]-cong₂ (pullˡ ((sym k-assoc) ○ extend-≈ (pullˡ k-identityʳ))) (pullˡ ((sym k-assoc) ○ extend-≈ (pullˡ k-identityʳ)))))) ⟩∘⟨refl ⟩ [ (idC +₁ σ) , i₂ ∘ [ extend' (out⁻¹ ∘ (idC +₁ σ)) ∘ D₁ i₁ ∘ τ , extend' (out⁻¹ ∘ (idC +₁ σ)) ∘ D₁ i₂ ∘ τ ] ] ∘ (distributeʳ⁻¹ +₁ distributeʳ⁻¹) ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ ([]-cong₂ (pullˡ ((sym k-assoc) ○ extend-≈ (pullˡ k-identityʳ))) (pullˡ ((sym k-assoc) ○ extend-≈ (pullˡ k-identityʳ)))))) ⟩∘⟨refl ⟩
[ (idC +₁ σ) , i₂ ∘ [ extend' ((out⁻¹ ∘ (idC +₁ σ)) ∘ i₁) ∘ τ , extend' ((out⁻¹ ∘ (idC +₁ σ)) ∘ i₂) ∘ τ ] ] ∘ (distributeʳ⁻¹ +₁ distributeʳ⁻¹) ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ ([]-cong₂ ((extend-≈ (pullʳ +₁∘i₁)) ⟩∘⟨refl) ((extend-≈ (pullʳ +₁∘i₂)) ⟩∘⟨refl)))) ⟩∘⟨refl ⟩ [ (idC +₁ σ) , i₂ ∘ [ extend' ((out⁻¹ ∘ (idC +₁ σ)) ∘ i₁) ∘ τ , extend' ((out⁻¹ ∘ (idC +₁ σ)) ∘ i₂) ∘ τ ] ] ∘ (distributeʳ⁻¹ +₁ distributeʳ⁻¹) ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ ([]-cong₂ ((extend-≈ (pullʳ +₁∘i₁)) ⟩∘⟨refl) ((extend-≈ (pullʳ +₁∘i₂)) ⟩∘⟨refl)))) ⟩∘⟨refl ⟩
[ (idC +₁ σ) , i₂ ∘ [ extend' (out⁻¹ ∘ i₁ ∘ idC) ∘ τ , extend' (out⁻¹ ∘ i₂ ∘ σ) ∘ τ ] ] ∘ (distributeʳ⁻¹ +₁ distributeʳ⁻¹) ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ ([]-cong₂ (elimˡ ((extend-≈ (refl⟩∘⟨ identityʳ)) ○ k-identityˡ)) ((extend-≈ sym-assoc) ⟩∘⟨refl)))) ⟩∘⟨refl ⟩ [ (idC +₁ σ) , i₂ ∘ [ extend' (out⁻¹ ∘ i₁ ∘ idC) ∘ τ , extend' (out⁻¹ ∘ i₂ ∘ σ) ∘ τ ] ] ∘ (distributeʳ⁻¹ +₁ distributeʳ⁻¹) ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ ([]-cong₂ (elimˡ ((extend-≈ (refl⟩∘⟨ identityʳ)) ○ k-identityˡ)) ((extend-≈ sym-assoc) ⟩∘⟨refl)))) ⟩∘⟨refl ⟩
[ (idC +₁ σ) , i₂ ∘ [ τ , extend' (▷ ∘ σ) ∘ τ ] ] ∘ (distributeʳ⁻¹ +₁ distributeʳ⁻¹) ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ ([]-cong₂ refl ((sym (▷∘extendˡ σ)) ⟩∘⟨refl ○ assoc)))) ⟩∘⟨refl ⟩ [ (idC +₁ σ) , i₂ ∘ [ τ , extend' (▷ ∘ σ) ∘ τ ] ] ∘ (distributeʳ⁻¹ +₁ distributeʳ⁻¹) ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ ([]-cong₂ refl ((sym (▷∘extendˡ σ)) ⟩∘⟨refl ○ assoc)))) ⟩∘⟨refl ⟩
[ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ extend σ ∘ τ ] ] ∘ w ≈˘⟨ cancelˡ (_≅_.isoʳ out-≅) ⟩ [ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ extend σ ∘ τ ] ] ∘ w ≈˘⟨ cancelˡ (_≅_.isoʳ out-≅) ⟩
out ∘ out⁻¹ ∘ [ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ extend σ ∘ τ ] ] ∘ w ∎) out ∘ out⁻¹ ∘ [ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ extend σ ∘ τ ] ] ∘ w ∎)
helper₂ : (D₁ distributeˡ⁻¹) ∘ σ ≈ [ D₁ i₁ ∘ σ , D₁ i₂ ∘ σ ] ∘ distributeˡ⁻¹ helper₂ : (D₁ distributeˡ⁻¹) ∘ σ ≈ [ D₁ i₁ ∘ σ , D₁ i₂ ∘ σ ] ∘ distributeˡ⁻¹
helper₂ = Iso⇒Epi (IsIso.iso isIsoˡ) ((D₁ distributeˡ⁻¹) ∘ σ) ([ D₁ i₁ ∘ σ , D₁ i₂ ∘ σ ] ∘ distributeˡ⁻¹) (begin helper₂ = Iso⇒Epi (IsIso.iso isIsoˡ) ((D₁ distributeˡ⁻¹) ∘ σ) ([ D₁ i₁ ∘ σ , D₁ i₂ ∘ σ ] ∘ distributeˡ⁻¹) (begin
((D₁ distributeˡ⁻¹) ∘ σ) ∘ distributeˡ ≈⟨ ∘[] ⟩ ((D₁ distributeˡ⁻¹) ∘ σ) ∘ distributeˡ ≈⟨ ∘[] ⟩
[ ((D₁ distributeˡ⁻¹) ∘ σ) ∘ (idC ⁂ i₁) , ((D₁ distributeˡ⁻¹) ∘ σ) ∘ (idC ⁂ i₂) ] ≈⟨ []-cong₂ (refl⟩∘⟨ ⁂-cong₂ (sym D-identity) refl) (refl⟩∘⟨ ⁂-cong₂ (sym D-identity) refl) ⟩ [ ((D₁ distributeˡ⁻¹) ∘ σ) ∘ (idC ⁂ i₁) , ((D₁ distributeˡ⁻¹) ∘ σ) ∘ (idC ⁂ i₂) ] ≈⟨ []-cong₂ (refl⟩∘⟨ ⁂-cong₂ (sym D-identity) refl) (refl⟩∘⟨ ⁂-cong₂ (sym D-identity) refl) ⟩
[ ((D₁ distributeˡ⁻¹) ∘ σ) ∘ (D₁ idC ⁂ i₁) , ((D₁ distributeˡ⁻¹) ∘ σ) ∘ (D₁ idC ⁂ i₂) ] ≈⟨ []-cong₂ (pullʳ (σ-commute idC i₁)) (pullʳ (σ-commute idC i₂)) ⟩ [ ((D₁ distributeˡ⁻¹) ∘ σ) ∘ (D₁ idC ⁂ i₁) , ((D₁ distributeˡ⁻¹) ∘ σ) ∘ (D₁ idC ⁂ i₂) ] ≈⟨ []-cong₂ (pullʳ (σ-commute idC i₁)) (pullʳ (σ-commute idC i₂)) ⟩
[ (D₁ distributeˡ⁻¹) ∘ D₁ (idC ⁂ i₁) ∘ σ , (D₁ distributeˡ⁻¹) ∘ D₁ (idC ⁂ i₂) ∘ σ ] ≈⟨ []-cong₂ (pullˡ (sym D-homomorphism)) (pullˡ (sym D-homomorphism)) ⟩ [ (D₁ distributeˡ⁻¹) ∘ D₁ (idC ⁂ i₁) ∘ σ , (D₁ distributeˡ⁻¹) ∘ D₁ (idC ⁂ i₂) ∘ σ ] ≈⟨ []-cong₂ (pullˡ (sym D-homomorphism)) (pullˡ (sym D-homomorphism)) ⟩
[ D₁ (distributeˡ⁻¹ ∘ (idC ⁂ i₁)) ∘ σ , D₁ (distributeˡ⁻¹ ∘ (idC ⁂ i₂)) ∘ σ ] ≈⟨ []-cong₂ (D-resp-≈ dstr-law₁ ⟩∘⟨refl) (D-resp-≈ dstr-law₂ ⟩∘⟨refl) ⟩ [ D₁ (distributeˡ⁻¹ ∘ (idC ⁂ i₁)) ∘ σ , D₁ (distributeˡ⁻¹ ∘ (idC ⁂ i₂)) ∘ σ ] ≈⟨ []-cong₂ (D-resp-≈ dstr-law₁ ⟩∘⟨refl) (D-resp-≈ dstr-law₂ ⟩∘⟨refl) ⟩
[ D₁ i₁ ∘ σ , D₁ i₂ ∘ σ ] ≈˘⟨ cancelʳ (IsIso.isoˡ isIsoˡ) ⟩ [ D₁ i₁ ∘ σ , D₁ i₂ ∘ σ ] ≈˘⟨ cancelʳ (IsIso.isoˡ isIsoˡ) ⟩
([ D₁ i₁ ∘ σ , D₁ i₂ ∘ σ ] ∘ distributeˡ⁻¹) ∘ distributeˡ ∎) ([ D₁ i₁ ∘ σ , D₁ i₂ ∘ σ ] ∘ distributeˡ⁻¹) ∘ distributeˡ ∎)
helper₃ = begin helper₃ = begin
[ idC +₁ τ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ] ∘ (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ (out {X} ⁂ out {Y}) ≈⟨ refl⟩∘⟨ helper ⟩ [ idC +₁ τ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ] ∘ (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ (out {X} ⁂ out {Y}) ≈⟨ refl⟩∘⟨ helper ⟩
[ idC +₁ τ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ] ∘ [ [ i₁ ∘ i₁ , i₂ ∘ i₁ ] , [ (i₁ ∘ i₂) , (i₂ ∘ i₂) ] ] ∘ w ≈⟨ pullˡ ∘[] ⟩ [ idC +₁ τ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ] ∘ [ [ i₁ ∘ i₁ , i₂ ∘ i₁ ] , [ (i₁ ∘ i₂) , (i₂ ∘ i₂) ] ] ∘ w ≈⟨ pullˡ ∘[] ⟩
[ [ idC +₁ τ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ] ∘ [ i₁ ∘ i₁ , i₂ ∘ i₁ ] , [ idC +₁ τ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ] ∘ [ (i₁ ∘ i₂) , (i₂ ∘ i₂) ] ] ∘ w ≈⟨ ([]-cong₂ ∘[] ∘[]) ⟩∘⟨refl ⟩ [ [ idC +₁ τ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ] ∘ [ i₁ ∘ i₁ , i₂ ∘ i₁ ]
[ [ [ idC +₁ τ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ] ∘ i₁ ∘ i₁ , [ idC +₁ τ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ] ∘ i₂ ∘ i₁ ] , [ [ idC +₁ τ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ] ∘ (i₁ ∘ i₂) , [ idC +₁ τ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ] ∘ (i₂ ∘ i₂) ] ] ∘ w ≈⟨ ([]-cong₂ ([]-cong₂ (pullˡ inject₁) (pullˡ inject₂ ○ assoc)) ([]-cong₂ (pullˡ inject₁) (pullˡ inject₂ ○ assoc))) ⟩∘⟨refl ⟩ , [ idC +₁ τ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ] ∘ [ (i₁ ∘ i₂) , (i₂ ∘ i₂) ] ] ∘ w ≈⟨ ([]-cong₂ ∘[] ∘[]) ⟩∘⟨refl ⟩
[ [ [ idC +₁ τ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ] ∘ i₁ ∘ i₁ , [ idC +₁ τ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ] ∘ i₂ ∘ i₁ ]
, [ [ idC +₁ τ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ] ∘ (i₁ ∘ i₂)
, [ idC +₁ τ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ] ∘ (i₂ ∘ i₂) ] ] ∘ w ≈⟨ ([]-cong₂ ([]-cong₂ (pullˡ inject₁) (pullˡ inject₂ ○ assoc)) ([]-cong₂ (pullˡ inject₁) (pullˡ inject₂ ○ assoc))) ⟩∘⟨refl ⟩
[ [ (idC +₁ τ) ∘ i₁ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ∘ i₁ ] , [ (idC +₁ τ) ∘ i₂ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ∘ i₂ ] ] ∘ w ≈⟨ ([]-cong₂ ([]-cong₂ (+₁∘i₁ ○ identityʳ) (refl⟩∘⟨ inject₁)) ([]-cong₂ +₁∘i₂ (refl⟩∘⟨ inject₂))) ⟩∘⟨refl ⟩ [ [ (idC +₁ τ) ∘ i₁ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ∘ i₁ ] , [ (idC +₁ τ) ∘ i₂ , i₂ ∘ [ D₁ swap ∘ τ ∘ swap , ▷ ∘ extend τ ∘ σ ] ∘ i₂ ] ] ∘ w ≈⟨ ([]-cong₂ ([]-cong₂ (+₁∘i₁ ○ identityʳ) (refl⟩∘⟨ inject₁)) ([]-cong₂ +₁∘i₂ (refl⟩∘⟨ inject₂))) ⟩∘⟨refl ⟩
[ [ i₁ , i₂ ∘ σ ] , [ i₂ ∘ τ , i₂ ∘ ▷ ∘ extend τ ∘ σ ] ] ∘ w ≈˘⟨ ([]-cong₂ ([]-cong₂ identityʳ refl) ∘[]) ⟩∘⟨refl ⟩ [ [ i₁ , i₂ ∘ σ ] , [ i₂ ∘ τ , i₂ ∘ ▷ ∘ extend τ ∘ σ ] ] ∘ w ≈˘⟨ ([]-cong₂ ([]-cong₂ identityʳ refl) ∘[]) ⟩∘⟨refl ⟩
[ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ extend τ ∘ σ ] ] ∘ w ∎ [ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ extend τ ∘ σ ] ] ∘ w
where where
helper : (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ (out {X} ⁂ out {Y}) ≈ [ [ i₁ ∘ i₁ , i₂ ∘ i₁ ] , [ (i₁ ∘ i₂) , (i₂ ∘ i₂) ] ] ∘ w helper : (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ (out {X} ⁂ out {Y}) ≈ [ [ i₁ ∘ i₁ , i₂ ∘ i₁ ] , [ (i₁ ∘ i₂) , (i₂ ∘ i₂) ] ] ∘ w
helper = sym (begin helper = sym (begin
[ [ i₁ ∘ i₁ , i₂ ∘ i₁ ] , [ (i₁ ∘ i₂) , (i₂ ∘ i₂) ] ] ∘ w ≈⟨ pullˡ []∘+₁ ⟩ [ [ i₁ ∘ i₁ , i₂ ∘ i₁ ] , [ (i₁ ∘ i₂) , (i₂ ∘ i₂) ] ] ∘ w ≈⟨ pullˡ []∘+₁ ⟩
[ (i₁ +₁ i₁) ∘ distributeʳ⁻¹ , (i₂ +₁ i₂) ∘ distributeʳ⁻¹ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ ((+₁-cong₂ (sym dstr-law₁) (sym dstr-law₁)) ⟩∘⟨refl) ((+₁-cong₂ (sym dstr-law₂) (sym dstr-law₂)) ⟩∘⟨refl)) ⟩∘⟨refl ⟩ [ (i₁ +₁ i₁) ∘ distributeʳ⁻¹ , (i₂ +₁ i₂) ∘ distributeʳ⁻¹ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ ((+₁-cong₂ (sym dstr-law₁) (sym dstr-law₁)) ⟩∘⟨refl) ((+₁-cong₂ (sym dstr-law₂) (sym dstr-law₂)) ⟩∘⟨refl)) ⟩∘⟨refl ⟩
[ (distributeˡ⁻¹ ∘ (idC ⁂ i₁) +₁ distributeˡ⁻¹ ∘ (idC ⁂ i₁)) ∘ distributeʳ⁻¹ , (distributeˡ⁻¹ ∘ (idC ⁂ i₂) +₁ distributeˡ⁻¹ ∘ (idC ⁂ i₂)) ∘ distributeʳ⁻¹ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ sym (([]-cong₂ (pullˡ +₁∘+₁) (pullˡ +₁∘+₁)) ⟩∘⟨refl) ⟩ [ (distributeˡ⁻¹ ∘ (idC ⁂ i₁) +₁ distributeˡ⁻¹ ∘ (idC ⁂ i₁)) ∘ distributeʳ⁻¹
[ (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ ((idC ⁂ i₁) +₁ (idC ⁂ i₁)) ∘ distributeʳ⁻¹ , (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ ((idC ⁂ i₂) +₁ (idC ⁂ i₂)) ∘ distributeʳ⁻¹ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ sym (pullˡ ∘[]) ⟩ , (distributeˡ⁻¹ ∘ (idC ⁂ i₂) +₁ distributeˡ⁻¹ ∘ (idC ⁂ i₂)) ∘ distributeʳ⁻¹ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ sym (([]-cong₂ (pullˡ +₁∘+₁) (pullˡ +₁∘+₁)) ⟩∘⟨refl) ⟩
[ (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ ((idC ⁂ i₁) +₁ (idC ⁂ i₁)) ∘ distributeʳ⁻¹
, (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ ((idC ⁂ i₂) +₁ (idC ⁂ i₂)) ∘ distributeʳ⁻¹ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ sym (pullˡ ∘[]) ⟩
(distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ [ ((idC ⁂ i₁) +₁ (idC ⁂ i₁)) ∘ distributeʳ⁻¹ , ((idC ⁂ i₂) +₁ (idC ⁂ i₂)) ∘ distributeʳ⁻¹ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ refl⟩∘⟨ []-cong₂ (distribute₁' i₁ idC idC) (distribute₁' i₂ idC idC) ⟩∘⟨refl ⟩ (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ [ ((idC ⁂ i₁) +₁ (idC ⁂ i₁)) ∘ distributeʳ⁻¹ , ((idC ⁂ i₂) +₁ (idC ⁂ i₂)) ∘ distributeʳ⁻¹ ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ refl⟩∘⟨ []-cong₂ (distribute₁' i₁ idC idC) (distribute₁' i₂ idC idC) ⟩∘⟨refl ⟩
(distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ [ distributeʳ⁻¹ ∘ ((idC +₁ idC) ⁂ i₁) , distributeʳ⁻¹ ∘ ((idC +₁ idC) ⁂ i₂) ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ refl⟩∘⟨ (sym (pullˡ ∘[])) ⟩ (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ [ distributeʳ⁻¹ ∘ ((idC +₁ idC) ⁂ i₁) , distributeʳ⁻¹ ∘ ((idC +₁ idC) ⁂ i₂) ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ refl⟩∘⟨ (sym (pullˡ ∘[])) ⟩
(distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ [ ((idC +₁ idC) ⁂ i₁) , ((idC +₁ idC) ⁂ i₂) ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ ([]-cong₂ (⁂-cong₂ ([]-unique id-comm-sym id-comm-sym) refl) (⁂-cong₂ ([]-unique id-comm-sym id-comm-sym) refl)) ⟩∘⟨refl ⟩ (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ [ ((idC +₁ idC) ⁂ i₁) , ((idC +₁ idC) ⁂ i₂) ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ ([]-cong₂ (⁂-cong₂ ([]-unique id-comm-sym id-comm-sym) refl) (⁂-cong₂ ([]-unique id-comm-sym id-comm-sym) refl)) ⟩∘⟨refl ⟩
(distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ [ (idC ⁂ i₁) , (idC ⁂ i₂) ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ cancelˡ (IsIso.isoʳ isIsoˡ) ⟩ (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ [ (idC ⁂ i₁) , (idC ⁂ i₂) ] ∘ distributeˡ⁻¹ ∘ (out ⁂ out) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ cancelˡ (IsIso.isoʳ isIsoˡ) ⟩
(distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ∎) (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ∎)
fixpoint₂ = Iso⇒Mono ((_≅_.iso out-≅)) (extend τ ∘ σ) (out⁻¹ ∘ [ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ extend τ ∘ σ ] ] ∘ w) (begin fixpoint₂ = Iso⇒Mono ((_≅_.iso out-≅)) (extend τ ∘ σ) (out⁻¹ ∘ [ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ extend τ ∘ σ ] ] ∘ w) (begin
out ∘ extend τ ∘ σ ≈⟨ pullˡ (extendlaw τ) ⟩ out ∘ extend τ ∘ σ ≈⟨ pullˡ (extendlaw τ) ⟩
([ out ∘ τ , i₂ ∘ extend τ ] ∘ out) ∘ σ ≈⟨ pullʳ (u-commutes (σ-coalg X (D₀ Y))) ⟩ ([ out ∘ τ , i₂ ∘ extend τ ] ∘ out) ∘ σ ≈⟨ pullʳ (u-commutes (σ-coalg X (D₀ Y))) ⟩
[ out ∘ τ , i₂ ∘ extend τ ] ∘ (idC +₁ σ) ∘ distributeʳ⁻¹ ∘ (out ⁂ idC) ≈⟨ pullˡ []∘+₁ ⟩ [ out ∘ τ , i₂ ∘ extend τ ] ∘ (idC +₁ σ) ∘ distributeʳ⁻¹ ∘ (out ⁂ idC) ≈⟨ pullˡ []∘+₁ ⟩
[ (out ∘ τ) ∘ idC , (i₂ ∘ extend τ) ∘ σ ] ∘ distributeʳ⁻¹ ∘ (out ⁂ idC) ≈⟨ ([]-cong₂ (identityʳ ○ τ-law (X , Y)) assoc) ⟩∘⟨refl ⟩ [ (out ∘ τ) ∘ idC , (i₂ ∘ extend τ) ∘ σ ] ∘ distributeʳ⁻¹ ∘ (out ⁂ idC) ≈⟨ ([]-cong₂ (identityʳ ○ τ-law (X , Y)) assoc) ⟩∘⟨refl ⟩
[ (idC +₁ τ) ∘ distributeˡ⁻¹ ∘ (idC ⁂ out) , i₂ ∘ extend τ ∘ σ ] ∘ distributeʳ⁻¹ ∘ (out ⁂ idC) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ sym (⁂∘⁂ ○ ⁂-cong₂ (elimˡ ([]-unique id-comm-sym id-comm-sym)) (_≅_.isoˡ out-≅)) ⟩ [ (idC +₁ τ) ∘ distributeˡ⁻¹ ∘ (idC ⁂ out) , i₂ ∘ extend τ ∘ σ ] ∘ distributeʳ⁻¹ ∘ (out ⁂ idC) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ sym (⁂∘⁂ ○ ⁂-cong₂ (elimˡ ([]-unique id-comm-sym id-comm-sym)) (_≅_.isoˡ out-≅)) ⟩
[ (idC +₁ τ) ∘ distributeˡ⁻¹ ∘ (idC ⁂ out) , i₂ ∘ extend τ ∘ σ ] ∘ distributeʳ⁻¹ ∘ ((idC +₁ idC) ⁂ out⁻¹) ∘ (out ⁂ out) ≈⟨ refl⟩∘⟨ (pullˡ (sym (distribute₁' out⁻¹ idC idC))) ⟩ [ (idC +₁ τ) ∘ distributeˡ⁻¹ ∘ (idC ⁂ out) , i₂ ∘ extend τ ∘ σ ] ∘ distributeʳ⁻¹ ∘ ((idC +₁ idC) ⁂ out⁻¹) ∘ (out ⁂ out) ≈⟨ refl⟩∘⟨ (pullˡ (sym (distribute₁' out⁻¹ idC idC))) ⟩
[ (idC +₁ τ) ∘ distributeˡ⁻¹ ∘ (idC ⁂ out) , i₂ ∘ extend τ ∘ σ ] ∘ (((idC ⁂ out⁻¹) +₁ (idC ⁂ out⁻¹)) ∘ distributeʳ⁻¹) ∘ (out ⁂ out) ≈⟨ pullˡ (pullˡ []∘+₁) ⟩ [ (idC +₁ τ) ∘ distributeˡ⁻¹ ∘ (idC ⁂ out) , i₂ ∘ extend τ ∘ σ ] ∘ (((idC ⁂ out⁻¹) +₁ (idC ⁂ out⁻¹)) ∘ distributeʳ⁻¹) ∘ (out ⁂ out) ≈⟨ pullˡ (pullˡ []∘+₁) ⟩
([ ((idC +₁ τ) ∘ distributeˡ⁻¹ ∘ (idC ⁂ out)) ∘ (idC ⁂ out⁻¹) , (i₂ ∘ extend τ ∘ σ) ∘ (idC ⁂ out⁻¹) ] ∘ distributeʳ⁻¹) ∘ (out ⁂ out) ≈⟨ assoc ○ ([]-cong₂ (pullʳ (cancelʳ (⁂∘⁂ ○ ⁂-cong₂ identity² (_≅_.isoʳ out-≅) ○ ⟨⟩-unique id-comm id-comm))) (refl⟩∘⟨ (⁂-cong₂ (sym D-identity) refl))) ⟩∘⟨refl ⟩ ([ ((idC +₁ τ) ∘ distributeˡ⁻¹ ∘ (idC ⁂ out)) ∘ (idC ⁂ out⁻¹) , (i₂ ∘ extend τ ∘ σ) ∘ (idC ⁂ out⁻¹) ] ∘ distributeʳ⁻¹) ∘ (out ⁂ out) ≈⟨ assoc ○ ([]-cong₂ (pullʳ (cancelʳ (⁂∘⁂ ○ ⁂-cong₂ identity² (_≅_.isoʳ out-≅) ○ ⟨⟩-unique id-comm id-comm))) (refl⟩∘⟨ (⁂-cong₂ (sym D-identity) refl))) ⟩∘⟨refl ⟩
[ (idC +₁ τ) ∘ distributeˡ⁻¹ , (i₂ ∘ extend τ ∘ σ) ∘ (D₁ idC ⁂ out⁻¹) ] ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (pullʳ (pullʳ (σ-commute idC out⁻¹)))) ⟩∘⟨refl ⟩ [ (idC +₁ τ) ∘ distributeˡ⁻¹ , (i₂ ∘ extend τ ∘ σ) ∘ (D₁ idC ⁂ out⁻¹) ] ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (pullʳ (pullʳ (σ-commute idC out⁻¹)))) ⟩∘⟨refl ⟩
[ (idC +₁ τ) ∘ distributeˡ⁻¹ , i₂ ∘ extend τ ∘ D₁ (idC ⁂ out⁻¹) ∘ σ ] ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ (pullˡ ((sym k-assoc) ○ extend-≈ (pullˡ k-identityʳ))))) ⟩∘⟨refl ⟩ [ (idC +₁ τ) ∘ distributeˡ⁻¹ , i₂ ∘ extend τ ∘ D₁ (idC ⁂ out⁻¹) ∘ σ ] ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ (pullˡ ((sym k-assoc) ○ extend-≈ (pullˡ k-identityʳ))))) ⟩∘⟨refl ⟩
[ (idC +₁ τ) ∘ distributeˡ⁻¹ , i₂ ∘ extend (τ ∘ (idC ⁂ out⁻¹)) ∘ σ ] ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ ((extend-≈ (τ-helper (X , Y))) ⟩∘⟨refl))) ⟩∘⟨refl ⟩ [ (idC +₁ τ) ∘ distributeˡ⁻¹ , i₂ ∘ extend (τ ∘ (idC ⁂ out⁻¹)) ∘ σ ] ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ ((extend-≈ (τ-helper (X , Y))) ⟩∘⟨refl))) ⟩∘⟨refl ⟩
[ (idC +₁ τ) ∘ distributeˡ⁻¹ , i₂ ∘ extend (out⁻¹ ∘ (idC +₁ τ) ∘ distributeˡ⁻¹) ∘ σ ] ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈˘⟨ ([]-cong₂ refl (refl⟩∘⟨ pullˡ ((sym k-assoc) ○ (extend-≈ (pullˡ k-identityʳ) ○ extend-≈ assoc)))) ⟩∘⟨refl ⟩ [ (idC +₁ τ) ∘ distributeˡ⁻¹ , i₂ ∘ extend (out⁻¹ ∘ (idC +₁ τ) ∘ distributeˡ⁻¹) ∘ σ ] ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈˘⟨ ([]-cong₂ refl (refl⟩∘⟨ pullˡ ((sym k-assoc) ○ (extend-≈ (pullˡ k-identityʳ) ○ extend-≈ assoc)))) ⟩∘⟨refl ⟩
[ (idC +₁ τ) ∘ distributeˡ⁻¹ , i₂ ∘ extend (out⁻¹ ∘ (idC +₁ τ)) ∘ D₁ distributeˡ⁻¹ ∘ σ ] ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl ((refl⟩∘⟨ (refl⟩∘⟨ helper₂)) ○ sym assoc²')) ⟩∘⟨refl ⟩ [ (idC +₁ τ) ∘ distributeˡ⁻¹ , i₂ ∘ extend (out⁻¹ ∘ (idC +₁ τ)) ∘ D₁ distributeˡ⁻¹ ∘ σ ] ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl ((refl⟩∘⟨ (refl⟩∘⟨ helper₂)) ○ sym assoc²')) ⟩∘⟨refl ⟩
[ (idC +₁ τ) ∘ distributeˡ⁻¹ , (i₂ ∘ extend (out⁻¹ ∘ (idC +₁ τ)) ∘ [ D₁ i₁ ∘ σ , D₁ i₂ ∘ σ ]) ∘ distributeˡ⁻¹ ] ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈˘⟨ []∘+₁ ⟩∘⟨refl ⟩ [ (idC +₁ τ) ∘ distributeˡ⁻¹ , (i₂ ∘ extend (out⁻¹ ∘ (idC +₁ τ)) ∘ [ D₁ i₁ ∘ σ , D₁ i₂ ∘ σ ]) ∘ distributeˡ⁻¹ ] ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈˘⟨ []∘+₁ ⟩∘⟨refl ⟩
([ (idC +₁ τ) , i₂ ∘ extend (out⁻¹ ∘ (idC +₁ τ)) ∘ [ D₁ i₁ ∘ σ , D₁ i₂ ∘ σ ] ] ∘ (distributeˡ⁻¹ +₁ distributeˡ⁻¹)) ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ assoc ○ ([]-cong₂ refl (refl⟩∘⟨ ∘[])) ⟩∘⟨refl ⟩ ([ (idC +₁ τ) , i₂ ∘ extend (out⁻¹ ∘ (idC +₁ τ)) ∘ [ D₁ i₁ ∘ σ , D₁ i₂ ∘ σ ] ] ∘ (distributeˡ⁻¹ +₁ distributeˡ⁻¹)) ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ assoc ○ ([]-cong₂ refl (refl⟩∘⟨ ∘[])) ⟩∘⟨refl ⟩
[ (idC +₁ τ) , i₂ ∘ [ extend (out⁻¹ ∘ (idC +₁ τ)) ∘ D₁ i₁ ∘ σ , extend (out⁻¹ ∘ (idC +₁ τ)) ∘ D₁ i₂ ∘ σ ] ] ∘ (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ ([]-cong₂ (pullˡ (sym k-assoc ○ extend-≈ (pullˡ k-identityʳ))) (pullˡ (sym k-assoc ○ extend-≈ (pullˡ k-identityʳ)))))) ⟩∘⟨refl ⟩ [ (idC +₁ τ) , i₂ ∘ [ extend (out⁻¹ ∘ (idC +₁ τ)) ∘ D₁ i₁ ∘ σ , extend (out⁻¹ ∘ (idC +₁ τ)) ∘ D₁ i₂ ∘ σ ] ] ∘ (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ ([]-cong₂ refl (refl⟩∘⟨ ([]-cong₂ (pullˡ (sym k-assoc ○ extend-≈ (pullˡ k-identityʳ))) (pullˡ (sym k-assoc ○ extend-≈ (pullˡ k-identityʳ)))))) ⟩∘⟨refl ⟩
[ (idC +₁ τ) , i₂ ∘ [ extend ((out⁻¹ ∘ (idC +₁ τ)) ∘ i₁) ∘ σ , extend ((out⁻¹ ∘ (idC +₁ τ)) ∘ i₂) ∘ σ ] ] ∘ (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ []-cong₂ refl (refl⟩∘⟨ []-cong₂ (extend-≈ (pullʳ inject₁) ⟩∘⟨refl) (extend-≈ (pullʳ inject₂) ⟩∘⟨refl)) ⟩∘⟨refl ⟩ [ (idC +₁ τ) , i₂ ∘ [ extend ((out⁻¹ ∘ (idC +₁ τ)) ∘ i₁) ∘ σ , extend ((out⁻¹ ∘ (idC +₁ τ)) ∘ i₂) ∘ σ ] ] ∘ (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ []-cong₂ refl (refl⟩∘⟨ []-cong₂ (extend-≈ (pullʳ inject₁) ⟩∘⟨refl) (extend-≈ (pullʳ inject₂) ⟩∘⟨refl)) ⟩∘⟨refl ⟩
[ (idC +₁ τ) , i₂ ∘ [ extend (out⁻¹ ∘ i₁ ∘ idC) ∘ σ , extend (out⁻¹ ∘ i₂ ∘ τ) ∘ σ ] ] ∘ (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ []-cong₂ refl (refl⟩∘⟨ []-cong₂ (elimˡ (extend-≈ (refl⟩∘⟨ identityʳ) ○ k-identityˡ)) (extend-≈ sym-assoc ⟩∘⟨refl ○ sym (pullˡ (▷∘extendˡ τ)))) ⟩∘⟨refl ⟩ [ (idC +₁ τ) , i₂ ∘ [ extend (out⁻¹ ∘ i₁ ∘ idC) ∘ σ , extend (out⁻¹ ∘ i₂ ∘ τ) ∘ σ ] ] ∘ (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ []-cong₂ refl (refl⟩∘⟨ []-cong₂ (elimˡ (extend-≈ (refl⟩∘⟨ identityʳ) ○ k-identityˡ)) (extend-≈ sym-assoc ⟩∘⟨refl ○ sym (pullˡ (▷∘extendˡ τ)))) ⟩∘⟨refl ⟩
[ (idC +₁ τ) , i₂ ∘ [ σ , ▷ ∘ extend τ ∘ σ ] ] ∘ (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ helper₃ ⟩ [ (idC +₁ τ) , i₂ ∘ [ σ , ▷ ∘ extend τ ∘ σ ] ] ∘ (distributeˡ⁻¹ +₁ distributeˡ⁻¹) ∘ distributeʳ⁻¹ ∘ (out ⁂ out) ≈⟨ helper₃ ⟩
[ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ extend τ ∘ σ ] ] ∘ w ≈˘⟨ cancelˡ (_≅_.isoʳ out-≅) ⟩ [ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ extend τ ∘ σ ] ] ∘ w ≈˘⟨ cancelˡ (_≅_.isoʳ out-≅) ⟩
out ∘ out⁻¹ ∘ [ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ extend τ ∘ σ ] ] ∘ w ∎) out ∘ out⁻¹ ∘ [ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ extend τ ∘ σ ] ] ∘ w ∎)
fixpoint-eq : ∀ {f : D₀ X × D₀ Y ⇒ D₀ (X × Y)} → f ≈ out⁻¹ ∘ [ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ f ] ] ∘ w → f ≈ extend [ now , f ] ∘ g fixpoint-eq : ∀ {f : D₀ X × D₀ Y ⇒ D₀ (X × Y)} → f ≈ out⁻¹ ∘ [ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ f ] ] ∘ w → f ≈ extend [ now , f ] ∘ g
fixpoint-eq {f} fix = begin fixpoint-eq {f} fix = begin
f ≈⟨ fix ⟩ f ≈⟨ fix ⟩
out⁻¹ ∘ [ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ f ] ] ∘ w ≈˘⟨ refl⟩∘⟨ []-cong₂ refl (refl⟩∘⟨ ([]-cong₂ refl (pullʳ inject₂))) ⟩∘⟨refl ⟩ out⁻¹ ∘ [ idC +₁ σ , i₂ ∘ [ τ , ▷ ∘ f ] ] ∘ w ≈˘⟨ refl⟩∘⟨ []-cong₂ refl (refl⟩∘⟨ ([]-cong₂ refl (pullʳ inject₂))) ⟩∘⟨refl ⟩
out⁻¹ ∘ [ idC +₁ σ , i₂ ∘ [ τ , (▷ ∘ [ now , f ]) ∘ i₂ ] ] ∘ w ≈˘⟨ refl⟩∘⟨ []-cong₂ refl (sym ∘[] ○ refl⟩∘⟨ []-cong₂ (elimˡ (extend-≈ inject₁ ○ k-identityˡ)) (pullˡ k-identityʳ)) ⟩∘⟨refl ⟩ out⁻¹ ∘ [ idC +₁ σ , i₂ ∘ [ τ , (▷ ∘ [ now , f ]) ∘ i₂ ] ] ∘ w ≈˘⟨ refl⟩∘⟨ []-cong₂ refl (sym ∘[] ○ refl⟩∘⟨ []-cong₂ (elimˡ (extend-≈ inject₁ ○ k-identityˡ)) (pullˡ k-identityʳ)) ⟩∘⟨refl ⟩
out⁻¹ ∘ [ idC +₁ σ , [ i₂ ∘ extend ([ now , f ] ∘ i₁) ∘ τ , i₂ ∘ extend (▷ ∘ [ now , f ]) ∘ now ∘ i₂ ] ] ∘ w ≈˘⟨ refl⟩∘⟨ []-cong₂ ([]-cong₂ (sym identityʳ) (refl⟩∘⟨ (elimˡ ((extend-≈ inject₁) ○ k-identityˡ)))) ([]-cong₂ (pullʳ (pullˡ ((sym k-assoc) ○ extend-≈ (pullˡ k-identityʳ)))) (pullʳ (pullˡ (▷∘extendʳ [ now , f ])))) ⟩∘⟨refl ⟩ out⁻¹ ∘ [ idC +₁ σ , [ i₂ ∘ extend ([ now , f ] ∘ i₁) ∘ τ , i₂ ∘ extend (▷ ∘ [ now , f ]) ∘ now ∘ i₂ ] ] ∘ w ≈˘⟨ refl⟩∘⟨ []-cong₂ ([]-cong₂ (sym identityʳ) (refl⟩∘⟨ (elimˡ ((extend-≈ inject₁) ○ k-identityˡ)))) ([]-cong₂ (pullʳ (pullˡ ((sym k-assoc) ○ extend-≈ (pullˡ k-identityʳ)))) (pullʳ (pullˡ (▷∘extendʳ [ now , f ])))) ⟩∘⟨refl ⟩
out⁻¹ ∘ [ [ i₁ , i₂ ∘ extend ([ now , f ] ∘ i₁) ∘ σ ] , [ (i₂ ∘ extend [ now , f ]) ∘ D₁ i₁ ∘ τ , (i₂ ∘ extend [ now , f ]) ∘ ▷ ∘ now ∘ i₂ ] ] ∘ w ≈˘⟨ refl⟩∘⟨ []-cong₂ ([]-cong₂ inject₁ (pullʳ (pullˡ ((sym k-assoc) ○ extend-≈ (pullˡ k-identityʳ))))) ∘[] ⟩∘⟨refl ⟩ out⁻¹ ∘ [ [ i₁ , i₂ ∘ extend ([ now , f ] ∘ i₁) ∘ σ ] , [ (i₂ ∘ extend [ now , f ]) ∘ D₁ i₁ ∘ τ , (i₂ ∘ extend [ now , f ]) ∘ ▷ ∘ now ∘ i₂ ] ] ∘ w ≈˘⟨ refl⟩∘⟨ []-cong₂ ([]-cong₂ inject₁ (pullʳ (pullˡ ((sym k-assoc) ○ extend-≈ (pullˡ k-identityʳ))))) ∘[] ⟩∘⟨refl ⟩
out⁻¹ ∘ [ [ [ i₁ , out ∘ f ] ∘ i₁ , (i₂ ∘ extend [ now , f ]) ∘ D₁ i₁ ∘ σ ] , (i₂ ∘ extend [ now , f ]) ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈˘⟨ refl⟩∘⟨ (([]-cong₂ []∘+₁ (pullˡ inject₂)) ⟩∘⟨refl) ⟩ out⁻¹ ∘ [ [ [ i₁ , out ∘ f ] ∘ i₁ , (i₂ ∘ extend [ now , f ]) ∘ D₁ i₁ ∘ σ ] , (i₂ ∘ extend [ now , f ]) ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈˘⟨ refl⟩∘⟨ (([]-cong₂ []∘+₁ (pullˡ inject₂)) ⟩∘⟨refl) ⟩
out⁻¹ ∘ [ [ [ i₁ , out ∘ f ] , i₂ ∘ extend [ now , f ] ] ∘ (i₁ +₁ D₁ i₁ ∘ σ) , [ [ i₁ , out ∘ f ] , i₂ ∘ extend [ now , f ] ] ∘ i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈˘⟨ refl⟩∘⟨ (pullˡ ∘[]) ⟩ out⁻¹ ∘ [ [ [ i₁ , out ∘ f ] , i₂ ∘ extend [ now , f ] ] ∘ (i₁ +₁ D₁ i₁ ∘ σ) , [ [ i₁ , out ∘ f ] , i₂ ∘ extend [ now , f ] ] ∘ i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈˘⟨ refl⟩∘⟨ (pullˡ ∘[]) ⟩
out⁻¹ ∘ [ [ i₁ , out ∘ f ] , i₂ ∘ extend [ now , f ] ] ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈⟨ Iso⇒Mono (_≅_.iso out-≅) (out⁻¹ ∘ [ [ i₁ , out ∘ f ] , i₂ ∘ extend [ now , f ] ] ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w) (extend [ now , f ] ∘ out⁻¹ ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w) helper ⟩ out⁻¹ ∘ [ [ i₁ , out ∘ f ] , i₂ ∘ extend [ now , f ] ] ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈⟨ Iso⇒Mono (_≅_.iso out-≅) (out⁻¹ ∘ [ [ i₁ , out ∘ f ] , i₂ ∘ extend [ now , f ] ] ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w) (extend [ now , f ] ∘ out⁻¹ ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w) helper ⟩
extend [ now , f ] ∘ out⁻¹ ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ∎ extend [ now , f ] ∘ out⁻¹ ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w
where where
helper = begin helper = begin
out ∘ out⁻¹ ∘ [ [ i₁ , out ∘ f ] , i₂ ∘ extend [ now , f ] ] ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈⟨ cancelˡ (_≅_.isoʳ out-≅) ⟩ out ∘ out⁻¹ ∘ [ [ i₁ , out ∘ f ] , i₂ ∘ extend [ now , f ] ] ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈⟨ cancelˡ (_≅_.isoʳ out-≅) ⟩
[ [ i₁ , out ∘ f ] , i₂ ∘ extend [ now , f ] ] ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈˘⟨ ([]-cong₂ (∘[] ○ []-cong₂ unitlaw refl) refl) ⟩∘⟨refl ⟩ [ [ i₁ , out ∘ f ] , i₂ ∘ extend [ now , f ] ] ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈˘⟨ ([]-cong₂ (∘[] ○ []-cong₂ unitlaw refl) refl) ⟩∘⟨refl ⟩
[ out ∘ [ now , f ] , i₂ ∘ extend [ now , f ] ] ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈˘⟨ pullʳ (cancelˡ (_≅_.isoʳ out-≅)) ⟩ [ out ∘ [ now , f ] , i₂ ∘ extend [ now , f ] ] ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈˘⟨ pullʳ (cancelˡ (_≅_.isoʳ out-≅)) ⟩
([ out ∘ [ now , f ] , i₂ ∘ extend [ now , f ] ] ∘ out) ∘ out⁻¹ ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈˘⟨ pullˡ (extendlaw [ now , f ]) ⟩ ([ out ∘ [ now , f ] , i₂ ∘ extend [ now , f ] ] ∘ out) ∘ out⁻¹ ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ≈˘⟨ pullˡ (extendlaw [ now , f ]) ⟩
out ∘ extend [ now , f ] ∘ out⁻¹ ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ∎ out ∘ extend [ now , f ] ∘ out⁻¹ ∘ [ i₁ +₁ D₁ i₁ ∘ σ , i₂ ∘ [ D₁ i₁ ∘ τ , ▷ ∘ now ∘ i₂ ] ] ∘ w ∎
{-
TODO there's an error in the paper, at the end of the proof of proposition two:
the last line of the 3 line calulation 'f = ....'
should be ▷ ∘ η ∘ inr, but is η ∘ inr!!
-}
``` ```