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 Function
|
|
|
|
|
|
2018-03-20 15:26:23 +00:00
|
|
|
|
open import Cubical
|
|
|
|
|
open import Cubical.NType.Properties using (lemPropF)
|
|
|
|
|
|
2018-01-08 21:54:53 +00:00
|
|
|
|
open import Cat.Category
|
|
|
|
|
|
2018-02-23 11:41:15 +00:00
|
|
|
|
open Category hiding (_∘_ ; raw ; IsIdentity)
|
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 ℓ
|
|
|
|
|
|
2018-03-05 09:28:16 +00:00
|
|
|
|
Omap = Object ℂ → Object 𝔻
|
|
|
|
|
Fmap : Omap → Set _
|
|
|
|
|
Fmap omap = ∀ {A B}
|
|
|
|
|
→ ℂ [ A , B ] → 𝔻 [ omap A , omap B ]
|
2018-02-06 13:24:34 +00:00
|
|
|
|
record RawFunctor : 𝓤 where
|
|
|
|
|
field
|
2018-03-08 10:03:56 +00:00
|
|
|
|
omap : Object ℂ → Object 𝔻
|
|
|
|
|
fmap : ∀ {A B} → ℂ [ A , B ] → 𝔻 [ omap A , omap B ]
|
2018-02-06 13:24:34 +00:00
|
|
|
|
|
2018-02-23 11:41:15 +00:00
|
|
|
|
IsIdentity : Set _
|
2018-03-08 10:03:56 +00:00
|
|
|
|
IsIdentity = {A : Object ℂ} → fmap (𝟙 ℂ {A}) ≡ 𝟙 𝔻 {omap A}
|
2018-02-23 11:41:15 +00:00
|
|
|
|
|
|
|
|
|
IsDistributive : Set _
|
|
|
|
|
IsDistributive = {A B C : Object ℂ} {f : ℂ [ A , B ]} {g : ℂ [ B , C ]}
|
2018-03-08 10:03:56 +00:00
|
|
|
|
→ fmap (ℂ [ g ∘ f ]) ≡ 𝔻 [ fmap g ∘ fmap f ]
|
2018-02-23 11:41:15 +00:00
|
|
|
|
|
2018-03-05 09:28:16 +00:00
|
|
|
|
-- | Equality principle for raw functors
|
|
|
|
|
--
|
2018-03-08 10:03:56 +00:00
|
|
|
|
-- The type of `fmap` depend on the value of `omap`. We can wrap this up
|
2018-03-05 09:28:16 +00:00
|
|
|
|
-- into an equality principle for this type like is done for e.g. `Σ` using
|
|
|
|
|
-- `pathJ`.
|
|
|
|
|
module _ {x y : RawFunctor} where
|
|
|
|
|
open RawFunctor
|
|
|
|
|
private
|
2018-03-08 10:03:56 +00:00
|
|
|
|
P : (omap' : Omap) → (eq : omap x ≡ omap') → Set _
|
2018-03-05 09:28:16 +00:00
|
|
|
|
P y eq = (fmap' : Fmap y) → (λ i → Fmap (eq i))
|
2018-03-08 10:03:56 +00:00
|
|
|
|
[ fmap x ≡ fmap' ]
|
2018-03-05 09:28:16 +00:00
|
|
|
|
module _
|
2018-03-08 10:03:56 +00:00
|
|
|
|
(eq : (λ i → Omap) [ omap x ≡ omap y ])
|
|
|
|
|
(kk : P (omap x) refl)
|
2018-03-05 09:28:16 +00:00
|
|
|
|
where
|
|
|
|
|
private
|
2018-03-08 10:03:56 +00:00
|
|
|
|
p : P (omap y) eq
|
|
|
|
|
p = pathJ P kk (omap y) eq
|
|
|
|
|
eq→ : (λ i → Fmap (eq i)) [ fmap x ≡ fmap y ]
|
|
|
|
|
eq→ = p (fmap y)
|
2018-03-05 09:28:16 +00:00
|
|
|
|
RawFunctor≡ : x ≡ y
|
2018-03-08 10:03:56 +00:00
|
|
|
|
omap (RawFunctor≡ i) = eq i
|
|
|
|
|
fmap (RawFunctor≡ i) = eq→ i
|
2018-03-05 09:28:16 +00:00
|
|
|
|
|
2018-02-06 13:24:34 +00:00
|
|
|
|
record IsFunctor (F : RawFunctor) : 𝓤 where
|
2018-02-23 11:21:16 +00:00
|
|
|
|
open RawFunctor F public
|
2018-01-30 15:23:36 +00:00
|
|
|
|
field
|
2018-03-08 00:09:40 +00:00
|
|
|
|
-- FIXME Really ought to be preserves identity or something like this.
|
2018-02-23 11:53:35 +00:00
|
|
|
|
isIdentity : IsIdentity
|
|
|
|
|
isDistributive : IsDistributive
|
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
|
|
|
|
|
|
2018-02-23 11:21:16 +00:00
|
|
|
|
open IsFunctor isFunctor public
|
2018-02-06 13:24:34 +00:00
|
|
|
|
|
2018-02-28 18:03:11 +00:00
|
|
|
|
EndoFunctor : ∀ {ℓa ℓb} (ℂ : Category ℓa ℓb) → Set _
|
|
|
|
|
EndoFunctor ℂ = Functor ℂ ℂ
|
|
|
|
|
|
2018-02-06 13:24:34 +00:00
|
|
|
|
module _
|
2018-03-21 16:52:32 +00:00
|
|
|
|
{ℓc ℓc' ℓd ℓd' : Level}
|
|
|
|
|
{ℂ : Category ℓc ℓc'} {𝔻 : Category ℓd ℓd'}
|
2018-02-23 09:36:59 +00:00
|
|
|
|
(F : RawFunctor ℂ 𝔻)
|
2018-02-06 13:24:34 +00:00
|
|
|
|
where
|
|
|
|
|
private
|
2018-02-23 15:41:17 +00:00
|
|
|
|
module 𝔻 = Category 𝔻
|
2018-02-06 13:24:34 +00:00
|
|
|
|
|
2018-02-20 13:08:47 +00:00
|
|
|
|
propIsFunctor : isProp (IsFunctor _ _ F)
|
|
|
|
|
propIsFunctor isF0 isF1 i = record
|
2018-03-20 15:26:23 +00:00
|
|
|
|
{ isIdentity = 𝔻.arrowsAreSets _ _ isF0.isIdentity isF1.isIdentity i
|
2018-02-23 11:53:35 +00:00
|
|
|
|
; isDistributive = 𝔻.arrowsAreSets _ _ isF0.isDistributive isF1.isDistributive 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 _
|
2018-03-21 16:52:32 +00:00
|
|
|
|
{ℓc ℓc' ℓd ℓd' : Level} {ℂ : Category ℓc ℓc'} {𝔻 : Category ℓd ℓd'}
|
2018-02-06 13:24:34 +00:00
|
|
|
|
{F : I → RawFunctor ℂ 𝔻}
|
|
|
|
|
where
|
|
|
|
|
private
|
2018-02-23 15:41:17 +00:00
|
|
|
|
module 𝔻 = Category 𝔻
|
2018-02-06 13:24:34 +00:00
|
|
|
|
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-23 09:36:59 +00:00
|
|
|
|
(\ F → propIsFunctor F) (\ i → F i)
|
2018-01-30 15:23:36 +00:00
|
|
|
|
|
2018-03-21 16:52:32 +00:00
|
|
|
|
module _ {ℓc ℓc' ℓd ℓd' : Level} {ℂ : Category ℓc ℓc'} {𝔻 : Category ℓd ℓd'} where
|
2018-03-20 15:26:23 +00:00
|
|
|
|
open Functor
|
2018-01-25 11:11:50 +00:00
|
|
|
|
Functor≡ : {F G : Functor ℂ 𝔻}
|
2018-03-20 15:26:23 +00:00
|
|
|
|
→ Functor.raw F ≡ Functor.raw G
|
2018-03-05 09:28:16 +00:00
|
|
|
|
→ F ≡ G
|
2018-03-20 15:26:23 +00:00
|
|
|
|
Functor.raw (Functor≡ eq i) = eq i
|
|
|
|
|
Functor.isFunctor (Functor≡ {F} {G} eq i)
|
2018-03-05 09:28:16 +00:00
|
|
|
|
= res i
|
|
|
|
|
where
|
|
|
|
|
res : (λ i → IsFunctor ℂ 𝔻 (eq i)) [ isFunctor F ≡ isFunctor G ]
|
|
|
|
|
res = IsFunctorIsProp' (isFunctor F) (isFunctor G)
|
|
|
|
|
|
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-03-20 15:26:23 +00:00
|
|
|
|
module F = Functor F
|
|
|
|
|
module G = Functor 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-03-20 15:26:23 +00:00
|
|
|
|
dist : (F.fmap ∘ G.fmap) (A [ α1 ∘ α0 ]) ≡ C [ (F.fmap ∘ G.fmap) α1 ∘ (F.fmap ∘ G.fmap) α0 ]
|
2018-01-08 21:54:53 +00:00
|
|
|
|
dist = begin
|
2018-03-20 15:26:23 +00:00
|
|
|
|
(F.fmap ∘ G.fmap) (A [ α1 ∘ α0 ])
|
|
|
|
|
≡⟨ refl ⟩
|
|
|
|
|
F.fmap (G.fmap (A [ α1 ∘ α0 ]))
|
|
|
|
|
≡⟨ cong F.fmap G.isDistributive ⟩
|
|
|
|
|
F.fmap (B [ G.fmap α1 ∘ G.fmap α0 ])
|
|
|
|
|
≡⟨ F.isDistributive ⟩
|
|
|
|
|
C [ (F.fmap ∘ G.fmap) α1 ∘ (F.fmap ∘ G.fmap) α0 ]
|
|
|
|
|
∎
|
|
|
|
|
|
|
|
|
|
raw : RawFunctor A C
|
|
|
|
|
RawFunctor.omap raw = F.omap ∘ G.omap
|
|
|
|
|
RawFunctor.fmap raw = F.fmap ∘ G.fmap
|
|
|
|
|
|
|
|
|
|
isFunctor : IsFunctor A C raw
|
|
|
|
|
isFunctor = record
|
|
|
|
|
{ isIdentity = begin
|
|
|
|
|
(F.fmap ∘ G.fmap) (𝟙 A) ≡⟨ refl ⟩
|
|
|
|
|
F.fmap (G.fmap (𝟙 A)) ≡⟨ cong F.fmap (G.isIdentity)⟩
|
|
|
|
|
F.fmap (𝟙 B) ≡⟨ F.isIdentity ⟩
|
|
|
|
|
𝟙 C ∎
|
|
|
|
|
; isDistributive = dist
|
|
|
|
|
}
|
2018-02-06 13:24:34 +00:00
|
|
|
|
|
2018-02-24 11:55:08 +00:00
|
|
|
|
F[_∘_] : Functor A C
|
2018-03-20 15:26:23 +00:00
|
|
|
|
Functor.raw F[_∘_] = raw
|
|
|
|
|
Functor.isFunctor F[_∘_] = isFunctor
|
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
|
2018-03-08 10:03:56 +00:00
|
|
|
|
{ omap = λ x → x
|
|
|
|
|
; fmap = λ x → x
|
2018-02-06 13:24:34 +00:00
|
|
|
|
}
|
2018-01-30 15:23:36 +00:00
|
|
|
|
; isFunctor = record
|
2018-02-23 11:49:41 +00:00
|
|
|
|
{ isIdentity = refl
|
2018-02-23 11:53:35 +00:00
|
|
|
|
; isDistributive = refl
|
2018-01-30 15:23:36 +00:00
|
|
|
|
}
|
2018-01-15 15:13:23 +00:00
|
|
|
|
}
|