Move identity functor laws to functor module...

and make progress on univalence in the functor category
This commit is contained in:
Frederik Hanghøj Iversen 2018-03-23 13:55:03 +01:00
parent a713d560d5
commit ef688202a2
8 changed files with 148 additions and 71 deletions

View file

@ -14,32 +14,12 @@ open import Cat.Categories.Fun
-- The category of categories -- The category of categories
module _ ( ' : Level) where module _ ( ' : Level) where
private
module _ {𝔸 𝔹 𝔻 : Category '} {F : Functor 𝔸 𝔹} {G : Functor 𝔹 } {H : Functor 𝔻} where
assc : F[ H F[ G F ] ] F[ F[ H G ] F ]
assc = Functor≡ refl
module _ { 𝔻 : Category '} {F : Functor 𝔻} where
ident-r : F[ F identity ] F
ident-r = Functor≡ refl
ident-l : F[ identity F ] F
ident-l = Functor≡ refl
RawCat : RawCategory (lsuc ( ')) ( ') RawCat : RawCategory (lsuc ( ')) ( ')
RawCategory.Object RawCat = Category ' RawCategory.Object RawCat = Category '
RawCategory.Arrow RawCat = Functor RawCategory.Arrow RawCat = Functor
RawCategory.𝟙 RawCat = identity RawCategory.𝟙 RawCat = Functors.identity
RawCategory._∘_ RawCat = F[_∘_] RawCategory._∘_ RawCat = F[_∘_]
private
open RawCategory RawCat
isAssociative : IsAssociative
isAssociative {f = F} {G} {H} = assc {F = F} {G = G} {H = H}
isIdentity : IsIdentity identity
isIdentity = ident-l , ident-r
-- NB! `ArrowsAreSets RawCat` is *not* provable. The type of functors, -- NB! `ArrowsAreSets RawCat` is *not* provable. The type of functors,
-- however, form a groupoid! Therefore there is no (1-)category of -- however, form a groupoid! Therefore there is no (1-)category of
-- categories. There does, however, exist a 2-category of 1-categories. -- categories. There does, however, exist a 2-category of 1-categories.
@ -283,7 +263,7 @@ module CatExponential { : Level} ( 𝔻 : Category ) where
: Functor 𝔸 object Functor : Functor 𝔸 object Functor
Functor (𝔸 ) (object ) Functor (𝔸 ) (object )
transpose : Functor 𝔸 object transpose : Functor 𝔸 object
eq : F[ eval (parallelProduct transpose (identity {C = })) ] F eq : F[ eval (parallelProduct transpose (Functors.identity { = })) ] F
-- eq : F[ :eval: ∘ {!!} ] ≡ F -- eq : F[ :eval: ∘ {!!} ] ≡ F
-- eq : Cat [ :eval: ∘ (HasProducts._|×|_ hasProducts transpose (𝟙 Cat {o = })) ] ≡ F -- eq : Cat [ :eval: ∘ (HasProducts._|×|_ hasProducts transpose (𝟙 Cat {o = })) ] ≡ F
-- eq' : (Cat [ :eval: ∘ -- eq' : (Cat [ :eval: ∘

View file

@ -4,7 +4,7 @@ module Cat.Categories.Fun where
open import Cat.Prelude open import Cat.Prelude
open import Cat.Category open import Cat.Category
open import Cat.Category.Functor hiding (identity) open import Cat.Category.Functor
open import Cat.Category.NaturalTransformation open import Cat.Category.NaturalTransformation
module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : Category d d') where module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : Category d d') where
@ -14,20 +14,18 @@ module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : C
module = Category module = Category
module 𝔻 = Category 𝔻 module 𝔻 = Category 𝔻
private private
module _ {A B C D : Functor 𝔻} {θ' : NaturalTransformation A B} module _ {A B C D : Functor 𝔻} {θNT : NaturalTransformation A B}
{η' : NaturalTransformation B C} {ζ' : NaturalTransformation C D} where {ηNT : NaturalTransformation B C} {ζNT : NaturalTransformation C D} where
θ = proj₁ θ' open Σ θNT renaming (proj₁ to θ ; proj₂ to θNat)
η = proj₁ η' open Σ ηNT renaming (proj₁ to η ; proj₂ to ηNat)
ζ = proj₁ ζ' open Σ ζNT renaming (proj₁ to ζ ; proj₂ to ζNat)
θNat = proj₂ θ' private
ηNat = proj₂ η' L : NaturalTransformation A D
ζNat = proj₂ ζ' L = (NT[_∘_] {A} {C} {D} ζNT (NT[_∘_] {A} {B} {C} ηNT θNT))
L : NaturalTransformation A D R : NaturalTransformation A D
L = (NT[_∘_] {A} {C} {D} ζ' (NT[_∘_] {A} {B} {C} η' θ')) R = (NT[_∘_] {A} {B} {D} (NT[_∘_] {B} {C} {D} ζNT ηNT) θNT)
R : NaturalTransformation A D _g⊕f_ = NT[_∘_] {A} {B} {C}
R = (NT[_∘_] {A} {B} {D} (NT[_∘_] {B} {C} {D} ζ' η') θ') _h⊕g_ = NT[_∘_] {B} {C} {D}
_g⊕f_ = NT[_∘_] {A} {B} {C}
_h⊕g_ = NT[_∘_] {B} {C} {D}
isAssociative : L R isAssociative : L R
isAssociative = lemSig (naturalIsProp {F = A} {D}) isAssociative = lemSig (naturalIsProp {F = A} {D})
L R (funExt (λ x 𝔻.isAssociative)) L R (funExt (λ x 𝔻.isAssociative))
@ -62,6 +60,71 @@ module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : C
open RawCategory RawFun open RawCategory RawFun
open Univalence (λ {A} {B} {f} isIdentity {A} {B} {f}) open Univalence (λ {A} {B} {f} isIdentity {A} {B} {f})
private
module _ (F : Functor 𝔻) where
center : Σ[ G Object ] (F G)
center = F , id-to-iso F F refl
open Σ center renaming (proj₂ to isoF)
module _ (cG : Σ[ G Object ] (F G)) where
open Σ cG renaming (proj₁ to G ; proj₂ to isoG)
module G = Functor G
open Σ isoG renaming (proj₁ to θNT ; proj₂ to invθNT)
open Σ invθNT renaming (proj₁ to ηNT ; proj₂ to areInv)
open Σ θNT renaming (proj₁ to θ ; proj₂ to θN)
open Σ ηNT renaming (proj₁ to η ; proj₂ to ηN)
open Σ areInv renaming (proj₁ to ve-re ; proj₂ to re-ve)
-- f ~ Transformation G G
-- f : (X : .Object) → 𝔻 [ G.omap X , G.omap X ]
-- f X = T[ θ ∘ η ] X
-- g = T[ η ∘ θ ] {!!}
ntF : NaturalTransformation F F
ntF = 𝟙 {A = F}
ntG : NaturalTransformation G G
ntG = 𝟙 {A = G}
idFunctor = Functors.identity
-- Dunno if this is the way to go, but if I can construct a an inverse of
-- G that is also inverse of F (possibly by being propositionally equal to
-- another functor F~)
postulate
G~ : Functor 𝔻
F~ : Functor 𝔻
F~ = G~
postulate
prop0 : F[ G~ G ] idFunctor
prop1 : F[ F G~ ] idFunctor
lem : F[ F F~ ] idFunctor
lem = begin
F[ F F~ ] ≡⟨⟩
F[ F G~ ] ≡⟨ prop1
idFunctor
open import Cubical.Univalence
p0 : F G
p0 = begin
F ≡⟨ sym Functors.rightIdentity
F[ F idFunctor ] ≡⟨ cong (λ φ F[ F φ ]) (sym prop0)
F[ F F[ G~ G ] ] ≡⟨ Functors.isAssociative {F = G} {G = G~} {H = F}
F[ F[ F G~ ] G ] ≡⟨⟩
F[ F[ F F~ ] G ] ≡⟨ cong (λ φ F[ φ G ]) lem
F[ idFunctor G ] ≡⟨ Functors.leftIdentity
G
p1 : (λ i Σ (Arrow F (p0 i)) (Isomorphism {A = F} {B = p0 i})) [ isoF isoG ]
p1 = {!!}
isContractible : (F , isoF) (G , isoG)
isContractible i = p0 i , p1 i
univalent[Contr] : isContr (Σ[ G Object ] (F G))
univalent[Contr] = center , isContractible
private private
module _ {A B : Functor 𝔻} where module _ {A B : Functor 𝔻} where

View file

@ -1,7 +1,7 @@
{-# OPTIONS --cubical #-} {-# OPTIONS --cubical #-}
module Cat.Category.Functor where module Cat.Category.Functor where
open import Agda.Primitive open import Cat.Prelude
open import Function open import Function
open import Cubical open import Cubical
@ -9,31 +9,31 @@ open import Cubical.NType.Properties using (lemPropF)
open import Cat.Category open import Cat.Category
open Category hiding (_∘_ ; raw ; IsIdentity)
module _ {c c' d d'} module _ {c c' d d'}
( : Category c c') ( : Category c c')
(𝔻 : Category d d') (𝔻 : Category d d')
where where
private private
module = Category
module 𝔻 = Category 𝔻
= c c' d d' = c c' d d'
𝓤 = Set 𝓤 = Set
Omap = Object Object 𝔻 Omap = .Object 𝔻.Object
Fmap : Omap Set _ Fmap : Omap Set _
Fmap omap = {A B} Fmap omap = {A B}
[ A , B ] 𝔻 [ omap A , omap B ] [ A , B ] 𝔻 [ omap A , omap B ]
record RawFunctor : 𝓤 where record RawFunctor : 𝓤 where
field field
omap : Object Object 𝔻 omap : .Object 𝔻.Object
fmap : {A B} [ A , B ] 𝔻 [ omap A , omap B ] fmap : {A B} [ A , B ] 𝔻 [ omap A , omap B ]
IsIdentity : Set _ IsIdentity : Set _
IsIdentity = {A : Object } fmap (𝟙 {A}) 𝟙 𝔻 {omap A} IsIdentity = {A : .Object} fmap (.𝟙 {A}) 𝔻.𝟙 {omap A}
IsDistributive : Set _ IsDistributive : Set _
IsDistributive = {A B C : Object } {f : [ A , B ]} {g : [ B , C ]} IsDistributive = {A B C : .Object} {f : [ A , B ]} {g : [ B , C ]}
fmap ( [ g f ]) 𝔻 [ fmap g fmap f ] fmap ( [ g f ]) 𝔻 [ fmap g fmap f ]
-- | Equality principle for raw functors -- | Equality principle for raw functors
@ -120,11 +120,18 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
res : (λ i IsFunctor 𝔻 (eq i)) [ isFunctor F isFunctor G ] res : (λ i IsFunctor 𝔻 (eq i)) [ isFunctor F isFunctor G ]
res = IsFunctorIsProp' (isFunctor F) (isFunctor G) res = IsFunctorIsProp' (isFunctor F) (isFunctor G)
module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : Functor A B) where 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 private
module A = Category A
module B = Category B
module C = Category C
module F = Functor F module F = Functor F
module G = Functor G module G = Functor G
module _ {a0 a1 a2 : Object A} {α0 : A [ a0 , a1 ]} {α1 : A [ a1 , a2 ]} where 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 : (F.fmap G.fmap) (A [ α1 α0 ]) C [ (F.fmap G.fmap) α1 (F.fmap G.fmap) α0 ]
dist = begin dist = begin
(F.fmap G.fmap) (A [ α1 α0 ]) (F.fmap G.fmap) (A [ α1 α0 ])
@ -143,10 +150,10 @@ module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : F
isFunctor : IsFunctor A C raw isFunctor : IsFunctor A C raw
isFunctor = record isFunctor = record
{ isIdentity = begin { isIdentity = begin
(F.fmap G.fmap) (𝟙 A) ≡⟨ refl (F.fmap G.fmap) A.𝟙 ≡⟨ refl
F.fmap (G.fmap (𝟙 A)) ≡⟨ cong F.fmap (G.isIdentity) F.fmap (G.fmap A.𝟙) ≡⟨ cong F.fmap (G.isIdentity)
F.fmap (𝟙 B) ≡⟨ F.isIdentity F.fmap B.𝟙 ≡⟨ F.isIdentity
𝟙 C C.𝟙
; isDistributive = dist ; isDistributive = dist
} }
@ -154,15 +161,42 @@ module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : F
Functor.raw F[_∘_] = raw Functor.raw F[_∘_] = raw
Functor.isFunctor F[_∘_] = isFunctor Functor.isFunctor F[_∘_] = isFunctor
-- The identity functor -- | The identity functor
identity : { '} {C : Category '} Functor C C module Functors where
identity = record module _ {c cc : Level} { : Category c cc} where
{ raw = record private
{ omap = λ x x raw : RawFunctor
; fmap = λ x x RawFunctor.omap raw = Function.id
} RawFunctor.fmap raw = Function.id
; isFunctor = record
{ isIdentity = refl isFunctor : IsFunctor raw
; isDistributive = refl 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

View file

@ -176,9 +176,9 @@ module _ {a b : Level} ( : Category a b) where
pureTEq : M.RawMonad.pureT (backRaw (forth m)) pureT pureTEq : M.RawMonad.pureT (backRaw (forth m)) pureT
pureTEq = funExt (λ X refl) pureTEq = funExt (λ X refl)
pureNTEq : (λ i NaturalTransformation F.identity (Req i)) pureNTEq : (λ i NaturalTransformation Functors.identity (Req i))
[ M.RawMonad.pureNT (backRaw (forth m)) pureNT ] [ M.RawMonad.pureNT (backRaw (forth m)) pureNT ]
pureNTEq = lemSigP (λ i propIsNatural F.identity (Req i)) _ _ pureTEq pureNTEq = lemSigP (λ i propIsNatural Functors.identity (Req i)) _ _ pureTEq
joinTEq : M.RawMonad.joinT (backRaw (forth m)) joinT joinTEq : M.RawMonad.joinT (backRaw (forth m)) joinT
joinTEq = funExt (λ X begin joinTEq = funExt (λ X begin

View file

@ -119,7 +119,7 @@ record IsMonad (raw : RawMonad) : Set where
open NaturalTransformation open NaturalTransformation
R⁰ : EndoFunctor R⁰ : EndoFunctor
R⁰ = F.identity R⁰ = Functors.identity
: EndoFunctor : EndoFunctor
= F[ R R ] = F[ R R ]
module R = Functor R module R = Functor R

View file

@ -22,14 +22,14 @@ open NaturalTransformation
record RawMonad : Set where record RawMonad : Set where
field field
R : EndoFunctor R : EndoFunctor
pureNT : NaturalTransformation F.identity R pureNT : NaturalTransformation Functors.identity R
joinNT : NaturalTransformation F[ R R ] R joinNT : NaturalTransformation F[ R R ] R
-- Note that `pureT` and `joinT` differs from their definition in the -- Note that `pureT` and `joinT` differs from their definition in the
-- kleisli formulation only by having an explicit parameter. -- kleisli formulation only by having an explicit parameter.
pureT : Transformation F.identity R pureT : Transformation Functors.identity R
pureT = proj₁ pureNT pureT = proj₁ pureNT
pureN : Natural F.identity R pureT pureN : Natural Functors.identity R pureT
pureN = proj₂ pureNT pureN = proj₂ pureNT
joinT : Transformation F[ R R ] R joinT : Transformation F[ R R ] R

View file

@ -50,9 +50,9 @@ module voe {a b : Level} ( : Category a b) where
pureT X = pure {X} pureT X = pure {X}
field field
pureN : Natural F.identity R pureT pureN : Natural Functors.identity R pureT
pureNT : NaturalTransformation F.identity R pureNT : NaturalTransformation Functors.identity R
pureNT = pureT , pureN pureNT = pureT , pureN
joinT : (A : Object) [ omap (omap A) , omap A ] joinT : (A : Object) [ omap (omap A) , omap A ]

View file

@ -26,7 +26,7 @@ open import Data.Nat using (_≤_ ; z≤n ; s≤s)
module Nat = Data.Nat module Nat = Data.Nat
open import Cat.Category open import Cat.Category
open import Cat.Category.Functor hiding (identity) open import Cat.Category.Functor
open import Cat.Wishlist open import Cat.Wishlist
module NaturalTransformation {c c' d d' : Level} module NaturalTransformation {c c' d d' : Level}