2018-01-15 15:13:23 +00:00
|
|
|
|
{-# 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
|
2018-01-15 15:13:23 +00:00
|
|
|
|
open import Data.Product
|
|
|
|
|
open import Data.Product renaming (proj₁ to fst ; proj₂ to snd)
|
|
|
|
|
|
|
|
|
|
open import Cat.Category
|
|
|
|
|
open import Cat.Functor
|
2018-01-21 13:31:37 +00:00
|
|
|
|
open Category
|
2017-11-15 20:51:10 +00:00
|
|
|
|
|
2018-01-21 00:11:08 +00:00
|
|
|
|
Sets : {ℓ : Level} → Category (lsuc ℓ) ℓ
|
2017-11-15 20:51:10 +00:00
|
|
|
|
Sets {ℓ} = record
|
|
|
|
|
{ Object = Set ℓ
|
2018-01-21 13:31:37 +00:00
|
|
|
|
; Arrow = λ T U → T → U
|
|
|
|
|
; 𝟙 = id
|
|
|
|
|
; _⊕_ = _∘′_
|
|
|
|
|
; isCategory = record { assoc = refl ; ident = funExt (λ _ → refl) , funExt (λ _ → refl) }
|
2017-11-15 20:51:10 +00:00
|
|
|
|
}
|
2018-01-21 13:31:37 +00:00
|
|
|
|
where
|
|
|
|
|
open import Function
|
2017-11-15 20:51:10 +00:00
|
|
|
|
|
2018-01-20 23:21:25 +00:00
|
|
|
|
-- Covariant Presheaf
|
2018-01-21 00:11:08 +00:00
|
|
|
|
Representable : {ℓ ℓ' : Level} → (ℂ : Category ℓ ℓ') → Set (ℓ ⊔ lsuc ℓ')
|
2018-01-17 11:16:07 +00:00
|
|
|
|
Representable {ℓ' = ℓ'} ℂ = Functor ℂ (Sets {ℓ'})
|
|
|
|
|
|
2018-01-20 23:21:25 +00:00
|
|
|
|
-- The "co-yoneda" embedding.
|
2018-01-21 00:11:08 +00:00
|
|
|
|
representable : ∀ {ℓ ℓ'} {ℂ : Category ℓ ℓ'} → Category.Object ℂ → Representable ℂ
|
2018-01-17 11:10:18 +00:00
|
|
|
|
representable {ℂ = ℂ} A = record
|
2018-01-21 13:31:37 +00:00
|
|
|
|
{ func* = λ B → ℂ .Arrow A B
|
|
|
|
|
; func→ = ℂ ._⊕_
|
|
|
|
|
; ident = funExt λ _ → snd ident
|
|
|
|
|
; distrib = funExt λ x → sym assoc
|
2018-01-15 15:13:23 +00:00
|
|
|
|
}
|
|
|
|
|
where
|
2018-01-21 13:31:37 +00:00
|
|
|
|
open IsCategory (ℂ .isCategory)
|
2018-01-15 15:13:23 +00:00
|
|
|
|
|
2018-01-20 23:21:25 +00:00
|
|
|
|
-- Contravariant Presheaf
|
2018-01-21 00:11:08 +00:00
|
|
|
|
Presheaf : ∀ {ℓ ℓ'} (ℂ : Category ℓ ℓ') → Set (ℓ ⊔ lsuc ℓ')
|
2018-01-17 11:16:07 +00:00
|
|
|
|
Presheaf {ℓ' = ℓ'} ℂ = Functor (Opposite ℂ) (Sets {ℓ'})
|
|
|
|
|
|
2018-01-20 23:21:25 +00:00
|
|
|
|
-- Alternate name: `yoneda`
|
2018-01-21 00:11:08 +00:00
|
|
|
|
presheaf : {ℓ ℓ' : Level} {ℂ : Category ℓ ℓ'} → Category.Object (Opposite ℂ) → Presheaf ℂ
|
2018-01-17 11:16:07 +00:00
|
|
|
|
presheaf {ℂ = ℂ} B = record
|
2018-01-21 13:31:37 +00:00
|
|
|
|
{ func* = λ A → ℂ .Arrow A B
|
|
|
|
|
; func→ = λ f g → ℂ ._⊕_ g f
|
|
|
|
|
; ident = funExt λ x → fst ident
|
|
|
|
|
; distrib = funExt λ x → assoc
|
2018-01-15 15:13:23 +00:00
|
|
|
|
}
|
|
|
|
|
where
|
2018-01-21 13:31:37 +00:00
|
|
|
|
open IsCategory (ℂ .isCategory)
|