cat/src/Cat/Functor.agda

63 lines
2.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
record Functor {c c' d d'} (C : Category c c') (D : Category d d')
2018-01-08 21:54:53 +00:00
: Set (c c' d d') where
2018-01-22 10:35:37 +00:00
open Category
2018-01-08 21:54:53 +00:00
field
2018-01-22 10:35:37 +00:00
func* : C .Object D .Object
func→ : {dom cod : C .Object} C .Arrow dom cod D .Arrow (func* dom) (func* cod)
ident : { c : C .Object } func→ (C .𝟙 {c}) D .𝟙 {func* c}
2018-01-08 21:54:53 +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
2018-01-22 10:35:37 +00:00
distrib : { c c' c'' : C .Object} {a : C .Arrow c c'} {a' : C .Arrow c' c''}
func→ (C ._⊕_ a' a) D ._⊕_ (func→ a') (func→ a)
2018-01-08 21:54:53 +00:00
module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : Functor A B) where
2018-01-21 14:19:15 +00:00
open Functor
open Category
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
F→ (G→ (α1 A⊕ α0)) ≡⟨ cong F→ (G .distrib)
F→ ((G→ α1) B⊕ (G→ α0)) ≡⟨ F .distrib
(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→
; ident = begin
2018-01-21 14:19:15 +00:00
(F→ G→) (A .𝟙) ≡⟨ refl
F→ (G→ (A .𝟙)) ≡⟨ cong F→ (G .ident)
F→ (B .𝟙) ≡⟨ F .ident
C .𝟙
2018-01-08 21:54:53 +00:00
; distrib = dist
}
-- The identity functor
identity : { '} {C : Category '} Functor C C
identity = record
{ func* = λ x x
; func→ = λ x x
; ident = refl
; distrib = refl
}