Have yoneda without having a category of categories

I did break some things in Cat.Categories.Cat but since this is
unprovable anyways it's not that big a deal.
This commit is contained in:
Frederik Hanghøj Iversen 2018-03-05 13:52:41 +01:00
parent 5c3616bca5
commit 1bf565b87a
5 changed files with 259 additions and 225 deletions

View file

@ -11,7 +11,7 @@ open import Data.Product renaming (proj₁ to fst ; proj₂ to snd)
open import Cat.Category
open import Cat.Category.Functor
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.Equality
@ -174,190 +174,211 @@ module _ { ' : Level} (unprovable : IsCategory (RawCat ')) where
hasProducts = record { product = product }
-- Basically proves that `Cat ` is cartesian closed.
module CatExponential { : Level} ( 𝔻 : Category ) where
open Data.Product
open import Cat.Categories.Fun
Category = Category
open Fun 𝔻 renaming (identity to idN)
private
:func*: : Functor 𝔻 × Object Object 𝔻
:func*: (F , A) = func* F A
prodObj : Category
prodObj = 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
: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
open CatProduct renaming (obj to _×p_) using ()
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} (𝟙 (prodObj ×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._∘_ (prodObj ×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 (CatProduct.obj prodObj ) 𝔻
-- :eval: : Functor (prodObj ×p ) 𝔻
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 ) 𝔻) where
-- open HasProducts (hasProducts {} {} unprovable) renaming (_|×|_ to parallelProduct)
postulate
parallelProduct
: Functor 𝔸 prodObj Functor
Functor (𝔸 ×p ) (prodObj ×p )
transpose : Functor 𝔸 prodObj
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
module _ ( : Level) (unprovable : IsCategory (RawCat )) where
private
open Data.Product
open import Cat.Categories.Fun
Cat : Category (lsuc ( )) ( )
Cat = Cat unprovable
module _ ( 𝔻 : Category ) where
open Fun 𝔻 renaming (identity to idN)
private
:obj: : Object Cat
:obj: = Fun
module _ ( 𝔻 : Category ) where
open CatExponential 𝔻 using (prodObj ; eval)
-- Putting in the type annotation causes Agda to loop indefinitely.
-- eval' : Functor (CatProduct.obj prodObj ) 𝔻
-- Likewise, using it below also results in this.
eval' : _
eval' = eval
-- private
-- -- module _ ( 𝔻 : Category ) where
-- postulate :isExponential: : IsExponential Cat 𝔻 prodObj :eval:
-- -- :isExponential: : IsExponential Cat 𝔻 :obj: :eval:
-- -- :isExponential: = {!catTranspose!}
-- -- where
-- -- open HasProducts (hasProducts {} {} unprovable) using (_|×|_)
-- -- :isExponential: = λ 𝔸 F → transpose 𝔸 F , eq' 𝔸 F
:func*: : Functor 𝔻 × Object Object 𝔻
:func*: (F , A) = func* F A
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
: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:
}
-- -- :exponent: : Exponential (Cat ) A B
exponent : Exponential Cat 𝔻
exponent = record
{ obj = prodObj
; eval = {!evalll'!}
; isExponential = {!:isExponential:!}
}
where
open HasProducts (hasProducts unprovable) renaming (_×_ to _×p_)
open import Cat.Categories.Fun
open Fun
-- _×p_ = CatProduct.obj -- prodObj
-- eval' : Functor CatP.obj 𝔻
hasExponentials : HasExponentials Cat
hasExponentials = record { exponent = :exponent: }
hasExponentials = record { exponent = exponent }

View file

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

@ -27,9 +27,10 @@ module _ (a b : Level) where
open Category category public
field
{{hasProducts}} : HasProducts category
mempty : Object
empty : Object
-- aka. tensor product, monoidal product.
mappend : Functor (category × category) category
append : Functor (category × category) category
open HasProducts hasProducts public
record MonoidalCategory : Set where
field
@ -40,10 +41,10 @@ module _ {a b : Level} ( : MonoidalCategory a b) where
private
= a b
module MC = MonoidalCategory
open HasProducts MC.hasProducts
open MonoidalCategory public
record Monoid : Set where
field
carrier : MC.Object
mempty : MC.Arrow (carrier × carrier) carrier
mappend : MC.Arrow MC.mempty carrier
carrier : Object
mempty : Arrow empty carrier
mappend : Arrow (carrier × carrier) carrier

View file

@ -31,6 +31,7 @@ record Product { ' : Level} { : Category '} (A B : Object ) :
proj₂ : [ obj , B ]
{{isProduct}} : IsProduct proj₁ proj₂
-- | Arrow product
_P[_×_] : {X} (π₁ : [ X , A ]) (π₂ : [ X , B ])
[ X , obj ]
_P[_×_] π₁ π₂ = proj₁ (isProduct π₁ π₂)
@ -39,16 +40,21 @@ record HasProducts { ' : Level} ( : Category ') : Set (
field
product : (A B : Object ) Product { = } A B
open Product
open Product hiding (obj)
_×_ : (A B : Object ) Object
A × B = Product.obj (product A B)
-- The product mentioned in awodey in Def 6.1 is not the regular product of arrows.
-- It's a "parallel" product
_|×|_ : {A A' B B' : Object } [ A , A' ] [ B , B' ]
[ A × B , A' × B' ]
_|×|_ {A = A} {A' = A'} {B = B} {B' = B'} a b
= product A' B'
P[ [ a (product A B) .proj₁ ]
× [ b (product A B) .proj₂ ]
module _ (A B : Object ) where
open Product (product A B)
_×_ : Object
_×_ = obj
-- | Parallel product of arrows
--
-- The product mentioned in awodey in Def 6.1 is not the regular product of
-- arrows. It's a "parallel" product
module _ {A A' B B' : Object } where
open Product (product A B) hiding (_P[_×_]) renaming (proj₁ to fst ; proj₂ to snd)
_|×|_ : [ A , A' ] [ B , B' ] [ A × B , A' × B' ]
a |×| b = product A' B'
P[ [ a fst ]
× [ b snd ]
]

View file

@ -15,7 +15,7 @@ open Equality.Data.Product
-- category of categories (since it doesn't exist).
open import Cat.Categories.Cat using (RawCat)
module _ { : Level} { : Category } (unprovable : IsCategory (RawCat )) where
module _ { : Level} { : Category } where
private
open import Cat.Categories.Fun
open import Cat.Categories.Sets
@ -24,15 +24,17 @@ module _ { : Level} { : Category } (unprovable : IsCategory (RawCat
open Functor
𝓢 = Sets
open Fun (opposite ) 𝓢
Cat : Category _ _
Cat = Cat.Cat unprovable
prshf = presheaf
module = Category
_⇑_ : (A B : Category.Object Cat) Category.Object Cat
A B = (exponent A B) .obj
where
open HasExponentials (Cat.hasExponentials unprovable)
-- There is no (small) category of categories. So we won't use _⇑_ from
-- `HasExponential`
--
-- open HasExponentials (Cat.hasExponentials unprovable) using (_⇑_)
--
-- In stead we'll use an ad-hoc definition -- which is definitionally
-- equivalent to that other one.
_⇑_ = Cat.CatExponential.prodObj
module _ {A B : .Object} (f : [ A , B ]) where
:func→: : NaturalTransformation (prshf A) (prshf B)