cat/src/Cat/Categories/Sets.agda

73 lines
2.3 KiB
Agda
Raw Normal View History

{-# OPTIONS --allow-unsolved-metas #-}
module Cat.Categories.Sets where
2017-11-15 20:51:10 +00:00
open import Cubical.PathPrelude
open import Agda.Primitive
open import Data.Product
open import Data.Product renaming (proj₁ to fst ; proj₂ to snd)
open import Cat.Category
open import Cat.Functor
open Category
2017-11-15 20:51:10 +00:00
2018-01-24 15:38:28 +00:00
module _ { : Level} where
Sets : Category (lsuc )
Sets = record
{ Object = Set
; Arrow = λ T U T U
; 𝟙 = id
; _⊕_ = _∘_
; isCategory = record { assoc = refl ; ident = funExt (λ _ refl) , funExt (λ _ refl) }
}
where
open import Function
private
module _ {X A B : Set } (f : X A) (g : X B) where
pair : (X A × B)
pair x = f x , g x
lem : Sets ._⊕_ proj₁ pair f × Sets ._⊕_ snd pair g
proj₁ lem = refl
snd lem = refl
instance
isProduct : {A B : Sets .Object} IsProduct Sets {A} {B} fst snd
isProduct f g = pair f g , lem f g
product : (A B : Sets .Object) Product { = Sets} A B
product A B = record { obj = A × B ; proj₁ = fst ; proj₂ = snd ; isProduct = {!!} }
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 .Arrow A B
; func→ = ._⊕_
; ident = funExt λ _ snd ident
; 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 .Arrow A B
; func→ = λ f g ._⊕_ g f
; ident = funExt λ x fst ident
; distrib = funExt λ x assoc
}
where
open IsCategory ( .isCategory)