cat/src/Cat/Categories/Sets.agda

85 lines
2.7 KiB
Agda
Raw Normal View History

2018-02-05 13:47:15 +00:00
{-# OPTIONS --allow-unsolved-metas --cubical #-}
module Cat.Categories.Sets where
2017-11-15 20:51:10 +00:00
2018-01-30 10:19:48 +00:00
open import Cubical
2017-11-15 20:51:10 +00:00
open import Agda.Primitive
open import Data.Product
import Function
open import Cat.Category
2018-02-05 13:59:53 +00:00
open import Cat.Category.Functor
open import Cat.Category.Product
open Category
2017-11-15 20:51:10 +00:00
2018-01-24 15:38:28 +00:00
module _ { : Level} where
SetsRaw : RawCategory (lsuc )
2018-02-05 13:47:15 +00:00
RawCategory.Object SetsRaw = Set
RawCategory.Arrow SetsRaw = λ T U T U
RawCategory.𝟙 SetsRaw = Function.id
RawCategory._∘_ SetsRaw = Function._∘_
2018-02-05 13:47:15 +00:00
open IsCategory
SetsIsCategory : IsCategory SetsRaw
2018-02-05 13:47:15 +00:00
assoc SetsIsCategory = refl
proj₁ (ident SetsIsCategory) = funExt λ _ refl
proj₂ (ident SetsIsCategory) = funExt λ _ refl
2018-02-06 09:34:43 +00:00
arrowIsSet SetsIsCategory = {!!}
2018-02-05 13:47:15 +00:00
univalent SetsIsCategory = {!!}
Sets : Category (lsuc )
2018-02-05 13:47:15 +00:00
raw Sets = SetsRaw
isCategory Sets = SetsIsCategory
2018-01-24 15:38:28 +00:00
private
module _ {X A B : Set } (f : X A) (g : X B) where
2018-01-25 11:01:37 +00:00
_&&&_ : (X A × B)
_&&&_ x = f x , g x
module _ {X A B : Set } (f : X A) (g : X B) where
lem : Sets [ proj₁ (f &&& g)] f × Sets [ proj₂ (f &&& g)] g
2018-01-24 15:38:28 +00:00
proj₁ lem = refl
2018-01-25 11:01:37 +00:00
proj₂ lem = refl
2018-01-24 15:38:28 +00:00
instance
isProduct : {A B : Object Sets} IsProduct Sets {A} {B} proj₁ proj₂
2018-01-25 11:01:37 +00:00
isProduct f g = f &&& g , lem f g
2018-01-24 15:38:28 +00:00
product : (A B : Object Sets) Product { = Sets} A B
product A B = record { obj = A × B ; proj₁ = proj₁ ; proj₂ = proj₂ ; isProduct = isProduct }
2018-01-24 15:38:28 +00:00
instance
SetsHasProducts : HasProducts Sets
SetsHasProducts = record { product = product }
2017-11-15 20:51:10 +00:00
-- Covariant Presheaf
Representable : { ' : Level} ( : Category ') Set ( lsuc ')
Representable {' = '} = Functor (Sets {'})
-- The "co-yoneda" embedding.
representable : { '} { : Category '} Category.Object Representable
2018-01-17 11:10:18 +00:00
representable { = } A = record
{ func* = λ B [ A , B ]
; func→ = [_∘_]
2018-01-30 15:23:36 +00:00
; isFunctor = record
{ ident = funExt λ _ proj₂ ident
2018-01-30 15:23:36 +00:00
; distrib = funExt λ x sym assoc
}
}
where
open IsCategory (isCategory )
-- Contravariant Presheaf
Presheaf : { '} ( : Category ') Set ( lsuc ')
Presheaf {' = '} = Functor (Opposite ) (Sets {'})
-- Alternate name: `yoneda`
presheaf : { ' : Level} { : Category '} Category.Object (Opposite ) Presheaf
presheaf { = } B = record
{ func* = λ A [ A , B ]
; func→ = λ f g [ g f ]
2018-01-30 15:23:36 +00:00
; isFunctor = record
{ ident = funExt λ x proj₁ ident
2018-01-30 15:23:36 +00:00
; distrib = funExt λ x assoc
}
}
where
open IsCategory (isCategory )