2018-02-24 14:13:25 +00:00
|
|
|
|
{-# OPTIONS --cubical --allow-unsolved-metas #-}
|
2018-02-23 16:33:09 +00:00
|
|
|
|
module Cat.Category.Monad where
|
|
|
|
|
|
2018-02-24 11:52:16 +00:00
|
|
|
|
open import Agda.Primitive
|
|
|
|
|
|
|
|
|
|
open import Data.Product
|
|
|
|
|
|
2018-02-23 16:33:09 +00:00
|
|
|
|
open import Cubical
|
2018-03-01 19:23:34 +00:00
|
|
|
|
open import Cubical.NType.Properties using (lemPropF)
|
2018-02-23 16:33:09 +00:00
|
|
|
|
|
2018-03-02 12:31:46 +00:00
|
|
|
|
open import Cat.Category
|
2018-02-24 11:52:16 +00:00
|
|
|
|
open import Cat.Category.Functor as F
|
|
|
|
|
open import Cat.Category.NaturalTransformation
|
2018-02-23 16:33:09 +00:00
|
|
|
|
open import Cat.Categories.Fun
|
2018-02-24 11:52:16 +00:00
|
|
|
|
|
2018-02-24 14:13:25 +00:00
|
|
|
|
-- "A monad in the monoidal form" [voe]
|
2018-02-24 11:52:16 +00:00
|
|
|
|
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 ; 𝟙 ; _∘_)
|
2018-02-24 11:52:16 +00:00
|
|
|
|
open NaturalTransformation ℂ ℂ
|
|
|
|
|
record RawMonad : Set ℓ where
|
|
|
|
|
field
|
2018-03-01 13:19:46 +00:00
|
|
|
|
-- TODO rename fields here
|
2018-02-24 19:37:21 +00:00
|
|
|
|
-- R ~ m
|
2018-02-28 18:03:11 +00:00
|
|
|
|
R : EndoFunctor ℂ
|
2018-02-24 19:37:21 +00:00
|
|
|
|
-- η ~ pure
|
2018-02-26 19:23:31 +00:00
|
|
|
|
ηNatTrans : NaturalTransformation F.identity R
|
2018-02-24 19:37:21 +00:00
|
|
|
|
-- μ ~ join
|
2018-02-26 19:23:31 +00:00
|
|
|
|
μNatTrans : NaturalTransformation F[ R ∘ R ] R
|
2018-02-24 11:52:16 +00:00
|
|
|
|
|
2018-02-24 14:13:25 +00:00
|
|
|
|
η : Transformation F.identity R
|
2018-02-26 19:23:31 +00:00
|
|
|
|
η = proj₁ ηNatTrans
|
2018-02-26 19:31:47 +00:00
|
|
|
|
ηNat : Natural F.identity R η
|
|
|
|
|
ηNat = proj₂ ηNatTrans
|
|
|
|
|
|
2018-02-24 14:13:25 +00:00
|
|
|
|
μ : Transformation F[ R ∘ R ] R
|
2018-02-26 19:23:31 +00:00
|
|
|
|
μ = proj₁ μNatTrans
|
2018-02-26 19:31:47 +00:00
|
|
|
|
μNat : Natural F[ R ∘ R ] R μ
|
|
|
|
|
μNat = proj₂ μNatTrans
|
2018-02-24 13:00:52 +00:00
|
|
|
|
|
2018-02-24 11:52:16 +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) ≡ 𝟙
|
2018-02-28 17:55:32 +00:00
|
|
|
|
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)
|
2018-02-24 11:52:16 +00:00
|
|
|
|
|
|
|
|
|
record IsMonad (raw : RawMonad) : Set ℓ where
|
|
|
|
|
open RawMonad raw public
|
|
|
|
|
field
|
|
|
|
|
isAssociative : IsAssociative
|
2018-02-26 19:36:39 +00:00
|
|
|
|
isInverse : IsInverse
|
2018-02-24 13:00:52 +00:00
|
|
|
|
|
2018-02-26 19:31:47 +00:00
|
|
|
|
private
|
|
|
|
|
module R = Functor R
|
|
|
|
|
module ℂ = Category ℂ
|
|
|
|
|
|
2018-02-28 17:55:32 +00:00
|
|
|
|
isNatural : IsNatural
|
|
|
|
|
isNatural {X} {Y} f = begin
|
2018-02-26 19:31:47 +00:00
|
|
|
|
μ 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 ∎
|
|
|
|
|
|
2018-02-28 17:55:32 +00:00
|
|
|
|
isDistributive : IsDistributive
|
|
|
|
|
isDistributive {X} {Y} {Z} g f = sym done
|
2018-02-26 19:36:39 +00:00
|
|
|
|
where
|
2018-02-28 17:55:32 +00:00
|
|
|
|
module R² = Functor F[ R ∘ R ]
|
2018-03-01 19:47:36 +00:00
|
|
|
|
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
|
2018-02-28 17:55:32 +00:00
|
|
|
|
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
|
2018-03-01 19:47:36 +00:00
|
|
|
|
μ 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)
|
|
|
|
|
∎
|
2018-02-26 19:36:39 +00:00
|
|
|
|
|
2018-02-24 13:01:57 +00:00
|
|
|
|
record Monad : Set ℓ where
|
|
|
|
|
field
|
|
|
|
|
raw : RawMonad
|
|
|
|
|
isMonad : IsMonad raw
|
|
|
|
|
open IsMonad isMonad public
|
|
|
|
|
|
2018-03-01 19:12:49 +00:00
|
|
|
|
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
|
2018-02-25 02:09:25 +00:00
|
|
|
|
|
2018-02-24 14:13:25 +00:00
|
|
|
|
-- "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
|
2018-03-01 13:58:01 +00:00
|
|
|
|
module ℂ = Category ℂ
|
|
|
|
|
open ℂ using (Arrow ; 𝟙 ; Object ; _∘_ ; _>>>_)
|
2018-02-26 19:08:48 +00:00
|
|
|
|
|
|
|
|
|
-- | Data for a monad.
|
|
|
|
|
--
|
|
|
|
|
-- Note that (>>=) is not expressible in a general category because objects
|
|
|
|
|
-- are not generally types.
|
2018-02-24 13:00:52 +00:00
|
|
|
|
record RawMonad : Set ℓ where
|
|
|
|
|
field
|
|
|
|
|
RR : Object → Object
|
2018-02-24 14:13:25 +00:00
|
|
|
|
-- 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 ]
|
2018-02-26 19:08:48 +00:00
|
|
|
|
|
|
|
|
|
-- | functor map
|
|
|
|
|
--
|
|
|
|
|
-- This should perhaps be defined in a "Klesli-version" of functors as well?
|
2018-02-25 02:09:25 +00:00
|
|
|
|
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 19:08:48 +00:00
|
|
|
|
|
|
|
|
|
-- | Composition of monads aka. the kleisli-arrow.
|
2018-02-24 18:08:20 +00:00
|
|
|
|
_>=>_ : {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)
|
2018-02-26 19:08:48 +00:00
|
|
|
|
|
|
|
|
|
-- | Flattening nested monads.
|
2018-02-26 18:57:05 +00:00
|
|
|
|
join : {A : Object} → ℂ [ RR (RR A) , RR A ]
|
|
|
|
|
join = bind 𝟙
|
2018-02-24 18:08:20 +00:00
|
|
|
|
|
2018-02-26 19:08:48 +00:00
|
|
|
|
------------------
|
|
|
|
|
-- * Monad laws --
|
|
|
|
|
------------------
|
|
|
|
|
|
|
|
|
|
-- There may be better names than what I've chosen here.
|
|
|
|
|
|
2018-02-24 13:00:52 +00:00
|
|
|
|
IsIdentity = {X : Object}
|
2018-02-26 18:57:05 +00:00
|
|
|
|
→ 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
|
|
|
|
→ pure >>> (bind f) ≡ f
|
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
|
|
|
|
→ (bind f) >>> (bind g) ≡ bind (f >=> g)
|
2018-02-26 19:08:48 +00:00
|
|
|
|
|
|
|
|
|
-- | 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?
|
2018-02-25 02:09:25 +00:00
|
|
|
|
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
|
|
|
|
|
2018-03-01 13:58:01 +00:00
|
|
|
|
-- 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 ≡ 𝟙
|
|
|
|
|
|
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
|
2018-02-26 19:23:31 +00:00
|
|
|
|
|
|
|
|
|
-- | Map fusion is admissable.
|
2018-02-25 02:09:25 +00:00
|
|
|
|
fusion : Fusion
|
|
|
|
|
fusion {g = g} {f} = begin
|
2018-02-26 19:23:31 +00:00
|
|
|
|
fmap (g ∘ f) ≡⟨⟩
|
2018-02-26 18:57:05 +00:00
|
|
|
|
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) ⟩
|
2018-02-26 19:23:31 +00:00
|
|
|
|
bind (fmap g ∘ (pure ∘ f)) ≡⟨ sym lem ⟩
|
2018-02-26 18:58:27 +00:00
|
|
|
|
bind (pure ∘ g) ∘ bind (pure ∘ f) ≡⟨⟩
|
2018-02-25 02:09:25 +00:00
|
|
|
|
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
|
|
|
|
|
2018-02-28 18:00:21 +00:00
|
|
|
|
-- | This formulation gives rise to the following endo-functor.
|
|
|
|
|
private
|
|
|
|
|
rawR : RawFunctor ℂ ℂ
|
2018-02-28 18:31:53 +00:00
|
|
|
|
RawFunctor.func* rawR = RR
|
|
|
|
|
RawFunctor.func→ rawR = fmap
|
2018-02-28 18:00:21 +00:00
|
|
|
|
|
|
|
|
|
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!
|
2018-02-28 18:03:11 +00:00
|
|
|
|
R : EndoFunctor ℂ
|
2018-02-28 18:00:21 +00:00
|
|
|
|
Functor.raw R = rawR
|
|
|
|
|
Functor.isFunctor R = isFunctorR
|
|
|
|
|
|
2018-02-28 18:31:53 +00:00
|
|
|
|
private
|
|
|
|
|
open NaturalTransformation ℂ ℂ
|
|
|
|
|
|
|
|
|
|
R⁰ : EndoFunctor ℂ
|
|
|
|
|
R⁰ = F.identity
|
|
|
|
|
R² : EndoFunctor ℂ
|
|
|
|
|
R² = F[ R ∘ R ]
|
|
|
|
|
module R = Functor R
|
|
|
|
|
module R⁰ = Functor R⁰
|
|
|
|
|
module R² = Functor R²
|
2018-03-05 09:28:16 +00:00
|
|
|
|
η : Transformation R⁰ R
|
|
|
|
|
η A = pure
|
|
|
|
|
ηNatural : Natural R⁰ R η
|
2018-02-28 18:31:53 +00:00
|
|
|
|
ηNatural {A} {B} f = begin
|
2018-03-05 09:28:16 +00:00
|
|
|
|
η B ∘ R⁰.func→ f ≡⟨⟩
|
2018-02-28 18:31:53 +00:00
|
|
|
|
pure ∘ f ≡⟨ sym (isNatural _) ⟩
|
|
|
|
|
bind (pure ∘ f) ∘ pure ≡⟨⟩
|
|
|
|
|
fmap f ∘ pure ≡⟨⟩
|
2018-03-05 09:28:16 +00:00
|
|
|
|
R.func→ f ∘ η A ∎
|
|
|
|
|
μ : Transformation R² R
|
|
|
|
|
μ C = join
|
|
|
|
|
μNatural : Natural R² R μ
|
2018-02-28 22:41:59 +00:00
|
|
|
|
μ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
|
2018-02-28 18:31:53 +00:00
|
|
|
|
|
|
|
|
|
ηNatTrans : NaturalTransformation R⁰ R
|
2018-03-05 09:28:16 +00:00
|
|
|
|
proj₁ ηNatTrans = η
|
2018-02-28 18:31:53 +00:00
|
|
|
|
proj₂ ηNatTrans = ηNatural
|
|
|
|
|
|
|
|
|
|
μNatTrans : NaturalTransformation R² R
|
2018-03-05 09:28:16 +00:00
|
|
|
|
proj₁ μNatTrans = μ
|
2018-02-28 18:31:53 +00:00
|
|
|
|
proj₂ μNatTrans = μNatural
|
|
|
|
|
|
2018-03-01 13:58:01 +00:00
|
|
|
|
isNaturalForeign : IsNaturalForeign
|
|
|
|
|
isNaturalForeign = begin
|
2018-03-01 17:00:51 +00:00
|
|
|
|
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 ∎
|
2018-03-01 13:58:01 +00:00
|
|
|
|
|
|
|
|
|
isInverse : IsInverse
|
|
|
|
|
isInverse = inv-l , inv-r
|
|
|
|
|
where
|
|
|
|
|
inv-l = begin
|
2018-03-01 16:50:06 +00:00
|
|
|
|
pure >>> join ≡⟨⟩
|
|
|
|
|
pure >>> bind 𝟙 ≡⟨ isNatural _ ⟩
|
2018-03-01 13:58:01 +00:00
|
|
|
|
𝟙 ∎
|
|
|
|
|
inv-r = begin
|
2018-03-01 16:50:06 +00:00
|
|
|
|
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 ⟩
|
2018-03-01 13:58:01 +00:00
|
|
|
|
𝟙 ∎
|
|
|
|
|
|
2018-02-24 13:00:52 +00:00
|
|
|
|
record Monad : Set ℓ where
|
|
|
|
|
field
|
|
|
|
|
raw : RawMonad
|
|
|
|
|
isMonad : IsMonad raw
|
|
|
|
|
open IsMonad isMonad public
|
2018-02-24 14:13:25 +00:00
|
|
|
|
|
2018-03-01 19:23:34 +00:00
|
|
|
|
module _ (raw : RawMonad) where
|
|
|
|
|
open RawMonad raw
|
2018-03-02 12:31:46 +00:00
|
|
|
|
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
|
|
|
|
|
|
2018-03-01 19:23:34 +00:00
|
|
|
|
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
|
2018-02-25 02:09:25 +00:00
|
|
|
|
|
2018-02-26 19:23:31 +00:00
|
|
|
|
-- | The monoidal- and kleisli presentation of monads are equivalent.
|
|
|
|
|
--
|
|
|
|
|
-- This is problem 2.3 in [voe].
|
2018-02-24 14:13:25 +00:00
|
|
|
|
module _ {ℓa ℓb : Level} {ℂ : Category ℓa ℓb} where
|
|
|
|
|
private
|
2018-03-05 09:28:16 +00:00
|
|
|
|
module ℂ = Category ℂ
|
|
|
|
|
open ℂ using (Object ; Arrow ; 𝟙 ; _∘_ ; _>>>_)
|
2018-02-24 14:13:25 +00:00
|
|
|
|
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!!
|
2018-02-24 14:13:25 +00:00
|
|
|
|
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-24 14:13:25 +00:00
|
|
|
|
|
2018-02-26 18:58:27 +00:00
|
|
|
|
pure : {X : Object} → ℂ [ X , RR X ]
|
|
|
|
|
pure {X} = η X
|
2018-02-24 14:13:25 +00:00
|
|
|
|
|
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
|
2018-02-24 14:13:25 +00:00
|
|
|
|
|
|
|
|
|
forthRaw : K.RawMonad
|
2018-02-26 19:23:31 +00:00
|
|
|
|
Kraw.RR forthRaw = RR
|
|
|
|
|
Kraw.pure forthRaw = pure
|
2018-02-26 18:57:05 +00:00
|
|
|
|
Kraw.bind forthRaw = bind
|
2018-02-24 14:13:25 +00:00
|
|
|
|
|
|
|
|
|
module _ {raw : M.RawMonad} (m : M.IsMonad raw) where
|
2018-03-01 13:19:46 +00:00
|
|
|
|
private
|
|
|
|
|
module MI = M.IsMonad m
|
|
|
|
|
module KI = K.IsMonad
|
2018-02-24 14:13:25 +00:00
|
|
|
|
forthIsMonad : K.IsMonad (forthRaw raw)
|
2018-02-28 17:55:32 +00:00
|
|
|
|
KI.isIdentity forthIsMonad = proj₂ MI.isInverse
|
|
|
|
|
KI.isNatural forthIsMonad = MI.isNatural
|
|
|
|
|
KI.isDistributive forthIsMonad = MI.isDistributive
|
2018-02-24 14:13:25 +00:00
|
|
|
|
|
|
|
|
|
forth : M.Monad → K.Monad
|
2018-02-25 02:09:25 +00:00
|
|
|
|
Kleisli.Monad.raw (forth m) = forthRaw (M.Monad.raw m)
|
2018-02-24 14:13:25 +00:00
|
|
|
|
Kleisli.Monad.isMonad (forth m) = forthIsMonad (M.Monad.isMonad m)
|
|
|
|
|
|
2018-02-25 02:09:25 +00:00
|
|
|
|
module _ (m : K.Monad) where
|
2018-03-01 13:58:01 +00:00
|
|
|
|
private
|
|
|
|
|
open K.Monad m
|
|
|
|
|
module MR = M.RawMonad
|
|
|
|
|
module MI = M.IsMonad
|
2018-02-25 02:09:25 +00:00
|
|
|
|
|
|
|
|
|
backRaw : M.RawMonad
|
2018-02-26 19:23:31 +00:00
|
|
|
|
MR.R backRaw = R
|
|
|
|
|
MR.ηNatTrans backRaw = ηNatTrans
|
|
|
|
|
MR.μNatTrans backRaw = μNatTrans
|
2018-02-25 02:09:25 +00:00
|
|
|
|
|
2018-03-01 13:58:01 +00:00
|
|
|
|
private
|
|
|
|
|
open MR backRaw
|
|
|
|
|
module R = Functor (MR.R backRaw)
|
|
|
|
|
|
2018-03-01 13:19:46 +00:00
|
|
|
|
backIsMonad : M.IsMonad backRaw
|
2018-03-01 13:58:01 +00:00
|
|
|
|
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 ⟩
|
|
|
|
|
𝟙 ∎
|
2018-02-25 02:09:25 +00:00
|
|
|
|
|
2018-02-24 18:07:58 +00:00
|
|
|
|
back : K.Monad → M.Monad
|
2018-02-25 02:09:25 +00:00
|
|
|
|
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
|
2018-03-05 09:28:16 +00:00
|
|
|
|
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
|
|
|
|
|
|
2018-02-25 02:12:23 +00:00
|
|
|
|
forthRawEq : forthRaw (backRaw m) ≡ K.Monad.raw m
|
2018-02-26 19:23:31 +00:00
|
|
|
|
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-03-05 09:28:16 +00:00
|
|
|
|
K.RawMonad.bind (forthRawEq i) = bindEq i
|
2018-02-24 18:07:58 +00:00
|
|
|
|
|
|
|
|
|
fortheq : (m : K.Monad) → forth (back m) ≡ m
|
2018-02-25 02:09:25 +00:00
|
|
|
|
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)
|
2018-03-05 09:28:16 +00:00
|
|
|
|
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 ℂ ℂ
|
2018-02-25 02:12:23 +00:00
|
|
|
|
backRawEq : backRaw (forth m) ≡ M.Monad.raw m
|
|
|
|
|
-- stuck
|
2018-03-05 09:28:16 +00:00
|
|
|
|
M.RawMonad.R (backRawEq i) = Req i
|
|
|
|
|
M.RawMonad.ηNatTrans (backRawEq i) = let t = NaturalTransformation≡ F.identity R ηeq in {!t i!}
|
2018-02-26 19:23:31 +00:00
|
|
|
|
M.RawMonad.μNatTrans (backRawEq i) = {!!}
|
2018-02-24 18:07:58 +00:00
|
|
|
|
|
|
|
|
|
backeq : (m : M.Monad) → back (forth m) ≡ m
|
2018-02-25 02:09:25 +00:00
|
|
|
|
backeq m = M.Monad≡ (backRawEq m)
|
2018-02-24 18:07:58 +00:00
|
|
|
|
|
|
|
|
|
open import Cubical.GradLemma
|
2018-02-24 14:13:25 +00:00
|
|
|
|
eqv : isEquiv M.Monad K.Monad forth
|
2018-02-24 18:07:58 +00:00
|
|
|
|
eqv = gradLemma forth back fortheq backeq
|
2018-02-24 14:13:25 +00:00
|
|
|
|
|
|
|
|
|
Monoidal≃Kleisli : M.Monad ≃ K.Monad
|
|
|
|
|
Monoidal≃Kleisli = forth , eqv
|