Merge branch 'dev'

This commit is contained in:
Frederik Hanghøj Iversen 2018-03-13 10:41:54 +01:00
commit 3ab88395dc
25 changed files with 1827 additions and 931 deletions

View file

@ -1,9 +1,18 @@
Backlog Backlog
======= =======
Prove univalence for various categories
Prove postulates in `Cat.Wishlist` Prove postulates in `Cat.Wishlist`
`ntypeCommulative` might be there as well.
Prove that the opposite category is a category.
Prove univalence for the category of
* sets
* functors and natural transformations
Prove:
* `isProp (Product ...)`
* `isProp (HasProducts ...)`
* Functor ✓ * Functor ✓
* Applicative Functor ✗ * Applicative Functor ✗
@ -11,4 +20,11 @@ Prove postulates in `Cat.Wishlist`
* Monoidal functor ✗ * Monoidal functor ✗
* Tensorial strength ✗ * Tensorial strength ✗
* Category ✓ * Category ✓
* Monoidal category ✗ * Monoidal category ✗
* Monad
* Monoidal monad ✓
* Kleisli monad ✓
* Problem 2.3 in voe
* 1st contruction ~ monoidal ✓
* 2nd contruction ~ klesli ✓
* 1st ≃ 2nd ✗

View file

@ -1,9 +1,27 @@
Changelog Changelog
========= =========
Version 1.4.0
-------------
Adds documentation to a number of modules.
Adds an "equality principle" for categories and monads.
Prove that `IsMonad` is a mere proposition.
Provides the yoneda embedding without relying on the existence of a category of
categories. This is acheived by providing some of the data needed to make a ccc
out of the category of categories without actually having such a category.
Renames functors object map and arrow map to `omap` and `fmap`.
Prove that kleisli- and monoidal- monads are equivalent!
[WIP] Started working on the proofs for univalence for the category of sets and
the category of functors.
Version 1.3.0 Version 1.3.0
------------- -------------
Removed unused modules and streamlined things more: All specific categories are Removed unused modules and streamlined things more: All specific categories are
in the namespace `Cat.Categories`. in the namespace `Cat.Categories`.

View file

@ -14,10 +14,12 @@ Dependencies
------------ ------------
To succesfully compile the following is needed: To succesfully compile the following is needed:
* Agda version >= [`707ce6042b6a3bdb26521f3fe8dfe5d8a8470a43`](https://github.com/agda/agda/commit/707ce6042b6a3bdb26521f3fe8dfe5d8a8470a43) * The Agda release candidate 2.5.4[^1]
* [Agda Standard Library](https://github.com/agda/agda-stdlib) * The experimental branch of [Agda Standard Library](https://github.com/agda/agda-stdlib)
* [Cubical](https://github.com/Saizan/cubical-demo/) * [Cubical](https://github.com/Saizan/cubical-demo/)
[^1]: At least version >= [`707ce6042b6a3bdb26521f3fe8dfe5d8a8470a43`](https://github.com/agda/agda/commit/707ce6042b6a3bdb26521f3fe8dfe5d8a8470a43)
It's important to have the right version of these - but which one is the right It's important to have the right version of these - but which one is the right
is in constant flux. It's most likely the newest one. is in constant flux. It's most likely the newest one.

@ -1 +1 @@
Subproject commit b9c8e02597751a1b15045cbc5108c221999bd540 Subproject commit fbd8ba7ea84c4b643fd08797b4031b18a59f561d

@ -1 +1 @@
Subproject commit 0d3f02e68297e940227137beac45fc1bce6e2bea Subproject commit 5b35333dbbd8fa523e478c1cfe60657321ca38fe

View file

@ -1,19 +1,20 @@
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
open import Cat.Category.Monad.Voevodsky
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

@ -4,20 +4,20 @@
module Cat.Categories.Cat where module Cat.Categories.Cat where
open import Agda.Primitive open import Agda.Primitive
open import Cubical
open import Function
open import Data.Product renaming (proj₁ to fst ; proj₂ to snd) open import Data.Product renaming (proj₁ to fst ; proj₂ to snd)
open import Cubical
open import Cubical.Sigma
open import Cat.Category open import Cat.Category
open import Cat.Category.Functor open import Cat.Category.Functor
open import Cat.Category.Product open import Cat.Category.Product
open import Cat.Category.Exponential open import Cat.Category.Exponential hiding (_×_ ; product)
open import Cat.Category.NaturalTransformation open import Cat.Category.NaturalTransformation
open import Cat.Equality open import Cat.Equality
open Equality.Data.Product open Equality.Data.Product
open Functor using (func→ ; func*)
open Category using (Object ; 𝟙) open Category using (Object ; 𝟙)
-- The category of categories -- The category of categories
@ -25,14 +25,14 @@ module _ ( ' : Level) where
private private
module _ {𝔸 𝔹 𝔻 : Category '} {F : Functor 𝔸 𝔹} {G : Functor 𝔹 } {H : Functor 𝔻} where module _ {𝔸 𝔹 𝔻 : Category '} {F : Functor 𝔸 𝔹} {G : Functor 𝔹 } {H : Functor 𝔻} where
assc : F[ H F[ G F ] ] F[ F[ H G ] F ] assc : F[ H F[ G F ] ] F[ F[ H G ] F ]
assc = Functor≡ refl refl assc = Functor≡ refl
module _ { 𝔻 : Category '} {F : Functor 𝔻} where module _ { 𝔻 : Category '} {F : Functor 𝔻} where
ident-r : F[ F identity ] F ident-r : F[ F identity ] F
ident-r = Functor≡ refl refl ident-r = Functor≡ refl
ident-l : F[ identity F ] F ident-l : F[ identity F ] F
ident-l = Functor≡ refl refl ident-l = Functor≡ refl
RawCat : RawCategory (lsuc ( ')) ( ') RawCat : RawCategory (lsuc ( ')) ( ')
RawCat = RawCat =
@ -46,313 +46,325 @@ module _ ( ' : Level) where
open RawCategory RawCat open RawCategory RawCat
isAssociative : IsAssociative isAssociative : IsAssociative
isAssociative {f = F} {G} {H} = assc {F = F} {G = G} {H = H} isAssociative {f = F} {G} {H} = assc {F = F} {G = G} {H = H}
-- TODO: Rename `ident'` to `ident` after changing how names are exposed in Functor. ident : IsIdentity identity
ident' : IsIdentity identity ident = ident-r , ident-l
ident' = ident-r , ident-l
-- NB! `ArrowsAreSets RawCat` is *not* provable. The type of functors,
-- however, form a groupoid! Therefore there is no (1-)category of
-- categories. There does, however, exist a 2-category of 1-categories.
-- Because of the note above there is not category of categories. -- NB! `ArrowsAreSets RawCat` is *not* provable. The type of functors,
-- however, form a groupoid! Therefore there is no (1-)category of
-- categories. There does, however, exist a 2-category of 1-categories.
--
-- Because of this there is no category of categories.
Cat : (unprovable : IsCategory RawCat) Category (lsuc ( ')) ( ') Cat : (unprovable : IsCategory RawCat) Category (lsuc ( ')) ( ')
Category.raw (Cat _) = RawCat Category.raw (Cat _) = RawCat
Category.isCategory (Cat unprovable) = unprovable Category.isCategory (Cat unprovable) = unprovable
-- Category.raw Cat _ = RawCat
-- Category.isCategory Cat unprovable = unprovable
-- The following to some extend depends on the category of categories being a -- | In the following we will pretend there is a category of categories when
-- category. In some places it may not actually be needed, however. -- e.g. talking about it being cartesian closed. It still makes sense to
-- construct these things even though that category does not exist.
--
-- If the notion of a category is later generalized to work on different
-- homotopy levels, then the proof that the category of categories is cartesian
-- closed will follow immediately from these constructions.
-- | the category of categories have products.
module CatProduct { ' : Level} ( 𝔻 : Category ') where
private
module = Category
module 𝔻 = Category 𝔻
Obj = Object × Object 𝔻
Arr : Obj Obj Set '
Arr (c , d) (c' , d') = [ c , c' ] × 𝔻 [ d , d' ]
𝟙' : {o : Obj} Arr o o
𝟙' = 𝟙 , 𝟙 𝔻
_∘_ :
{a b c : Obj}
Arr b c
Arr a b
Arr a c
_∘_ = λ { (bc∈C , bc∈D) (ab∈C , ab∈D) [ bc∈C ab∈C ] , 𝔻 [ bc∈D ab∈D ]}
rawProduct : RawCategory '
RawCategory.Object rawProduct = Obj
RawCategory.Arrow rawProduct = Arr
RawCategory.𝟙 rawProduct = 𝟙'
RawCategory._∘_ rawProduct = _∘_
open RawCategory rawProduct
arrowsAreSets : ArrowsAreSets
arrowsAreSets = setSig {sA = .arrowsAreSets} {sB = λ x 𝔻.arrowsAreSets}
isIdentity : IsIdentity 𝟙'
isIdentity
= Σ≡ (fst .isIdentity) (fst 𝔻.isIdentity)
, Σ≡ (snd .isIdentity) (snd 𝔻.isIdentity)
postulate univalent : Univalence.Univalent rawProduct isIdentity
instance
isCategory : IsCategory rawProduct
IsCategory.isAssociative isCategory = Σ≡ .isAssociative 𝔻.isAssociative
IsCategory.isIdentity isCategory = isIdentity
IsCategory.arrowsAreSets isCategory = arrowsAreSets
IsCategory.univalent isCategory = univalent
object : Category '
Category.raw object = rawProduct
proj₁ : Functor object
proj₁ = record
{ raw = record
{ omap = fst ; fmap = fst }
; isFunctor = record
{ isIdentity = refl ; isDistributive = refl }
}
proj₂ : Functor object 𝔻
proj₂ = record
{ raw = record
{ omap = snd ; fmap = snd }
; isFunctor = record
{ isIdentity = refl ; isDistributive = refl }
}
module _ {X : Category '} (x₁ : Functor X ) (x₂ : Functor X 𝔻) where
private
x : Functor X object
x = record
{ raw = record
{ omap = λ x x₁.omap x , x₂.omap x
; fmap = λ x x₁.fmap x , x₂.fmap x
}
; isFunctor = record
{ isIdentity = Σ≡ x₁.isIdentity x₂.isIdentity
; isDistributive = Σ≡ x₁.isDistributive x₂.isDistributive
}
}
where
open module x = Functor x₁
open module x = Functor x₂
isUniqL : F[ proj₁ x ] x₁
isUniqL = Functor≡ refl
isUniqR : F[ proj₂ x ] x₂
isUniqR = Functor≡ refl
isUniq : F[ proj₁ x ] x₁ × F[ proj₂ x ] x₂
isUniq = isUniqL , isUniqR
isProduct : ∃![ x ] (F[ proj₁ x ] x₁ × F[ proj₂ x ] x₂)
isProduct = x , isUniq
module _ { ' : Level} (unprovable : IsCategory (RawCat ')) where module _ { ' : Level} (unprovable : IsCategory (RawCat ')) where
private
Cat = Cat ' unprovable
module _ ( 𝔻 : Category ') where module _ ( 𝔻 : Category ') where
private private
Catt = Cat ' unprovable module P = CatProduct 𝔻
:Object: = Object × Object 𝔻
:Arrow: : :Object: :Object: Set '
:Arrow: (c , d) (c' , d') = [ c , c' ] × 𝔻 [ d , d' ]
:𝟙: : {o : :Object:} :Arrow: o o
:𝟙: = 𝟙 , 𝟙 𝔻
_:⊕:_ :
{a b c : :Object:}
:Arrow: b c
:Arrow: a b
:Arrow: a c
_:⊕:_ = λ { (bc∈C , bc∈D) (ab∈C , ab∈D) [ bc∈C ab∈C ] , 𝔻 [ bc∈D ab∈D ]}
:rawProduct: : RawCategory ' rawProduct : RawProduct Cat 𝔻
RawCategory.Object :rawProduct: = :Object: RawProduct.object rawProduct = P.object
RawCategory.Arrow :rawProduct: = :Arrow: RawProduct.proj₁ rawProduct = P.proj₁
RawCategory.𝟙 :rawProduct: = :𝟙: RawProduct.proj₂ rawProduct = P.proj₂
RawCategory._∘_ :rawProduct: = _:⊕:_
open RawCategory :rawProduct:
module C = Category isProduct : IsProduct Cat _ _ rawProduct
module D = Category 𝔻 IsProduct.isProduct isProduct = P.isProduct
open import Cubical.Sigma
issSet : {A B : RawCategory.Object :rawProduct:} isSet (Arrow A B)
issSet = setSig {sA = C.arrowsAreSets} {sB = λ x D.arrowsAreSets}
ident' : IsIdentity :𝟙:
ident'
= Σ≡ (fst C.isIdentity) (fst D.isIdentity)
, Σ≡ (snd C.isIdentity) (snd D.isIdentity)
postulate univalent : Univalence.Univalent :rawProduct: ident'
instance
:isCategory: : IsCategory :rawProduct:
IsCategory.isAssociative :isCategory: = Σ≡ C.isAssociative D.isAssociative
IsCategory.isIdentity :isCategory: = ident'
IsCategory.arrowsAreSets :isCategory: = issSet
IsCategory.univalent :isCategory: = univalent
:product: : Category ' product : Product Cat 𝔻
Category.raw :product: = :rawProduct: Product.raw product = rawProduct
Product.isProduct product = isProduct
proj₁ : Catt [ :product: , ]
proj₁ = record
{ raw = record { func* = fst ; func→ = fst }
; isFunctor = record { isIdentity = refl ; isDistributive = refl }
}
proj₂ : Catt [ :product: , 𝔻 ]
proj₂ = record
{ raw = record { func* = snd ; func→ = snd }
; isFunctor = record { isIdentity = refl ; isDistributive = refl }
}
module _ {X : Object Catt} (x₁ : Catt [ X , ]) (x₂ : Catt [ X , 𝔻 ]) where
x : Functor X :product:
x = record
{ raw = record
{ func* = λ x x₁ .func* x , x₂ .func* x
; func→ = λ x func→ x₁ x , func→ x₂ x
}
; isFunctor = record
{ isIdentity = Σ≡ x₁.isIdentity x₂.isIdentity
; isDistributive = Σ≡ x₁.isDistributive x₂.isDistributive
}
}
where
open module x = Functor x₁
open module x = Functor x₂
isUniqL : Catt [ proj₁ x ] x₁
isUniqL = Functor≡ eq* eq→
where
eq* : (Catt [ proj₁ x ]) .func* x₁ .func*
eq* = refl
eq→ : (λ i {A : Object X} {B : Object X} X [ A , B ] [ eq* i A , eq* i B ])
[ (Catt [ proj₁ x ]) .func→ x₁ .func→ ]
eq→ = refl
isUniqR : Catt [ proj₂ x ] x₂
isUniqR = Functor≡ refl refl
isUniq : Catt [ proj₁ x ] x₁ × Catt [ proj₂ x ] x₂
isUniq = isUniqL , isUniqR
uniq : ∃![ x ] (Catt [ proj₁ x ] x₁ × Catt [ proj₂ x ] x₂)
uniq = x , isUniq
instance
isProduct : IsProduct Catt proj₁ proj₂
isProduct = uniq
product : Product { = Catt} 𝔻
product = record
{ obj = :product:
; proj₁ = proj₁
; proj₂ = proj₂
}
module _ { ' : Level} (unprovable : IsCategory (RawCat ')) where
Catt = Cat ' unprovable
instance instance
hasProducts : HasProducts Catt hasProducts : HasProducts Cat
hasProducts = record { product = product unprovable } hasProducts = record { product = product }
-- Basically proves that `Cat ` is cartesian closed. -- | The category of categories have expoentntials - and because it has products
module _ ( : Level) (unprovable : IsCategory (RawCat )) where -- it is therefory also cartesian closed.
module CatExponential { : Level} ( 𝔻 : Category ) where
private private
open Data.Product open Data.Product
open import Cat.Categories.Fun open import Cat.Categories.Fun
module = Category
module 𝔻 = Category 𝔻
Category = Category
open Fun 𝔻 renaming (identity to idN)
omap : Functor 𝔻 × Object Object 𝔻
omap (F , A) = Functor.omap F A
-- The exponential object
object : Category
object = Fun
module _ {dom cod : Functor 𝔻 × Object } where
private
F : Functor 𝔻
F = proj₁ dom
A : Object
A = proj₂ dom
G : Functor 𝔻
G = proj₁ cod
B : Object
B = proj₂ cod
module F = Functor F
module G = Functor G
fmap : (pobj : NaturalTransformation F G × [ A , B ])
𝔻 [ F.omap A , G.omap B ]
fmap ((θ , θNat) , f) = result
where
θA : 𝔻 [ F.omap A , G.omap A ]
θA = θ A
θB : 𝔻 [ F.omap B , G.omap B ]
θB = θ B
F→f : 𝔻 [ F.omap A , F.omap B ]
F→f = F.fmap f
G→f : 𝔻 [ G.omap A , G.omap B ]
G→f = G.fmap f
l : 𝔻 [ F.omap A , G.omap B ]
l = 𝔻 [ θB F.fmap f ]
r : 𝔻 [ F.omap A , G.omap B ]
r = 𝔻 [ G.fmap f θA ]
result : 𝔻 [ F.omap A , G.omap B ]
result = l
open CatProduct renaming (object to _⊗_) using ()
module _ {c : Functor 𝔻 × Object } where
private
F : Functor 𝔻
F = proj₁ c
C : Object
C = proj₂ c
ident : fmap {c} {c} (NT.identity F , 𝟙 {A = proj₂ c}) 𝟙 𝔻
ident = begin
fmap {c} {c} (𝟙 (object ) {c}) ≡⟨⟩
fmap {c} {c} (idN F , 𝟙 ) ≡⟨⟩
𝔻 [ identityTrans F C F.fmap (𝟙 )] ≡⟨⟩
𝔻 [ 𝟙 𝔻 F.fmap (𝟙 )] ≡⟨ proj₂ 𝔻.isIdentity
F.fmap (𝟙 ) ≡⟨ F.isIdentity
𝟙 𝔻
where
module F = Functor F
module _ {F×A G×B H×C : Functor 𝔻 × Object } where
private
F = F×A .proj₁
A = F×A .proj₂
G = G×B .proj₁
B = G×B .proj₂
H = H×C .proj₁
C = H×C .proj₂
module F = Functor F
module G = Functor G
module H = Functor H
module _
-- NaturalTransformation F G × .Arrow A B
{θ×f : NaturalTransformation F G × [ A , B ]}
{η×g : NaturalTransformation G H × [ B , C ]} where
private
θ : Transformation F G
θ = proj₁ (proj₁ θ×f)
θNat : Natural F G θ
θNat = proj₂ (proj₁ θ×f)
f : [ A , B ]
f = proj₂ θ×f
η : Transformation G H
η = proj₁ (proj₁ η×g)
ηNat : Natural G H η
ηNat = proj₂ (proj₁ η×g)
g : [ B , C ]
g = proj₂ η×g
ηθNT : NaturalTransformation F H
ηθNT = Category._∘_ Fun {F} {G} {H} (η , ηNat) (θ , θNat)
ηθ = proj₁ ηθNT
ηθNat = proj₂ ηθNT
isDistributive :
𝔻 [ 𝔻 [ η C θ C ] F.fmap ( [ g f ] ) ]
𝔻 [ 𝔻 [ η C G.fmap g ] 𝔻 [ θ B F.fmap f ] ]
isDistributive = begin
𝔻 [ (ηθ C) F.fmap ( [ g f ]) ]
≡⟨ ηθNat ( [ g f ])
𝔻 [ H.fmap ( [ g f ]) (ηθ A) ]
≡⟨ cong (λ φ 𝔻 [ φ ηθ A ]) (H.isDistributive)
𝔻 [ 𝔻 [ H.fmap g H.fmap f ] (ηθ A) ]
≡⟨ sym 𝔻.isAssociative
𝔻 [ H.fmap g 𝔻 [ H.fmap f ηθ A ] ]
≡⟨ cong (λ φ 𝔻 [ H.fmap g φ ]) 𝔻.isAssociative
𝔻 [ H.fmap g 𝔻 [ 𝔻 [ H.fmap f η A ] θ A ] ]
≡⟨ cong (λ φ 𝔻 [ H.fmap g φ ]) (cong (λ φ 𝔻 [ φ θ A ]) (sym (ηNat f)))
𝔻 [ H.fmap g 𝔻 [ 𝔻 [ η B G.fmap f ] θ A ] ]
≡⟨ cong (λ φ 𝔻 [ H.fmap g φ ]) (sym 𝔻.isAssociative)
𝔻 [ H.fmap g 𝔻 [ η B 𝔻 [ G.fmap f θ A ] ] ]
≡⟨ 𝔻.isAssociative
𝔻 [ 𝔻 [ H.fmap g η B ] 𝔻 [ G.fmap f θ A ] ]
≡⟨ cong (λ φ 𝔻 [ φ 𝔻 [ G.fmap f θ A ] ]) (sym (ηNat g))
𝔻 [ 𝔻 [ η C G.fmap g ] 𝔻 [ G.fmap f θ A ] ]
≡⟨ cong (λ φ 𝔻 [ 𝔻 [ η C G.fmap g ] φ ]) (sym (θNat f))
𝔻 [ 𝔻 [ η C G.fmap g ] 𝔻 [ θ B F.fmap f ] ]
eval : Functor (CatProduct.object object ) 𝔻
eval = record
{ raw = record
{ omap = omap
; fmap = λ {dom} {cod} fmap {dom} {cod}
}
; isFunctor = record
{ isIdentity = λ {o} ident {o}
; isDistributive = λ {f u n k y} isDistributive {f} {u} {n} {k} {y}
}
}
module _ (𝔸 : Category ) (F : Functor (𝔸 ) 𝔻) where
postulate
parallelProduct
: Functor 𝔸 object Functor
Functor (𝔸 ) (object )
transpose : Functor 𝔸 object
eq : F[ eval (parallelProduct transpose (identity {C = })) ] F
-- eq : F[ :eval: ∘ {!!} ] ≡ F
-- eq : Cat [ :eval: ∘ (HasProducts._|×|_ hasProducts transpose (𝟙 Cat {o = })) ] ≡ F
-- eq' : (Cat [ :eval: ∘
-- (record { product = product } HasProducts.|×| transpose)
-- (𝟙 Cat)
-- ])
-- ≡ F
-- For some reason after `e8215b2c051062c6301abc9b3f6ec67106259758`
-- `catTranspose` makes Agda hang. catTranspose : ∃![ F~ ] (Cat [
-- :eval: (parallelProduct F~ (𝟙 Cat {o = }))] F) catTranspose =
-- transpose , eq
-- We don't care about filling out the holes below since they are anyways hidden
-- behind an unprovable statement.
module _ ( : Level) (unprovable : IsCategory (RawCat )) where
private
Cat : Category (lsuc ( )) ( ) Cat : Category (lsuc ( )) ( )
Cat = Cat unprovable Cat = Cat unprovable
module _ ( 𝔻 : Category ) where module _ ( 𝔻 : Category ) where
open Fun 𝔻 renaming (identity to idN) module CatExp = CatExponential 𝔻
private _⊗_ = CatProduct.object
:obj: : Object Cat
:obj: = Fun
:func*: : Functor 𝔻 × Object Object 𝔻 -- Filling the hole causes Agda to loop indefinitely.
:func*: (F , A) = func* F A eval : Functor (CatExp.object ) 𝔻
eval = {!CatExp.eval!}
module _ {dom cod : Functor 𝔻 × Object } where isExponential : IsExponential Cat 𝔻 CatExp.object eval
private isExponential = {!CatExp.isExponential!}
F : Functor 𝔻
F = proj₁ dom
A : Object
A = proj₂ dom
G : Functor 𝔻 exponent : Exponential Cat 𝔻
G = proj₁ cod exponent = record
B : Object { obj = CatExp.object
B = proj₂ cod ; eval = eval
; isExponential = isExponential
:func→: : (pobj : NaturalTransformation F G × [ A , B ])
𝔻 [ func* F A , func* G B ]
:func→: ((θ , θNat) , f) = result
where
θA : 𝔻 [ func* F A , func* G A ]
θA = θ A
θB : 𝔻 [ func* F B , func* G B ]
θB = θ B
F→f : 𝔻 [ func* F A , func* F B ]
F→f = func→ F f
G→f : 𝔻 [ func* G A , func* G B ]
G→f = func→ G f
l : 𝔻 [ func* F A , func* G B ]
l = 𝔻 [ θB F→f ]
r : 𝔻 [ func* F A , func* G B ]
r = 𝔻 [ G→f θA ]
-- There are two choices at this point,
-- but I suppose the whole point is that
-- by `θNat f` we have `l ≡ r`
-- lem : 𝔻 [ θ B ∘ F .func→ f ] ≡ 𝔻 [ G .func→ f ∘ θ A ]
-- lem = θNat f
result : 𝔻 [ func* F A , func* G B ]
result = l
_×p_ = product unprovable
module _ {c : Functor 𝔻 × Object } where
private
F : Functor 𝔻
F = proj₁ c
C : Object
C = proj₂ c
-- NaturalTransformation F G × .Arrow A B
-- :ident: : :func→: {c} {c} (identityNat F , .𝟙) 𝔻 .𝟙
-- :ident: = trans (proj₂ 𝔻.isIdentity) (F .isIdentity)
-- where
-- open module 𝔻 = IsCategory (𝔻 .isCategory)
-- Unfortunately the equational version has some ambigous arguments.
:ident: : :func→: {c} {c} (NT.identity F , 𝟙 {A = proj₂ c}) 𝟙 𝔻
:ident: = begin
:func→: {c} {c} (𝟙 (Product.obj (:obj: ×p )) {c}) ≡⟨⟩
:func→: {c} {c} (idN F , 𝟙 ) ≡⟨⟩
𝔻 [ identityTrans F C func→ F (𝟙 )] ≡⟨⟩
𝔻 [ 𝟙 𝔻 func→ F (𝟙 )] ≡⟨ proj₂ 𝔻.isIdentity
func→ F (𝟙 ) ≡⟨ F.isIdentity
𝟙 𝔻
where
open module 𝔻 = Category 𝔻
open module F = Functor F
module _ {F×A G×B H×C : Functor 𝔻 × Object } where
F = F×A .proj₁
A = F×A .proj₂
G = G×B .proj₁
B = G×B .proj₂
H = H×C .proj₁
C = H×C .proj₂
-- Not entirely clear what this is at this point:
_P⊕_ = Category._∘_ (Product.obj (:obj: ×p )) {F×A} {G×B} {H×C}
module _
-- NaturalTransformation F G × .Arrow A B
{θ×f : NaturalTransformation F G × [ A , B ]}
{η×g : NaturalTransformation G H × [ B , C ]} where
private
θ : Transformation F G
θ = proj₁ (proj₁ θ×f)
θNat : Natural F G θ
θNat = proj₂ (proj₁ θ×f)
f : [ A , B ]
f = proj₂ θ×f
η : Transformation G H
η = proj₁ (proj₁ η×g)
ηNat : Natural G H η
ηNat = proj₂ (proj₁ η×g)
g : [ B , C ]
g = proj₂ η×g
ηθNT : NaturalTransformation F H
ηθNT = Category._∘_ Fun {F} {G} {H} (η , ηNat) (θ , θNat)
ηθ = proj₁ ηθNT
ηθNat = proj₂ ηθNT
:isDistributive: :
𝔻 [ 𝔻 [ η C θ C ] func→ F ( [ g f ] ) ]
𝔻 [ 𝔻 [ η C func→ G g ] 𝔻 [ θ B func→ F f ] ]
:isDistributive: = begin
𝔻 [ (ηθ C) func→ F ( [ g f ]) ]
≡⟨ ηθNat ( [ g f ])
𝔻 [ func→ H ( [ g f ]) (ηθ A) ]
≡⟨ cong (λ φ 𝔻 [ φ ηθ A ]) (H.isDistributive)
𝔻 [ 𝔻 [ func→ H g func→ H f ] (ηθ A) ]
≡⟨ sym isAssociative
𝔻 [ func→ H g 𝔻 [ func→ H f ηθ A ] ]
≡⟨ cong (λ φ 𝔻 [ func→ H g φ ]) isAssociative
𝔻 [ func→ H g 𝔻 [ 𝔻 [ func→ H f η A ] θ A ] ]
≡⟨ cong (λ φ 𝔻 [ func→ H g φ ]) (cong (λ φ 𝔻 [ φ θ A ]) (sym (ηNat f)))
𝔻 [ func→ H g 𝔻 [ 𝔻 [ η B func→ G f ] θ A ] ]
≡⟨ cong (λ φ 𝔻 [ func→ H g φ ]) (sym isAssociative)
𝔻 [ func→ H g 𝔻 [ η B 𝔻 [ func→ G f θ A ] ] ]
≡⟨ isAssociative
𝔻 [ 𝔻 [ func→ H g η B ] 𝔻 [ func→ G f θ A ] ]
≡⟨ cong (λ φ 𝔻 [ φ 𝔻 [ func→ G f θ A ] ]) (sym (ηNat g))
𝔻 [ 𝔻 [ η C func→ G g ] 𝔻 [ func→ G f θ A ] ]
≡⟨ cong (λ φ 𝔻 [ 𝔻 [ η C func→ G g ] φ ]) (sym (θNat f))
𝔻 [ 𝔻 [ η C func→ G g ] 𝔻 [ θ B func→ F f ] ]
where
open Category 𝔻
module H = Functor H
:eval: : Functor ((:obj: ×p ) .Product.obj) 𝔻
:eval: = record
{ raw = record
{ func* = :func*:
; func→ = λ {dom} {cod} :func→: {dom} {cod}
}
; isFunctor = record
{ isIdentity = λ {o} :ident: {o}
; isDistributive = λ {f u n k y} :isDistributive: {f} {u} {n} {k} {y}
}
}
module _ (𝔸 : Category ) (F : Functor ((𝔸 ×p ) .Product.obj) 𝔻) where
open HasProducts (hasProducts {} {} unprovable) renaming (_|×|_ to parallelProduct)
postulate
transpose : Functor 𝔸 :obj:
eq : Cat [ :eval: (parallelProduct transpose (𝟙 Cat {A = })) ] F
-- eq : Cat [ :eval: ∘ (HasProducts._|×|_ hasProducts transpose (𝟙 Cat {o = })) ] ≡ F
-- eq' : (Cat [ :eval: ∘
-- (record { product = product } HasProducts.|×| transpose)
-- (𝟙 Cat)
-- ])
-- ≡ F
-- For some reason after `e8215b2c051062c6301abc9b3f6ec67106259758`
-- `catTranspose` makes Agda hang. catTranspose : ∃![ F~ ] (Cat [
-- :eval: (parallelProduct F~ (𝟙 Cat {o = }))] F) catTranspose =
-- transpose , eq
postulate :isExponential: : IsExponential Cat 𝔻 :obj: :eval:
-- :isExponential: : IsExponential Cat 𝔻 :obj: :eval:
-- :isExponential: = {!catTranspose!}
-- where
-- open HasProducts (hasProducts {} {} unprovable) using (_|×|_)
-- :isExponential: = λ 𝔸 F transpose 𝔸 F , eq' 𝔸 F
-- :exponent: : Exponential (Cat ) A B
:exponent: : Exponential Cat 𝔻
:exponent: = record
{ obj = :obj:
; eval = :eval:
; isExponential = :isExponential:
} }
hasExponentials : HasExponentials Cat hasExponentials : HasExponentials Cat
hasExponentials = record { exponent = :exponent: } hasExponentials = record { exponent = exponent }

View file

@ -26,24 +26,24 @@ open Category hiding (_∘_)
open Functor open Functor
module _ { ' : Level} (Ns : Set ) where module _ { ' : Level} (Ns : Set ) where
-- Ns is the "namespace" private
o = (suc zero ) -- Ns is the "namespace"
o = (suc zero )
FiniteDecidableSubset : Set FiniteDecidableSubset : Set
FiniteDecidableSubset = Ns Dec FiniteDecidableSubset = Ns Dec
isTrue : Bool Set isTrue : Bool Set
isTrue false = isTrue false =
isTrue true = isTrue true =
elmsof : FiniteDecidableSubset Set elmsof : FiniteDecidableSubset Set
elmsof P = Σ Ns (λ σ True (P σ)) -- (σ : Ns) → isTrue (P σ) elmsof P = Σ Ns (λ σ True (P σ)) -- (σ : Ns) → isTrue (P σ)
𝟚 : Set 𝟚 : Set
𝟚 = 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
@ -63,18 +63,18 @@ module _ { ' : Level} (Ns : Set ) where
; (inj₂ _) Lift ; (inj₂ _) Lift
} }
Hom = Σ Hom' rules Hom = Σ Hom' rules
module Raw = RawCategory module Raw = RawCategory
-- The category of names and substitutions -- The category of names and substitutions
Raw : RawCategory -- o (lsuc lzero ⊔ o) Raw : RawCategory -- o (lsuc lzero ⊔ o)
Raw.Object Raw = FiniteDecidableSubset Raw.Object Raw = FiniteDecidableSubset
Raw.Arrow Raw = Hom Raw.Arrow Raw = Hom
Raw.𝟙 Raw {o} = inj₁ , λ { (i , ii) (j , jj) eq Σ≡ eq {!refl!} } Raw.𝟙 Raw {o} = inj₁ , λ { (i , ii) (j , jj) eq Σ≡ eq {!refl!} }
Raw._∘_ Raw = {!!} Raw._∘_ Raw = {!!}
postulate IsCategory : IsCategory Raw postulate IsCategory : IsCategory Raw
: Category : Category
raw = Raw raw = Raw
isCategory = IsCategory isCategory = IsCategory

View file

@ -28,20 +28,20 @@ module _ {a b : Level} where
private private
module T = Functor T module T = Functor T
Type : (Γ : Object ) Set a Type : (Γ : Object ) Set a
Type Γ = proj₁ (proj₁ (T.func* Γ)) Type Γ = proj₁ (proj₁ (T.omap Γ))
module _ {Γ : Object } {A : Type Γ} where module _ {Γ : Object } {A : Type Γ} where
-- module _ {A B : Object } {γ : [ A , B ]} where -- module _ {A B : Object } {γ : [ A , B ]} where
-- k : Σ (proj₁ (func* T B) → proj₁ (func* T A)) -- k : Σ (proj₁ (omap T B) → proj₁ (omap T A))
-- (λ f → -- (λ f →
-- {x : proj₁ (func* T B)} → -- {x : proj₁ (omap T B)} →
-- proj₂ (func* T B) x → proj₂ (func* T A) (f x)) -- proj₂ (omap T B) x → proj₂ (omap T A) (f x))
-- k = T.func→ γ -- k = T.fmap γ
-- k₁ : proj₁ (func* T B) → proj₁ (func* T A) -- k₁ : proj₁ (omap T B) → proj₁ (omap T A)
-- k₁ = proj₁ k -- k₁ = proj₁ k
-- k₂ : ({x : proj₁ (func* T B)} → -- k₂ : ({x : proj₁ (omap T B)} →
-- proj₂ (func* T B) x → proj₂ (func* T A) (k₁ x)) -- proj₂ (omap T B) x → proj₂ (omap T A) (k₁ x))
-- k₂ = proj₂ k -- k₂ = proj₂ k
record ContextComprehension : Set (a b) where record ContextComprehension : Set (a b) where
@ -51,7 +51,7 @@ module _ {a b : Level} where
-- proj2 : ???? -- proj2 : ????
-- if γ : [ A , B ] -- if γ : [ A , B ]
-- then T .func→ γ (written T[γ]) interpret substitutions in types and terms respectively. -- then T .fmap γ (written T[γ]) interpret substitutions in types and terms respectively.
-- field -- field
-- ump : {Δ : .Object} → (γ : [ Δ , Γ ]) -- ump : {Δ : .Object} → (γ : [ Δ , Γ ])
-- → (a : {!!}) → {!!} -- → (a : {!!}) → {!!}

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
module = Category
open Category
private private
module = Category
open Category
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

@ -4,50 +4,24 @@ module Cat.Categories.Fun where
open import Agda.Primitive open import Agda.Primitive
open import Data.Product open import Data.Product
open import Data.Nat using (_≤_ ; z≤n ; s≤s)
module Nat = Data.Nat
open import Data.Product
open import Cubical open import Cubical
open import Cubical.Sigma open import Cubical.GradLemma
open import Cubical.NType.Properties 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)
open import Cat.Category.NaturalTransformation open import Cat.Category.NaturalTransformation
open import Cat.Wishlist
open import Cat.Equality
import Cat.Category.NaturalTransformation
open Equality.Data.Product
module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : Category d d') where module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : Category d d') where
open Category using (Object ; 𝟙)
module NT = NaturalTransformation 𝔻 module NT = NaturalTransformation 𝔻
open NT public open NT public
private private
module = Category
module 𝔻 = Category 𝔻 module 𝔻 = Category 𝔻
private
module _ {F G : Functor 𝔻} where module _ {A B C D : Functor 𝔻} {θ' : NaturalTransformation A B}
transformationIsSet : isSet (Transformation F G) {η' : NaturalTransformation B C} {ζ' : NaturalTransformation C D} where
transformationIsSet _ _ p q i j C = 𝔻.arrowsAreSets _ _ (λ l p l C) (λ l q l C) i j
naturalIsProp : (θ : Transformation F G) isProp (Natural F G θ)
naturalIsProp θ θNat θNat' = lem
where
lem : (λ _ Natural F G θ) [ (λ f θNat f) (λ f θNat' f) ]
lem = λ i f 𝔻.arrowsAreSets _ _ (θNat f) (θNat' f) i
naturalTransformationIsSets : isSet (NaturalTransformation F G)
naturalTransformationIsSets = sigPresSet transformationIsSet
λ θ ntypeCommulative
(s≤s {n = Nat.suc Nat.zero} z≤n)
(naturalIsProp θ)
module _ {A B C D : Functor 𝔻} {θ' : NaturalTransformation A B}
{η' : NaturalTransformation B C} {ζ' : NaturalTransformation C D} where
private
θ = proj₁ θ' θ = proj₁ θ'
η = proj₁ η' η = proj₁ η'
ζ = proj₁ ζ' ζ = proj₁ ζ'
@ -58,13 +32,11 @@ module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : C
L = (NT[_∘_] {A} {C} {D} ζ' (NT[_∘_] {A} {B} {C} η' θ')) L = (NT[_∘_] {A} {C} {D} ζ' (NT[_∘_] {A} {B} {C} η' θ'))
R : NaturalTransformation A D R : NaturalTransformation A D
R = (NT[_∘_] {A} {B} {D} (NT[_∘_] {B} {C} {D} ζ' η') θ') R = (NT[_∘_] {A} {B} {D} (NT[_∘_] {B} {C} {D} ζ' η') θ')
_g⊕f_ = NT[_∘_] {A} {B} {C} _g⊕f_ = NT[_∘_] {A} {B} {C}
_h⊕g_ = NT[_∘_] {B} {C} {D} _h⊕g_ = NT[_∘_] {B} {C} {D}
:isAssociative: : L R isAssociative : L R
:isAssociative: = lemSig (naturalIsProp {F = A} {D}) isAssociative = lemSig (naturalIsProp {F = A} {D})
L R (funExt (λ x isAssociative)) L R (funExt (λ x 𝔻.isAssociative))
where
open Category 𝔻
private private
module _ {A B : Functor 𝔻} {f : NaturalTransformation A B} where module _ {A B : Functor 𝔻} {f : NaturalTransformation A B} where
@ -94,27 +66,120 @@ module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : C
; _∘_ = λ {F G H} NT[_∘_] {F} {G} {H} ; _∘_ = λ {F G H} NT[_∘_] {F} {G} {H}
} }
open RawCategory RawFun
open Univalence RawFun
module _ {A B : Functor 𝔻} where
module A = Functor A
module B = Functor B
module _ (p : A B) where
omapP : A.omap B.omap
omapP i = Functor.omap (p i)
coerceAB : {X} 𝔻 [ A.omap X , A.omap X ] 𝔻 [ A.omap X , B.omap X ]
coerceAB {X} = cong (λ φ 𝔻 [ A.omap X , φ X ]) omapP
-- The transformation will be the identity on 𝔻. Such an arrow has the
-- type `A.omap A → A.omap A`. Which we can coerce to have the type
-- `A.omap → B.omap` since `A` and `B` are equal.
coe𝟙 : Transformation A B
coe𝟙 X = coe coerceAB 𝔻.𝟙
module _ {a b : .Object} (f : [ a , b ]) where
nat' : 𝔻 [ coe𝟙 b A.fmap f ] 𝔻 [ B.fmap f coe𝟙 a ]
nat' = begin
(𝔻 [ coe𝟙 b A.fmap f ]) ≡⟨ {!!}
(𝔻 [ B.fmap f coe𝟙 a ])
transs : (i : I) Transformation A (p i)
transs = {!!}
natt : (i : I) Natural A (p i) {!!}
natt = {!!}
t : Natural A B coe𝟙
t = coe c (identityNatural A)
where
c : Natural A A (identityTrans A) Natural A B coe𝟙
c = begin
Natural A A (identityTrans A) ≡⟨ (λ x {!natt ?!})
Natural A B coe𝟙
-- cong (λ φ → {!Natural A A (identityTrans A)!}) {!!}
k : Natural A A (identityTrans A) Natural A B coe𝟙
k n {a} {b} f = res
where
res : (𝔻 [ coe𝟙 b A.fmap f ]) (𝔻 [ B.fmap f coe𝟙 a ])
res = {!!}
nat : Natural A B coe𝟙
nat = nat'
fromEq : NaturalTransformation A B
fromEq = coe𝟙 , nat
module _ {A B : Functor 𝔻} where
obverse : A B A B
obverse p = res
where
ob : Arrow A B
ob = fromEq p
re : Arrow B A
re = fromEq (sym p)
vr : _∘_ {A = A} {B} {A} re ob 𝟙 {A}
vr = {!!}
rv : _∘_ {A = B} {A} {B} ob re 𝟙 {B}
rv = {!!}
isInverse : IsInverseOf {A} {B} ob re
isInverse = vr , rv
iso : Isomorphism {A} {B} ob
iso = re , isInverse
res : A B
res = ob , iso
reverse : A B A B
reverse iso = {!!}
ve-re : (y : A B) obverse (reverse y) y
ve-re = {!!}
re-ve : (x : A B) reverse (obverse x) x
re-ve = {!!}
done : isEquiv (A B) (A B) (id-to-iso (λ { {A} {B} isIdentity {A} {B}}) A B)
done = {!gradLemma obverse reverse ve-re re-ve!}
univalent : Univalent (λ{ {A} {B} isIdentity {A} {B}})
univalent = done
instance instance
:isCategory: : IsCategory RawFun isCategory : IsCategory RawFun
:isCategory: = record isCategory = record
{ isAssociative = λ {A B C D} :isAssociative: {A} {B} {C} {D} { isAssociative = λ {A B C D} isAssociative {A} {B} {C} {D}
; isIdentity = λ {A B} isIdentity {A} {B} ; isIdentity = λ {A B} isIdentity {A} {B}
; arrowsAreSets = λ {F} {G} naturalTransformationIsSets {F} {G} ; arrowsAreSets = λ {F} {G} naturalTransformationIsSet {F} {G}
; univalent = {!!} ; univalent = univalent
} }
Fun : Category (c c' d d') (c c' d') Fun : Category (c c' d d') (c c' d')
Category.raw Fun = RawFun Category.raw Fun = RawFun
module _ { ' : Level} ( : Category ') where module _ { ' : Level} ( : Category ') where
open import Cat.Categories.Sets private
open NaturalTransformation (opposite ) (𝓢𝓮𝓽 ') open import Cat.Categories.Sets
open NaturalTransformation (opposite ) (𝓢𝓮𝓽 ')
-- Restrict the functors to Presheafs. -- Restrict the functors to Presheafs.
RawPresh : RawCategory ( lsuc ') ( ') rawPresh : RawCategory ( lsuc ') ( ')
RawPresh = record rawPresh = record
{ Object = Presheaf { Object = Presheaf
; Arrow = NaturalTransformation ; Arrow = NaturalTransformation
; 𝟙 = λ {F} identity F ; 𝟙 = λ {F} identity F
; _∘_ = λ {F G H} NT[_∘_] {F = F} {G = G} {H = H} ; _∘_ = λ {F G H} NT[_∘_] {F = F} {G = G} {H = H}
} }
instance
isCategory : IsCategory rawPresh
isCategory = Fun.isCategory _ _
Presh : Category ( lsuc ') ( ')
Category.raw Presh = rawPresh
Category.isCategory Presh = isCategory

View file

@ -56,7 +56,6 @@ module _ {A B : Set} {S : Subset (A × B)} (ab : A × B) where
backwards (a' , (a=a' , a'b∈S)) = subst (sym a=a') a'b∈S backwards (a' , (a=a' , a'b∈S)) = subst (sym a=a') a'b∈S
fwd-bwd : (x : (a , b) S) (backwards forwards) x x fwd-bwd : (x : (a , b) S) (backwards forwards) x x
-- isbijective x = pathJ (λ y x₁ → (backwards ∘ forwards) x ≡ x) {!!} {!!} {!!}
fwd-bwd x = pathJprop (λ y _ y) x fwd-bwd x = pathJprop (λ y _ y) x
bwd-fwd : (x : Σ[ a' A ] (a , a') Diag A × (a' , b) S) bwd-fwd : (x : Σ[ a' A ] (a , a') Diag A × (a' , b) S)

View file

@ -1,35 +1,172 @@
-- | The category of homotopy sets
{-# OPTIONS --allow-unsolved-metas --cubical #-} {-# OPTIONS --allow-unsolved-metas --cubical #-}
module Cat.Categories.Sets where module Cat.Categories.Sets where
open import Cubical
open import Agda.Primitive open import Agda.Primitive
open import Data.Product open import Data.Product
import Function open import Function using (_∘_)
open import Cubical hiding (_≃_ ; inverse)
open import Cubical.Equivalence
renaming
( _≅_ to _A≅_ )
using
(_≃_ ; con ; AreInverses)
open import Cubical.Univalence
open import Cubical.GradLemma
open import Cat.Category open import Cat.Category
open import Cat.Category.Functor open import Cat.Category.Functor
open import Cat.Category.Product open import Cat.Category.Product
open import Cat.Wishlist
module _ ( : Level) where module _ ( : Level) where
private private
open RawCategory
open IsCategory
open import Cubical.Univalence open import Cubical.Univalence
open import Cubical.NType.Properties open import Cubical.NType.Properties
open import Cubical.Universe open import Cubical.Universe
SetsRaw : RawCategory (lsuc ) SetsRaw : RawCategory (lsuc )
Object SetsRaw = hSet RawCategory.Object SetsRaw = hSet
Arrow SetsRaw (T , _) (U , _) = T U RawCategory.Arrow SetsRaw (T , _) (U , _) = T U
𝟙 SetsRaw = Function.id RawCategory.𝟙 SetsRaw = Function.id
_∘_ SetsRaw = Function._∘_ RawCategory._∘_ SetsRaw = Function._∘_
open RawCategory SetsRaw hiding (_∘_)
open Univalence SetsRaw
isIdentity : IsIdentity Function.id
proj₁ isIdentity = funExt λ _ refl
proj₂ isIdentity = funExt λ _ refl
arrowsAreSets : ArrowsAreSets
arrowsAreSets {B = (_ , s)} = setPi λ _ s
module _ {hA hB : Object} where
private
A = proj₁ hA
isSetA : isSet A
isSetA = proj₂ hA
B = proj₁ hB
isSetB : isSet B
isSetB = proj₂ hB
toIsomorphism : A B hA hB
toIsomorphism e = obverse , inverse , verso-recto , recto-verso
where
open _≃_ e
fromIsomorphism : hA hB A B
fromIsomorphism iso = con obverse (gradLemma obverse inverse recto-verso verso-recto)
where
obverse : A B
obverse = proj₁ iso
inverse : B A
inverse = proj₁ (proj₂ iso)
-- FIXME IsInverseOf should change name to AreInverses and the
-- ordering should be swapped.
areInverses : IsInverseOf {A = hA} {hB} obverse inverse
areInverses = proj₂ (proj₂ iso)
verso-recto : a (inverse obverse) a a
verso-recto a i = proj₁ areInverses i a
recto-verso : b (obverse Function.∘ inverse) b b
recto-verso b i = proj₂ areInverses i b
private
univIso : (A B) A≅ (A B)
univIso = _≃_.toIsomorphism univalence
obverse' : A B A B
obverse' = proj₁ univIso
inverse' : A B A B
inverse' = proj₁ (proj₂ univIso)
-- Drop proof of being a set from both sides of an equality.
dropP : hA hB A B
dropP eq i = proj₁ (eq i)
-- Add proof of being a set to both sides of a set-theoretic equivalence
-- returning a category-theoretic equivalence.
addE : A A≅ B hA hB
addE eqv = proj₁ eqv , (proj₁ (proj₂ eqv)) , asPair
where
areeqv = proj₂ (proj₂ eqv)
asPair =
let module Inv = AreInverses areeqv
in Inv.verso-recto , Inv.recto-verso
obverse : hA hB hA hB
obverse = addE _≃_.toIsomorphism obverse' dropP
-- Drop proof of being a set form both sides of a category-theoretic
-- equivalence returning a set-theoretic equivalence.
dropE : hA hB A A≅ B
dropE eqv = obv , inv , asAreInverses
where
obv = proj₁ eqv
inv = proj₁ (proj₂ eqv)
areEq = proj₂ (proj₂ eqv)
asAreInverses : AreInverses A B obv inv
asAreInverses = record { verso-recto = proj₁ areEq ; recto-verso = proj₂ areEq }
-- Dunno if this is a thing.
isoToEquiv : A A≅ B A B
isoToEquiv = {!!}
-- Add proof of being a set to both sides of an equality.
addP : A B hA hB
addP p = lemSig (λ X propPi λ x propPi (λ y propIsProp)) hA hB p
inverse : hA hB hA hB
inverse = addP inverse' isoToEquiv dropE
-- open AreInverses (proj₂ (proj₂ univIso)) renaming
-- ( verso-recto to verso-recto'
-- ; recto-verso to recto-verso'
-- )
-- I can just open them but I wanna be able to see the type annotations.
verso-recto' : inverse' obverse' Function.id
verso-recto' = AreInverses.verso-recto (proj₂ (proj₂ univIso))
recto-verso' : obverse' inverse' Function.id
recto-verso' = AreInverses.recto-verso (proj₂ (proj₂ univIso))
verso-recto : (iso : hA hB) obverse (inverse iso) iso
verso-recto iso = begin
obverse (inverse iso) ≡⟨⟩
( addE _≃_.toIsomorphism
obverse' dropP addP
inverse' isoToEquiv
dropE) iso
≡⟨⟩
( addE _≃_.toIsomorphism
obverse'
inverse' isoToEquiv
dropE) iso
≡⟨ {!!} -- obverse' inverse' are inverses
( addE _≃_.toIsomorphism isoToEquiv dropE) iso
≡⟨ {!!} -- should be easy to prove
-- _≃_.toIsomorphism ∘ isoToEquiv ≡ id
(addE dropE) iso
≡⟨⟩
iso
-- Similar to above.
recto-verso : (eq : hA hB) inverse (obverse eq) eq
recto-verso eq = begin
inverse (obverse eq) ≡⟨ {!!}
eq
-- Use the fact that being an h-level is a mere proposition.
-- This is almost provable using `Wishlist.isSetIsProp` - although
-- this creates homogenous paths.
isSetEq : (p : A B) (λ i isSet (p i)) [ isSetA isSetB ]
isSetEq = {!!}
res : hA hB
proj₁ (res i) = {!!}
proj₂ (res i) = isSetEq {!!} i
univalent : isEquiv (hA hB) (hA hB) (id-to-iso (λ {A} {B} isIdentity {A} {B}) hA hB)
univalent = {!gradLemma obverse inverse verso-recto recto-verso!}
SetsIsCategory : IsCategory SetsRaw SetsIsCategory : IsCategory SetsRaw
isAssociative SetsIsCategory = refl IsCategory.isAssociative SetsIsCategory = refl
proj₁ (isIdentity SetsIsCategory) = funExt λ _ refl IsCategory.isIdentity SetsIsCategory {A} {B} = isIdentity {A} {B}
proj₂ (isIdentity SetsIsCategory) = funExt λ _ refl IsCategory.arrowsAreSets SetsIsCategory {A} {B} = arrowsAreSets {A} {B}
arrowsAreSets SetsIsCategory {B = (_ , s)} = setPi λ _ s IsCategory.univalent SetsIsCategory = univalent
univalent SetsIsCategory = {!!}
𝓢𝓮𝓽 Sets : Category (lsuc ) 𝓢𝓮𝓽 Sets : Category (lsuc )
Category.raw 𝓢𝓮𝓽 = SetsRaw Category.raw 𝓢𝓮𝓽 = SetsRaw
@ -64,58 +201,57 @@ module _ { : Level} where
lem : proj₁ Function.∘′ (f &&& g) f × proj₂ Function.∘′ (f &&& g) g lem : proj₁ Function.∘′ (f &&& g) f × proj₂ Function.∘′ (f &&& g) g
proj₁ lem = refl proj₁ lem = refl
proj₂ lem = refl proj₂ lem = refl
instance
isProduct : IsProduct 𝓢 {0A} {0B} {0A×0B} proj₁ proj₂
isProduct {X = X} f g = (f &&& g) , lem {0X = X} f g
product : Product { = 𝓢} 0A 0B rawProduct : RawProduct 𝓢 0A 0B
product = record RawProduct.object rawProduct = 0A×0B
{ obj = 0A×0B RawProduct.proj₁ rawProduct = Data.Product.proj₁
; proj₁ = Data.Product.proj₁ RawProduct.proj₂ rawProduct = Data.Product.proj₂
; proj₂ = Data.Product.proj₂
; isProduct = λ { {X} isProduct {X = X}} isProduct : IsProduct 𝓢 _ _ rawProduct
} IsProduct.isProduct isProduct {X = X} f g
= (f &&& g) , lem {0X = X} f g
product : Product 𝓢 0A 0B
Product.raw product = rawProduct
Product.isProduct product = isProduct
instance instance
SetsHasProducts : HasProducts 𝓢 SetsHasProducts : HasProducts 𝓢
SetsHasProducts = record { product = product } SetsHasProducts = record { product = product }
module _ {a b : Level} where module _ {a b : Level} ( : Category a b) where
module _ ( : Category a b) where -- Covariant Presheaf
-- Covariant Presheaf Representable : Set (a lsuc b)
Representable : Set (a lsuc b) Representable = Functor (𝓢𝓮𝓽 b)
Representable = Functor (𝓢𝓮𝓽 b)
-- Contravariant Presheaf -- Contravariant Presheaf
Presheaf : Set (a lsuc b) Presheaf : Set (a lsuc b)
Presheaf = Functor (opposite ) (𝓢𝓮𝓽 b) Presheaf = Functor (opposite ) (𝓢𝓮𝓽 b)
open Category
-- The "co-yoneda" embedding. -- The "co-yoneda" embedding.
representable : { : Category a b} Category.Object Representable representable : Category.Object Representable
representable { = } A = record representable A = record
{ raw = record { raw = record
{ func* = λ B [ A , B ] , arrowsAreSets { omap = λ B [ A , B ] , arrowsAreSets
; func→ = [_∘_] ; fmap = [_∘_]
} }
; isFunctor = record ; isFunctor = record
{ isIdentity = funExt λ _ proj₂ isIdentity { isIdentity = funExt λ _ proj₂ isIdentity
; isDistributive = funExt λ x sym isAssociative ; isDistributive = funExt λ x sym isAssociative
} }
} }
where
open Category
-- Alternate name: `yoneda` -- Alternate name: `yoneda`
presheaf : { : Category a b} Category.Object (opposite ) Presheaf presheaf : Category.Object (opposite ) Presheaf
presheaf { = } B = record presheaf B = record
{ raw = record { raw = record
{ func* = λ A [ A , B ] , arrowsAreSets { omap = λ A [ A , B ] , arrowsAreSets
; func→ = λ f g [ g f ] ; fmap = λ f g [ g f ]
} }
; isFunctor = record ; isFunctor = record
{ isIdentity = funExt λ x proj₁ isIdentity { isIdentity = funExt λ x proj₁ isIdentity
; isDistributive = funExt λ x isAssociative ; isDistributive = funExt λ x isAssociative
} }
} }
where
open Category

View file

@ -24,9 +24,6 @@
-- ------ -- ------
-- --
-- Propositionality for all laws about the category. -- Propositionality for all laws about the category.
--
-- TODO: An equality principle for categories that focuses on the pure data-part.
--
{-# OPTIONS --allow-unsolved-metas --cubical #-} {-# OPTIONS --allow-unsolved-metas --cubical #-}
module Cat.Category where module Cat.Category where
@ -41,7 +38,7 @@ open import Data.Product renaming
open import Data.Empty open import Data.Empty
import Function import Function
open import Cubical open import Cubical
open import Cubical.NType.Properties using ( propIsEquiv ) open import Cubical.NType.Properties using ( propIsEquiv ; lemPropF )
open import Cat.Wishlist open import Cat.Wishlist
@ -76,7 +73,7 @@ record RawCategory (a b : Level) : Set (lsuc (a ⊔ b)) where
𝟙 : {A : Object} Arrow A A 𝟙 : {A : Object} Arrow A A
_∘_ : {A B C : Object} Arrow B C Arrow A B Arrow A C _∘_ : {A B C : Object} Arrow B C Arrow A B Arrow A C
infixl 10 _∘_ infixl 10 _∘_ _>>>_
-- | Operations on data -- | Operations on data
@ -86,9 +83,12 @@ record RawCategory (a b : Level) : Set (lsuc (a ⊔ b)) where
codomain : { a b : Object } Arrow a b Object codomain : { a b : Object } Arrow a b Object
codomain {b = b} _ = b codomain {b = b} _ = b
_>>>_ : {A B C : Object} (Arrow A B) (Arrow B C) Arrow A C
f >>> g = g f
-- | Laws about the data -- | Laws about the data
-- TODO: It seems counter-intuitive that the normal-form is on the -- FIXME It seems counter-intuitive that the normal-form is on the
-- right-hand-side. -- right-hand-side.
IsAssociative : Set (a b) IsAssociative : Set (a b)
IsAssociative = {A B C D} {f : Arrow A B} {g : Arrow B C} {h : Arrow C D} IsAssociative = {A B C D} {f : Arrow A B} {g : Arrow B C} {h : Arrow C D}
@ -129,7 +129,9 @@ record RawCategory (a b : Level) : Set (lsuc (a ⊔ b)) where
Terminal : Set (a b) Terminal : Set (a b)
Terminal = Σ Object IsTerminal Terminal = Σ Object IsTerminal
-- Univalence is indexed by a raw category as well as an identity proof. -- | Univalence is indexed by a raw category as well as an identity proof.
--
-- FIXME Put this in `RawCategory` and index it on the witness to `isIdentity`.
module Univalence {a b : Level} ( : RawCategory a b) where module Univalence {a b : Level} ( : RawCategory a b) where
open RawCategory open RawCategory
module _ (isIdentity : IsIdentity 𝟙) where module _ (isIdentity : IsIdentity 𝟙) where
@ -150,9 +152,11 @@ module Univalence {a b : Level} ( : RawCategory a b) where
-- iso-is-epi : Isomorphism f → Epimorphism {X = X} f -- iso-is-epi : Isomorphism f → Epimorphism {X = X} f
-- iso-is-mono : Isomorphism f → Monomorphism {X = X} f -- iso-is-mono : Isomorphism f → Monomorphism {X = X} f
-- --
-- Sans `univalent` this would be what is referred to as a pre-category in
-- [HoTT].
record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc (a b)) where record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc (a b)) where
open RawCategory public open RawCategory public
open Univalence public open Univalence public
field field
isAssociative : IsAssociative isAssociative : IsAssociative
isIdentity : IsIdentity 𝟙 isIdentity : IsIdentity 𝟙
@ -191,9 +195,9 @@ record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc
-- --
-- Proves that all projections of `IsCategory` are mere propositions as well as -- Proves that all projections of `IsCategory` are mere propositions as well as
-- `IsCategory` itself being a mere proposition. -- `IsCategory` itself being a mere proposition.
module _ {a b : Level} {C : RawCategory a b} where module Propositionality {a b : Level} ( : RawCategory a b) where
open RawCategory C open RawCategory
module _ ( : IsCategory C) where module _ ( : IsCategory ) where
open IsCategory using (isAssociative ; arrowsAreSets ; isIdentity ; Univalent) open IsCategory using (isAssociative ; arrowsAreSets ; isIdentity ; Univalent)
open import Cubical.NType open import Cubical.NType
open import Cubical.NType.Properties open import Cubical.NType.Properties
@ -237,52 +241,71 @@ module _ {a b : Level} {C : RawCategory a b} where
propUnivalent a b i = propPi (λ iso propHasLevel ⟨-2⟩) a b i propUnivalent a b i = propPi (λ iso propHasLevel ⟨-2⟩) a b i
private private
module _ (x y : IsCategory C) where module _ (x y : IsCategory ) where
module IC = IsCategory module IC = IsCategory
module X = IsCategory x module X = IsCategory x
module Y = IsCategory y module Y = IsCategory y
open Univalence C open Univalence
-- In a few places I use the result of propositionality of the various -- In a few places I use the result of propositionality of the various
-- projections of `IsCategory` - I've arbitrarily chosed to use this -- projections of `IsCategory` - I've arbitrarily chosed to use this
-- result from `x : IsCategory C`. I don't know which (if any) possibly -- result from `x : IsCategory C`. I don't know which (if any) possibly
-- adverse effects this may have. -- adverse effects this may have.
isIdentity : (λ _ IsIdentity 𝟙) [ X.isIdentity Y.isIdentity ] isIdentity : (λ _ IsIdentity 𝟙) [ X.isIdentity Y.isIdentity ]
isIdentity = propIsIdentity x X.isIdentity Y.isIdentity isIdentity = propIsIdentity x X.isIdentity Y.isIdentity
done : x y
U : {a : IsIdentity 𝟙} U : {a : IsIdentity 𝟙}
(λ _ IsIdentity 𝟙) [ X.isIdentity a ] (λ _ IsIdentity 𝟙) [ X.isIdentity a ]
(b : Univalent a) (b : Univalent a)
Set _ Set _
U eqwal bbb = U eqwal univ =
(λ i Univalent (eqwal i)) (λ i Univalent (eqwal i))
[ X.univalent bbb ] [ X.univalent univ ]
P : (y : IsIdentity 𝟙) P : (y : IsIdentity 𝟙)
(λ _ IsIdentity 𝟙) [ X.isIdentity y ] Set _ (λ _ IsIdentity 𝟙) [ X.isIdentity y ] Set _
P y eq = (b' : Univalent y) U eq b' P y eq = (univ : Univalent y) U eq univ
helper : (b' : Univalent X.isIdentity) p : (b' : Univalent X.isIdentity)
(λ _ Univalent X.isIdentity) [ X.univalent b' ] (λ _ Univalent X.isIdentity) [ X.univalent b' ]
helper univ = propUnivalent x X.univalent univ p univ = propUnivalent x X.univalent univ
foo = pathJ P helper Y.isIdentity isIdentity helper : P Y.isIdentity isIdentity
helper = pathJ P p Y.isIdentity isIdentity
eqUni : U isIdentity Y.univalent eqUni : U isIdentity Y.univalent
eqUni = foo Y.univalent eqUni = helper Y.univalent
IC.isAssociative (done i) = propIsAssociative x X.isAssociative Y.isAssociative i done : x y
IC.isIdentity (done i) = isIdentity i IC.isAssociative (done i) = propIsAssociative x X.isAssociative Y.isAssociative i
IC.isIdentity (done i) = isIdentity i
IC.arrowsAreSets (done i) = propArrowIsSet x X.arrowsAreSets Y.arrowsAreSets i IC.arrowsAreSets (done i) = propArrowIsSet x X.arrowsAreSets Y.arrowsAreSets i
IC.univalent (done i) = eqUni i IC.univalent (done i) = eqUni i
propIsCategory : isProp (IsCategory C) propIsCategory : isProp (IsCategory )
propIsCategory = done propIsCategory = done
-- | Univalent categories -- | Univalent categories
-- --
-- Just bundles up the data with witnesses inhabting the propositions. -- Just bundles up the data with witnesses inhabiting the propositions.
record Category (a b : Level) : Set (lsuc (a b)) where record Category (a b : Level) : Set (lsuc (a b)) where
field field
raw : RawCategory a b raw : RawCategory a b
{{isCategory}} : IsCategory raw {{isCategory}} : IsCategory raw
open IsCategory isCategory public open IsCategory isCategory public
-- The fact that being a category is a mere proposition gives rise to this
-- equality principle for categories.
module _ {a b : Level} { 𝔻 : Category a b} where
private
module = Category
module 𝔻 = Category 𝔻
module _ (rawEq : .raw 𝔻.raw) where
private
isCategoryEq : (λ i IsCategory (rawEq i)) [ .isCategory 𝔻.isCategory ]
isCategoryEq = lemPropF Propositionality.propIsCategory rawEq
Category≡ : 𝔻
Category≡ i = record
{ raw = rawEq i
; isCategory = isCategoryEq i
}
-- | Syntax for arrows- and composition in a given category. -- | Syntax for arrows- and composition in a given category.
module _ {a b : Level} ( : Category a b) where module _ {a b : Level} ( : Category a b) where
open Category open Category
@ -298,23 +321,42 @@ module _ {a b : Level} ( : Category a b) where
-- flipped. -- flipped.
module Opposite {a b : Level} where module Opposite {a b : Level} where
module _ ( : Category a b) where module _ ( : Category a b) where
open Category
private private
module = Category
opRaw : RawCategory a b opRaw : RawCategory a b
RawCategory.Object opRaw = Object RawCategory.Object opRaw = .Object
RawCategory.Arrow opRaw = Function.flip Arrow RawCategory.Arrow opRaw = Function.flip .Arrow
RawCategory.𝟙 opRaw = 𝟙 RawCategory.𝟙 opRaw = .𝟙
RawCategory._∘_ opRaw = Function.flip _∘_ RawCategory._∘_ opRaw = Function.flip ._∘_
opIsCategory : IsCategory opRaw open RawCategory opRaw
IsCategory.isAssociative opIsCategory = sym isAssociative open Univalence opRaw
IsCategory.isIdentity opIsCategory = swap isIdentity
IsCategory.arrowsAreSets opIsCategory = arrowsAreSets isIdentity : IsIdentity 𝟙
IsCategory.univalent opIsCategory = {!!} isIdentity = swap .isIdentity
module _ {A B : .Object} where
univalent : isEquiv (A B) (A B)
(id-to-iso (swap .isIdentity) A B)
fst (univalent iso) = flipFiber (fst (.univalent (flipIso iso)))
where
flipIso : A B B .≅ A
flipIso (f , f~ , iso) = f , f~ , swap iso
flipFiber
: fiber (.id-to-iso .isIdentity B A) (flipIso iso)
fiber ( id-to-iso isIdentity A B) iso
flipFiber (eq , eqIso) = sym eq , {!!}
snd (univalent iso) = {!!}
isCategory : IsCategory opRaw
IsCategory.isAssociative isCategory = sym .isAssociative
IsCategory.isIdentity isCategory = isIdentity
IsCategory.arrowsAreSets isCategory = .arrowsAreSets
IsCategory.univalent isCategory = univalent
opposite : Category a b opposite : Category a b
raw opposite = opRaw Category.raw opposite = opRaw
Category.isCategory opposite = opIsCategory Category.isCategory opposite = isCategory
-- As demonstrated here a side-effect of having no-eta-equality on constructors -- As demonstrated here a side-effect of having no-eta-equality on constructors
-- means that we need to pick things apart to show that things are indeed -- means that we need to pick things apart to show that things are indeed
@ -331,9 +373,7 @@ module Opposite {a b : Level} where
RawCategory.𝟙 (rawInv _) = 𝟙 RawCategory.𝟙 (rawInv _) = 𝟙
RawCategory._∘_ (rawInv _) = _∘_ RawCategory._∘_ (rawInv _) = _∘_
-- TODO: Define and use Monad≡
oppositeIsInvolution : opposite (opposite ) oppositeIsInvolution : opposite (opposite )
Category.raw (oppositeIsInvolution i) = rawInv i oppositeIsInvolution = Category≡ rawInv
Category.isCategory (oppositeIsInvolution x) = {!!}
open Opposite public open Opposite public

View file

@ -1,40 +1,44 @@
module Cat.Category.Exponential where module Cat.Category.Exponential where
open import Agda.Primitive open import Agda.Primitive
open import Data.Product open import Data.Product hiding (_×_)
open import Cubical open import Cubical
open import Cat.Category open import Cat.Category
open import Cat.Category.Product open import Cat.Category.Product
open Category
module _ { '} ( : Category ') {{hasProducts : HasProducts }} where module _ { '} ( : Category ') {{hasProducts : HasProducts }} where
open HasProducts hasProducts open Category
open Product hiding (obj) open HasProducts hasProducts public
private
_×p_ : (A B : Object ) Object
_×p_ A B = Product.obj (product A B)
module _ (B C : Object ) where module _ (B C : Object) where
IsExponential : (Cᴮ : Object ) [ Cᴮ ×p B , C ] Set ( ') record IsExponential'
IsExponential Cᴮ eval = (A : Object ) (f : [ A ×p B , C ]) (Cᴮ : Object)
(eval : [ Cᴮ × B , C ]) : Set ( ') where
field
uniq
: (A : Object) (f : [ A × B , C ])
∃![ f~ ] ( [ eval f~ |×| Category.𝟙 ] f)
IsExponential : (Cᴮ : Object) [ Cᴮ × B , C ] Set ( ')
IsExponential Cᴮ eval = (A : Object) (f : [ A × B , C ])
∃![ f~ ] ( [ eval f~ |×| Category.𝟙 ] f) ∃![ f~ ] ( [ eval f~ |×| Category.𝟙 ] f)
record Exponential : Set ( ') where record Exponential : Set ( ') where
field field
-- obj ≡ Cᴮ -- obj ≡ Cᴮ
obj : Object obj : Object
eval : [ obj ×p B , C ] eval : [ obj × B , C ]
{{isExponential}} : IsExponential obj eval {{isExponential}} : IsExponential obj eval
-- If I make this an instance-argument then the instance resolution
-- algorithm goes into an infinite loop. Why? transpose : (A : Object) [ A × B , C ] [ A , obj ]
exponentialsHaveProducts : HasProducts
exponentialsHaveProducts = hasProducts
transpose : (A : Object ) [ A ×p B , C ] [ A , obj ]
transpose A f = proj₁ (isExponential A f) transpose A f = proj₁ (isExponential A f)
record HasExponentials { ' : Level} ( : Category ') {{_ : HasProducts }} : Set ( ') where record HasExponentials { ' : Level} ( : Category ') {{_ : HasProducts }} : Set ( ') where
open Category
open Exponential public open Exponential public
field field
exponent : (A B : Object ) Exponential A B exponent : (A B : Object) Exponential A B
_⇑_ : (A B : Object) Object
A B = (exponent A B) .obj

View file

@ -18,21 +18,50 @@ 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 𝔻 omap : Object Object 𝔻
func→ : {A B} [ A , B ] 𝔻 [ func* A , func* B ] fmap : {A B} [ A , B ] 𝔻 [ omap A , omap B ]
IsIdentity : Set _ IsIdentity : Set _
IsIdentity = {A : Object } func→ (𝟙 {A}) 𝟙 𝔻 {func* A} IsIdentity = {A : Object } fmap (𝟙 {A}) 𝟙 𝔻 {omap A}
IsDistributive : Set _ IsDistributive : Set _
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 ] fmap ( [ g f ]) 𝔻 [ fmap g fmap f ]
-- | Equality principle for raw functors
--
-- The type of `fmap` depend on the value of `omap`. 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 : omap x omap') Set _
P y eq = (fmap' : Fmap y) (λ i Fmap (eq i))
[ fmap x fmap' ]
module _
(eq : (λ i Omap) [ omap x omap y ])
(kk : P (omap x) refl)
where
private
p : P (omap y) eq
p = pathJ P kk (omap y) eq
eq→ : (λ i Fmap (eq i)) [ fmap x fmap y ]
eq→ = p (fmap y)
RawFunctor≡ : x y
omap (RawFunctor≡ i) = eq i
fmap (RawFunctor≡ i) = eq→ i
record IsFunctor (F : RawFunctor) : 𝓤 where record IsFunctor (F : RawFunctor) : 𝓤 where
open RawFunctor F public open RawFunctor F public
field field
-- FIXME Really ought to be preserves identity or something like this.
isIdentity : IsIdentity isIdentity : IsIdentity
isDistributive : IsDistributive isDistributive : IsDistributive
@ -45,6 +74,9 @@ module _ {c c' d d'}
open Functor open Functor
EndoFunctor : {a b} ( : Category a b) Set _
EndoFunctor = Functor
module _ module _
{a b : Level} {a b : Level}
{ 𝔻 : Category a b} { 𝔻 : Category a b}
@ -81,26 +113,21 @@ module _
module _ { ' : Level} { 𝔻 : Category '} where module _ { ' : Level} { 𝔻 : Category '} where
Functor≡ : {F G : Functor 𝔻} Functor≡ : {F G : Functor 𝔻}
(eq* : func* F func* G) raw F raw G
(eq→ : (λ i {x y} [ x , y ] 𝔻 [ eq* i x , eq* i y ])
[ func→ F func→ G ])
F G F G
Functor≡ {F} {G} eq* eq→ i = record raw (Functor≡ eq i) = eq i
{ raw = eqR i isFunctor (Functor≡ {F} {G} eq i)
; isFunctor = eqIsF i = res i
}
where where
eqR : raw F raw G res : (λ i IsFunctor 𝔻 (eq i)) [ isFunctor F isFunctor G ]
eqR i = record { func* = eq* i ; func→ = eq→ i } res = IsFunctorIsProp' (isFunctor F) (isFunctor G)
eqIsF : (λ i IsFunctor 𝔻 (eqR i)) [ isFunctor F isFunctor G ]
eqIsF = 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* = omap F
F→ = func→ F F→ = fmap F
G* = func* G G* = omap G
G→ = func→ G G→ = fmap G
module _ {a0 a1 a2 : Object A} {α0 : A [ a0 , a1 ]} {α1 : A [ a1 , a2 ]} where module _ {a0 a1 a2 : Object A} {α0 : A [ a0 , a1 ]} {α1 : A [ a1 , a2 ]} where
dist : (F→ G→) (A [ α1 α0 ]) C [ (F→ G→) α1 (F→ G→) α0 ] dist : (F→ G→) (A [ α1 α0 ]) C [ (F→ G→) α1 (F→ G→) α0 ]
@ -111,8 +138,8 @@ module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : F
C [ (F→ G→) α1 (F→ G→) α0 ] C [ (F→ G→) α1 (F→ G→) α0 ]
_∘fr_ : RawFunctor A C _∘fr_ : RawFunctor A C
RawFunctor.func* _∘fr_ = F* G* RawFunctor.omap _∘fr_ = F* G*
RawFunctor.func→ _∘fr_ = F→ G→ RawFunctor.fmap _∘fr_ = F→ G→
instance instance
isFunctor' : IsFunctor A C _∘fr_ isFunctor' : IsFunctor A C _∘fr_
isFunctor' = record isFunctor' = record
@ -131,8 +158,8 @@ module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : F
identity : { '} {C : Category '} Functor C C identity : { '} {C : Category '} Functor C C
identity = record identity = record
{ raw = record { raw = record
{ func* = λ x x { omap = λ x x
; func→ = λ x x ; fmap = λ x x
} }
; isFunctor = record ; isFunctor = record
{ isIdentity = refl { isIdentity = refl

View file

@ -1,3 +1,22 @@
{---
Monads
This module presents two formulations of monads:
* The standard monoidal presentation
* Kleisli's presentation
The first one defines a monad in terms of an endofunctor and two natural
transformations. The second defines it in terms of a function on objects and a
pair of arrows.
These two formulations are proven to be equivalent:
Monoidal.Monad Kleisli.Monad
The monoidal representation is exposed by default from this module.
---}
{-# OPTIONS --cubical --allow-unsolved-metas #-} {-# OPTIONS --cubical --allow-unsolved-metas #-}
module Cat.Category.Monad where module Cat.Category.Monad where
@ -6,326 +25,185 @@ open import Agda.Primitive
open import Data.Product open import Data.Product
open import Cubical 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
open import Cat.Category.Functor as F open import Cat.Category.Functor as F
open import Cat.Category.NaturalTransformation open import Cat.Category.NaturalTransformation
open import Cat.Category.Monad.Monoidal as Monoidal public
open import Cat.Category.Monad.Kleisli as Kleisli
open import Cat.Categories.Fun open import Cat.Categories.Fun
-- "A monad in the monoidal form" [voe] -- | The monoidal- and kleisli presentation of monads are equivalent.
module Monoidal {a b : Level} ( : Category a b) where
private
= a b
open Category using (Object ; Arrow ; 𝟙 ; _∘_)
open NaturalTransformation
record RawMonad : Set where
field
-- R ~ m
R : Functor
-- η ~ pure
ηNat : NaturalTransformation F.identity R
-- μ ~ join
μNat : NaturalTransformation F[ R R ] R
η : Transformation F.identity R
η = proj₁ ηNat
μ : Transformation F[ R R ] R
μ = proj₁ μNat
private
module R = Functor R
module RR = Functor F[ R R ]
module _ {X : Object} where
IsAssociative' : Set _
IsAssociative' = μ X R.func→ (μ X) μ X μ (R.func* X)
IsInverse' : Set _
IsInverse'
= μ X η (R.func* X) 𝟙
× μ X R.func→ (η X) 𝟙
-- 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
record Monad : Set where
field
raw : RawMonad
isMonad : IsMonad raw
open IsMonad isMonad public
postulate propIsMonad : {raw} isProp (IsMonad raw)
Monad≡ : {m n : Monad} Monad.raw m Monad.raw n m n
Monad.raw (Monad≡ eq i) = eq i
Monad.isMonad (Monad≡ {m} {n} eq i) = res i
where
-- TODO: PathJ nightmare + `propIsMonad`.
res : (λ i IsMonad (eq i)) [ Monad.isMonad m Monad.isMonad n ]
res = {!!}
-- "A monad in the Kleisli form" [voe]
module Kleisli {a b : Level} ( : Category a b) where
private
= a b
open Category using (Arrow ; 𝟙 ; Object ; _∘_)
record RawMonad : Set where
field
RR : Object Object
-- Note name-change from [voe]
ζ : {X : Object} [ X , RR X ]
rr : {X Y : Object} [ X , RR Y ] [ RR X , RR Y ]
-- Note the correspondance with Haskell:
--
-- RR ~ m
-- ζ ~ pure
-- rr ~ flip (>>=)
--
-- Where those things have these types:
--
-- m : 𝓤𝓤
-- pure : x → m x
-- flip (>>=) :: (a → m b) → m a → m b
--
pure : {X : Object} [ X , RR X ]
pure = ζ
fmap : {A B} [ A , B ] [ RR A , RR B ]
fmap f = rr (ζ f)
-- Why is (>>=) not implementable?
--
-- (>>=) : m a -> (a -> m b) -> m b
-- (>=>) : (a -> m b) -> (b -> m c) -> a -> m c
_>=>_ : {A B C : Object} [ A , RR B ] [ B , RR C ] [ A , RR C ]
f >=> g = rr g f
-- fmap id ≡ id
IsIdentity = {X : Object}
rr ζ 𝟙 {RR X}
IsNatural = {X Y : Object} (f : [ X , RR Y ])
rr f ζ f
IsDistributive = {X Y Z : Object} (g : [ Y , RR Z ]) (f : [ X , RR Y ])
rr g rr f rr (rr g f)
Fusion = {X Y Z : Object} {g : [ Y , Z ]} {f : [ X , Y ]}
fmap (g f) fmap g fmap f
record IsMonad (raw : RawMonad) : Set where
open RawMonad raw public
field
isIdentity : IsIdentity
isNatural : IsNatural
isDistributive : IsDistributive
fusion : Fusion
fusion {g = g} {f} = begin
fmap (g f) ≡⟨⟩
rr (ζ (g f)) ≡⟨ {!!}
rr (rr (ζ g) (ζ f)) ≡⟨ sym lem
rr (ζ g) rr (ζ f) ≡⟨⟩
fmap g fmap f
where
lem : rr (ζ g) rr (ζ f) rr (rr (ζ g) (ζ f))
lem = isDistributive (ζ g) (ζ f)
record Monad : Set where
field
raw : RawMonad
isMonad : IsMonad raw
open IsMonad isMonad public
postulate propIsMonad : {raw} isProp (IsMonad raw)
Monad≡ : {m n : Monad} Monad.raw m Monad.raw n m n
Monad.raw (Monad≡ eq i) = eq i
Monad.isMonad (Monad≡ {m} {n} eq i) = res i
where
-- TODO: PathJ nightmare + `propIsMonad`.
res : (λ i IsMonad (eq i)) [ Monad.isMonad m Monad.isMonad n ]
res = {!!}
-- Problem 2.3
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 Functor using (func* ; func→) open using (Object ; Arrow ; 𝟙 ; _∘_ ; _>>>_)
module M = Monoidal module M = Monoidal
module K = Kleisli module K = Kleisli
-- Note similarity with locally defined things in Kleisly.RawMonad!!
module _ (m : M.RawMonad) where module _ (m : M.RawMonad) where
private open M.RawMonad m
open M.RawMonad m
module Kraw = K.RawMonad
RR : Object Object
RR = func* R
ζ : {X : Object} [ X , RR X ]
ζ {X} = η X
rr : {X Y : Object} [ X , RR Y ] [ RR X , RR Y ]
rr {X} {Y} f = μ Y func→ R f
forthRaw : K.RawMonad forthRaw : K.RawMonad
Kraw.RR forthRaw = RR K.RawMonad.omap forthRaw = Romap
Kraw.ζ forthRaw = ζ K.RawMonad.pure forthRaw = pureT _
Kraw.rr forthRaw = rr K.RawMonad.bind forthRaw = bind
module _ {raw : M.RawMonad} (m : M.IsMonad raw) where module _ {raw : M.RawMonad} (m : M.IsMonad raw) where
private private
open M.IsMonad m module MI = M.IsMonad m
open K.RawMonad (forthRaw raw)
module Kis = K.IsMonad
isIdentity : IsIdentity
isIdentity {X} = begin
rr ζ ≡⟨⟩
rr (η X) ≡⟨⟩
μ X func→ R (η X) ≡⟨ proj₂ isInverse
𝟙
module R = Functor R
isNatural : IsNatural
isNatural {X} {Y} f = begin
rr f ζ ≡⟨⟩
rr f η X ≡⟨⟩
μ Y R.func→ f η X ≡⟨ sym .isAssociative
μ Y (R.func→ f η X) ≡⟨ cong (λ φ μ Y φ) (sym (ηN f))
μ Y (η (R.func* Y) f) ≡⟨ .isAssociative
μ Y η (R.func* Y) f ≡⟨ cong (λ φ φ f) (proj₁ isInverse)
𝟙 f ≡⟨ proj₂ .isIdentity
f
where
open NaturalTransformation
module = Category
ηN : Natural F.identity R η
ηN = proj₂ ηNat
isDistributive : IsDistributive
isDistributive {X} {Y} {Z} g f = begin
rr g rr f ≡⟨⟩
μ Z R.func→ g (μ Y R.func→ f) ≡⟨ sym lem2
μ Z R.func→ (μ Z R.func→ g f) ≡⟨⟩
μ Z R.func→ (rr g f)
where
-- Proved it in reverse here... otherwise it could be neatly inlined.
lem2
: μ Z R.func→ (μ Z R.func→ g f)
μ Z R.func→ g (μ Y R.func→ f)
lem2 = 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) RR.func→ g R.func→ f) ≡⟨ {!!} -- ●-solver?
(μ Z R.func→ (μ Z)) (RR.func→ g R.func→ f) ≡⟨ cong (λ φ φ (RR.func→ g R.func→ f)) lemmm
(μ Z μ (R.func* Z)) (RR.func→ g R.func→ f) ≡⟨ {!!} -- ●-solver?
μ Z μ (R.func* Z) RR.func→ g R.func→ f ≡⟨ {!!} -- ●-solver + lem4
μ Z R.func→ g μ Y R.func→ f ≡⟨ sym (Category.isAssociative )
μ Z R.func→ g (μ Y R.func→ f)
where
module RR = 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 = {!!}
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 = {!!}
μN = proj₂ μNat
lemmm : μ Z R.func→ (μ Z) μ Z μ (R.func* Z)
lemmm = isAssociative
lem4 : μ (R.func* Z) RR.func→ g R.func→ g μ Y
lem4 = μN g
forthIsMonad : K.IsMonad (forthRaw raw) forthIsMonad : K.IsMonad (forthRaw raw)
Kis.isIdentity forthIsMonad = isIdentity K.IsMonad.isIdentity forthIsMonad = proj₂ MI.isInverse
Kis.isNatural forthIsMonad = isNatural K.IsMonad.isNatural forthIsMonad = MI.isNatural
Kis.isDistributive forthIsMonad = isDistributive K.IsMonad.isDistributive forthIsMonad = MI.isDistributive
forth : M.Monad K.Monad forth : M.Monad K.Monad
Kleisli.Monad.raw (forth m) = forthRaw (M.Monad.raw m) Kleisli.Monad.raw (forth m) = forthRaw (M.Monad.raw m)
Kleisli.Monad.isMonad (forth m) = forthIsMonad (M.Monad.isMonad m) Kleisli.Monad.isMonad (forth m) = forthIsMonad (M.Monad.isMonad m)
module _ (m : K.Monad) where module _ (m : K.Monad) where
private open K.Monad m
module = Category
open K.Monad m
module Mraw = M.RawMonad
open NaturalTransformation
rawR : RawFunctor
RawFunctor.func* rawR = RR
RawFunctor.func→ rawR f = rr (ζ f)
isFunctorR : IsFunctor rawR
IsFunctor.isIdentity isFunctorR = begin
rr (ζ 𝟙) ≡⟨ cong rr (proj₁ .isIdentity)
rr ζ ≡⟨ isIdentity
𝟙
IsFunctor.isDistributive isFunctorR {f = f} {g} = begin
rr (ζ (g f)) ≡⟨⟩
fmap (g f) ≡⟨ fusion
fmap g fmap f ≡⟨⟩
rr (ζ g) rr (ζ f)
R : Functor
Functor.raw R = rawR
Functor.isFunctor R = isFunctorR
R2 : Functor
R2 = F[ R R ]
ηNat : NaturalTransformation F.identity R
ηNat = {!!}
μNat : NaturalTransformation R2 R
μNat = {!!}
backRaw : M.RawMonad backRaw : M.RawMonad
Mraw.R backRaw = R M.RawMonad.R backRaw = R
Mraw.ηNat backRaw = ηNat M.RawMonad.pureNT backRaw = pureNT
Mraw.μNat backRaw = μNat M.RawMonad.joinNT backRaw = joinNT
module _ (m : K.Monad) where private
open K.Monad m open M.RawMonad backRaw
open M.RawMonad (backRaw m) module R = Functor (M.RawMonad.R backRaw)
module Mis = M.IsMonad
backIsMonad : M.IsMonad (backRaw m) backIsMonad : M.IsMonad backRaw
backIsMonad = {!!} M.IsMonad.isAssociative backIsMonad {X} = begin
joinT X R.fmap (joinT X) ≡⟨⟩
join fmap (joinT X) ≡⟨⟩
join fmap join ≡⟨ isNaturalForeign
join join ≡⟨⟩
joinT X joinT (R.omap X)
M.IsMonad.isInverse backIsMonad {X} = inv-l , inv-r
where
inv-l = begin
joinT X pureT (R.omap X) ≡⟨⟩
join pure ≡⟨ proj₁ isInverse
𝟙
inv-r = begin
joinT X R.fmap (pureT X) ≡⟨⟩
join fmap pure ≡⟨ proj₂ isInverse
𝟙
back : K.Monad M.Monad back : K.Monad M.Monad
Monoidal.Monad.raw (back m) = backRaw m Monoidal.Monad.raw (back m) = backRaw m
Monoidal.Monad.isMonad (back m) = backIsMonad m Monoidal.Monad.isMonad (back m) = backIsMonad m
-- 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) private
open K.Monad 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 join fmap f) ≡⟨⟩
(λ f bind (f >>> pure) >>> bind 𝟙) ≡⟨ funExt lem
(λ f bind f) ≡⟨⟩
bind
where
lem : (f : Arrow X (omap 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
forthRawEq : forthRaw (backRaw m) K.Monad.raw m forthRawEq : forthRaw (backRaw m) K.Monad.raw m
K.RawMonad.RR (forthRawEq _) = RR K.RawMonad.omap (forthRawEq _) = omap
K.RawMonad.ζ (forthRawEq _) = ζ K.RawMonad.pure (forthRawEq _) = pure
-- stuck K.RawMonad.bind (forthRawEq i) = bindEq i
K.RawMonad.rr (forthRawEq 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) private
open M.Monad m
module KM = K.Monad (forth m)
module R = Functor R
omapEq : KM.omap Romap
omapEq = refl
bindEq : {X Y} {f : Arrow X (Romap Y)} KM.bind f bind f
bindEq {X} {Y} {f} = begin
KM.bind f ≡⟨⟩
joinT Y Rfmap f ≡⟨⟩
bind f
joinEq : {X} KM.join joinT X
joinEq {X} = begin
KM.join ≡⟨⟩
KM.bind 𝟙 ≡⟨⟩
bind 𝟙 ≡⟨⟩
joinT X Rfmap 𝟙 ≡⟨ cong (λ φ _ φ) R.isIdentity
joinT X 𝟙 ≡⟨ proj₁ .isIdentity
joinT X
fmapEq : {A B} KM.fmap {A} {B} Rfmap
fmapEq {A} {B} = funExt (λ f begin
KM.fmap f ≡⟨⟩
KM.bind (f >>> KM.pure) ≡⟨⟩
bind (f >>> pureT _) ≡⟨⟩
Rfmap (f >>> pureT B) >>> joinT B ≡⟨⟩
Rfmap (f >>> pureT B) >>> joinT B ≡⟨ cong (λ φ φ >>> joinT B) R.isDistributive
Rfmap f >>> Rfmap (pureT B) >>> joinT B ≡⟨ .isAssociative
joinT B Rfmap (pureT B) Rfmap f ≡⟨ cong (λ φ φ Rfmap f) (proj₂ isInverse)
𝟙 Rfmap f ≡⟨ proj₂ .isIdentity
Rfmap f
)
rawEq : Functor.raw KM.R Functor.raw R
RawFunctor.omap (rawEq i) = omapEq i
RawFunctor.fmap (rawEq i) = fmapEq i
Req : M.RawMonad.R (backRaw (forth m)) R
Req = Functor≡ rawEq
open NaturalTransformation
pureTEq : M.RawMonad.pureT (backRaw (forth m)) pureT
pureTEq = funExt (λ X refl)
pureNTEq : (λ i NaturalTransformation F.identity (Req i))
[ M.RawMonad.pureNT (backRaw (forth m)) pureNT ]
pureNTEq = lemSigP (λ i propIsNatural F.identity (Req i)) _ _ pureTEq
joinTEq : M.RawMonad.joinT (backRaw (forth m)) joinT
joinTEq = funExt (λ X begin
M.RawMonad.joinT (backRaw (forth m)) X ≡⟨⟩
KM.join ≡⟨⟩
joinT X Rfmap 𝟙 ≡⟨ cong (λ φ joinT X φ) R.isIdentity
joinT X 𝟙 ≡⟨ proj₁ .isIdentity
joinT X )
joinNTEq : (λ i NaturalTransformation F[ Req i Req i ] (Req i))
[ M.RawMonad.joinNT (backRaw (forth m)) joinNT ]
joinNTEq = lemSigP (λ i propIsNatural F[ Req i Req i ] (Req i)) _ _ joinTEq
backRawEq : backRaw (forth m) M.Monad.raw m backRawEq : backRaw (forth m) M.Monad.raw m
-- stuck M.RawMonad.R (backRawEq i) = Req i
M.RawMonad.R (backRawEq i) = {!!} M.RawMonad.pureNT (backRawEq i) = pureNTEq i
M.RawMonad.ηNat (backRawEq i) = {!!} M.RawMonad.joinNT (backRawEq i) = joinNTEq i
M.RawMonad.μNat (backRawEq i) = {!!}
backeq : (m : M.Monad) back (forth m) m backeq : (m : M.Monad) back (forth m) m
backeq m = M.Monad≡ (backRawEq m) backeq m = M.Monad≡ (backRawEq m)
open import Cubical.GradLemma
eqv : isEquiv M.Monad K.Monad forth eqv : isEquiv M.Monad K.Monad forth
eqv = gradLemma forth back fortheq backeq eqv = gradLemma forth back fortheq backeq

View file

@ -0,0 +1,253 @@
{---
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 (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)
-- 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 (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.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 (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
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

View file

@ -0,0 +1,154 @@
{---
Monoidal 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
module Cat.Category.Monad.Monoidal {a b : Level} ( : Category a b) where
-- "A monad in the monoidal form" [voe]
private
= a b
open Category using (Object ; Arrow ; 𝟙 ; _∘_)
open NaturalTransformation
record RawMonad : Set where
field
R : EndoFunctor
pureNT : NaturalTransformation F.identity R
joinNT : NaturalTransformation F[ R R ] R
-- Note that `pureT` and `joinT` differs from their definition in the
-- kleisli formulation only by having an explicit parameter.
pureT : Transformation F.identity R
pureT = proj₁ pureNT
pureN : Natural F.identity R pureT
pureN = proj₂ pureNT
joinT : Transformation F[ R R ] R
joinT = proj₁ joinNT
joinN : Natural F[ R R ] R joinT
joinN = proj₂ joinNT
Romap = Functor.omap R
Rfmap = Functor.fmap R
bind : {X Y : Object} [ X , Romap Y ] [ Romap X , Romap Y ]
bind {X} {Y} f = joinT Y Rfmap f
IsAssociative : Set _
IsAssociative = {X : Object}
joinT X Rfmap (joinT X) joinT X joinT (Romap X)
IsInverse : Set _
IsInverse = {X : Object}
joinT X pureT (Romap X) 𝟙
× joinT X Rfmap (pureT X) 𝟙
IsNatural = {X Y} f joinT Y Rfmap f pureT X f
IsDistributive = {X Y Z} (g : Arrow Y (Romap Z)) (f : Arrow X (Romap Y))
joinT Z Rfmap g (joinT Y Rfmap f)
joinT Z Rfmap (joinT Z Rfmap 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
joinT Y R.fmap f pureT X ≡⟨ sym .isAssociative
joinT Y (R.fmap f pureT X) ≡⟨ cong (λ φ joinT Y φ) (sym (pureN f))
joinT Y (pureT (R.omap Y) f) ≡⟨ .isAssociative
joinT Y pureT (R.omap Y) f ≡⟨ cong (λ φ φ f) (proj₁ isInverse)
𝟙 f ≡⟨ proj₂ .isIdentity
f
isDistributive : IsDistributive
isDistributive {X} {Y} {Z} g f = sym aux
where
module R² = Functor F[ R R ]
distrib3 : {A B C D} {a : Arrow C D} {b : Arrow B C} {c : Arrow A B}
R.fmap (a b c)
R.fmap a R.fmap b R.fmap c
distrib3 {a = a} {b} {c} = begin
R.fmap (a b c) ≡⟨ R.isDistributive
R.fmap (a b) R.fmap c ≡⟨ cong (_∘ _) R.isDistributive
R.fmap a R.fmap b R.fmap c
aux = begin
joinT Z R.fmap (joinT Z R.fmap g f)
≡⟨ cong (λ φ joinT Z φ) distrib3
joinT Z (R.fmap (joinT Z) R.fmap (R.fmap g) R.fmap f)
≡⟨⟩
joinT Z (R.fmap (joinT Z) R².fmap g R.fmap f)
≡⟨ cong (_∘_ (joinT Z)) (sym .isAssociative)
joinT Z (R.fmap (joinT Z) (R².fmap g R.fmap f))
≡⟨ .isAssociative
(joinT Z R.fmap (joinT Z)) (R².fmap g R.fmap f)
≡⟨ cong (λ φ φ (R².fmap g R.fmap f)) isAssociative
(joinT Z joinT (R.omap Z)) (R².fmap g R.fmap f)
≡⟨ .isAssociative
joinT Z joinT (R.omap Z) R².fmap g R.fmap f
≡⟨⟩
((joinT Z joinT (R.omap Z)) R².fmap g) R.fmap f
≡⟨ cong (_∘ R.fmap f) (sym .isAssociative)
(joinT Z (joinT (R.omap Z) R².fmap g)) R.fmap f
≡⟨ cong (λ φ φ R.fmap f) (cong (_∘_ (joinT Z)) (joinN g))
(joinT Z (R.fmap g joinT Y)) R.fmap f
≡⟨ cong (_∘ R.fmap f) .isAssociative
joinT Z R.fmap g joinT Y R.fmap f
≡⟨ sym (Category.isAssociative )
joinT Z R.fmap g (joinT Y R.fmap 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
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

View file

@ -0,0 +1,211 @@
{-
This module provides construction 2.3 in [voe]
-}
{-# OPTIONS --cubical --allow-unsolved-metas #-}
module Cat.Category.Monad.Voevodsky where
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.Category.Monad using (Monoidal≃Kleisli)
import Cat.Category.Monad.Monoidal as Monoidal
import Cat.Category.Monad.Kleisli as Kleisli
open import Cat.Categories.Fun
module voe {a b : Level} ( : Category a b) where
private
= a b
module = Category
open using (Object ; Arrow)
open NaturalTransformation
open import Function using (_∘_ ; _$_)
module M = Monoidal
module K = Kleisli
module §2-3 (omap : Omap ) (pure : {X : Object} Arrow X (omap X)) where
record §1 : Set where
open M
field
fmap : Fmap omap
join : {A : Object} [ omap (omap A) , omap A ]
Rraw : RawFunctor
Rraw = record
{ omap = omap
; fmap = fmap
}
field
RisFunctor : IsFunctor Rraw
R : EndoFunctor
R = record
{ raw = Rraw
; isFunctor = RisFunctor
}
pureT : (X : Object) Arrow X (omap X)
pureT X = pure {X}
field
pureN : Natural F.identity R pureT
pureNT : NaturalTransformation F.identity R
pureNT = pureT , pureN
joinT : (A : Object) [ omap (omap A) , omap A ]
joinT A = join {A}
field
joinN : Natural F[ R R ] R joinT
joinNT : NaturalTransformation F[ R R ] R
joinNT = joinT , joinN
rawMnd : RawMonad
rawMnd = record
{ R = R
; pureNT = pureNT
; joinNT = joinNT
}
field
isMnd : IsMonad rawMnd
toMonad : Monad
toMonad = record
{ raw = rawMnd
; isMonad = isMnd
}
record §2 : Set where
open K
field
bind : {X Y : Object} [ X , omap Y ] [ omap X , omap Y ]
rawMnd : RawMonad
rawMnd = record
{ omap = omap
; pure = pure
; bind = bind
}
field
isMnd : IsMonad rawMnd
toMonad : Monad
toMonad = record
{ raw = rawMnd
; isMonad = isMnd
}
§1-fromMonad : (m : M.Monad) §2-3.§1 (M.Monad.Romap m) (λ {X} M.Monad.pureT m X)
-- voe-2-3-1-fromMonad : (m : M.Monad) → voe.§2-3.§1 (M.Monad.Romap m) (λ {X} → M.Monad.pureT m X)
§1-fromMonad m = record
{ fmap = Functor.fmap R
; RisFunctor = Functor.isFunctor R
; pureN = pureN
; join = λ {X} joinT X
; joinN = joinN
; isMnd = M.Monad.isMonad m
}
where
raw = M.Monad.raw m
R = M.RawMonad.R raw
pureT = M.RawMonad.pureT raw
pureN = M.RawMonad.pureN raw
joinT = M.RawMonad.joinT raw
joinN = M.RawMonad.joinN raw
§2-fromMonad : (m : K.Monad) §2-3.§2 (K.Monad.omap m) (K.Monad.pure m)
§2-fromMonad m = record
{ bind = K.Monad.bind m
; isMnd = K.Monad.isMonad m
}
module _ (omap : Omap ) (pure : {X : Object} Arrow X (omap X)) where
private
Monoidal→Kleisli : M.Monad K.Monad
Monoidal→Kleisli = proj₁ Monoidal≃Kleisli
Kleisli→Monoidal : K.Monad M.Monad
Kleisli→Monoidal = inverse Monoidal≃Kleisli
forth : §2-3.§1 omap pure §2-3.§2 omap pure
forth = §2-fromMonad Monoidal→Kleisli §2-3.§1.toMonad
back : §2-3.§2 omap pure §2-3.§1 omap pure
back = §1-fromMonad Kleisli→Monoidal §2-3.§2.toMonad
forthEq : m _ _
forthEq m = begin
(forth back) m ≡⟨⟩
-- In full gory detail:
( §2-fromMonad
Monoidal→Kleisli
§2-3.§1.toMonad
§1-fromMonad
Kleisli→Monoidal
§2-3.§2.toMonad
) m ≡⟨⟩ -- fromMonad and toMonad are inverses
( §2-fromMonad
Monoidal→Kleisli
Kleisli→Monoidal
§2-3.§2.toMonad
) m ≡⟨ u
-- Monoidal→Kleisli and Kleisli→Monoidal are inverses
-- I should be able to prove this using congruence and `lem` below.
( §2-fromMonad
§2-3.§2.toMonad
) m ≡⟨⟩
( §2-fromMonad
§2-3.§2.toMonad
) m ≡⟨⟩ -- fromMonad and toMonad are inverses
m
where
lem : Monoidal→Kleisli Kleisli→Monoidal Function.id
lem = {!!} -- verso-recto Monoidal≃Kleisli
t : (§2-fromMonad (Monoidal→Kleisli Kleisli→Monoidal) §2-3.§2.toMonad)
(§2-fromMonad §2-3.§2.toMonad)
t = cong (λ φ §2-fromMonad (λ{ {ω} φ {{!????!}}}) §2-3.§2.toMonad) {!lem!}
u : (§2-fromMonad (Monoidal→Kleisli Kleisli→Monoidal) §2-3.§2.toMonad) m
(§2-fromMonad §2-3.§2.toMonad) m
u = cong (λ φ φ m) t
backEq : m (back forth) m m
backEq m = begin
(back forth) m ≡⟨⟩
( §1-fromMonad
Kleisli→Monoidal
§2-3.§2.toMonad
§2-fromMonad
Monoidal→Kleisli
§2-3.§1.toMonad
) m ≡⟨⟩ -- fromMonad and toMonad are inverses
( §1-fromMonad
Kleisli→Monoidal
Monoidal→Kleisli
§2-3.§1.toMonad
) m ≡⟨ cong (λ φ φ m) t -- Monoidal→Kleisli and Kleisli→Monoidal are inverses
( §1-fromMonad
§2-3.§1.toMonad
) m ≡⟨⟩ -- fromMonad and toMonad are inverses
m
where
t = {!!} -- cong (λ φ → voe-2-3-1-fromMonad ∘ φ ∘ voe-2-3.voe-2-3-1.toMonad) (recto-verso Monoidal≃Kleisli)
voe-isEquiv : isEquiv (§2-3.§1 omap pure) (§2-3.§2 omap pure) forth
voe-isEquiv = gradLemma forth back forthEq backEq
equiv-2-3 : §2-3.§1 omap pure §2-3.§2 omap pure
equiv-2-3 = forth , voe-isEquiv

View file

@ -12,10 +12,14 @@ module _ (a b : Level) where
private private
= lsuc (a b) = lsuc (a b)
-- Might not need this to be able to form products of categories! -- *If* the category of categories existed `_×_` would be equivalent to the
postulate unprovable : IsCategory (Cat.RawCat a b) -- one brought into scope by doing:
--
open HasProducts (Cat.hasProducts unprovable) -- open HasProducts (Cat.hasProducts unprovable) using (_×_)
--
-- Since it doesn't we'll make the following (definitionally equivalent) ad-hoc definition.
_×_ : {a b} Category a b Category a b Category a b
× 𝔻 = Cat.CatProduct.obj 𝔻
record RawMonoidalCategory : Set where record RawMonoidalCategory : Set where
field field
@ -23,9 +27,10 @@ module _ (a b : Level) where
open Category category public open Category category public
field field
{{hasProducts}} : HasProducts category {{hasProducts}} : HasProducts category
mempty : Object empty : Object
-- aka. tensor product, monoidal product. -- aka. tensor product, monoidal product.
mappend : Functor (category × category) category append : Functor (category × category) category
open HasProducts hasProducts public
record MonoidalCategory : Set where record MonoidalCategory : Set where
field field
@ -36,10 +41,10 @@ module _ {a b : Level} ( : MonoidalCategory a b) where
private private
= a b = a b
module MC = MonoidalCategory open MonoidalCategory public
open HasProducts MC.hasProducts
record Monoid : Set where record Monoid : Set where
field field
carrier : MC.Object carrier : Object
mempty : MC.Arrow (carrier × carrier) carrier mempty : Arrow empty carrier
mappend : MC.Arrow MC.mempty carrier mappend : Arrow (carrier × carrier) carrier

View file

@ -21,15 +21,24 @@
module Cat.Category.NaturalTransformation where module Cat.Category.NaturalTransformation where
open import Agda.Primitive open import Agda.Primitive
open import Data.Product open import Data.Product
open import Data.Nat using (_≤_ ; z≤n ; s≤s)
module Nat = Data.Nat
open import Cubical open import Cubical
open import Cubical.Sigma
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)
open import Cat.Wishlist
module NaturalTransformation {c c' d d' : Level} module NaturalTransformation {c c' d d' : Level}
( : Category c c') (𝔻 : Category d d') where ( : Category c c') (𝔻 : Category d d') where
open Category using (Object ; 𝟙) open Category using (Object ; 𝟙)
private
module = Category
module 𝔻 = Category 𝔻
module _ (F G : Functor 𝔻) where module _ (F G : Functor 𝔻) where
private private
@ -37,28 +46,25 @@ module NaturalTransformation {c c' d d' : Level}
module G = Functor G module G = Functor G
-- What do you call a non-natural tranformation? -- What do you call a non-natural tranformation?
Transformation : Set (c d') Transformation : Set (c d')
Transformation = (C : Object ) 𝔻 [ F.func* C , G.func* C ] Transformation = (C : Object ) 𝔻 [ F.omap C , G.omap C ]
Natural : Transformation Set (c (c' d')) Natural : Transformation Set (c (c' d'))
Natural θ Natural θ
= {A B : Object } = {A B : Object }
(f : [ A , B ]) (f : [ A , B ])
𝔻 [ θ B F.func→ f ] 𝔻 [ G.func→ f θ A ] 𝔻 [ θ B F.fmap f ] 𝔻 [ G.fmap f θ A ]
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. propIsNatural : (θ : _) isProp (Natural θ)
propIsNatural θ x y i {A} {B} f = 𝔻.arrowsAreSets _ _ (x f) (y f) i
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 = 𝟙 𝔻
@ -72,8 +78,7 @@ module NaturalTransformation {c c' d d' : Level}
𝔻 [ F→ f identityTrans F A ] 𝔻 [ F→ f identityTrans F A ]
where where
module F = Functor F module F = Functor F
F→ = F.func→ F→ = F.fmap
module 𝔻 = Category 𝔻
identity : (F : Functor 𝔻) NaturalTransformation F F identity : (F : Functor 𝔻) NaturalTransformation F F
identity F = identityTrans F , identityNatural F identity F = identityTrans F , identityNatural F
@ -89,13 +94,27 @@ module NaturalTransformation {c c' d d' : Level}
NT[_∘_] : NaturalTransformation G H NaturalTransformation F G NaturalTransformation F H NT[_∘_] : NaturalTransformation G H NaturalTransformation F G NaturalTransformation F H
proj₁ NT[ (θ , _) (η , _) ] = T[ θ η ] proj₁ NT[ (θ , _) (η , _) ] = T[ θ η ]
proj₂ NT[ (θ , θNat) (η , ηNat) ] {A} {B} f = begin proj₂ NT[ (θ , θNat) (η , ηNat) ] {A} {B} f = begin
𝔻 [ T[ θ η ] B F.func→ f ] ≡⟨⟩ 𝔻 [ T[ θ η ] B F.fmap f ] ≡⟨⟩
𝔻 [ 𝔻 [ θ B η B ] F.func→ f ] ≡⟨ sym isAssociative 𝔻 [ 𝔻 [ θ B η B ] F.fmap f ] ≡⟨ sym 𝔻.isAssociative
𝔻 [ θ B 𝔻 [ η B F.func→ f ] ] ≡⟨ cong (λ φ 𝔻 [ θ B φ ]) (ηNat f) 𝔻 [ θ B 𝔻 [ η B F.fmap f ] ] ≡⟨ cong (λ φ 𝔻 [ θ B φ ]) (ηNat f)
𝔻 [ θ B 𝔻 [ G.func→ f η A ] ] ≡⟨ isAssociative 𝔻 [ θ B 𝔻 [ G.fmap f η A ] ] ≡⟨ 𝔻.isAssociative
𝔻 [ 𝔻 [ θ B G.func→ f ] η A ] ≡⟨ cong (λ φ 𝔻 [ φ η A ]) (θNat f) 𝔻 [ 𝔻 [ θ B G.fmap f ] η A ] ≡⟨ cong (λ φ 𝔻 [ φ η A ]) (θNat f)
𝔻 [ 𝔻 [ H.func→ f θ A ] η A ] ≡⟨ sym isAssociative 𝔻 [ 𝔻 [ H.fmap f θ A ] η A ] ≡⟨ sym 𝔻.isAssociative
𝔻 [ H.func→ f 𝔻 [ θ A η A ] ] ≡⟨⟩ 𝔻 [ H.fmap f 𝔻 [ θ A η A ] ] ≡⟨⟩
𝔻 [ H.func→ f T[ θ η ] A ] 𝔻 [ H.fmap f T[ θ η ] A ]
module _ {F G : Functor 𝔻} where
transformationIsSet : isSet (Transformation F G)
transformationIsSet _ _ p q i j C = 𝔻.arrowsAreSets _ _ (λ l p l C) (λ l q l C) i j
naturalIsProp : (θ : Transformation F G) isProp (Natural F G θ)
naturalIsProp θ θNat θNat' = lem
where where
open Category 𝔻 lem : (λ _ Natural F G θ) [ (λ f θNat f) (λ f θNat' f) ]
lem = λ i f 𝔻.arrowsAreSets _ _ (θNat f) (θNat' f) i
naturalTransformationIsSet : isSet (NaturalTransformation F G)
naturalTransformationIsSet = sigPresSet transformationIsSet
λ θ ntypeCommulative
(s≤s {n = Nat.suc Nat.zero} z≤n)
(naturalIsProp θ)

View file

@ -1,54 +1,68 @@
{-# OPTIONS --allow-unsolved-metas #-}
module Cat.Category.Product where module Cat.Category.Product where
open import Agda.Primitive open import Agda.Primitive
open import Cubical open import Cubical
open import Data.Product as P hiding (_×_) open import Data.Product as P hiding (_×_ ; proj₁ ; proj₂)
open import Cat.Category open import Cat.Category hiding (module Propositionality)
open Category module _ {a b : Level} ( : Category a b) where
module _ { ' : Level} ( : Category ') {A B obj : Object } where open Category
IsProduct : (π₁ : [ obj , A ]) (π₂ : [ obj , B ]) Set ( ')
IsProduct π₁ π₂
= {X : Object } (x₁ : [ X , A ]) (x₂ : [ X , B ])
∃![ x ] ( [ π₁ x ] x₁ P.× [ π₂ x ] x₂)
-- Tip from Andrea; Consider this style for efficiency: module _ (A B : Object) where
-- record IsProduct {a b : Level} ( : Category a b) record RawProduct : Set (a b) where
-- {A B obj : Object } (π₁ : Arrow obj A) (π₂ : Arrow obj B) : Set (a ⊔ b) where no-eta-equality
-- field field
-- issProduct : ∀ {X : Object } (x₁ : [ X , A ]) (x₂ : [ X , B ]) object : Object
-- → ∃![ x ] ( [ π₁ ∘ x ] ≡ x₁ P.× [ π₂ ∘ x ] ≡ x₂) proj₁ : [ object , A ]
proj₂ : [ object , B ]
-- open IsProduct -- FIXME Not sure this is actually a proposition - so this name is
-- misleading.
record IsProduct (raw : RawProduct) : Set (a b) where
open RawProduct raw public
field
isProduct : {X : Object} (f : [ X , A ]) (g : [ X , B ])
∃![ f×g ] ( [ proj₁ f×g ] f P.× [ proj₂ f×g ] g)
record Product { ' : Level} { : Category '} (A B : Object ) : Set ( ') where -- | Arrow product
no-eta-equality _P[_×_] : {X} (π₁ : [ X , A ]) (π₂ : [ X , B ])
field [ X , object ]
obj : Object _P[_×_] π₁ π₂ = P.proj₁ (isProduct π₁ π₂)
proj₁ : [ obj , A ]
proj₂ : [ obj , B ]
{{isProduct}} : IsProduct proj₁ proj₂
_P[_×_] : {X} (π₁ : [ X , A ]) (π₂ : [ X , B ]) record Product : Set (a b) where
[ X , obj ] field
_P[_×_] π₁ π₂ = proj₁ (isProduct π₁ π₂) raw : RawProduct
isProduct : IsProduct raw
record HasProducts { ' : Level} ( : Category ') : Set ( ') where open IsProduct isProduct public
field
product : (A B : Object ) Product { = } A B
open Product record HasProducts : Set (a b) where
field
product : (A B : Object) Product A B
_×_ : (A B : Object ) Object _×_ : Object Object Object
A × B = Product.obj (product A B) A × B = Product.object (product A B)
-- The product mentioned in awodey in Def 6.1 is not the regular product of arrows.
-- It's a "parallel" product -- | Parallel product of arrows
_|×|_ : {A A' B B' : Object } [ A , A' ] [ B , B' ] --
[ A × B , A' × B' ] -- The product mentioned in awodey in Def 6.1 is not the regular product of
_|×|_ {A = A} {A' = A'} {B = B} {B' = B'} a b -- arrows. It's a "parallel" product
= product A' B' module _ {A A' B B' : Object} where
P[ [ a (product A B) .proj₁ ] open Product
× [ b (product A B) .proj₂ ] open Product (product A B) hiding (_P[_×_]) renaming (proj₁ to fst ; proj₂ to snd)
] _|×|_ : [ A , A' ] [ B , B' ] [ A × B , A' × B' ]
f |×| g = product A' B'
P[ [ f fst ]
× [ g snd ]
]
module Propositionality {a b : Level} { : Category a b} {A B : Category.Object } where
-- TODO I'm not sure this is actually provable. Check with Thierry.
propProduct : isProp (Product A B)
propProduct = {!!}
propHasProducts : isProp (HasProducts )
propHasProducts = {!!}

View file

@ -5,60 +5,76 @@ module Cat.Category.Yoneda where
open import Agda.Primitive 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 open import Cat.Category.Functor
open import Cat.Equality open import Cat.Equality
open Equality.Data.Product
-- TODO: We want to avoid defining the yoneda embedding going through the open import Cat.Categories.Fun
-- category of categories (since it doesn't exist). open import Cat.Categories.Sets
open import Cat.Categories.Cat using (RawCat) open import Cat.Categories.Cat
module _ { : Level} { : Category } (unprovable : IsCategory (RawCat )) where module _ { : Level} { : Category } where
open import Cat.Categories.Fun
open import Cat.Categories.Sets
module Cat = Cat.Categories.Cat
open import Cat.Category.Exponential
open Functor
𝓢 = Sets
open Fun (opposite ) 𝓢
private private
Cat : Category _ _ 𝓢 = Sets
Cat = record { raw = RawCat ; isCategory = unprovable} open Fun (opposite ) 𝓢
prshf = presheaf { = } prshf = presheaf
module = Category module = Category
_⇑_ : (A B : Category.Object Cat) Category.Object Cat -- There is no (small) category of categories. So we won't use _⇑_ from
A B = (exponent A B) .obj -- `HasExponential`
where --
open HasExponentials (Cat.hasExponentials unprovable) -- open HasExponentials (Cat.hasExponentials unprovable) using (_⇑_)
--
-- In stead we'll use an ad-hoc definition -- which is definitionally
-- equivalent to that other one.
_⇑_ = CatExponential.object
module _ {A B : .Object} (f : [ A , B ]) where module _ {A B : .Object} (f : [ A , B ]) where
:func→: : NaturalTransformation (prshf A) (prshf B) fmap : Transformation (prshf A) (prshf B)
:func→: = (λ C x [ f x ]) , λ f₁ funExt λ _ .isAssociative fmap C x = [ f x ]
module _ {c : Category.Object } where fmapNatural : Natural (prshf A) (prshf B) fmap
eqTrans : (λ _ Transformation (prshf c) (prshf c)) fmapNatural g = funExt λ _ .isAssociative
[ (λ _ x [ .𝟙 x ]) identityTrans (prshf c) ]
eqTrans = funExt λ x funExt λ x .isIdentity .proj₂
open import Cubical.NType.Properties fmapNT : NaturalTransformation (prshf A) (prshf B)
open import Cat.Categories.Fun fmapNT = fmap , fmapNatural
:ident: : :func→: (.𝟙 {c}) Category.𝟙 Fun {A = prshf c}
:ident: = lemSig (naturalIsProp {F = prshf c} {prshf c}) _ _ eq rawYoneda : RawFunctor Fun
where RawFunctor.omap rawYoneda = prshf
eq : (λ C x [ .𝟙 x ]) identityTrans (prshf c) RawFunctor.fmap rawYoneda = fmapNT
eq = funExt λ A funExt λ B proj₂ .isIdentity open RawFunctor rawYoneda hiding (fmap)
isIdentity : IsIdentity
isIdentity {c} = lemSig (naturalIsProp {F = prshf c} {prshf c}) _ _ eq
where
eq : (λ C x [ .𝟙 x ]) identityTrans (prshf c)
eq = funExt λ A funExt λ B proj₂ .isIdentity
isDistributive : IsDistributive
isDistributive {A} {B} {C} {f = f} {g}
= lemSig (propIsNatural (prshf A) (prshf C)) _ _ eq
where
T[_∘_]' = T[_∘_] {F = prshf A} {prshf B} {prshf C}
eqq : (X : .Object) (x : [ X , A ])
fmap ( [ g f ]) X x T[ fmap g fmap f ]' X x
eqq X x = begin
fmap ( [ g f ]) X x ≡⟨⟩
[ [ g f ] x ] ≡⟨ sym .isAssociative
[ g [ f x ] ] ≡⟨⟩
[ g fmap f X x ] ≡⟨⟩
T[ fmap g fmap f ]' X x
eq : fmap ( [ g f ]) T[ fmap g fmap f ]'
eq = begin
fmap ( [ g f ]) ≡⟨ funExt (λ X funExt λ α eqq X α)
T[ fmap g fmap f ]'
instance
isFunctor : IsFunctor Fun rawYoneda
IsFunctor.isIdentity isFunctor = isIdentity
IsFunctor.isDistributive isFunctor = isDistributive
yoneda : Functor Fun yoneda : Functor Fun
yoneda = record Functor.raw yoneda = rawYoneda
{ raw = record Functor.isFunctor yoneda = isFunctor
{ func* = prshf
; func→ = :func→:
}
; isFunctor = record
{ isIdentity = :ident:
; isDistributive = {!!}
}
}

View file

@ -1,15 +1,41 @@
{-# OPTIONS --allow-unsolved-metas #-}
module Cat.Wishlist where module Cat.Wishlist where
open import Level open import Level hiding (suc)
open import Cubical
open import Cubical.NType open import Cubical.NType
open import Data.Nat using (_≤_ ; z≤n ; s≤s) open import Data.Nat using (_≤_ ; z≤n ; s≤s ; zero ; suc)
open import Agda.Builtin.Sigma
postulate ntypeCommulative : { n m} {A : Set } n m HasLevel n ⟩₋₂ A HasLevel m ⟩₋₂ A open import Cubical.NType.Properties
step : {} {A : Set } isContr A (x y : A) isContr (x y)
step (a , contr) x y = {!p , c!}
-- where
-- p : x ≡ y
-- p = begin
-- x ≡⟨ sym (contr x) ⟩
-- a ≡⟨ contr y ⟩
-- y ∎
-- c : (q : x ≡ y) → p ≡ q
-- c q i j = contr (p {!!}) {!!}
-- Contractible types have any given homotopy level.
contrInitial : { : Level} {A : Set } n isContr A HasLevel n A
contrInitial ⟨-2⟩ contr = contr
-- lem' (S ⟨-2⟩) (a , contr) = {!step!}
contrInitial (S ⟨-2⟩) (a , contr) x y = begin
x ≡⟨ sym (contr x)
a ≡⟨ contr y
y
contrInitial (S (S n)) contr x y = {!lvl!} -- Why is this not well-founded?
where
c : isContr (x y)
c = step contr x y
lvl : HasLevel (S n) (x y)
lvl = contrInitial {A = x y} (S n) c
module _ { : Level} {A : Set } where module _ { : Level} {A : Set } where
-- This is §7.1.10 in [HoTT]. Andrea says the proof is in `cubical` but I ntypeCommulative : {n m} n m HasLevel n ⟩₋₂ A HasLevel m ⟩₋₂ A
-- can't find it. ntypeCommulative {n = zero} {m} z≤n lvl = {!contrInitial ⟨ m ⟩₋₂ lvl!}
postulate propHasLevel : n isProp (HasLevel n A) ntypeCommulative {n = .(suc _)} {.(suc _)} (s≤s x) lvl = {!!}
isSetIsProp : isProp (isSet A)
isSetIsProp = propHasLevel (S (S ⟨-2⟩))