agda-kleisli/ExtensionSystem.agda
2023-07-15 12:29:26 +02:00

128 lines
No EOL
6.4 KiB
Agda
Raw Permalink Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

module ExtensionSystem where
open import Level
open import Function using (_$_)
open import Data.Product renaming (_×_ to _∧_)
open import Categories.Category
open import Categories.Functor renaming (id to Id)
open import Categories.Monad.Relative renaming (Monad to RMonad)
open import Categories.Monad
open import Categories.NaturalTransformation
-- formalizing the notion 'extension system', formerly known as 'kleisli-triple' and proving that it is equivalent to a monad
-- https://ncatlab.org/nlab/show/extension+system
private
variable
o e : Level
-- typedef, an extensionsystem is a specilization of RMonad
ExtensionSystem : (𝒞 : Category o e) Set (o e)
ExtensionSystem 𝒞 = (RMonad (Id {C = 𝒞}))
--*
-- Proposition: ExtensionSystem is an equivalent definition of monad
--*
-- '→'
ExtensionSystem→Monad : {𝒞 : Category o e} ExtensionSystem 𝒞 Monad 𝒞
ExtensionSystem→Monad {𝒞 = 𝒞} 𝐾 = record
{ F = T
; η = η'
; μ = μ'
-- M3
; identityˡ = Identityˡ
-- M2
; identityʳ = K2
-- M1
; assoc = assoc'
; sym-assoc = sym assoc'
}
where
open RMonad 𝐾 using (unit) renaming (extend to _ᵀ; F₀ to T₀; extend-≈ to ᵀ-≈; identityˡ to K1; identityʳ to K2; assoc to K3)
open Category 𝒞 renaming (id to idC)
open HomReasoning
T : Endofunctor 𝒞
T = RMonad⇒Functor 𝐾
open Functor T renaming (F₁ to T₁)
open Equiv
-- constructing the natural transformation η from the given family of morphisms 'unit'
η' = ntHelper {F = Id} {G = T} record
{ η = λ X unit
; commute = λ {X} {Y} f sym K2
}
-- constructing the natural transformation μ
μ' = ntHelper {F = T ∘F T} {G = T} record
{ η = λ X (idC {A = T₀ X})
; commute = λ {X} {Y} f begin
((idC ) (unit (unit f))) ≈⟨ (sym $ K3)
(((idC ) unit ((unit f) )) ) ≈⟨ ᵀ-≈ sym-assoc
((((idC ) unit) ((unit f) )) ) ≈⟨ ᵀ-≈ (∘-resp-≈ˡ K2)
((idC ((unit f) )) ) ≈⟨ ᵀ-≈ identityˡ
(((unit f) ) ) ≈⟨ ᵀ-≈ (sym identityʳ)
((((unit f) ) idC) ) ≈⟨ K3
(unit f) (idC )
}
open NaturalTransformation η' using () renaming (η to η)
open NaturalTransformation μ' using () renaming (η to μ)
Identityˡ = λ {X} begin
(μ X ((η (T₀ X)) (η X))) ≈⟨ (sym $ K3)
(((idC ) η (T₀ X) η X) ) ≈⟨ ᵀ-≈ sym-assoc
((((idC ) η (T₀ X)) η X) ) ≈⟨ ᵀ-≈ (∘-resp-≈ˡ K2)
((idC η X) ) ≈⟨ ᵀ-≈ identityˡ
(η X ) ≈⟨ K1
idC
assoc' = λ {X} begin
idC (η _ idC ) ≈⟨ sym K3
(((idC ) η (T₀ X) (idC )) ) ≈⟨ ᵀ-≈ sym-assoc
((((idC ) η (T₀ X)) (idC )) ) ≈⟨ ᵀ-≈ (∘-resp-≈ˡ K2)
((idC (idC )) ) ≈⟨ ᵀ-≈ identityˡ
((idC ) ) ≈⟨ ᵀ-≈ (sym identityʳ)
(((idC ) idC) ) ≈⟨ K3
idC idC
-- '←'
Monad→ExtensionSystem : {𝒞 : Category o e} Monad 𝒞 ExtensionSystem 𝒞
Monad→ExtensionSystem {𝒞 = 𝒞} 𝑀 = record
{ F₀ = T₀
; unit = η _
; extend = λ {X} {Y} f μ Y T₁ f
-- K1
; identityˡ = M3
-- K2
; identityʳ = λ {X} {Y} {f} begin
(μ Y T₁ f) η X ≈⟨ assoc
(μ Y T₁ f η X) ≈⟨ ∘-resp-≈ʳ (sym (η-commute _))
(μ Y η (T₀ Y) f) ≈⟨ sym-assoc
((μ Y η (T₀ Y)) f) ≈⟨ ∘-resp-≈ˡ M2
(idC f) ≈⟨ identityˡ
f
-- K3
; assoc = λ {X} {Y} {Z} {f} {g} begin
μ Z T₁ ((μ Z T₁ g) f) ≈⟨ ∘-resp-≈ʳ homomorphism
(μ Z (T₁ (μ Z T₁ g) T₁ f)) ≈⟨ ∘-resp-≈ʳ (∘-resp-≈ˡ homomorphism)
(μ Z (T₁ (μ Z) T₁ (T₁ g)) T₁ f) ≈⟨ sym-assoc
((μ Z T₁ (μ Z) T₁ (T₁ g)) T₁ f) ≈⟨ ∘-resp-≈ˡ sym-assoc
(((μ Z T₁ (μ Z)) T₁ (T₁ g)) T₁ f) ≈⟨ ∘-resp-≈ˡ (∘-resp-≈ˡ M1)
(((μ Z μ (T₀ Z)) T₁ (T₁ g)) T₁ f) ≈⟨ ∘-resp-≈ˡ assoc
((μ Z μ (T₀ Z) T₁ (T₁ g)) T₁ f) ≈⟨ ∘-resp-≈ˡ (∘-resp-≈ʳ (μ-commute g))
((μ Z T₁ g μ Y) T₁ f) ≈⟨ assoc
(μ Z (T₁ g μ Y) T₁ f) ≈⟨ ∘-resp-≈ʳ assoc
(μ Z T₁ g μ Y T₁ f) ≈⟨ sym-assoc
(μ Z T₁ g) μ Y T₁ f
; extend-≈ = λ fg ∘-resp-≈ʳ (T-resp-≈ fg)
}
where
open Category 𝒞 renaming (id to idC)
open Monad 𝑀 using () renaming (F to T; η to η'; μ to μ'; assoc to M1; identityʳ to M2; identityˡ to M3)
open Functor T renaming (F₀ to T₀; F₁ to T₁; F-resp-≈ to T-resp-≈)
open NaturalTransformation η' using (η) renaming (commute to η-commute)
open NaturalTransformation μ' renaming (η to μ; commute to μ-commute)
open HomReasoning
open Equiv
-- proof:
ExtensionSystem↔Monad : {𝒞 : Category o e} (Monad 𝒞 ExtensionSystem 𝒞) (ExtensionSystem 𝒞 Monad 𝒞)
ExtensionSystem↔Monad {𝒞} = (Monad→ExtensionSystem , ExtensionSystem→Monad)