cat/src/Cat/Category/Yoneda.agda

85 lines
2.9 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

{-# OPTIONS --cubical #-}
module Cat.Category.Yoneda where
open import Cat.Prelude
open import Cat.Category
open import Cat.Category.Functor
open import Cat.Category.NaturalTransformation
renaming (module Properties to F)
using ()
open import Cat.Categories.Fun using (module Fun)
open import Cat.Categories.Sets hiding (presheaf)
-- There is no (small) category of categories. So we won't use _⇑_ from
-- `HasExponential`
--
-- open HasExponentials (Cat.hasExponentials unprovable) using (_⇑_)
--
-- In stead we'll use an ad-hoc definition -- which is definitionally equivalent
-- to that other one - even without mentioning the category of categories.
_⇑_ : { : Level} Category Category Category
_⇑_ = Fun.Fun
module _ { : Level} { : Category } where
private
𝓢 = Sets
open Fun (opposite ) 𝓢
module = Category
presheaf : .Object Presheaf
presheaf = Cat.Categories.Sets.presheaf
module _ {A B : .Object} (f : [ A , B ]) where
fmap : Transformation (presheaf A) (presheaf B)
fmap C x = [ f x ]
fmapNatural : Natural (presheaf A) (presheaf B) fmap
fmapNatural g = funExt λ _ .isAssociative
fmapNT : NaturalTransformation (presheaf A) (presheaf B)
fmapNT = fmap , fmapNatural
rawYoneda : RawFunctor Fun
RawFunctor.omap rawYoneda = presheaf
RawFunctor.fmap rawYoneda = fmapNT
open RawFunctor rawYoneda hiding (fmap)
isIdentity : IsIdentity
isIdentity {c} = lemSig prp _ _ eq
where
eq : (λ C x [ .identity x ]) identityTrans (presheaf c)
eq = funExt λ A funExt λ B .leftIdentity
prp = F.naturalIsProp _ _ {F = presheaf c} {presheaf c}
isDistributive : IsDistributive
isDistributive {A} {B} {C} {f = f} {g}
= lemSig (propIsNatural (presheaf A) (presheaf C)) _ _ eq
where
T[_∘_]' = T[_∘_] {F = presheaf A} {presheaf B} {presheaf C}
eqq : (X : .Object) (x : [ X , A ])
fmap ( [ g f ]) X x T[ fmap g fmap f ]' X x
eqq X x = begin
fmap ( [ g f ]) X x ≡⟨⟩
[ [ g f ] x ] ≡⟨ sym .isAssociative
[ g [ f x ] ] ≡⟨⟩
[ g fmap f X x ] ≡⟨⟩
T[ fmap g fmap f ]' X x
eq : fmap ( [ g f ]) T[ fmap g fmap f ]'
eq = begin
fmap ( [ g f ]) ≡⟨ funExt (λ X funExt λ α eqq X α)
T[ fmap g fmap f ]'
instance
isFunctor : IsFunctor Fun rawYoneda
IsFunctor.isIdentity isFunctor = isIdentity
IsFunctor.isDistributive isFunctor = isDistributive
yoneda : Functor Fun
Functor.raw yoneda = rawYoneda
Functor.isFunctor yoneda = isFunctor