cat/src/Cat/Category/Monad.agda

327 lines
12 KiB
Agda
Raw Normal View History

{-# OPTIONS --cubical --allow-unsolved-metas #-}
module Cat.Category.Monad where
open import Agda.Primitive
open import Data.Product
open import Cubical
open import Cat.Category
open import Cat.Category.Functor as F
open import Cat.Category.NaturalTransformation
open import Cat.Categories.Fun
-- "A monad in the monoidal form" [voe]
module Monoidal {a b : Level} ( : Category a b) where
private
= a b
2018-02-25 00:27:20 +00:00
open Category using (Object ; Arrow ; 𝟙 ; _∘_)
open NaturalTransformation
record RawMonad : Set where
field
-- R ~ m
R : Functor
-- η ~ pure
ηNat : NaturalTransformation F.identity R
-- μ ~ join
μNat : NaturalTransformation F[ R R ] R
η : Transformation F.identity R
η = proj₁ ηNat
μ : Transformation F[ R R ] R
μ = proj₁ μNat
2018-02-24 13:00:52 +00:00
private
2018-02-24 13:00:52 +00:00
module R = Functor R
2018-02-25 18:03:30 +00:00
IsAssociative : Set _
IsAssociative = {X : Object}
μ X R.func→ (μ X) μ X μ (R.func* X)
IsInverse : Set _
IsInverse = {X : Object}
μ X η (R.func* X) 𝟙
× μ X R.func→ (η X) 𝟙
record IsMonad (raw : RawMonad) : Set where
open RawMonad raw public
field
isAssociative : IsAssociative
isInverse : IsInverse
2018-02-24 13:00:52 +00:00
2018-02-24 13:01:57 +00:00
record Monad : Set where
field
raw : RawMonad
isMonad : IsMonad raw
open IsMonad isMonad public
postulate propIsMonad : {raw} isProp (IsMonad raw)
Monad≡ : {m n : Monad} Monad.raw m Monad.raw n m n
Monad.raw (Monad≡ eq i) = eq i
Monad.isMonad (Monad≡ {m} {n} eq i) = res i
where
-- TODO: PathJ nightmare + `propIsMonad`.
res : (λ i IsMonad (eq i)) [ Monad.isMonad m Monad.isMonad n ]
res = {!!}
-- "A monad in the Kleisli form" [voe]
2018-02-24 13:00:52 +00:00
module Kleisli {a b : Level} ( : Category a b) where
private
= a b
open Category using (Arrow ; 𝟙 ; Object ; _∘_ ; _>>>_)
2018-02-24 13:00:52 +00:00
record RawMonad : Set where
field
RR : Object Object
-- Note name-change from [voe]
2018-02-26 18:58:27 +00:00
pure : {X : Object} [ X , RR X ]
2018-02-26 18:57:05 +00:00
bind : {X Y : Object} [ X , RR Y ] [ RR X , RR Y ]
fmap : {A B} [ A , B ] [ RR A , RR B ]
2018-02-26 18:58:27 +00:00
fmap f = bind (pure f)
2018-02-26 18:57:05 +00:00
-- Why is (>>=) not implementable? - Because in e.g. the category of sets is
-- `m a` a set. This is not necessarily the case.
_>=>_ : {A B C : Object} [ A , RR B ] [ B , RR C ] [ A , RR C ]
2018-02-26 18:57:05 +00:00
f >=> g = f >>> (bind g)
-- _>>=_ : {A B C : Object} {m : RR A} → [ A , RR B ] → RR C
-- m >>= f = ?
join : {A : Object} [ RR (RR A) , RR A ]
join = bind 𝟙
-- fmap id ≡ id
2018-02-24 13:00:52 +00:00
IsIdentity = {X : Object}
2018-02-26 18:57:05 +00:00
-- aka. `>>= pure ≡ 𝟙`
bind pure 𝟙 {RR X}
2018-02-24 13:00:52 +00:00
IsNatural = {X Y : Object} (f : [ X , RR Y ])
2018-02-26 18:57:05 +00:00
-- aka. `pure >>= f ≡ f`
pure >>> (bind f) f
-- Not stricly a distributive law, since ∘ becomes >=>
2018-02-24 13:00:52 +00:00
IsDistributive = {X Y Z : Object} (g : [ Y , RR Z ]) (f : [ X , RR Y ])
2018-02-26 18:57:05 +00:00
-- `>>= g . >>= f ≡ >>= (>>= g . f) ≡ >>= (\x -> (f x) >>= g)`
(bind f) >>> (bind g) bind (f >=> g)
Fusion = {X Y Z : Object} {g : [ Y , Z ]} {f : [ X , Y ]}
fmap (g f) fmap g fmap f
2018-02-24 13:00:52 +00:00
record IsMonad (raw : RawMonad) : Set where
open RawMonad raw public
field
isIdentity : IsIdentity
isNatural : IsNatural
isDistributive : IsDistributive
fusion : Fusion
fusion {g = g} {f} = begin
fmap (g f) ≡⟨⟩
2018-02-26 18:57:05 +00:00
-- f >=> g = >>= g ∘ f
bind ((f >>> g) >>> pure) ≡⟨ cong bind isAssociative
bind (f >>> (g >>> pure)) ≡⟨ cong (λ φ bind (f >>> φ)) (sym (isNatural _))
bind (f >>> (pure >>> (bind (g >>> pure)))) ≡⟨⟩
bind (f >>> (pure >>> fmap g)) ≡⟨⟩
bind ((fmap g pure) f) ≡⟨ cong bind (sym isAssociative)
bind
(fmap g (pure f)) ≡⟨ sym lem
2018-02-26 18:58:27 +00:00
bind (pure g) bind (pure f) ≡⟨⟩
fmap g fmap f
where
2018-02-26 18:57:05 +00:00
open Category using (isAssociative)
lem : fmap g fmap f bind (fmap g (pure f))
2018-02-26 18:58:27 +00:00
lem = isDistributive (pure g) (pure f)
2018-02-24 13:00:52 +00:00
record Monad : Set where
field
raw : RawMonad
isMonad : IsMonad raw
open IsMonad isMonad public
postulate propIsMonad : {raw} isProp (IsMonad raw)
Monad≡ : {m n : Monad} Monad.raw m Monad.raw n m n
Monad.raw (Monad≡ eq i) = eq i
Monad.isMonad (Monad≡ {m} {n} eq i) = res i
where
-- TODO: PathJ nightmare + `propIsMonad`.
res : (λ i IsMonad (eq i)) [ Monad.isMonad m Monad.isMonad n ]
res = {!!}
-- Problem 2.3
module _ {a b : Level} { : Category a b} where
private
2018-02-24 18:07:58 +00:00
open Category using (Object ; Arrow ; 𝟙 ; _∘_)
open Functor using (func* ; func→)
module M = Monoidal
module K = Kleisli
2018-02-24 18:07:58 +00:00
-- Note similarity with locally defined things in Kleisly.RawMonad!!
module _ (m : M.RawMonad) where
private
open M.RawMonad m
module Kraw = K.RawMonad
2018-02-24 18:07:58 +00:00
RR : Object Object
RR = func* R
2018-02-26 18:58:27 +00:00
pure : {X : Object} [ X , RR X ]
pure {X} = η X
2018-02-26 18:57:05 +00:00
bind : {X Y : Object} [ X , RR Y ] [ RR X , RR Y ]
bind {X} {Y} f = μ Y func→ R f
forthRaw : K.RawMonad
Kraw.RR forthRaw = RR
2018-02-26 18:58:27 +00:00
Kraw.pure forthRaw = pure
2018-02-26 18:57:05 +00:00
Kraw.bind forthRaw = bind
module _ {raw : M.RawMonad} (m : M.IsMonad raw) where
private
open M.IsMonad m
open K.RawMonad (forthRaw raw)
module Kis = K.IsMonad
isIdentity : IsIdentity
isIdentity {X} = begin
2018-02-26 18:58:27 +00:00
bind pure ≡⟨⟩
2018-02-26 18:57:05 +00:00
bind (η X) ≡⟨⟩
μ X func→ R (η X) ≡⟨ proj₂ isInverse
𝟙
module R = Functor R
isNatural : IsNatural
isNatural {X} {Y} f = begin
2018-02-26 18:58:27 +00:00
bind f pure ≡⟨⟩
2018-02-26 18:57:05 +00:00
bind f η X ≡⟨⟩
μ Y R.func→ f η X ≡⟨ sym .isAssociative
μ Y (R.func→ f η X) ≡⟨ cong (λ φ μ Y φ) (sym (ηN f))
μ Y (η (R.func* Y) f) ≡⟨ .isAssociative
μ Y η (R.func* Y) f ≡⟨ cong (λ φ φ f) (proj₁ isInverse)
𝟙 f ≡⟨ proj₂ .isIdentity
f
where
open NaturalTransformation
module = Category
ηN : Natural F.identity R η
ηN = proj₂ ηNat
isDistributive : IsDistributive
isDistributive {X} {Y} {Z} g f = begin
2018-02-26 18:57:05 +00:00
bind g bind f ≡⟨⟩
μ Z R.func→ g (μ Y R.func→ f) ≡⟨ sym lem2
μ Z R.func→ (μ Z R.func→ g f) ≡⟨⟩
2018-02-26 18:57:05 +00:00
μ Z R.func→ (bind g f)
where
-- Proved it in reverse here... otherwise it could be neatly inlined.
lem2
: μ Z R.func→ (μ Z R.func→ g f)
μ Z R.func→ g (μ Y R.func→ f)
lem2 = begin
μ Z R.func→ (μ Z R.func→ g f) ≡⟨ cong (λ φ μ Z φ) distrib
μ Z (R.func→ (μ Z) R.func→ (R.func→ g) R.func→ f) ≡⟨⟩
μ Z (R.func→ (μ Z) RR.func→ g R.func→ f) ≡⟨ {!!} -- ●-solver?
(μ Z R.func→ (μ Z)) (RR.func→ g R.func→ f) ≡⟨ cong (λ φ φ (RR.func→ g R.func→ f)) lemmm
(μ Z μ (R.func* Z)) (RR.func→ g R.func→ f) ≡⟨ {!!} -- ●-solver?
μ Z μ (R.func* Z) RR.func→ g R.func→ f ≡⟨ {!!} -- ●-solver + lem4
μ Z R.func→ g μ Y R.func→ f ≡⟨ sym (Category.isAssociative )
μ Z R.func→ g (μ Y R.func→ f)
where
module RR = Functor F[ R R ]
distrib : {A B C D} {a : Arrow C D} {b : Arrow B C} {c : Arrow A B}
R.func→ (a b c)
R.func→ a R.func→ b R.func→ c
distrib = {!!}
comm : {A B C D E}
{a : Arrow D E} {b : Arrow C D} {c : Arrow B C} {d : Arrow A B}
a (b c d) a b c d
comm = {!!}
μN = proj₂ μNat
lemmm : μ Z R.func→ (μ Z) μ Z μ (R.func* Z)
lemmm = isAssociative
lem4 : μ (R.func* Z) RR.func→ g R.func→ g μ Y
lem4 = μN g
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)
module _ (m : K.Monad) where
private
module = Category
open K.Monad m
module Mraw = M.RawMonad
open NaturalTransformation
rawR : RawFunctor
RawFunctor.func* rawR = RR
2018-02-26 18:58:27 +00:00
RawFunctor.func→ rawR f = bind (pure f)
isFunctorR : IsFunctor rawR
IsFunctor.isIdentity isFunctorR = begin
2018-02-26 18:58:27 +00:00
bind (pure 𝟙) ≡⟨ cong bind (proj₁ .isIdentity)
bind pure ≡⟨ isIdentity
𝟙
IsFunctor.isDistributive isFunctorR {f = f} {g} = begin
2018-02-26 18:58:27 +00:00
bind (pure (g f)) ≡⟨⟩
fmap (g f) ≡⟨ fusion
fmap g fmap f ≡⟨⟩
2018-02-26 18:58:27 +00:00
bind (pure g) bind (pure f)
R : Functor
Functor.raw R = rawR
Functor.isFunctor R = isFunctorR
R2 : Functor
R2 = F[ R R ]
ηNat : NaturalTransformation F.identity R
ηNat = {!!}
μNat : NaturalTransformation R2 R
μNat = {!!}
backRaw : M.RawMonad
Mraw.R backRaw = R
Mraw.ηNat backRaw = ηNat
Mraw.μNat backRaw = μNat
module _ (m : K.Monad) where
open K.Monad m
open M.RawMonad (backRaw m)
module Mis = M.IsMonad
backIsMonad : M.IsMonad (backRaw m)
backIsMonad = {!!}
2018-02-24 18:07:58 +00:00
back : K.Monad M.Monad
Monoidal.Monad.raw (back m) = backRaw m
Monoidal.Monad.isMonad (back m) = backIsMonad m
2018-02-25 02:12:23 +00:00
-- I believe all the proofs here should be `refl`.
module _ (m : K.Monad) where
open K.RawMonad (K.Monad.raw m)
forthRawEq : forthRaw (backRaw m) K.Monad.raw m
K.RawMonad.RR (forthRawEq _) = RR
2018-02-26 18:58:27 +00:00
K.RawMonad.pure (forthRawEq _) = pure
2018-02-25 02:12:23 +00:00
-- stuck
2018-02-26 18:57:05 +00:00
K.RawMonad.bind (forthRawEq i) = {!!}
2018-02-24 18:07:58 +00:00
fortheq : (m : K.Monad) forth (back m) m
fortheq m = K.Monad≡ (forthRawEq m)
2018-02-25 02:12:23 +00:00
module _ (m : M.Monad) where
open M.RawMonad (M.Monad.raw m)
backRawEq : backRaw (forth m) M.Monad.raw m
-- stuck
M.RawMonad.R (backRawEq i) = {!!}
M.RawMonad.ηNat (backRawEq i) = {!!}
M.RawMonad.μNat (backRawEq i) = {!!}
2018-02-24 18:07:58 +00:00
backeq : (m : M.Monad) back (forth m) m
backeq m = M.Monad≡ (backRawEq m)
2018-02-24 18:07:58 +00:00
open import Cubical.GradLemma
eqv : isEquiv M.Monad K.Monad forth
2018-02-24 18:07:58 +00:00
eqv = gradLemma forth back fortheq backeq
Monoidal≃Kleisli : M.Monad K.Monad
Monoidal≃Kleisli = forth , eqv