cat/src/Cat/Category/Monad.agda

220 lines
7.9 KiB
Agda
Raw Normal View History

{---
Monads
This module presents two formulations of monads:
* The standard monoidal presentation
* Kleisli's presentation
The first one defines a monad in terms of an endofunctor and two natural
transformations. The second defines it in terms of a function on objects and a
pair of arrows.
These two formulations are proven to be equivalent:
Monoidal.Monad Kleisli.Monad
The monoidal representation is exposed by default from this module.
---}
{-# OPTIONS --cubical --allow-unsolved-metas #-}
module Cat.Category.Monad where
open import Agda.Primitive
open import Data.Product
open import Cubical
open import Cubical.NType.Properties using (lemPropF ; lemSig ; lemSigP)
2018-03-06 09:16:42 +00:00
open import Cubical.GradLemma using (gradLemma)
2018-03-02 12:31:46 +00:00
open import Cat.Category
open import Cat.Category.Functor as F
open import Cat.Category.NaturalTransformation
import Cat.Category.Monad.Monoidal
import Cat.Category.Monad.Kleisli
open import Cat.Categories.Fun
module Monoidal = Cat.Category.Monad.Monoidal
module Kleisli = Cat.Category.Monad.Kleisli
2018-02-26 19:23:31 +00:00
-- | The monoidal- and kleisli presentation of monads are equivalent.
module _ {a b : Level} ( : Category a b) where
private
module = Category
open using (Object ; Arrow ; 𝟙 ; _∘_ ; _>>>_)
module M = Monoidal
2018-03-06 08:55:18 +00:00
module K = Kleisli
module _ (m : M.RawMonad) where
2018-03-06 08:52:01 +00:00
open M.RawMonad m
forthRaw : K.RawMonad
2018-03-06 09:16:42 +00:00
K.RawMonad.omap forthRaw = Romap
2018-03-06 08:52:01 +00:00
K.RawMonad.pure forthRaw = pureT _
K.RawMonad.bind forthRaw = bind
module _ {raw : M.RawMonad} (m : M.IsMonad raw) where
2018-03-01 13:19:46 +00:00
private
module MI = M.IsMonad m
forthIsMonad : K.IsMonad (forthRaw raw)
2018-03-06 08:55:18 +00:00
K.IsMonad.isIdentity forthIsMonad = proj₂ MI.isInverse
K.IsMonad.isNatural forthIsMonad = MI.isNatural
K.IsMonad.isDistributive forthIsMonad = MI.isDistributive
forth : M.Monad K.Monad
2018-03-06 09:16:42 +00:00
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
2018-03-06 09:16:42 +00:00
open K.Monad m
backRaw : M.RawMonad
2018-03-06 09:16:42 +00:00
M.RawMonad.R backRaw = R
M.RawMonad.pureNT backRaw = pureNT
M.RawMonad.joinNT backRaw = joinNT
2018-03-01 13:58:01 +00:00
private
2018-03-06 09:16:42 +00:00
open M.RawMonad backRaw
module R = Functor (M.RawMonad.R backRaw)
2018-03-01 13:58:01 +00:00
2018-03-01 13:19:46 +00:00
backIsMonad : M.IsMonad backRaw
2018-03-06 09:16:42 +00:00
M.IsMonad.isAssociative backIsMonad {X} = begin
joinT X R.fmap (joinT X) ≡⟨⟩
2018-03-06 09:16:42 +00:00
join fmap (joinT X) ≡⟨⟩
join fmap join ≡⟨ isNaturalForeign
join join ≡⟨⟩
joinT X joinT (R.omap X)
2018-03-06 09:16:42 +00:00
M.IsMonad.isInverse backIsMonad {X} = inv-l , inv-r
2018-03-01 13:58:01 +00:00
where
inv-l = begin
joinT X pureT (R.omap X) ≡⟨⟩
2018-03-06 09:16:42 +00:00
join pure ≡⟨ proj₁ isInverse
𝟙
2018-03-01 13:58:01 +00:00
inv-r = begin
joinT X R.fmap (pureT X) ≡⟨⟩
2018-03-06 09:16:42 +00:00
join fmap pure ≡⟨ proj₂ isInverse
𝟙
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
module _ (m : K.Monad) where
2018-03-06 14:52:22 +00:00
private
open K.Monad m
bindEq : {X Y}
K.RawMonad.bind (forthRaw (backRaw m)) {X} {Y}
K.RawMonad.bind (K.Monad.raw m)
bindEq {X} {Y} = begin
K.RawMonad.bind (forthRaw (backRaw m)) ≡⟨⟩
(λ f join fmap f) ≡⟨⟩
(λ f bind (f >>> pure) >>> bind 𝟙) ≡⟨ funExt lem
(λ f bind f) ≡⟨⟩
bind
where
lem : (f : Arrow X (omap Y))
bind (f >>> pure) >>> bind 𝟙
bind f
lem f = begin
bind (f >>> pure) >>> bind 𝟙
≡⟨ isDistributive _ _
bind ((f >>> pure) >>> bind 𝟙)
≡⟨ cong bind .isAssociative
bind (f >>> (pure >>> bind 𝟙))
≡⟨ cong (λ φ bind (f >>> φ)) (isNatural _)
bind (f >>> 𝟙)
≡⟨ cong bind .leftIdentity
2018-03-06 14:52:22 +00:00
bind f
2018-02-25 02:12:23 +00:00
forthRawEq : forthRaw (backRaw m) K.Monad.raw m
2018-03-06 09:05:35 +00:00
K.RawMonad.omap (forthRawEq _) = omap
2018-02-26 18:58:27 +00:00
K.RawMonad.pure (forthRawEq _) = pure
K.RawMonad.bind (forthRawEq i) = bindEq 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
2018-03-06 14:52:22 +00:00
private
open M.Monad m
module KM = K.Monad (forth m)
module R = Functor R
omapEq : KM.omap Romap
omapEq = refl
bindEq : {X Y} {f : Arrow X (Romap Y)} KM.bind f bind f
bindEq {X} {Y} {f} = begin
KM.bind f ≡⟨⟩
joinT Y Rfmap f ≡⟨⟩
bind f
joinEq : {X} KM.join joinT X
joinEq {X} = begin
KM.join ≡⟨⟩
KM.bind 𝟙 ≡⟨⟩
bind 𝟙 ≡⟨⟩
joinT X Rfmap 𝟙 ≡⟨ cong (λ φ _ φ) R.isIdentity
joinT X 𝟙 ≡⟨ .rightIdentity
2018-03-06 14:52:22 +00:00
joinT X
fmapEq : {A B} KM.fmap {A} {B} Rfmap
fmapEq {A} {B} = funExt (λ f begin
KM.fmap f ≡⟨⟩
KM.bind (f >>> KM.pure) ≡⟨⟩
bind (f >>> pureT _) ≡⟨⟩
Rfmap (f >>> pureT B) >>> joinT B ≡⟨⟩
Rfmap (f >>> pureT B) >>> joinT B ≡⟨ cong (λ φ φ >>> joinT B) R.isDistributive
Rfmap f >>> Rfmap (pureT B) >>> joinT B ≡⟨ .isAssociative
joinT B Rfmap (pureT B) Rfmap f ≡⟨ cong (λ φ φ Rfmap f) (proj₂ isInverse)
𝟙 Rfmap f ≡⟨ .leftIdentity
2018-03-06 14:52:22 +00:00
Rfmap f
)
rawEq : Functor.raw KM.R Functor.raw R
RawFunctor.omap (rawEq i) = omapEq i
RawFunctor.fmap (rawEq i) = fmapEq i
2018-03-05 16:31:13 +00:00
Req : M.RawMonad.R (backRaw (forth m)) R
2018-03-05 16:10:41 +00:00
Req = Functor≡ rawEq
open NaturalTransformation
2018-03-07 14:23:07 +00:00
pureTEq : M.RawMonad.pureT (backRaw (forth m)) pureT
pureTEq = funExt (λ X refl)
pureNTEq : (λ i NaturalTransformation F.identity (Req i))
[ M.RawMonad.pureNT (backRaw (forth m)) pureNT ]
pureNTEq = lemSigP (λ i propIsNatural F.identity (Req i)) _ _ pureTEq
2018-03-07 14:23:07 +00:00
joinTEq : M.RawMonad.joinT (backRaw (forth m)) joinT
joinTEq = funExt (λ X begin
M.RawMonad.joinT (backRaw (forth m)) X ≡⟨⟩
KM.join ≡⟨⟩
joinT X Rfmap 𝟙 ≡⟨ cong (λ φ joinT X φ) R.isIdentity
joinT X 𝟙 ≡⟨ .rightIdentity
2018-03-07 14:23:07 +00:00
joinT X )
joinNTEq : (λ i NaturalTransformation F[ Req i Req i ] (Req i))
[ M.RawMonad.joinNT (backRaw (forth m)) joinNT ]
joinNTEq = lemSigP (λ i propIsNatural F[ Req i Req i ] (Req i)) _ _ joinTEq
2018-03-07 14:23:07 +00:00
2018-02-25 02:12:23 +00:00
backRawEq : backRaw (forth m) M.Monad.raw m
2018-03-06 09:16:42 +00:00
M.RawMonad.R (backRawEq i) = Req i
2018-03-06 22:18:23 +00:00
M.RawMonad.pureNT (backRawEq i) = pureNTEq i
2018-03-06 14:55:03 +00:00
M.RawMonad.joinNT (backRawEq i) = joinNTEq 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
eqv : isEquiv M.Monad K.Monad forth
2018-02-24 18:07:58 +00:00
eqv = gradLemma forth back fortheq backeq
open import Cat.Equivalence
Monoidal≅Kleisli : M.Monad K.Monad
Monoidal≅Kleisli = forth , (back , (record { verso-recto = funExt backeq ; recto-verso = funExt fortheq }))
Monoidal≃Kleisli : M.Monad K.Monad
Monoidal≃Kleisli = forth , eqv