cat/src/Cat/Category/Monad/Kleisli.agda
2018-03-21 11:58:50 +01:00

254 lines
8.7 KiB
Agda
Raw 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.

{---
The Kleisli formulation of monads
---}
{-# OPTIONS --cubical --allow-unsolved-metas #-}
open import Agda.Primitive
open import Data.Product
open import Cubical
open import Cubical.NType.Properties using (lemPropF ; lemSig ; lemSigP)
open import Cubical.GradLemma using (gradLemma)
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 Kleisli form" [voe]
module Cat.Category.Monad.Kleisli {a b : Level} ( : Category a b) where
private
= a b
module = Category
open using (Arrow ; 𝟙 ; Object ; _∘_ ; _>>>_)
-- | Data for a monad.
--
-- Note that (>>=) is not expressible in a general category because objects
-- are not generally types.
record RawMonad : Set where
field
omap : Object Object
pure : {X : Object} [ X , omap X ]
bind : {X Y : Object} [ X , omap Y ] [ omap X , omap Y ]
-- | functor map
--
-- This should perhaps be defined in a "Klesli-version" of functors as well?
fmap : {A B} [ A , B ] [ omap A , omap B ]
fmap f = bind (pure f)
-- | Composition of monads aka. the kleisli-arrow.
_>=>_ : {A B C : Object} [ A , omap B ] [ B , omap C ] [ A , omap C ]
f >=> g = f >>> (bind g)
-- | Flattening nested monads.
join : {A : Object} [ omap (omap A) , omap A ]
join = bind 𝟙
------------------
-- * Monad laws --
------------------
-- There may be better names than what I've chosen here.
IsIdentity = {X : Object}
bind pure 𝟙 {omap X}
IsNatural = {X Y : Object} (f : [ X , omap Y ])
pure >>> (bind f) f
IsDistributive = {X Y Z : Object} (g : [ Y , omap Z ]) (f : [ X , omap Y ])
(bind f) >>> (bind g) bind (f >=> g)
-- | Functor map fusion.
--
-- This is really a functor law. Should we have a kleisli-representation of
-- functors as well and make them a super-class?
Fusion = {X Y Z : Object} {g : [ Y , Z ]} {f : [ X , Y ]}
fmap (g f) fmap g fmap f
-- In the ("foreign") formulation of a monad `IsNatural`'s analogue here would be:
IsNaturalForeign : Set _
IsNaturalForeign = {X : Object} join {X} fmap join join join
IsInverse : Set _
IsInverse = {X : Object} join {X} pure 𝟙 × join {X} fmap pure 𝟙
record IsMonad (raw : RawMonad) : Set where
open RawMonad raw public
field
isIdentity : IsIdentity
isNatural : IsNatural
isDistributive : IsDistributive
-- | Map fusion is admissable.
fusion : Fusion
fusion {g = g} {f} = begin
fmap (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 distrib
bind (pure g) bind (pure f) ≡⟨⟩
fmap g fmap f
where
distrib : fmap g fmap f bind (fmap g (pure f))
distrib = isDistributive (pure g) (pure f)
-- | This formulation gives rise to the following endo-functor.
private
rawR : RawFunctor
RawFunctor.omap rawR = omap
RawFunctor.fmap rawR = fmap
isFunctorR : IsFunctor rawR
IsFunctor.isIdentity isFunctorR = begin
bind (pure 𝟙) ≡⟨ cong bind (.rightIdentity)
bind pure ≡⟨ isIdentity
𝟙
IsFunctor.isDistributive isFunctorR {f = f} {g} = begin
bind (pure (g f)) ≡⟨⟩
fmap (g f) ≡⟨ fusion
fmap g fmap f ≡⟨⟩
bind (pure g) bind (pure f)
-- FIXME Naming!
R : EndoFunctor
Functor.raw R = rawR
Functor.isFunctor R = isFunctorR
private
open NaturalTransformation
R⁰ : EndoFunctor
R⁰ = F.identity
: EndoFunctor
= F[ R R ]
module R = Functor R
module R = Functor R⁰
module R² = Functor
pureT : Transformation R⁰ R
pureT A = pure
pureN : Natural R⁰ R pureT
pureN {A} {B} f = begin
pureT B R⁰.fmap f ≡⟨⟩
pure f ≡⟨ sym (isNatural _)
bind (pure f) pure ≡⟨⟩
fmap f pure ≡⟨⟩
R.fmap f pureT A
joinT : Transformation R
joinT C = join
joinN : Natural R joinT
joinN f = begin
join R².fmap f ≡⟨⟩
bind 𝟙 R².fmap f ≡⟨⟩
R².fmap f >>> bind 𝟙 ≡⟨⟩
fmap (fmap f) >>> bind 𝟙 ≡⟨⟩
fmap (bind (f >>> pure)) >>> bind 𝟙 ≡⟨⟩
bind (bind (f >>> pure) >>> pure) >>> bind 𝟙
≡⟨ isDistributive _ _
bind ((bind (f >>> pure) >>> pure) >=> 𝟙)
≡⟨⟩
bind ((bind (f >>> pure) >>> pure) >>> bind 𝟙)
≡⟨ cong bind .isAssociative
bind (bind (f >>> pure) >>> (pure >>> bind 𝟙))
≡⟨ cong (λ φ bind (bind (f >>> pure) >>> φ)) (isNatural _)
bind (bind (f >>> pure) >>> 𝟙)
≡⟨ cong bind .leftIdentity
bind (bind (f >>> pure))
≡⟨ cong bind (sym .rightIdentity)
bind (𝟙 >>> bind (f >>> pure)) ≡⟨⟩
bind (𝟙 >=> (f >>> pure))
≡⟨ sym (isDistributive _ _)
bind 𝟙 >>> bind (f >>> pure) ≡⟨⟩
bind 𝟙 >>> fmap f ≡⟨⟩
bind 𝟙 >>> R.fmap f ≡⟨⟩
R.fmap f bind 𝟙 ≡⟨⟩
R.fmap f join
pureNT : NaturalTransformation R⁰ R
proj₁ pureNT = pureT
proj₂ pureNT = pureN
joinNT : NaturalTransformation R
proj₁ joinNT = joinT
proj₂ joinNT = joinN
isNaturalForeign : IsNaturalForeign
isNaturalForeign = begin
fmap join >>> join ≡⟨⟩
bind (join >>> pure) >>> bind 𝟙
≡⟨ isDistributive _ _
bind ((join >>> pure) >>> bind 𝟙)
≡⟨ cong bind .isAssociative
bind (join >>> (pure >>> bind 𝟙))
≡⟨ cong (λ φ bind (join >>> φ)) (isNatural _)
bind (join >>> 𝟙)
≡⟨ cong bind .leftIdentity
bind join ≡⟨⟩
bind (bind 𝟙)
≡⟨ cong bind (sym .rightIdentity)
bind (𝟙 >>> bind 𝟙) ≡⟨⟩
bind (𝟙 >=> 𝟙) ≡⟨ sym (isDistributive _ _)
bind 𝟙 >>> bind 𝟙 ≡⟨⟩
join >>> join
isInverse : IsInverse
isInverse = inv-l , inv-r
where
inv-l = begin
pure >>> join ≡⟨⟩
pure >>> bind 𝟙 ≡⟨ isNatural _
𝟙
inv-r = begin
fmap pure >>> join ≡⟨⟩
bind (pure >>> pure) >>> bind 𝟙
≡⟨ isDistributive _ _
bind ((pure >>> pure) >=> 𝟙) ≡⟨⟩
bind ((pure >>> pure) >>> bind 𝟙)
≡⟨ cong bind .isAssociative
bind (pure >>> (pure >>> bind 𝟙))
≡⟨ cong (λ φ bind (pure >>> φ)) (isNatural _)
bind (pure >>> 𝟙)
≡⟨ cong bind .leftIdentity
bind pure ≡⟨ isIdentity
𝟙
record Monad : Set where
field
raw : RawMonad
isMonad : IsMonad raw
open IsMonad isMonad public
private
module _ (raw : RawMonad) where
open RawMonad raw
propIsIdentity : isProp IsIdentity
propIsIdentity x y i = .arrowsAreSets _ _ x y i
propIsNatural : isProp IsNatural
propIsNatural x y i = λ f
.arrowsAreSets _ _ (x f) (y f) i
propIsDistributive : isProp IsDistributive
propIsDistributive x y i = λ g f
.arrowsAreSets _ _ (x g f) (y g f) i
open IsMonad
propIsMonad : (raw : _) isProp (IsMonad raw)
IsMonad.isIdentity (propIsMonad raw x y i)
= propIsIdentity raw (isIdentity x) (isIdentity y) i
IsMonad.isNatural (propIsMonad raw x y i)
= propIsNatural raw (isNatural x) (isNatural y) i
IsMonad.isDistributive (propIsMonad raw x y i)
= propIsDistributive raw (isDistributive x) (isDistributive y) i
module _ {m n : Monad} (eq : Monad.raw m Monad.raw n) where
private
eqIsMonad : (λ i IsMonad (eq i)) [ Monad.isMonad m Monad.isMonad n ]
eqIsMonad = lemPropF propIsMonad eq
Monad≡ : m n
Monad.raw (Monad≡ i) = eq i
Monad.isMonad (Monad≡ i) = eqIsMonad i