2018-02-07 19:19:17 +00:00
|
|
|
|
{-# OPTIONS --cubical #-}
|
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
|
|
|
|
|
|
2018-02-06 13:24:34 +00:00
|
|
|
|
open Category hiding (_∘_ ; raw)
|
2018-01-30 15:23:36 +00:00
|
|
|
|
|
2018-02-06 13:24:34 +00:00
|
|
|
|
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
|
2018-01-30 15:23:36 +00:00
|
|
|
|
field
|
2018-02-05 11:21:39 +00:00
|
|
|
|
ident : {c : Object ℂ} → func→ (𝟙 ℂ {c}) ≡ 𝟙 𝔻 {func* c}
|
|
|
|
|
distrib : {A B C : Object ℂ} {f : ℂ [ A , B ]} {g : ℂ [ B , C ]}
|
2018-01-30 17:26:11 +00:00
|
|
|
|
→ func→ (ℂ [ g ∘ f ]) ≡ 𝔻 [ func→ g ∘ func→ f ]
|
2018-01-30 15:23:36 +00:00
|
|
|
|
|
|
|
|
|
record Functor : Set (ℓc ⊔ ℓc' ⊔ ℓd ⊔ ℓd') where
|
|
|
|
|
field
|
2018-02-06 13:24:34 +00:00
|
|
|
|
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→
|
2018-01-08 21:54:53 +00:00
|
|
|
|
|
2018-01-30 15:23:36 +00:00
|
|
|
|
open IsFunctor
|
2018-01-25 11:11:50 +00:00
|
|
|
|
open Functor
|
|
|
|
|
|
2018-02-06 13:24:34 +00:00
|
|
|
|
module _
|
|
|
|
|
{ℓa ℓb : Level}
|
|
|
|
|
{ℂ 𝔻 : Category ℓa ℓb}
|
|
|
|
|
{F : RawFunctor ℂ 𝔻}
|
|
|
|
|
where
|
|
|
|
|
private
|
|
|
|
|
module 𝔻 = IsCategory (isCategory 𝔻)
|
|
|
|
|
|
2018-02-20 13:08:47 +00:00
|
|
|
|
propIsFunctor : isProp (IsFunctor _ _ F)
|
|
|
|
|
propIsFunctor isF0 isF1 i = record
|
2018-02-09 11:09:59 +00:00
|
|
|
|
{ ident = 𝔻.arrowIsSet _ _ isF0.ident isF1.ident i
|
|
|
|
|
; distrib = 𝔻.arrowIsSet _ _ isF0.distrib isF1.distrib i
|
2018-02-06 13:24:34 +00:00
|
|
|
|
}
|
|
|
|
|
where
|
|
|
|
|
module isF0 = IsFunctor isF0
|
|
|
|
|
module isF1 = IsFunctor isF1
|
|
|
|
|
|
2018-02-09 11:09:59 +00:00
|
|
|
|
-- Alternate version of above where `F` is indexed by an interval
|
2018-02-06 13:24:34 +00:00
|
|
|
|
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 ]
|
2018-01-30 15:23:36 +00:00
|
|
|
|
|
2018-02-07 19:19:17 +00:00
|
|
|
|
IsFunctorIsProp' : IsProp' λ i → IsFunctor _ _ (F i)
|
|
|
|
|
IsFunctorIsProp' isF0 isF1 = lemPropF {B = IsFunctor ℂ 𝔻}
|
2018-02-20 13:08:47 +00:00
|
|
|
|
(\ F → propIsFunctor {F = F}) (\ i → F i)
|
2018-02-07 19:19:17 +00:00
|
|
|
|
where
|
2018-02-16 10:36:44 +00:00
|
|
|
|
open import Cubical.NType.Properties using (lemPropF)
|
2018-01-30 15:23:36 +00:00
|
|
|
|
|
2018-02-06 13:24:34 +00:00
|
|
|
|
module _ {ℓ ℓ' : Level} {ℂ 𝔻 : Category ℓ ℓ'} where
|
2018-01-25 11:11:50 +00:00
|
|
|
|
Functor≡ : {F G : Functor ℂ 𝔻}
|
2018-02-06 13:24:34 +00:00
|
|
|
|
→ (eq* : func* F ≡ func* G)
|
2018-01-31 13:39:54 +00:00
|
|
|
|
→ (eq→ : (λ i → ∀ {x y} → ℂ [ x , y ] → 𝔻 [ eq* i x , eq* i y ])
|
2018-02-06 13:24:34 +00:00
|
|
|
|
[ func→ F ≡ func→ G ])
|
2018-01-25 11:11:50 +00:00
|
|
|
|
→ F ≡ G
|
2018-02-06 13:24:34 +00:00
|
|
|
|
Functor≡ {F} {G} eq* eq→ i = record
|
|
|
|
|
{ raw = eqR i
|
2018-02-07 19:19:17 +00:00
|
|
|
|
; isFunctor = eqIsF i
|
2018-02-06 13:24:34 +00:00
|
|
|
|
}
|
|
|
|
|
where
|
|
|
|
|
eqR : raw F ≡ raw G
|
|
|
|
|
eqR i = record { func* = eq* i ; func→ = eq→ i }
|
2018-02-07 19:19:17 +00:00
|
|
|
|
eqIsF : (λ i → IsFunctor ℂ 𝔻 (eqR i)) [ isFunctor F ≡ isFunctor G ]
|
|
|
|
|
eqIsF = IsFunctorIsProp' (isFunctor F) (isFunctor G)
|
2018-01-25 11:11:50 +00:00
|
|
|
|
|
2018-01-21 00:11:08 +00:00
|
|
|
|
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-02-06 13:24:34 +00:00
|
|
|
|
F* = func* F
|
|
|
|
|
F→ = func→ F
|
|
|
|
|
G* = func* G
|
|
|
|
|
G→ = func→ G
|
2018-02-05 11:21:39 +00:00
|
|
|
|
module _ {a0 a1 a2 : Object A} {α0 : A [ a0 , a1 ]} {α1 : A [ a1 , a2 ]} where
|
2018-01-08 21:54:53 +00:00
|
|
|
|
|
2018-01-30 17:26:11 +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
|
2018-01-30 17:26:11 +00:00
|
|
|
|
(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
|
|
|
|
|
2018-02-06 13:24:34 +00:00
|
|
|
|
_∘fr_ : RawFunctor A C
|
|
|
|
|
RawFunctor.func* _∘fr_ = F* ∘ G*
|
|
|
|
|
RawFunctor.func→ _∘fr_ = F→ ∘ G→
|
|
|
|
|
instance
|
|
|
|
|
isFunctor' : IsFunctor A C _∘fr_
|
|
|
|
|
isFunctor' = record
|
2018-01-30 15:23:36 +00:00
|
|
|
|
{ ident = begin
|
2018-02-05 11:21:39 +00:00
|
|
|
|
(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-02-06 13:24:34 +00:00
|
|
|
|
|
|
|
|
|
_∘f_ : Functor A C
|
|
|
|
|
raw _∘f_ = _∘fr_
|
2018-01-08 21:54:53 +00:00
|
|
|
|
|
|
|
|
|
-- The identity functor
|
2018-01-21 00:11:08 +00:00
|
|
|
|
identity : ∀ {ℓ ℓ'} → {C : Category ℓ ℓ'} → Functor C C
|
2018-01-15 15:13:23 +00:00
|
|
|
|
identity = record
|
2018-02-06 13:24:34 +00:00
|
|
|
|
{ raw = record
|
|
|
|
|
{ func* = λ x → x
|
|
|
|
|
; func→ = λ x → x
|
|
|
|
|
}
|
2018-01-30 15:23:36 +00:00
|
|
|
|
; isFunctor = record
|
|
|
|
|
{ ident = refl
|
|
|
|
|
; distrib = refl
|
|
|
|
|
}
|
2018-01-15 15:13:23 +00:00
|
|
|
|
}
|