cat/src/Cat/Category/Monad.agda

570 lines
21 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.

{-# 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)
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
open Category using (Object ; Arrow ; 𝟙 ; _∘_)
open NaturalTransformation
record RawMonad : Set where
field
-- TODO rename fields here
-- R ~ m
R : EndoFunctor
-- η ~ pure
ηNatTrans : NaturalTransformation F.identity R
-- μ ~ join
μNatTrans : NaturalTransformation F[ R R ] R
η : Transformation F.identity R
η = proj₁ ηNatTrans
ηNat : Natural F.identity R η
ηNat = proj₂ ηNatTrans
μ : Transformation F[ R R ] R
μ = proj₁ μNatTrans
μNat : Natural F[ R R ] R μ
μNat = proj₂ μNatTrans
private
module R = Functor R
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) 𝟙
IsNatural = {X Y} f μ Y R.func→ f η X f
IsDistributive = {X Y Z} (g : Arrow Y (R.func* Z)) (f : Arrow X (R.func* Y))
μ Z R.func→ g (μ Y R.func→ f)
μ Z R.func→ (μ Z R.func→ g f)
record IsMonad (raw : RawMonad) : Set where
open RawMonad raw public
field
isAssociative : IsAssociative
isInverse : IsInverse
private
module R = Functor R
module = Category
isNatural : IsNatural
isNatural {X} {Y} f = begin
μ Y R.func→ f η X ≡⟨ sym .isAssociative
μ Y (R.func→ f η X) ≡⟨ cong (λ φ μ Y φ) (sym (ηNat f))
μ Y (η (R.func* Y) f) ≡⟨ .isAssociative
μ Y η (R.func* Y) f ≡⟨ cong (λ φ φ f) (proj₁ isInverse)
𝟙 f ≡⟨ proj₂ .isIdentity
f
isDistributive : IsDistributive
isDistributive {X} {Y} {Z} g f = sym done
where
module R² = 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 {a = a} {b} {c} = begin
R.func→ (a b c) ≡⟨ distr
R.func→ (a b) R.func→ c ≡⟨ cong (_∘ _) distr
R.func→ a R.func→ b R.func→ c
where
distr = R.isDistributive
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 {a = a} {b} {c} {d} = begin
a (b c d) ≡⟨⟩
a ((b c) d) ≡⟨ cong (_∘_ a) (sym asc)
a (b (c d)) ≡⟨ asc
(a b) (c d) ≡⟨ asc
((a b) c) d ≡⟨⟩
a b c d
where
asc = .isAssociative
lemmm : μ Z R.func→ (μ Z) μ Z μ (R.func* Z)
lemmm = isAssociative
lem4 : μ (R.func* Z) R².func→ g R.func→ g μ Y
lem4 = μNat g
done = 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) R².func→ g R.func→ f)
≡⟨ cong (_∘_ (μ Z)) (sym .isAssociative) -- ●-solver?
μ Z (R.func→ (μ Z) (R².func→ g R.func→ f))
≡⟨ .isAssociative
(μ Z R.func→ (μ Z)) (R².func→ g R.func→ f)
≡⟨ cong (λ φ φ (R².func→ g R.func→ f)) isAssociative
(μ Z μ (R.func* Z)) (R².func→ g R.func→ f)
≡⟨ .isAssociative -- ●-solver?
μ Z μ (R.func* Z) R².func→ g R.func→ f
≡⟨⟩ -- ●-solver + lem4
((μ Z μ (R.func* Z)) R².func→ g) R.func→ f
≡⟨ cong (_∘ R.func→ f) (sym .isAssociative)
(μ Z (μ (R.func* Z) R².func→ g)) R.func→ f
≡⟨ cong (λ φ φ R.func→ f) (cong (_∘_ (μ Z)) lem4)
(μ Z (R.func→ g μ Y)) R.func→ f ≡⟨ cong (_∘ R.func→ f) .isAssociative
μ Z R.func→ g μ Y R.func→ f
≡⟨ sym (Category.isAssociative )
μ Z R.func→ g (μ Y R.func→ f)
record Monad : Set where
field
raw : RawMonad
isMonad : IsMonad raw
open IsMonad isMonad public
private
module _ {m : RawMonad} where
open RawMonad m
propIsAssociative : isProp IsAssociative
propIsAssociative x y i {X}
= Category.arrowsAreSets _ _ (x {X}) (y {X}) i
propIsInverse : isProp IsInverse
propIsInverse x y i {X} = e1 i , e2 i
where
xX = x {X}
yX = y {X}
e1 = Category.arrowsAreSets _ _ (proj₁ xX) (proj₁ yX)
e2 = Category.arrowsAreSets _ _ (proj₂ xX) (proj₂ yX)
open IsMonad
propIsMonad : (raw : _) isProp (IsMonad raw)
IsMonad.isAssociative (propIsMonad raw a b i) j
= propIsAssociative {raw}
(isAssociative a) (isAssociative b) i j
IsMonad.isInverse (propIsMonad raw a b i)
= propIsInverse {raw}
(isInverse a) (isInverse b) i
module _ {m n : Monad} (eq : Monad.raw m Monad.raw n) where
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
-- "A monad in the Kleisli form" [voe]
module 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
RR : Object Object
-- Note name-change from [voe]
pure : {X : Object} [ X , RR X ]
bind : {X Y : Object} [ X , RR Y ] [ RR X , RR Y ]
-- | functor map
--
-- This should perhaps be defined in a "Klesli-version" of functors as well?
fmap : {A B} [ A , B ] [ RR A , RR B ]
fmap f = bind (pure f)
-- | Composition of monads aka. the kleisli-arrow.
_>=>_ : {A B C : Object} [ A , RR B ] [ B , RR C ] [ A , RR C ]
f >=> g = f >>> (bind g)
-- | Flattening nested monads.
join : {A : Object} [ RR (RR A) , RR A ]
join = bind 𝟙
------------------
-- * Monad laws --
------------------
-- There may be better names than what I've chosen here.
IsIdentity = {X : Object}
bind pure 𝟙 {RR X}
IsNatural = {X Y : Object} (f : [ X , RR Y ])
pure >>> (bind f) f
IsDistributive = {X Y Z : Object} (g : [ Y , RR Z ]) (f : [ X , RR 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 lem
bind (pure g) bind (pure f) ≡⟨⟩
fmap g fmap f
where
open Category using (isAssociative)
lem : fmap g fmap f bind (fmap g (pure f))
lem = isDistributive (pure g) (pure f)
-- | This formulation gives rise to the following endo-functor.
private
rawR : RawFunctor
RawFunctor.func* rawR = RR
RawFunctor.func→ rawR = fmap
isFunctorR : IsFunctor rawR
IsFunctor.isIdentity isFunctorR = begin
bind (pure 𝟙) ≡⟨ cong bind (proj₁ .isIdentity)
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)
-- TODO: 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
η : Transformation R⁰ R
η A = pure
ηNatural : Natural R⁰ R η
ηNatural {A} {B} f = begin
η B R⁰.func→ f ≡⟨⟩
pure f ≡⟨ sym (isNatural _)
bind (pure f) pure ≡⟨⟩
fmap f pure ≡⟨⟩
R.func→ f η A
μ : Transformation R
μ C = join
μNatural : Natural R μ
μNatural f = begin
join R².func→ f ≡⟨⟩
bind 𝟙 R².func→ f ≡⟨⟩
R².func→ 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 (proj₂ .isIdentity)
bind (bind (f >>> pure))
≡⟨ cong bind (sym (proj₁ .isIdentity))
bind (𝟙 >>> bind (f >>> pure)) ≡⟨⟩
bind (𝟙 >=> (f >>> pure))
≡⟨ sym (isDistributive _ _)
bind 𝟙 >>> bind (f >>> pure) ≡⟨⟩
bind 𝟙 >>> fmap f ≡⟨⟩
bind 𝟙 >>> R.func→ f ≡⟨⟩
R.func→ f bind 𝟙 ≡⟨⟩
R.func→ f join
where
ηNatTrans : NaturalTransformation R⁰ R
proj₁ ηNatTrans = η
proj₂ ηNatTrans = ηNatural
μNatTrans : NaturalTransformation R
proj₁ μNatTrans = μ
proj₂ μNatTrans = μNatural
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 (proj₂ .isIdentity)
bind join ≡⟨⟩
bind (bind 𝟙)
≡⟨ cong bind (sym (proj₁ .isIdentity))
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 (proj₂ .isIdentity)
bind pure ≡⟨ isIdentity
𝟙
record Monad : Set where
field
raw : RawMonad
isMonad : IsMonad raw
open IsMonad isMonad public
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
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
-- | The monoidal- and kleisli presentation of monads are equivalent.
--
-- This is problem 2.3 in [voe].
module _ {a b : Level} { : Category a b} where
private
module = Category
open using (Object ; Arrow ; 𝟙 ; _∘_ ; _>>>_)
open Functor using (func* ; func→)
module M = Monoidal
module K = Kleisli
-- Note similarity with locally defined things in Kleisly.RawMonad!!
module _ (m : M.RawMonad) where
private
open M.RawMonad m
module Kraw = K.RawMonad
RR : Object Object
RR = func* R
pure : {X : Object} [ X , RR X ]
pure {X} = η X
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
Kraw.pure forthRaw = pure
Kraw.bind forthRaw = bind
module _ {raw : M.RawMonad} (m : M.IsMonad raw) where
private
module MI = M.IsMonad m
module KI = K.IsMonad
forthIsMonad : K.IsMonad (forthRaw raw)
KI.isIdentity forthIsMonad = proj₂ MI.isInverse
KI.isNatural forthIsMonad = MI.isNatural
KI.isDistributive forthIsMonad = MI.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
open K.Monad m
module MR = M.RawMonad
module MI = M.IsMonad
backRaw : M.RawMonad
MR.R backRaw = R
MR.ηNatTrans backRaw = ηNatTrans
MR.μNatTrans backRaw = μNatTrans
private
open MR backRaw
module R = Functor (MR.R backRaw)
backIsMonad : M.IsMonad backRaw
MI.isAssociative backIsMonad {X} = begin
μ X R.func→ (μ X) ≡⟨⟩
join fmap (μ X) ≡⟨⟩
join fmap join ≡⟨ isNaturalForeign
join join ≡⟨⟩
μ X μ (R.func* X)
MI.isInverse backIsMonad {X} = inv-l , inv-r
where
inv-l = begin
μ X η (R.func* X) ≡⟨⟩
join pure ≡⟨ proj₁ isInverse
𝟙
inv-r = begin
μ X R.func→ (η X) ≡⟨⟩
join fmap pure ≡⟨ proj₂ isInverse
𝟙
back : K.Monad M.Monad
Monoidal.Monad.raw (back m) = backRaw m
Monoidal.Monad.isMonad (back m) = backIsMonad m
-- I believe all the proofs here should be `refl`.
module _ (m : K.Monad) where
open K.Monad m
-- open K.RawMonad (K.Monad.raw 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 μ Y func→ R f) ≡⟨⟩
(λ f join fmap f) ≡⟨⟩
(λ f bind (f >>> pure) >>> bind 𝟙) ≡⟨ funExt lem
(λ f bind f) ≡⟨⟩
bind
where
μ = proj₁ μNatTrans
lem : (f : Arrow X (RR 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 (proj₂ .isIdentity)
bind f
_&_ : {a b} {A : Set a} {B : Set b} A (A B) B
x & f = f x
forthRawEq : forthRaw (backRaw m) K.Monad.raw m
K.RawMonad.RR (forthRawEq _) = RR
K.RawMonad.pure (forthRawEq _) = pure
-- stuck
K.RawMonad.bind (forthRawEq i) = bindEq i
fortheq : (m : K.Monad) forth (back m) m
fortheq m = K.Monad≡ (forthRawEq m)
module _ (m : M.Monad) where
open M.RawMonad (M.Monad.raw m)
rawEq* : Functor.func* (K.Monad.R (forth m)) Functor.func* R
rawEq* = refl
left = Functor.raw (K.Monad.R (forth m))
right = Functor.raw R
P : (omap : Omap )
(eq : RawFunctor.func* left omap)
(fmap' : Fmap omap)
Set _
P _ eq fmap' = (λ i Fmap (eq i))
[ RawFunctor.func→ left fmap' ]
-- rawEq→ : (λ i → Fmap (refl i)) [ Functor.func→ (K.Monad.R (forth m)) ≡ Functor.func→ R ]
rawEq→ : P (RawFunctor.func* right) refl (RawFunctor.func→ right)
-- rawEq→ : (fmap' : Fmap {!!}) → RawFunctor.func→ left ≡ fmap'
rawEq→ = begin
(λ {A} {B} RawFunctor.func→ left) ≡⟨ {!!}
(λ {A} {B} RawFunctor.func→ right)
-- destfmap =
source = (Functor.raw (K.Monad.R (forth m)))
-- p : (fmap' : Fmap (RawFunctor.func* source)) → (λ i → Fmap (refl i)) [ func→ source ≡ fmap' ]
-- p = {!!}
rawEq : Functor.raw (K.Monad.R (forth m)) Functor.raw R
rawEq = RawFunctor≡ {x = left} {right} refl λ fmap' {!rawEq→!}
Req : M.RawMonad.R (backRaw (forth m)) R
Req = FunctorEq rawEq
ηeq : M.RawMonad.η (backRaw (forth m)) η
ηeq = {!!}
postulate ηNatTransEq : {!!} [ M.RawMonad.ηNatTrans (backRaw (forth m)) ηNatTrans ]
open NaturalTransformation
backRawEq : backRaw (forth m) M.Monad.raw m
-- stuck
M.RawMonad.R (backRawEq i) = Req i
M.RawMonad.ηNatTrans (backRawEq i) = let t = NaturalTransformation≡ F.identity R ηeq in {!t i!}
M.RawMonad.μNatTrans (backRawEq i) = {!!}
backeq : (m : M.Monad) back (forth m) m
backeq m = M.Monad≡ (backRawEq m)
open import Cubical.GradLemma
eqv : isEquiv M.Monad K.Monad forth
eqv = gradLemma forth back fortheq backeq
Monoidal≃Kleisli : M.Monad K.Monad
Monoidal≃Kleisli = forth , eqv