cat/src/Cat/Category/Functor.agda

144 lines
4.3 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.Functor where
open import Agda.Primitive
open import Cubical
open import Function
open import Cat.Category
open Category hiding (_∘_ ; raw)
module _ {c c' d d'}
( : Category c c')
(𝔻 : Category d d')
where
private
= c c' d d'
𝓤 = Set
record RawFunctor : 𝓤 where
field
func* : Object Object 𝔻
func→ : {A B} [ A , B ] 𝔻 [ func* A , func* B ]
record IsFunctor (F : RawFunctor) : 𝓤 where
open RawFunctor F
field
ident : {c : Object } func→ (𝟙 {c}) 𝟙 𝔻 {func* c}
distrib : {A B C : Object } {f : [ A , B ]} {g : [ B , C ]}
func→ ( [ g f ]) 𝔻 [ func→ g func→ f ]
record Functor : Set (c c' d d') where
field
raw : RawFunctor
{{isFunctor}} : IsFunctor raw
private
module R = RawFunctor raw
func* : Object Object 𝔻
func* = R.func*
func→ : {A B} [ A , B ] 𝔻 [ func* A , func* B ]
func→ = R.func→
open IsFunctor
open Functor
module _
{a b : Level}
{ 𝔻 : Category a b}
{F : RawFunctor 𝔻}
where
private
module 𝔻 = IsCategory (isCategory 𝔻)
propIsFunctor : isProp (IsFunctor _ _ F)
propIsFunctor isF0 isF1 i = record
{ ident = 𝔻.arrowIsSet _ _ isF0.ident isF1.ident i
; distrib = 𝔻.arrowIsSet _ _ isF0.distrib isF1.distrib i
}
where
module isF0 = IsFunctor isF0
module isF1 = IsFunctor isF1
-- Alternate version of above where `F` is indexed by an interval
module _
{a b : Level}
{ 𝔻 : Category a b}
{F : I RawFunctor 𝔻}
where
private
module 𝔻 = IsCategory (isCategory 𝔻)
IsProp' : { : Level} (A : I Set ) Set
IsProp' A = (a0 : A i0) (a1 : A i1) A [ a0 a1 ]
IsFunctorIsProp' : IsProp' λ i IsFunctor _ _ (F i)
IsFunctorIsProp' isF0 isF1 = lemPropF {B = IsFunctor 𝔻}
(\ F propIsFunctor {F = F}) (\ i F i)
where
open import Cubical.NType.Properties using (lemPropF)
module _ { ' : Level} { 𝔻 : Category '} where
Functor≡ : {F G : Functor 𝔻}
(eq* : func* F func* G)
(eq→ : (λ i {x y} [ x , y ] 𝔻 [ eq* i x , eq* i y ])
[ func→ F func→ G ])
F G
Functor≡ {F} {G} eq* eq→ i = record
{ raw = eqR i
; isFunctor = eqIsF i
}
where
eqR : raw F raw G
eqR i = record { func* = eq* i ; func→ = eq→ i }
eqIsF : (λ i IsFunctor 𝔻 (eqR i)) [ isFunctor F isFunctor G ]
eqIsF = IsFunctorIsProp' (isFunctor F) (isFunctor G)
module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : Functor A B) where
private
F* = func* F
F→ = func→ F
G* = func* G
G→ = func→ G
module _ {a0 a1 a2 : Object A} {α0 : A [ a0 , a1 ]} {α1 : A [ a1 , a2 ]} where
dist : (F→ G→) (A [ α1 α0 ]) C [ (F→ G→) α1 (F→ G→) α0 ]
dist = begin
(F→ G→) (A [ α1 α0 ]) ≡⟨ refl
F→ (G→ (A [ α1 α0 ])) ≡⟨ cong F→ (G .isFunctor .distrib)
F→ (B [ G→ α1 G→ α0 ]) ≡⟨ F .isFunctor .distrib
C [ (F→ G→) α1 (F→ G→) α0 ]
_∘fr_ : RawFunctor A C
RawFunctor.func* _∘fr_ = F* G*
RawFunctor.func→ _∘fr_ = F→ G→
instance
isFunctor' : IsFunctor A C _∘fr_
isFunctor' = record
{ ident = begin
(F→ G→) (𝟙 A) ≡⟨ refl
F→ (G→ (𝟙 A)) ≡⟨ cong F→ (G .isFunctor .ident)
F→ (𝟙 B) ≡⟨ F .isFunctor .ident
𝟙 C
; distrib = dist
}
_∘f_ : Functor A C
raw _∘f_ = _∘fr_
-- The identity functor
identity : { '} {C : Category '} Functor C C
identity = record
{ raw = record
{ func* = λ x x
; func→ = λ x x
}
; isFunctor = record
{ ident = refl
; distrib = refl
}
}