Rename functor composition - implement monads...

In their monoidal form.
This commit is contained in:
Frederik Hanghøj Iversen 2018-02-24 12:52:16 +01:00
parent cb8533b84a
commit 8527fe0df4
2 changed files with 67 additions and 3 deletions

View file

@ -124,8 +124,9 @@ module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : F
; isDistributive = dist
}
_∘f_ : Functor A C
raw _∘f_ = _∘fr_
F[_∘_] _∘f_ : Functor A C
raw F[_∘_] = _∘fr_
_∘f_ = F[_∘_]
-- The identity functor
identity : { '} {C : Category '} Functor C C

View file

@ -1,8 +1,71 @@
{-# OPTIONS --cubical #-}
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
open import Cat.Category.Functor as F
open import Cat.Category.NaturalTransformation
open import Cat.Categories.Fun
-- "A monad in the monoidal form" [vlad]
module Monoidal {a b : Level} ( : Category a b) where
private
= a b
open Category hiding (IsAssociative)
open NaturalTransformation
record RawMonad : Set where
field
R : Functor
-- pure
ηNat : NaturalTransformation F.identity R
-- (>=>)
μNat : NaturalTransformation F[ R R ] R
module R = Functor R
module RR = Functor F[ R R ]
private
module _ {X : Object} where
-- module IdRX = Functor (F.identity {C = RX})
η : Transformation F.identity R
η = proj₁ ηNat
ηX : [ X , R.func* X ]
ηX = η X
RηX : [ R.func* X , R.func* (R.func* X) ] -- [ R.func* X , {!R.func* (R.func* X))!} ]
RηX = R.func→ ηX
ηRX = η (R.func* X)
IdRX : Arrow (R.func* X) (R.func* X)
IdRX = 𝟙 {R.func* X}
μ : Transformation F[ R R ] R
μ = proj₁ μNat
μX : [ RR.func* X , R.func* X ]
μX = μ X
RμX : [ R.func* (RR.func* X) , RR.func* X ]
RμX = R.func→ μX
μRX : [ RR.func* (R.func* X) , R.func* (R.func* X) ]
μRX = μ (R.func* X)
IsAssociative' : Set _
IsAssociative' = [ μX RμX ] [ μX μRX ]
IsInverse' : Set _
IsInverse'
= [ μX ηRX ] IdRX
× [ μX RηX ] IdRX
-- We don't want the objects to be indexes of the type, but rather just
-- universally quantify over *all* objects of the category.
IsAssociative = {X : Object} IsAssociative' {X}
IsInverse = {X : Object} IsInverse' {X}
record IsMonad (raw : RawMonad) : Set where
open RawMonad raw public
field
isAssociative : IsAssociative
isInverse : IsInverse