cat/src/Cat/Category/Functor.agda

201 lines
6.4 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

{-# OPTIONS --cubical #-}
module Cat.Category.Functor where
open import Cat.Prelude
open import Cubical
open import Cat.Category
module _ {c c' d d'}
( : Category c c')
(𝔻 : Category d d')
where
private
module = Category
module 𝔻 = Category 𝔻
= c c' d d'
𝓤 = Set
Omap = .Object 𝔻.Object
Fmap : Omap Set _
Fmap omap = {A B}
[ A , B ] 𝔻 [ omap A , omap B ]
record RawFunctor : 𝓤 where
field
omap : .Object 𝔻.Object
fmap : {A B} [ A , B ] 𝔻 [ omap A , omap B ]
IsIdentity : Set _
IsIdentity = {A : .Object} fmap (.identity {A}) 𝔻.identity {omap A}
IsDistributive : Set _
IsDistributive = {A B C : .Object} {f : [ A , B ]} {g : [ B , C ]}
fmap ( [ g f ]) 𝔻 [ fmap g fmap f ]
-- | Equality principle for raw functors
--
-- The type of `fmap` depend on the value of `omap`. We can wrap this up
-- into an equality principle for this type like is done for e.g. `Σ` using
-- `pathJ`.
module _ {x y : RawFunctor} where
open RawFunctor
private
P : (omap' : Omap) (eq : omap x omap') Set _
P y eq = (fmap' : Fmap y) (λ i Fmap (eq i))
[ fmap x fmap' ]
module _
(eq : (λ i Omap) [ omap x omap y ])
(kk : P (omap x) refl)
where
private
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)
RawFunctor≡ : x y
omap (RawFunctor≡ i) = eq i
fmap (RawFunctor≡ i) = eq→ i
record IsFunctor (F : RawFunctor) : 𝓤 where
open RawFunctor F public
field
-- FIXME Really ought to be preserves identity or something like this.
isIdentity : IsIdentity
isDistributive : IsDistributive
record Functor : Set (c c' d d') where
field
raw : RawFunctor
{{isFunctor}} : IsFunctor raw
open IsFunctor isFunctor public
EndoFunctor : {a b} ( : Category a b) Set _
EndoFunctor = Functor
module _
{c c' d d' : Level}
{ : Category c c'} {𝔻 : Category d d'}
(F : RawFunctor 𝔻)
where
private
module 𝔻 = Category 𝔻
propIsFunctor : isProp (IsFunctor _ _ F)
propIsFunctor isF0 isF1 i = record
{ isIdentity = 𝔻.arrowsAreSets _ _ isF0.isIdentity isF1.isIdentity i
; isDistributive = 𝔻.arrowsAreSets _ _ isF0.isDistributive isF1.isDistributive i
}
where
module isF0 = IsFunctor isF0
module isF1 = IsFunctor isF1
-- Alternate version of above where `F` is indexed by an interval
module _
{c c' d d' : Level} { : Category c c'} {𝔻 : Category d d'}
{F : I RawFunctor 𝔻}
where
private
module 𝔻 = Category 𝔻
IsProp' : { : Level} (A : I Set ) Set
IsProp' A = (a0 : A i0) (a1 : A i1) A [ a0 a1 ]
IsFunctorIsProp' : IsProp' λ i IsFunctor _ _ (F i)
IsFunctorIsProp' isF0 isF1 = lemPropF {B = IsFunctor 𝔻}
(\ F propIsFunctor F) (\ i F i)
module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Category d d'} where
open Functor
Functor≡ : {F G : Functor 𝔻}
Functor.raw F Functor.raw G
F G
Functor.raw (Functor≡ eq i) = eq i
Functor.isFunctor (Functor≡ {F} {G} eq i)
= res i
where
res : (λ i IsFunctor 𝔻 (eq i)) [ isFunctor F isFunctor G ]
res = IsFunctorIsProp' (isFunctor F) (isFunctor G)
module _ {0 1 2 3 4 5 : Level}
{A : Category 0 1}
{B : Category 2 3}
{C : Category 4 5}
(F : Functor B C) (G : Functor A B) where
private
module A = Category A
module B = Category B
module C = Category C
module F = Functor F
module G = Functor G
module _ {a0 a1 a2 : A.Object} {α0 : A [ a0 , a1 ]} {α1 : A [ a1 , a2 ]} where
dist : (F.fmap G.fmap) (A [ α1 α0 ]) C [ (F.fmap G.fmap) α1 (F.fmap G.fmap) α0 ]
dist = begin
(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.identity ≡⟨ refl
F.fmap (G.fmap A.identity) ≡⟨ cong F.fmap (G.isIdentity)
F.fmap B.identity ≡⟨ F.isIdentity
C.identity
; isDistributive = dist
}
F[_∘_] : Functor A C
Functor.raw F[_∘_] = raw
Functor.isFunctor F[_∘_] = isFunctor
-- | The identity functor
module Functors where
module _ {c cc : Level} { : Category c cc} where
private
raw : RawFunctor
RawFunctor.omap raw = idFun _
RawFunctor.fmap raw = idFun _
isFunctor : IsFunctor raw
IsFunctor.isIdentity isFunctor = refl
IsFunctor.isDistributive isFunctor = refl
identity : Functor
Functor.raw identity = raw
Functor.isFunctor identity = isFunctor
module _
{a aa b bb c cc d dd : Level}
{𝔸 : Category a aa}
{𝔹 : Category b bb}
{ : Category c cc}
{𝔻 : Category d dd}
{F : Functor 𝔸 𝔹} {G : Functor 𝔹 } {H : Functor 𝔻} where
isAssociative : F[ H F[ G F ] ] F[ F[ H G ] F ]
isAssociative = Functor≡ refl
module _
{c cc d dd : Level}
{ : Category c cc}
{𝔻 : Category d dd}
{F : Functor 𝔻} where
leftIdentity : F[ identity F ] F
leftIdentity = Functor≡ refl
rightIdentity : F[ F identity ] F
rightIdentity = Functor≡ refl
isIdentity : F[ identity F ] F × F[ F identity ] F
isIdentity = leftIdentity , rightIdentity