[WIP] equivalence of kleisli- resp. monoidal- representation of monad

This commit is contained in:
Frederik Hanghøj Iversen 2018-02-24 15:13:25 +01:00
parent 3e12331294
commit e4e327d1d2

View file

@ -1,4 +1,4 @@
{-# OPTIONS --cubical #-}
{-# OPTIONS --cubical --allow-unsolved-metas #-}
module Cat.Category.Monad where
open import Agda.Primitive
@ -12,7 +12,7 @@ open import Cat.Category.Functor as F
open import Cat.Category.NaturalTransformation
open import Cat.Categories.Fun
-- "A monad in the monoidal form" [vlad]
-- "A monad in the monoidal form" [voe]
module Monoidal {a b : Level} ( : Category a b) where
private
= a b
@ -27,15 +27,16 @@ module Monoidal {a b : Level} ( : Category a b) where
-- (>=>)
μNat : NaturalTransformation F[ R R ] R
η : Transformation F.identity R
η = proj₁ ηNat
μ : Transformation F[ R R ] R
μ = proj₁ μNat
private
module R = Functor R
module RR = Functor F[ R R ]
module _ {X : Object} where
-- module IdRX = Functor (F.identity {C = RX})
η : Transformation F.identity R
η = proj₁ ηNat
ηX : [ X , R.func* X ]
ηX = η X
RηX : [ R.func* X , R.func* (R.func* X) ] -- [ R.func* X , {!R.func* (R.func* X))!} ]
@ -44,8 +45,6 @@ module Monoidal {a b : Level} ( : Category a b) where
IdRX : Arrow (R.func* X) (R.func* X)
IdRX = 𝟙 {R.func* X}
μ : Transformation F[ R R ] R
μ = proj₁ μNat
μX : [ RR.func* X , R.func* X ]
μX = μ X
RμX : [ R.func* (RR.func* X) , RR.func* X ]
@ -77,7 +76,7 @@ module Monoidal {a b : Level} ( : Category a b) where
isMonad : IsMonad raw
open IsMonad isMonad public
-- "A monad in the Kleisli form" [vlad]
-- "A monad in the Kleisli form" [voe]
module Kleisli {a b : Level} ( : Category a b) where
private
= a b
@ -86,13 +85,14 @@ module Kleisli {a b : Level} ( : Category a b) where
record RawMonad : Set where
field
RR : Object Object
η : {X : Object} [ X , RR X ]
-- Note name-change from [voe]
ζ : {X : Object} [ X , RR X ]
rr : {X Y : Object} [ X , RR Y ] [ RR X , RR Y ]
-- Name suggestions are welcome!
IsIdentity = {X : Object}
rr η 𝟙 {RR X}
rr ζ 𝟙 {RR X}
IsNatural = {X Y : Object} (f : [ X , RR Y ])
( [ rr f η ]) f
( [ rr f ζ ]) f
IsDistributive = {X Y Z : Object} (g : [ Y , RR Z ]) (f : [ X , RR Y ])
[ rr g rr f ] rr ( [ rr g f ])
@ -108,3 +108,67 @@ module Kleisli {a b : Level} ( : Category a b) where
raw : RawMonad
isMonad : IsMonad raw
open IsMonad isMonad public
-- Problem 2.3
module _ {a b : Level} { : Category a b} where
private
open Category using (Object ; Arrow ; 𝟙)
open Functor using (func* ; func→)
module M = Monoidal
module K = Kleisli
module _ (m : M.RawMonad) where
private
open M.RawMonad m
module Kraw = K.RawMonad
RR : Object Object
RR = func* R
R→ : {A B : Object} [ A , B ] [ RR A , RR B ]
R→ = func→ R
ζ : {X : Object} [ X , RR X ]
ζ = {!!}
rr : {X Y : Object} [ X , RR Y ] [ RR X , RR Y ]
-- Order is different now!
rr {X} {Y} f = [ f {!!} ]
where
μY : [ func* F[ R R ] Y , func* R Y ]
μY = μ Y
ζY : [ Y , RR Y ]
ζY = ζ {Y}
forthRaw : K.RawMonad
Kraw.RR forthRaw = RR
Kraw.ζ forthRaw = ζ
Kraw.rr forthRaw = rr
module _ {raw : M.RawMonad} (m : M.IsMonad raw) where
open M.IsMonad m
module Kraw = K.RawMonad (forthRaw raw)
module Kis = K.IsMonad
isIdentity : Kraw.IsIdentity
isIdentity = {!!}
isNatural : Kraw.IsNatural
isNatural = {!!}
isDistributive : Kraw.IsDistributive
isDistributive = {!!}
forthIsMonad : K.IsMonad (forthRaw raw)
Kis.isIdentity forthIsMonad = isIdentity
Kis.isNatural forthIsMonad = isNatural
Kis.isDistributive forthIsMonad = isDistributive
forth : M.Monad K.Monad
Kleisli.Monad.raw (forth m) = forthRaw (M.Monad.raw m)
Kleisli.Monad.isMonad (forth m) = forthIsMonad (M.Monad.isMonad m)
eqv : isEquiv M.Monad K.Monad forth
eqv = {!!}
Monoidal≃Kleisli : M.Monad K.Monad
Monoidal≃Kleisli = forth , eqv