cat/src/Cat/Categories/Sets.agda

156 lines
4.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.

-- | The category of homotopy sets
{-# OPTIONS --cubical --caching #-}
module Cat.Categories.Sets where
open import Cat.Prelude as P
open import Cat.Category
open import Cat.Category.Functor
open import Cat.Category.Product
open import Cat.Equivalence
_⊙_ : {a b c : Level} {A : Set a} {B : Set b} {C : Set c} (A B) (B C) A C
eqA eqB = Equivalence.compose eqA eqB
sym≃ : {a b} {A : Set a} {B : Set b} A B B A
sym≃ = Equivalence.symmetry
infixl 10 _⊙_
module _ ( : Level) where
private
SetsRaw : RawCategory (lsuc )
RawCategory.Object SetsRaw = hSet
RawCategory.Arrow SetsRaw (T , _) (U , _) = T U
RawCategory.identity SetsRaw = idFun _
RawCategory._<<<_ SetsRaw = _∘_
module _ where
private
open RawCategory SetsRaw hiding (_<<<_)
isIdentity : IsIdentity (idFun _)
fst isIdentity = funExt λ _ refl
snd isIdentity = funExt λ _ refl
arrowsAreSets : ArrowsAreSets
arrowsAreSets {B = (_ , s)} = setPi λ _ s
isPreCat : IsPreCategory SetsRaw
IsPreCategory.isAssociative isPreCat = refl
IsPreCategory.isIdentity isPreCat {A} {B} = isIdentity {A} {B}
IsPreCategory.arrowsAreSets isPreCat {A} {B} = arrowsAreSets {A} {B}
open IsPreCategory isPreCat
module _ {hA hB : Object} where
open Σ hA renaming (fst to A ; snd to sA)
open Σ hB renaming (fst to B ; snd to sB)
univ≃ : (hA hB) (hA hB)
univ≃
= equivSigProp (λ A isSetIsProp)
univalence
equivSig {P = isEquiv A B} {Q = TypeIsomorphism} (equiv≃iso sA sB)
univalent : Univalent
univalent = univalenceFrom≃ univ≃
SetsIsCategory : IsCategory SetsRaw
IsCategory.isPreCategory SetsIsCategory = isPreCat
IsCategory.univalent SetsIsCategory = univalent
𝓢𝓮𝓽 Sets : Category (lsuc )
Category.raw 𝓢𝓮𝓽 = SetsRaw
Category.isCategory 𝓢𝓮𝓽 = SetsIsCategory
Sets = 𝓢𝓮𝓽
module _ { : Level} where
private
𝓢 = 𝓢𝓮𝓽
open Category 𝓢
module _ (hA hB : Object) where
open Σ hA renaming (fst to A ; snd to sA)
open Σ hB renaming (fst to B ; snd to sB)
private
productObject : Object
productObject = (A × B) , sigPresSet sA λ _ sB
module _ {X A B : Set } (f : X A) (g : X B) where
_&&&_ : (X A × B)
_&&&_ x = f x , g x
module _ (hX : Object) where
open Σ hX renaming (fst to X)
module _ (f : X A ) (g : X B) where
ump : fst ∘′ (f &&& g) f × snd ∘′ (f &&& g) g
fst ump = refl
snd ump = refl
rawProduct : RawProduct 𝓢 hA hB
RawProduct.object rawProduct = productObject
RawProduct.fst rawProduct = fst
RawProduct.snd rawProduct = snd
isProduct : IsProduct 𝓢 _ _ rawProduct
IsProduct.ump isProduct {X = hX} f g
= f &&& g , ump hX f g , λ eq funExt (umpUniq eq)
where
open Σ hX renaming (fst to X) using ()
module _ {y : X A × B} (eq : fst ∘′ y f × snd ∘′ y g) (x : X) where
p1 : fst ((f &&& g) x) fst (y x)
p1 = begin
fst ((f &&& g) x) ≡⟨⟩
f x ≡⟨ (λ i sym (fst eq) i x)
fst (y x)
p2 : snd ((f &&& g) x) snd (y x)
p2 = λ i sym (snd eq) i x
umpUniq : (f &&& g) x y x
umpUniq i = p1 i , p2 i
product : Product 𝓢 hA hB
Product.raw product = rawProduct
Product.isProduct product = isProduct
instance
SetsHasProducts : HasProducts 𝓢
SetsHasProducts = record { product = product }
module _ {a b : Level} ( : Category a b) where
open Category
-- Covariant Presheaf
Representable : Set (a lsuc b)
Representable = Functor (𝓢𝓮𝓽 b)
-- Contravariant Presheaf
Presheaf : Set (a lsuc b)
Presheaf = Functor (opposite ) (𝓢𝓮𝓽 b)
-- The "co-yoneda" embedding.
representable : Category.Object Representable
representable A = record
{ raw = record
{ omap = λ B [ A , B ] , arrowsAreSets
; fmap = [_∘_]
}
; isFunctor = record
{ isIdentity = funExt λ _ leftIdentity
; isDistributive = funExt λ x sym isAssociative
}
}
-- Alternate name: `yoneda`
presheaf : Category.Object (opposite ) Presheaf
presheaf B = record
{ raw = record
{ omap = λ A [ A , B ] , arrowsAreSets
; fmap = λ f g [ g f ]
}
; isFunctor = record
{ isIdentity = funExt λ x rightIdentity
; isDistributive = funExt λ x isAssociative
}
}