cat/src/Cat/Categories/Sets.agda

61 lines
1.7 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)
}
module _ { ' : Level} { : Category {} {}} where
private
C-Obj = Object
_+_ = Arrow
RepFunctor : Functor Sets
RepFunctor =
record
{ func* = λ A (B : C-Obj) Hom { = } A B
; func→ = λ { {c} {c'} f g {!HomFromArrow { = {!!}} c' g!} }
; ident = {!!}
; distrib = {!!}
}
Hom0 : { ' : Level} { : Category {} {'}} Category.Object Functor (Sets {'})
Hom0 { = } 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
Hom1 : { ' : Level} { : Category {} {'}} Category.Object Functor (Opposite ) (Sets {'})
Hom1 { = } B = record
{ func* = λ A .Arrow A B
; func→ = λ f g {!!} .⊕ {!!}
; ident = {!!}
; distrib = {!!}
}
where
open module = Category