More stuff about kleisli \equiv monoidal

This commit is contained in:
Frederik Hanghøj Iversen 2018-03-05 10:28:16 +01:00
parent b079f5e426
commit 8f8800cb67
7 changed files with 166 additions and 73 deletions

View file

@ -1,19 +1,19 @@
module Cat where module Cat where
import Cat.Category open import Cat.Category
import Cat.Category.Functor open import Cat.Category.Functor
import Cat.Category.Product open import Cat.Category.Product
import Cat.Category.Exponential open import Cat.Category.Exponential
import Cat.Category.CartesianClosed open import Cat.Category.CartesianClosed
import Cat.Category.NaturalTransformation open import Cat.Category.NaturalTransformation
import Cat.Category.Yoneda open import Cat.Category.Yoneda
import Cat.Category.Monad open import Cat.Category.Monad
import Cat.Categories.Sets open import Cat.Categories.Sets
import Cat.Categories.Cat open import Cat.Categories.Cat
import Cat.Categories.Rel open import Cat.Categories.Rel
import Cat.Categories.Free open import Cat.Categories.Free
import Cat.Categories.Fun open import Cat.Categories.Fun
import Cat.Categories.Cube open import Cat.Categories.Cube
import Cat.Categories.CwF open import Cat.Categories.CwF

View file

@ -26,6 +26,7 @@ open Category hiding (_∘_)
open Functor open Functor
module _ { ' : Level} (Ns : Set ) where module _ { ' : Level} (Ns : Set ) where
private
-- Ns is the "namespace" -- Ns is the "namespace"
o = (suc zero ) o = (suc zero )
@ -43,7 +44,6 @@ module _ { ' : Level} (Ns : Set ) where
𝟚 = Bool 𝟚 = Bool
module _ (I J : FiniteDecidableSubset) where module _ (I J : FiniteDecidableSubset) where
private
Hom' : Set Hom' : Set
Hom' = elmsof I elmsof J 𝟚 Hom' = elmsof I elmsof J 𝟚
isInl : {a b : Level} {A : Set a} {B : Set b} A B Set isInl : {a b : Level} {A : Set a} {B : Set b} A B Set

View file

@ -20,10 +20,10 @@ singleton : ∀ {} {𝓤 : Set } {r} {R : 𝓤𝓤 → Set r} {
singleton f = cons f empty singleton f = cons f empty
module _ { ' : Level} ( : Category ') where module _ { ' : Level} ( : Category ') where
private
module = Category module = Category
open Category open Category
private
p-isAssociative : {A B C D : Object} {r : Path Arrow A B} {q : Path Arrow B C} {p : Path Arrow C D} p-isAssociative : {A B C D : Object} {r : Path Arrow A B} {q : Path Arrow B C} {p : Path Arrow C D}
p ++ (q ++ r) (p ++ q) ++ r p ++ (q ++ r) (p ++ q) ++ r
p-isAssociative {r = r} {q} {empty} = refl p-isAssociative {r = r} {q} {empty} = refl

View file

@ -18,6 +18,10 @@ module _ {c c' d d'}
= c c' d d' = c c' d d'
𝓤 = Set 𝓤 = Set
Omap = Object Object 𝔻
Fmap : Omap Set _
Fmap omap = {A B}
[ A , B ] 𝔻 [ omap A , omap B ]
record RawFunctor : 𝓤 where record RawFunctor : 𝓤 where
field field
func* : Object Object 𝔻 func* : Object Object 𝔻
@ -30,6 +34,30 @@ module _ {c c' d d'}
IsDistributive = {A B C : Object } {f : [ A , B ]} {g : [ B , C ]} IsDistributive = {A B C : Object } {f : [ A , B ]} {g : [ B , C ]}
func→ ( [ g f ]) 𝔻 [ func→ g func→ f ] func→ ( [ g f ]) 𝔻 [ func→ g func→ f ]
-- | Equality principle for raw functors
--
-- The type of `func→` depend on the value of `func*`. We can wrap this up
-- into an equality principle for this type like is done for e.g. `Σ` using
-- `pathJ`.
module _ {x y : RawFunctor} where
open RawFunctor
private
P : (omap : Omap) (eq : func* x omap) Set _
P y eq = (fmap' : Fmap y) (λ i Fmap (eq i))
[ func→ x fmap' ]
module _
(eq : (λ i Omap) [ func* x func* y ])
(kk : P (func* x) refl)
where
private
p : P (func* y) eq
p = pathJ P kk (func* y) eq
eq→ : (λ i Fmap (eq i)) [ func→ x func→ y ]
eq→ = p (func→ y)
RawFunctor≡ : x y
func* (RawFunctor≡ i) = eq i
func→ (RawFunctor≡ i) = eq→ i
record IsFunctor (F : RawFunctor) : 𝓤 where record IsFunctor (F : RawFunctor) : 𝓤 where
open RawFunctor F public open RawFunctor F public
field field
@ -98,6 +126,16 @@ module _ { ' : Level} { 𝔻 : Category '} where
eqIsF : (λ i IsFunctor 𝔻 (eqR i)) [ isFunctor F isFunctor G ] eqIsF : (λ i IsFunctor 𝔻 (eqR i)) [ isFunctor F isFunctor G ]
eqIsF = IsFunctorIsProp' (isFunctor F) (isFunctor G) eqIsF = IsFunctorIsProp' (isFunctor F) (isFunctor G)
FunctorEq : {F G : Functor 𝔻}
raw F raw G
F G
raw (FunctorEq eq i) = eq i
isFunctor (FunctorEq {F} {G} eq i)
= res i
where
res : (λ i IsFunctor 𝔻 (eq i)) [ isFunctor F isFunctor G ]
res = IsFunctorIsProp' (isFunctor F) (isFunctor G)
module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : Functor A B) where module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : Functor A B) where
private private
F* = func* F F* = func* F

View file

@ -279,18 +279,18 @@ module Kleisli {a b : Level} ( : Category a b) where
module R = Functor R module R = Functor R
module R = Functor R⁰ module R = Functor R⁰
module R² = Functor module R² = Functor
ηTrans : Transformation R⁰ R η : Transformation R⁰ R
ηTrans A = pure η A = pure
ηNatural : Natural R⁰ R ηTrans ηNatural : Natural R⁰ R η
ηNatural {A} {B} f = begin ηNatural {A} {B} f = begin
ηTrans B R⁰.func→ f ≡⟨⟩ η B R⁰.func→ f ≡⟨⟩
pure f ≡⟨ sym (isNatural _) pure f ≡⟨ sym (isNatural _)
bind (pure f) pure ≡⟨⟩ bind (pure f) pure ≡⟨⟩
fmap f pure ≡⟨⟩ fmap f pure ≡⟨⟩
R.func→ f ηTrans A R.func→ f η A
μTrans : Transformation R μ : Transformation R
μTrans C = join μ C = join
μNatural : Natural R μTrans μNatural : Natural R μ
μNatural f = begin μNatural f = begin
join R².func→ f ≡⟨⟩ join R².func→ f ≡⟨⟩
bind 𝟙 R².func→ f ≡⟨⟩ bind 𝟙 R².func→ f ≡⟨⟩
@ -320,11 +320,11 @@ module Kleisli {a b : Level} ( : Category a b) where
where where
ηNatTrans : NaturalTransformation R⁰ R ηNatTrans : NaturalTransformation R⁰ R
proj₁ ηNatTrans = ηTrans proj₁ ηNatTrans = η
proj₂ ηNatTrans = ηNatural proj₂ ηNatTrans = ηNatural
μNatTrans : NaturalTransformation R μNatTrans : NaturalTransformation R
proj₁ μNatTrans = μTrans proj₁ μNatTrans = μ
proj₂ μNatTrans = μNatural proj₂ μNatTrans = μNatural
isNaturalForeign : IsNaturalForeign isNaturalForeign : IsNaturalForeign
@ -405,7 +405,8 @@ module Kleisli {a b : Level} ( : Category a b) where
-- This is problem 2.3 in [voe]. -- This is problem 2.3 in [voe].
module _ {a b : Level} { : Category a b} where module _ {a b : Level} { : Category a b} where
private private
open Category using (Object ; Arrow ; 𝟙 ; _∘_) module = Category
open using (Object ; Arrow ; 𝟙 ; _∘_ ; _>>>_)
open Functor using (func* ; func→) open Functor using (func* ; func→)
module M = Monoidal module M = Monoidal
module K = Kleisli module K = Kleisli
@ -482,22 +483,79 @@ module _ {a b : Level} { : Category a b} where
-- I believe all the proofs here should be `refl`. -- I believe all the proofs here should be `refl`.
module _ (m : K.Monad) where module _ (m : K.Monad) where
open K.RawMonad (K.Monad.raw m) 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 forthRawEq : forthRaw (backRaw m) K.Monad.raw m
K.RawMonad.RR (forthRawEq _) = RR K.RawMonad.RR (forthRawEq _) = RR
K.RawMonad.pure (forthRawEq _) = pure K.RawMonad.pure (forthRawEq _) = pure
-- stuck -- stuck
K.RawMonad.bind (forthRawEq i) = {!!} K.RawMonad.bind (forthRawEq i) = bindEq i
fortheq : (m : K.Monad) forth (back m) m fortheq : (m : K.Monad) forth (back m) m
fortheq m = K.Monad≡ (forthRawEq m) fortheq m = K.Monad≡ (forthRawEq m)
module _ (m : M.Monad) where module _ (m : M.Monad) where
open M.RawMonad (M.Monad.raw m) 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 backRawEq : backRaw (forth m) M.Monad.raw m
-- stuck -- stuck
M.RawMonad.R (backRawEq i) = {!!} M.RawMonad.R (backRawEq i) = Req i
M.RawMonad.ηNatTrans (backRawEq i) = {!!} M.RawMonad.ηNatTrans (backRawEq i) = let t = NaturalTransformation≡ F.identity R ηeq in {!t i!}
M.RawMonad.μNatTrans (backRawEq i) = {!!} M.RawMonad.μNatTrans (backRawEq i) = {!!}
backeq : (m : M.Monad) back (forth m) m backeq : (m : M.Monad) back (forth m) m

View file

@ -23,6 +23,7 @@ open import Agda.Primitive
open import Data.Product open import Data.Product
open import Cubical open import Cubical
open import Cubical.NType.Properties
open import Cat.Category open import Cat.Category
open import Cat.Category.Functor hiding (identity) open import Cat.Category.Functor hiding (identity)
@ -48,17 +49,13 @@ module NaturalTransformation {c c' d d' : Level}
NaturalTransformation : Set (c c' d') NaturalTransformation : Set (c c' d')
NaturalTransformation = Σ Transformation Natural NaturalTransformation = Σ Transformation Natural
-- TODO: Since naturality is a mere proposition this principle can be -- Think I need propPi and that arrows are sets
-- simplified. postulate propIsNatural : (θ : _) isProp (Natural θ)
NaturalTransformation≡ : {α β : NaturalTransformation} NaturalTransformation≡ : {α β : NaturalTransformation}
(eq₁ : α .proj₁ β .proj₁) (eq₁ : α .proj₁ β .proj₁)
(eq₂ : PathP
(λ i {A B : Object } (f : [ A , B ])
𝔻 [ eq₁ i B F.func→ f ]
𝔻 [ G.func→ f eq₁ i A ])
(α .proj₂) (β .proj₂))
α β α β
NaturalTransformation≡ eq eq₂ i = eq₁ i , eq₂ i NaturalTransformation≡ eq = lemSig propIsNatural _ _ eq
identityTrans : (F : Functor 𝔻) Transformation F F identityTrans : (F : Functor 𝔻) Transformation F F
identityTrans F C = 𝟙 𝔻 identityTrans F C = 𝟙 𝔻

View file

@ -16,6 +16,7 @@ open Equality.Data.Product
open import Cat.Categories.Cat using (RawCat) open import Cat.Categories.Cat using (RawCat)
module _ { : Level} { : Category } (unprovable : IsCategory (RawCat )) where module _ { : Level} { : Category } (unprovable : IsCategory (RawCat )) where
private
open import Cat.Categories.Fun open import Cat.Categories.Fun
open import Cat.Categories.Sets open import Cat.Categories.Sets
module Cat = Cat.Categories.Cat module Cat = Cat.Categories.Cat
@ -23,7 +24,6 @@ module _ { : Level} { : Category } (unprovable : IsCategory (RawCat
open Functor open Functor
𝓢 = Sets 𝓢 = Sets
open Fun (opposite ) 𝓢 open Fun (opposite ) 𝓢
private
Cat : Category _ _ Cat : Category _ _
Cat = record { raw = RawCat ; isCategory = unprovable} Cat = record { raw = RawCat ; isCategory = unprovable}
prshf = presheaf { = } prshf = presheaf { = }