Naming and formatting

This commit is contained in:
Frederik Hanghøj Iversen 2018-03-20 16:26:23 +01:00
parent 63a51fbfdc
commit b6a9befd9c
2 changed files with 40 additions and 36 deletions

View file

@ -2,9 +2,11 @@
module Cat.Category.Functor where module Cat.Category.Functor where
open import Agda.Primitive open import Agda.Primitive
open import Cubical
open import Function open import Function
open import Cubical
open import Cubical.NType.Properties using (lemPropF)
open import Cat.Category open import Cat.Category
open Category hiding (_∘_ ; raw ; IsIdentity) open Category hiding (_∘_ ; raw ; IsIdentity)
@ -72,8 +74,6 @@ module _ {c c' d d'}
open IsFunctor isFunctor public open IsFunctor isFunctor public
open Functor
EndoFunctor : {a b} ( : Category a b) Set _ EndoFunctor : {a b} ( : Category a b) Set _
EndoFunctor = Functor EndoFunctor = Functor
@ -87,7 +87,7 @@ module _
propIsFunctor : isProp (IsFunctor _ _ F) propIsFunctor : isProp (IsFunctor _ _ F)
propIsFunctor isF0 isF1 i = record propIsFunctor isF0 isF1 i = record
{ isIdentity = 𝔻.arrowsAreSets _ _ isF0.isIdentity isF1.isIdentity i { isIdentity = 𝔻.arrowsAreSets _ _ isF0.isIdentity isF1.isIdentity i
; isDistributive = 𝔻.arrowsAreSets _ _ isF0.isDistributive isF1.isDistributive i ; isDistributive = 𝔻.arrowsAreSets _ _ isF0.isDistributive isF1.isDistributive i
} }
where where
@ -108,15 +108,14 @@ module _
IsFunctorIsProp' : IsProp' λ i IsFunctor _ _ (F i) IsFunctorIsProp' : IsProp' λ i IsFunctor _ _ (F i)
IsFunctorIsProp' isF0 isF1 = lemPropF {B = IsFunctor 𝔻} IsFunctorIsProp' isF0 isF1 = lemPropF {B = IsFunctor 𝔻}
(\ F propIsFunctor F) (\ i F i) (\ F propIsFunctor F) (\ i F i)
where
open import Cubical.NType.Properties using (lemPropF)
module _ { ' : Level} { 𝔻 : Category '} where module _ { ' : Level} { 𝔻 : Category '} where
open Functor
Functor≡ : {F G : Functor 𝔻} Functor≡ : {F G : Functor 𝔻}
raw F raw G Functor.raw F Functor.raw G
F G F G
raw (Functor≡ eq i) = eq i Functor.raw (Functor≡ eq i) = eq i
isFunctor (Functor≡ {F} {G} eq i) Functor.isFunctor (Functor≡ {F} {G} eq i)
= res i = res i
where where
res : (λ i IsFunctor 𝔻 (eq i)) [ isFunctor F isFunctor G ] res : (λ i IsFunctor 𝔻 (eq i)) [ isFunctor F isFunctor G ]
@ -124,35 +123,37 @@ module _ { ' : Level} { 𝔻 : Category '} where
module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : Functor A B) where module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : Functor A B) where
private private
F* = omap F module F = Functor F
F→ = fmap F module G = Functor G
G* = omap G
G→ = fmap G
module _ {a0 a1 a2 : Object A} {α0 : A [ a0 , a1 ]} {α1 : A [ a1 , a2 ]} where module _ {a0 a1 a2 : Object A} {α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→ G→) (A [ α1 α0 ]) C [ (F→ G→) α1 (F→ G→) α0 ]
dist = begin dist = begin
(F→ G→) (A [ α1 α0 ]) ≡⟨ refl (F.fmap G.fmap) (A [ α1 α0 ])
F→ (G→ (A [ α1 α0 ])) ≡⟨ cong F→ (isDistributive G) ≡⟨ refl
F→ (B [ G→ α1 G→ α0 ]) ≡⟨ isDistributive F F.fmap (G.fmap (A [ α1 α0 ]))
C [ (F→ G→) α1 (F→ G→) α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 ]
_∘fr_ : RawFunctor A C raw : RawFunctor A C
RawFunctor.omap _∘fr_ = F* G* RawFunctor.omap raw = F.omap G.omap
RawFunctor.fmap _∘fr_ = F→ G→ RawFunctor.fmap raw = F.fmap G.fmap
instance
isFunctor' : IsFunctor A C _∘fr_ isFunctor : IsFunctor A C raw
isFunctor' = record isFunctor = record
{ isIdentity = begin { isIdentity = begin
(F→ G→) (𝟙 A) ≡⟨ refl (F.fmap G.fmap) (𝟙 A) ≡⟨ refl
F→ (G→ (𝟙 A)) ≡⟨ cong F (isIdentity G) F.fmap (G.fmap (𝟙 A)) ≡⟨ cong F.fmap (G.isIdentity)
F→ (𝟙 B) ≡⟨ isIdentity F F.fmap (𝟙 B) ≡⟨ F.isIdentity
𝟙 C 𝟙 C
; isDistributive = dist ; isDistributive = dist
} }
F[_∘_] : Functor A C F[_∘_] : Functor A C
raw F[_∘_] = _∘fr_ Functor.raw F[_∘_] = raw
Functor.isFunctor F[_∘_] = isFunctor
-- The identity functor -- The identity functor
identity : { '} {C : Category '} Functor C C identity : { '} {C : Category '} Functor C C

View file

@ -113,8 +113,11 @@ module NaturalTransformation {c c' d d' : Level}
lem : (λ _ Natural F G θ) [ (λ f θNat f) (λ f θNat' f) ] lem : (λ _ Natural F G θ) [ (λ f θNat f) (λ f θNat' f) ]
lem = λ i f 𝔻.arrowsAreSets _ _ (θNat f) (θNat' f) i lem = λ i f 𝔻.arrowsAreSets _ _ (θNat f) (θNat' f) i
naturalTransformationIsSet : isSet (NaturalTransformation F G) naturalIsSet : (θ : Transformation F G) isSet (Natural F G θ)
naturalTransformationIsSet = sigPresSet transformationIsSet naturalIsSet θ =
λ θ ntypeCommulative ntypeCommulative
(s≤s {n = Nat.suc Nat.zero} z≤n) (s≤s {n = Nat.suc Nat.zero} z≤n)
(naturalIsProp θ) (naturalIsProp θ)
naturalTransformationIsSet : isSet (NaturalTransformation F G)
naturalTransformationIsSet = sigPresSet transformationIsSet naturalIsSet