cat/src/Cat/Category/Functor.agda

98 lines
3.7 KiB
Agda
Raw Normal View History

2018-02-05 13:59:53 +00:00
module Cat.Category.Functor where
2018-01-08 21:54:53 +00:00
open import Agda.Primitive
open import Cubical
open import Function
open import Cat.Category
open Category hiding (_∘_)
2018-01-30 15:23:36 +00:00
module _ {c c' d d'} ( : Category c c') (𝔻 : Category d d') where
record IsFunctor
(func* : Object Object 𝔻)
(func→ : {A B : Object } [ A , B ] 𝔻 [ func* A , func* B ])
: Set (c c' d d') where
2018-01-30 15:23:36 +00:00
field
ident : {c : Object } func→ (𝟙 {c}) 𝟙 𝔻 {func* c}
2018-01-30 15:23:36 +00:00
-- TODO: Avoid use of ugly explicit arguments somehow.
-- This guy managed to do it:
-- https://github.com/copumpkin/categories/blob/master/Categories/Functor/Core.agda
distrib : {A B C : Object } {f : [ A , B ]} {g : [ B , C ]}
func→ ( [ g f ]) 𝔻 [ func→ g func→ f ]
2018-01-30 15:23:36 +00:00
record Functor : Set (c c' d d') where
field
func* : Object Object 𝔻
func→ : {A B} [ A , B ] 𝔻 [ func* A , func* B ]
2018-01-30 15:23:36 +00:00
{{isFunctor}} : IsFunctor func* func→
2018-01-08 21:54:53 +00:00
2018-01-30 15:23:36 +00:00
open IsFunctor
open Functor
module _ { ' : Level} { 𝔻 : Category '} where
2018-01-30 15:23:36 +00:00
IsFunctor≡
: {func* : Object Object 𝔻}
{func→ : {A B : Object } [ A , B ] 𝔻 [ func* A , func* B ]}
{F G : IsFunctor 𝔻 func* func→}
(eqI
: (λ i {A} func→ (𝟙 {A}) 𝟙 𝔻 {func* A})
[ F .ident G .ident ])
(eqD :
(λ i {A B C} {f : [ A , B ]} {g : [ B , C ]}
func→ ( [ g f ]) 𝔻 [ func→ g func→ f ])
[ F .distrib G .distrib ])
(λ _ IsFunctor 𝔻 (λ i func* i) func→) [ F G ]
IsFunctor≡ eqI eqD i = record { ident = eqI i ; distrib = eqD i }
2018-01-30 15:23:36 +00:00
Functor≡ : {F G : Functor 𝔻}
(eq* : F .func* G .func*)
2018-01-31 13:39:54 +00:00
(eq→ : (λ i {x y} [ x , y ] 𝔻 [ eq* i x , eq* i y ])
[ F .func→ G .func→ ])
2018-01-30 15:23:36 +00:00
-- → (eqIsF : PathP (λ i → IsFunctor 𝔻 (eq* i) (eq→ i)) (F .isFunctor) (G .isFunctor))
(eqIsFunctor : (λ i IsFunctor 𝔻 (eq* i) (eq→ i)) [ F .isFunctor G .isFunctor ])
F G
Functor≡ eq* eq→ eqIsFunctor i = record { func* = eq* i ; func→ = eq→ i ; isFunctor = eqIsFunctor i }
module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : Functor A B) where
2018-01-08 21:54:53 +00:00
private
2018-01-21 14:19:15 +00:00
F* = F .func*
F→ = F .func→
G* = G .func*
G→ = G .func→
module _ {a0 a1 a2 : Object A} {α0 : A [ a0 , a1 ]} {α1 : A [ a1 , a2 ]} where
2018-01-08 21:54:53 +00:00
dist : (F→ G→) (A [ α1 α0 ]) C [ (F→ G→) α1 (F→ G→) α0 ]
2018-01-08 21:54:53 +00:00
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 ]
2018-01-08 21:54:53 +00:00
_∘f_ : Functor A C
_∘f_ =
2018-01-08 21:54:53 +00:00
record
{ func* = F* G*
; func→ = F→ G→
2018-01-30 15:23:36 +00:00
; isFunctor = record
{ ident = begin
(F→ G→) (𝟙 A) ≡⟨ refl
F→ (G→ (𝟙 A)) ≡⟨ cong F→ (G .isFunctor .ident)
F→ (𝟙 B) ≡⟨ F .isFunctor .ident
𝟙 C
2018-01-30 15:23:36 +00:00
; distrib = dist
}
2018-01-08 21:54:53 +00:00
}
-- The identity functor
identity : { '} {C : Category '} Functor C C
identity = record
{ func* = λ x x
; func→ = λ x x
2018-01-30 15:23:36 +00:00
; isFunctor = record
{ ident = refl
; distrib = refl
}
}