cat/src/Cat/Categories/Sets.agda

47 lines
1.4 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
2017-11-15 20:51:10 +00:00
-- Sets are built-in to Agda. The set of all small sets is called Set.
Fun : { : Level} ( T U : Set ) Set
Fun T U = T U
Sets : { : Level} Category {lsuc } {}
Sets {} = record
{ Object = Set
; Arrow = λ T U Fun {} T U
; 𝟙 = λ x x
; _⊕_ = λ g f x g ( f x )
; assoc = refl
; ident = funExt (λ x refl) , funExt (λ x refl)
}
2018-01-17 11:10:18 +00:00
representable : { ' : Level} { : Category {} {'}} Category.Object Functor (Sets {'})
representable { = } A = record
{ func* = λ B .Arrow A B
; func→ = λ f g f .⊕ g
; ident = funExt λ _ snd .ident
; distrib = funExt λ x sym .assoc
}
where
open module = Category
2018-01-17 11:10:18 +00:00
coRepresentable : { ' : Level} { : Category {} {'}} Category.Object (Opposite ) Functor (Opposite ) (Sets {'})
coRepresentable { = } B = record
{ func* = λ A .Arrow A B
2018-01-17 11:10:18 +00:00
; func→ = λ f g g .⊕ f
; ident = funExt λ x fst .ident
; distrib = funExt λ x .assoc
}
where
open module = Category