cat/src/Cat/Functor.agda

102 lines
4.1 KiB
Agda
Raw Normal View History

2018-01-08 21:54:53 +00:00
module Cat.Functor where
open import Agda.Primitive
open import Cubical
open import Function
open import Cat.Category
2018-01-30 15:23:36 +00:00
open Category
module _ {c c' d d'} ( : Category c c') (𝔻 : Category d d') where
record IsFunctor
(func* : .Object 𝔻 .Object)
(func→ : {A B : .Object} .Arrow A B 𝔻 .Arrow (func* A) (func* B))
: Set (c c' d d') where
field
ident : { c : .Object } func→ ( .𝟙 {c}) 𝔻 .𝟙 {func* c}
-- 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 : .Arrow A B} {g : .Arrow B C}
func→ ( ._⊕_ g f) 𝔻 ._⊕_ (func→ g) (func→ f)
record Functor : Set (c c' d d') where
field
func* : .Object 𝔻 .Object
func→ : {A B : .Object} .Arrow A B 𝔻 .Arrow (func* A) (func* B)
{{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
private
_⊕_ = ._⊕_
2018-01-30 15:23:36 +00:00
-- IsFunctor≡ : ∀ {A B : .Object} {func* : .Object → 𝔻 .Object} {func→ : {A B : .Object} → .Arrow A B → 𝔻 .Arrow (func* A) (func* B)} {F G : IsFunctor 𝔻 func* func→}
-- → (eqI : PathP (λ i → ∀ {A : .Object} → func→ ( .𝟙 {A}) ≡ 𝔻 .𝟙 {func* A})
-- (F .ident) (G .ident))
-- → (eqD : PathP (λ i → {A B C : .Object} {f : .Arrow A B} {g : .Arrow B C}
-- → func→ ( ._⊕_ g f) ≡ 𝔻 ._⊕_ (func→ g) (func→ f))
-- (F .distrib) (G .distrib))
-- → F ≡ G
-- IsFunctor≡ eqI eqD i = record { ident = eqI i ; distrib = eqD i }
Functor≡ : {F G : Functor 𝔻}
(eq* : F .func* G .func*)
(eq→ : PathP (λ i {x y} .Arrow x y 𝔻 .Arrow (eq* i x) (eq* i y))
2018-01-25 11:44:47 +00:00
(F .func→) (G .func→))
2018-01-30 15:23:36 +00:00
-- → (eqIsF : PathP (λ i → IsFunctor 𝔻 (eq* i) (eq→ i)) (F .isFunctor) (G .isFunctor))
(eqI : PathP (λ i {A : .Object} eq→ i ( .𝟙 {A}) 𝔻 .𝟙 {eq* i A})
2018-01-30 15:23:36 +00:00
(F .isFunctor .ident) (G .isFunctor .ident))
(eqD : PathP (λ i {A B C : .Object} {f : .Arrow A B} {g : .Arrow B C}
2018-01-25 11:44:47 +00:00
eq→ i ( ._⊕_ g f) 𝔻 ._⊕_ (eq→ i g) (eq→ i f))
2018-01-30 15:23:36 +00:00
(F .isFunctor .distrib) (G .isFunctor .distrib))
F G
2018-01-30 15:23:36 +00:00
Functor≡ eq* eq→ eqI eqD i = record { func* = eq* i ; func→ = eq→ i ; isFunctor = record { ident = eqI i ; distrib = eqD 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→
_A⊕_ = A ._⊕_
_B⊕_ = B ._⊕_
_C⊕_ = C ._⊕_
module _ {a0 a1 a2 : A .Object} {α0 : A .Arrow a0 a1} {α1 : A .Arrow a1 a2} where
2018-01-08 21:54:53 +00:00
2018-01-21 14:19:15 +00:00
dist : (F→ G→) (α1 A⊕ α0) (F→ G→) α1 C⊕ (F→ G→) α0
2018-01-08 21:54:53 +00:00
dist = begin
2018-01-21 14:19:15 +00:00
(F→ G→) (α1 A⊕ α0) ≡⟨ refl
2018-01-30 15:23:36 +00:00
F→ (G→ (α1 A⊕ α0)) ≡⟨ cong F→ (G .isFunctor .distrib)
F→ ((G→ α1) B⊕ (G→ α0)) ≡⟨ F .isFunctor .distrib
2018-01-21 14:19:15 +00:00
(F→ G→) α1 C⊕ (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 .𝟙
; 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
}
}