cat/src/Cat/Category/Functor.agda

201 lines
6.4 KiB
Agda
Raw Normal View History

{-# 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 Cat.Prelude
2018-01-08 21:54:53 +00:00
2018-03-20 15:26:23 +00:00
open import Cubical
2018-01-08 21:54:53 +00:00
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 ]
2018-02-23 11:41:15 +00:00
IsIdentity : Set _
IsIdentity = {A : .Object} fmap (.identity {A}) 𝔻.identity {omap A}
2018-02-23 11:41:15 +00:00
IsDistributive : Set _
IsDistributive = {A B C : .Object} {f : [ A , B ]} {g : [ B , C ]}
fmap ( [ g f ]) 𝔻 [ fmap g fmap f ]
2018-02-23 11:41:15 +00:00
-- | 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
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
raw : RawFunctor
{{isFunctor}} : IsFunctor raw
2018-02-23 11:21:16 +00:00
open IsFunctor isFunctor public
2018-02-28 18:03:11 +00:00
EndoFunctor : {a b} ( : Category a b) Set _
EndoFunctor = Functor
module _
{c c' d d' : Level}
{ : Category c c'} {𝔻 : Category d d'}
2018-02-23 09:36:59 +00:00
(F : RawFunctor 𝔻)
where
private
2018-02-23 15:41:17 +00:00
module 𝔻 = Category 𝔻
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
}
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
2018-02-23 15:41:17 +00:00
module 𝔻 = Category 𝔻
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
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
module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Category d d'} where
2018-03-20 15:26:23 +00:00
open Functor
Functor≡ : {F G : Functor 𝔻}
2018-03-20 15:26:23 +00:00
Functor.raw F Functor.raw G
F G
2018-03-20 15:26:23 +00:00
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
2018-01-08 21:54:53 +00:00
private
module A = Category A
module B = Category B
module C = Category C
2018-03-20 15:26:23 +00:00
module F = Functor F
module G = Functor G
module _ {a0 a1 a2 : A.Object} {α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.identity ≡⟨ refl
F.fmap (G.fmap A.identity) ≡⟨ cong F.fmap (G.isIdentity)
F.fmap B.identity ≡⟨ F.isIdentity
C.identity
2018-03-20 15:26:23 +00:00
; isDistributive = dist
}
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
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