Rename ident to isIdentity

This commit is contained in:
Frederik Hanghøj Iversen 2018-02-23 12:49:41 +01:00
parent 5cbc409770
commit 6446435a49
9 changed files with 55 additions and 55 deletions

View file

@ -91,13 +91,13 @@ module _ { ' : Level} (unprovable : IsCategory (RawCat ')) where
issSet = setSig {sA = C.arrowIsSet} {sB = λ x D.arrowIsSet} issSet = setSig {sA = C.arrowIsSet} {sB = λ x D.arrowIsSet}
ident' : IsIdentity :𝟙: ident' : IsIdentity :𝟙:
ident' ident'
= Σ≡ (fst C.ident) (fst D.ident) = Σ≡ (fst C.isIdentity) (fst D.isIdentity)
, Σ≡ (snd C.ident) (snd D.ident) , Σ≡ (snd C.isIdentity) (snd D.isIdentity)
postulate univalent : Univalence.Univalent :rawProduct: ident' postulate univalent : Univalence.Univalent :rawProduct: ident'
instance instance
:isCategory: : IsCategory :rawProduct: :isCategory: : IsCategory :rawProduct:
IsCategory.isAssociative :isCategory: = Σ≡ C.isAssociative D.isAssociative IsCategory.isAssociative :isCategory: = Σ≡ C.isAssociative D.isAssociative
IsCategory.ident :isCategory: = ident' IsCategory.isIdentity :isCategory: = ident'
IsCategory.arrowIsSet :isCategory: = issSet IsCategory.arrowIsSet :isCategory: = issSet
IsCategory.univalent :isCategory: = univalent IsCategory.univalent :isCategory: = univalent
@ -107,13 +107,13 @@ module _ { ' : Level} (unprovable : IsCategory (RawCat ')) where
proj₁ : Catt [ :product: , ] proj₁ : Catt [ :product: , ]
proj₁ = record proj₁ = record
{ raw = record { func* = fst ; func→ = fst } { raw = record { func* = fst ; func→ = fst }
; isFunctor = record { ident = refl ; distrib = refl } ; isFunctor = record { isIdentity = refl ; distrib = refl }
} }
proj₂ : Catt [ :product: , 𝔻 ] proj₂ : Catt [ :product: , 𝔻 ]
proj₂ = record proj₂ = record
{ raw = record { func* = snd ; func→ = snd } { raw = record { func* = snd ; func→ = snd }
; isFunctor = record { ident = refl ; distrib = refl } ; isFunctor = record { isIdentity = refl ; distrib = refl }
} }
module _ {X : Object Catt} (x₁ : Catt [ X , ]) (x₂ : Catt [ X , 𝔻 ]) where module _ {X : Object Catt} (x₁ : Catt [ X , ]) (x₂ : Catt [ X , 𝔻 ]) where
@ -124,7 +124,7 @@ module _ { ' : Level} (unprovable : IsCategory (RawCat ')) where
; func→ = λ x func→ x₁ x , func→ x₂ x ; func→ = λ x func→ x₁ x , func→ x₂ x
} }
; isFunctor = record ; isFunctor = record
{ ident = Σ≡ x₁.ident x₂.ident { isIdentity = Σ≡ x₁.isIdentity x₂.isIdentity
; distrib = Σ≡ x₁.distrib x₂.distrib ; distrib = Σ≡ x₁.distrib x₂.distrib
} }
} }
@ -230,7 +230,7 @@ module _ ( : Level) (unprovable : IsCategory (RawCat )) where
-- NaturalTransformation F G × .Arrow A B -- NaturalTransformation F G × .Arrow A B
-- :ident: : :func→: {c} {c} (identityNat F , .𝟙) 𝔻 .𝟙 -- :ident: : :func→: {c} {c} (identityNat F , .𝟙) 𝔻 .𝟙
-- :ident: = trans (proj₂ 𝔻.ident) (F .ident) -- :ident: = trans (proj₂ 𝔻.isIdentity) (F .isIdentity)
-- where -- where
-- open module 𝔻 = IsCategory (𝔻 .isCategory) -- open module 𝔻 = IsCategory (𝔻 .isCategory)
-- Unfortunately the equational version has some ambigous arguments. -- Unfortunately the equational version has some ambigous arguments.
@ -239,8 +239,8 @@ module _ ( : Level) (unprovable : IsCategory (RawCat )) where
:func→: {c} {c} (𝟙 (Product.obj (:obj: ×p )) {c}) ≡⟨⟩ :func→: {c} {c} (𝟙 (Product.obj (:obj: ×p )) {c}) ≡⟨⟩
:func→: {c} {c} (identityNat F , 𝟙 ) ≡⟨⟩ :func→: {c} {c} (identityNat F , 𝟙 ) ≡⟨⟩
𝔻 [ identityTrans F C func→ F (𝟙 )] ≡⟨⟩ 𝔻 [ identityTrans F C func→ F (𝟙 )] ≡⟨⟩
𝔻 [ 𝟙 𝔻 func→ F (𝟙 )] ≡⟨ proj₂ 𝔻.ident 𝔻 [ 𝟙 𝔻 func→ F (𝟙 )] ≡⟨ proj₂ 𝔻.isIdentity
func→ F (𝟙 ) ≡⟨ F.ident func→ F (𝟙 ) ≡⟨ F.isIdentity
𝟙 𝔻 𝟙 𝔻
where where
open module 𝔻 = Category 𝔻 open module 𝔻 = Category 𝔻
@ -313,7 +313,7 @@ module _ ( : Level) (unprovable : IsCategory (RawCat )) where
; func→ = λ {dom} {cod} :func→: {dom} {cod} ; func→ = λ {dom} {cod} :func→: {dom} {cod}
} }
; isFunctor = record ; isFunctor = record
{ ident = λ {o} :ident: {o} { isIdentity = λ {o} :ident: {o}
; distrib = λ {f u n k y} :distrib: {f} {u} {n} {k} {y} ; distrib = λ {f u n k y} :distrib: {f} {u} {n} {k} {y}
} }
} }

View file

@ -29,8 +29,8 @@ module _ (a b : Level) where
isAssociative = Σ≡ refl refl isAssociative = Σ≡ refl refl
module _ {A B : Obj'} {f : Arr A B} where module _ {A B : Obj'} {f : Arr A B} where
ident : B f one f × B one {B} f f isIdentity : B f one f × B one {B} f f
ident = (Σ≡ refl refl) , Σ≡ refl refl isIdentity = (Σ≡ refl refl) , Σ≡ refl refl
RawFam : RawCategory (lsuc (a b)) (a b) RawFam : RawCategory (lsuc (a b)) (a b)
@ -45,7 +45,7 @@ module _ (a b : Level) where
isCategory : IsCategory RawFam isCategory : IsCategory RawFam
isCategory = record isCategory = record
{ isAssociative = λ {A} {B} {C} {D} {f} {g} {h} isAssociative {D = D} {f} {g} {h} { isAssociative = λ {A} {B} {C} {D} {f} {g} {h} isAssociative {D = D} {f} {g} {h}
; ident = λ {A} {B} {f} ident {A} {B} {f = f} ; isIdentity = λ {A} {B} {f} isIdentity {A} {B} {f = f}
; arrowIsSet = {!!} ; arrowIsSet = {!!}
; univalent = {!!} ; univalent = {!!}
} }

View file

@ -58,7 +58,7 @@ module _ { ' : Level} ( : Category ') where
RawIsCategoryFree : IsCategory RawFree RawIsCategoryFree : IsCategory RawFree
RawIsCategoryFree = record RawIsCategoryFree = record
{ isAssociative = λ { {f = f} {g} {h} p-isAssociative {r = f} {g} {h}} { isAssociative = λ { {f = f} {g} {h} p-isAssociative {r = f} {g} {h}}
; ident = ident-r , ident-l ; isIdentity = ident-r , ident-l
; arrowIsSet = {!!} ; arrowIsSet = {!!}
; univalent = {!!} ; univalent = {!!}
} }

View file

@ -60,8 +60,8 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
identityNatural : (F : Functor 𝔻) Natural F F (identityTrans F) identityNatural : (F : Functor 𝔻) Natural F F (identityTrans F)
identityNatural F {A = A} {B = B} f = begin identityNatural F {A = A} {B = B} f = begin
𝔻 [ identityTrans F B F→ f ] ≡⟨⟩ 𝔻 [ identityTrans F B F→ f ] ≡⟨⟩
𝔻 [ 𝟙 𝔻 F→ f ] ≡⟨ proj₂ 𝔻.ident 𝔻 [ 𝟙 𝔻 F→ f ] ≡⟨ proj₂ 𝔻.isIdentity
F→ f ≡⟨ sym (proj₁ 𝔻.ident) F→ f ≡⟨ sym (proj₁ 𝔻.isIdentity)
𝔻 [ F→ f 𝟙 𝔻 ] ≡⟨⟩ 𝔻 [ F→ f 𝟙 𝔻 ] ≡⟨⟩
𝔻 [ F→ f identityTrans F A ] 𝔻 [ F→ f identityTrans F A ]
where where
@ -143,10 +143,10 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
eq-r : C (𝔻 [ f' C identityTrans A C ]) f' C eq-r : C (𝔻 [ f' C identityTrans A C ]) f' C
eq-r C = begin eq-r C = begin
𝔻 [ f' C identityTrans A C ] ≡⟨⟩ 𝔻 [ f' C identityTrans A C ] ≡⟨⟩
𝔻 [ f' C 𝔻.𝟙 ] ≡⟨ proj₁ 𝔻.ident 𝔻 [ f' C 𝔻.𝟙 ] ≡⟨ proj₁ 𝔻.isIdentity
f' C f' C
eq-l : C (𝔻 [ identityTrans B C f' C ]) f' C eq-l : C (𝔻 [ identityTrans B C f' C ]) f' C
eq-l C = proj₂ 𝔻.ident eq-l C = proj₂ 𝔻.isIdentity
ident-r : (_:⊕:_ {A} {A} {B} f (identityNat A)) f ident-r : (_:⊕:_ {A} {A} {B} f (identityNat A)) f
ident-r = lemSig allNatural _ _ (funExt eq-r) ident-r = lemSig allNatural _ _ (funExt eq-r)
ident-l : (_:⊕:_ {A} {B} {B} (identityNat B) f) f ident-l : (_:⊕:_ {A} {B} {B} (identityNat B) f) f
@ -169,7 +169,7 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
:isCategory: : IsCategory RawFun :isCategory: : IsCategory RawFun
:isCategory: = record :isCategory: = record
{ isAssociative = λ {A B C D} :isAssociative: {A} {B} {C} {D} { isAssociative = λ {A B C D} :isAssociative: {A} {B} {C} {D}
; ident = λ {A B} :ident: {A} {B} ; isIdentity = λ {A B} :ident: {A} {B}
; arrowIsSet = λ {F} {G} naturalTransformationIsSets {F} {G} ; arrowIsSet = λ {F} {G} naturalTransformationIsSets {F} {G}
; univalent = {!!} ; univalent = {!!}
} }

View file

@ -165,7 +165,7 @@ RawRel = record
RawIsCategoryRel : IsCategory RawRel RawIsCategoryRel : IsCategory RawRel
RawIsCategoryRel = record RawIsCategoryRel = record
{ isAssociative = funExt is-isAssociative { isAssociative = funExt is-isAssociative
; ident = funExt ident-l , funExt ident-r ; isIdentity = funExt ident-l , funExt ident-r
; arrowIsSet = {!!} ; arrowIsSet = {!!}
; univalent = {!!} ; univalent = {!!}
} }

View file

@ -26,8 +26,8 @@ module _ ( : Level) where
SetsIsCategory : IsCategory SetsRaw SetsIsCategory : IsCategory SetsRaw
isAssociative SetsIsCategory = refl isAssociative SetsIsCategory = refl
proj₁ (ident SetsIsCategory) = funExt λ _ refl proj₁ (isIdentity SetsIsCategory) = funExt λ _ refl
proj₂ (ident SetsIsCategory) = funExt λ _ refl proj₂ (isIdentity SetsIsCategory) = funExt λ _ refl
arrowIsSet SetsIsCategory {B = (_ , s)} = setPi λ _ s arrowIsSet SetsIsCategory {B = (_ , s)} = setPi λ _ s
univalent SetsIsCategory = {!!} univalent SetsIsCategory = {!!}
@ -98,7 +98,7 @@ module _ {a b : Level} where
; func→ = [_∘_] ; func→ = [_∘_]
} }
; isFunctor = record ; isFunctor = record
{ ident = funExt λ _ proj₂ ident { isIdentity = funExt λ _ proj₂ isIdentity
; distrib = funExt λ x sym isAssociative ; distrib = funExt λ x sym isAssociative
} }
} }
@ -113,7 +113,7 @@ module _ {a b : Level} where
; func→ = λ f g [ g f ] ; func→ = λ f g [ g f ]
} }
; isFunctor = record ; isFunctor = record
{ ident = funExt λ x proj₁ ident { isIdentity = funExt λ x proj₁ isIdentity
; distrib = funExt λ x isAssociative ; distrib = funExt λ x isAssociative
} }
} }

View file

@ -83,9 +83,9 @@ record RawCategory (a b : Level) : Set (lsuc (a ⊔ b)) where
-- Univalence is indexed by a raw category as well as an identity proof. -- Univalence is indexed by a raw category as well as an identity proof.
module Univalence {a b : Level} ( : RawCategory a b) where module Univalence {a b : Level} ( : RawCategory a b) where
open RawCategory open RawCategory
module _ (ident : IsIdentity 𝟙) where module _ (isIdentity : IsIdentity 𝟙) where
idIso : (A : Object) A A idIso : (A : Object) A A
idIso A = 𝟙 , (𝟙 , ident) idIso A = 𝟙 , (𝟙 , isIdentity)
-- Lemma 9.1.4 in [HoTT] -- Lemma 9.1.4 in [HoTT]
id-to-iso : (A B : Object) A B A B id-to-iso : (A B : Object) A B A B
@ -102,9 +102,9 @@ record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc
open Univalence public open Univalence public
field field
isAssociative : IsAssociative isAssociative : IsAssociative
ident : IsIdentity 𝟙 isIdentity : IsIdentity 𝟙
arrowIsSet : ArrowsAreSets arrowIsSet : ArrowsAreSets
univalent : Univalent ident univalent : Univalent isIdentity
-- `IsCategory` is a mere proposition. -- `IsCategory` is a mere proposition.
module _ {a b : Level} {C : RawCategory a b} where module _ {a b : Level} {C : RawCategory a b} where
@ -142,14 +142,14 @@ module _ {a b : Level} {C : RawCategory a b} where
open Cubical.NType.Properties open Cubical.NType.Properties
geq : g g' geq : g g'
geq = begin geq = begin
g ≡⟨ sym (fst ident) g ≡⟨ sym (fst isIdentity)
g 𝟙 ≡⟨ cong (λ φ g φ) (sym ε') g 𝟙 ≡⟨ cong (λ φ g φ) (sym ε')
g (f g') ≡⟨ isAssociative g (f g') ≡⟨ isAssociative
(g f) g' ≡⟨ cong (λ φ φ g') η (g f) g' ≡⟨ cong (λ φ φ g') η
𝟙 g' ≡⟨ snd ident 𝟙 g' ≡⟨ snd isIdentity
g' g'
propUnivalent : isProp (Univalent ident) propUnivalent : isProp (Univalent isIdentity)
propUnivalent a b i = propPi (λ iso propHasLevel ⟨-2⟩) a b i propUnivalent a b i = propPi (λ iso propHasLevel ⟨-2⟩) a b i
private private
@ -162,27 +162,27 @@ module _ {a b : Level} {C : RawCategory a b} where
-- projections of `IsCategory` - I've arbitrarily chosed to use this -- projections of `IsCategory` - I've arbitrarily chosed to use this
-- result from `x : IsCategory C`. I don't know which (if any) possibly -- result from `x : IsCategory C`. I don't know which (if any) possibly
-- adverse effects this may have. -- adverse effects this may have.
ident : (λ _ IsIdentity 𝟙) [ X.ident Y.ident ] isIdentity : (λ _ IsIdentity 𝟙) [ X.isIdentity Y.isIdentity ]
ident = propIsIdentity x X.ident Y.ident isIdentity = propIsIdentity x X.isIdentity Y.isIdentity
done : x y done : x y
U : {a : IsIdentity 𝟙} U : {a : IsIdentity 𝟙}
(λ _ IsIdentity 𝟙) [ X.ident a ] (λ _ IsIdentity 𝟙) [ X.isIdentity a ]
(b : Univalent a) (b : Univalent a)
Set _ Set _
U eqwal bbb = U eqwal bbb =
(λ i Univalent (eqwal i)) (λ i Univalent (eqwal i))
[ X.univalent bbb ] [ X.univalent bbb ]
P : (y : IsIdentity 𝟙) P : (y : IsIdentity 𝟙)
(λ _ IsIdentity 𝟙) [ X.ident y ] Set _ (λ _ IsIdentity 𝟙) [ X.isIdentity y ] Set _
P y eq = (b' : Univalent y) U eq b' P y eq = (b' : Univalent y) U eq b'
helper : (b' : Univalent X.ident) helper : (b' : Univalent X.isIdentity)
(λ _ Univalent X.ident) [ X.univalent b' ] (λ _ Univalent X.isIdentity) [ X.univalent b' ]
helper univ = propUnivalent x X.univalent univ helper univ = propUnivalent x X.univalent univ
foo = pathJ P helper Y.ident ident foo = pathJ P helper Y.isIdentity isIdentity
eqUni : U ident Y.univalent eqUni : U isIdentity Y.univalent
eqUni = foo Y.univalent eqUni = foo Y.univalent
IC.isAssociative (done i) = propIsAssociative x X.isAssociative Y.isAssociative i IC.isAssociative (done i) = propIsAssociative x X.isAssociative Y.isAssociative i
IC.ident (done i) = ident i IC.isIdentity (done i) = isIdentity i
IC.arrowIsSet (done i) = propArrowIsSet x X.arrowIsSet Y.arrowIsSet i IC.arrowIsSet (done i) = propArrowIsSet x X.arrowIsSet Y.arrowIsSet i
IC.univalent (done i) = eqUni i IC.univalent (done i) = eqUni i
@ -217,7 +217,7 @@ module _ {a b : Level} ( : Category a b) where
OpIsCategory : IsCategory OpRaw OpIsCategory : IsCategory OpRaw
IsCategory.isAssociative OpIsCategory = sym isAssociative IsCategory.isAssociative OpIsCategory = sym isAssociative
IsCategory.ident OpIsCategory = swap ident IsCategory.isIdentity OpIsCategory = swap isIdentity
IsCategory.arrowIsSet OpIsCategory = arrowIsSet IsCategory.arrowIsSet OpIsCategory = arrowIsSet
IsCategory.univalent OpIsCategory = {!!} IsCategory.univalent OpIsCategory = {!!}
@ -243,7 +243,7 @@ module _ {a b : Level} { : Category a b} where
module IsCat = IsCategory ( .isCategory) module IsCat = IsCategory ( .isCategory)
rawIsCat : (i : I) IsCategory (rawOp i) rawIsCat : (i : I) IsCategory (rawOp i)
isAssociative (rawIsCat i) = IsCat.isAssociative isAssociative (rawIsCat i) = IsCat.isAssociative
ident (rawIsCat i) = IsCat.ident isIdentity (rawIsCat i) = IsCat.isIdentity
arrowIsSet (rawIsCat i) = IsCat.arrowIsSet arrowIsSet (rawIsCat i) = IsCat.arrowIsSet
univalent (rawIsCat i) = IsCat.univalent univalent (rawIsCat i) = IsCat.univalent

View file

@ -33,7 +33,7 @@ module _ {c c' d d'}
record IsFunctor (F : RawFunctor) : 𝓤 where record IsFunctor (F : RawFunctor) : 𝓤 where
open RawFunctor F public open RawFunctor F public
field field
ident : IsIdentity isIdentity : IsIdentity
distrib : IsDistributive distrib : IsDistributive
record Functor : Set (c c' d d') where record Functor : Set (c c' d d') where
@ -55,7 +55,7 @@ module _
propIsFunctor : isProp (IsFunctor _ _ F) propIsFunctor : isProp (IsFunctor _ _ F)
propIsFunctor isF0 isF1 i = record propIsFunctor isF0 isF1 i = record
{ ident = 𝔻.arrowIsSet _ _ isF0.ident isF1.ident i { isIdentity = 𝔻.arrowIsSet _ _ isF0.isIdentity isF1.isIdentity i
; distrib = 𝔻.arrowIsSet _ _ isF0.distrib isF1.distrib i ; distrib = 𝔻.arrowIsSet _ _ isF0.distrib isF1.distrib i
} }
where where
@ -116,10 +116,10 @@ module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : F
instance instance
isFunctor' : IsFunctor A C _∘fr_ isFunctor' : IsFunctor A C _∘fr_
isFunctor' = record isFunctor' = record
{ ident = begin { isIdentity = begin
(F→ G→) (𝟙 A) ≡⟨ refl (F→ G→) (𝟙 A) ≡⟨ refl
F→ (G→ (𝟙 A)) ≡⟨ cong F→ (ident G) F→ (G→ (𝟙 A)) ≡⟨ cong F→ (isIdentity G)
F→ (𝟙 B) ≡⟨ ident F F→ (𝟙 B) ≡⟨ isIdentity F
𝟙 C 𝟙 C
; distrib = dist ; distrib = dist
} }
@ -135,7 +135,7 @@ identity = record
; func→ = λ x x ; func→ = λ x x
} }
; isFunctor = record ; isFunctor = record
{ ident = refl { isIdentity = refl
; distrib = refl ; distrib = refl
} }
} }

View file

@ -17,25 +17,25 @@ module _ { ' : Level} { : Category '} { A B : Category.Object
iso-is-epi : Isomorphism f Epimorphism {X = X} f iso-is-epi : Isomorphism f Epimorphism {X = X} f
iso-is-epi (f- , left-inv , right-inv) g₀ g₁ eq = begin iso-is-epi (f- , left-inv , right-inv) g₀ g₁ eq = begin
g₀ ≡⟨ sym (proj₁ ident) g₀ ≡⟨ sym (proj₁ isIdentity)
g₀ 𝟙 ≡⟨ cong (_∘_ g₀) (sym right-inv) g₀ 𝟙 ≡⟨ cong (_∘_ g₀) (sym right-inv)
g₀ (f f-) ≡⟨ isAssociative g₀ (f f-) ≡⟨ isAssociative
(g₀ f) f- ≡⟨ cong (λ φ φ f-) eq (g₀ f) f- ≡⟨ cong (λ φ φ f-) eq
(g₁ f) f- ≡⟨ sym isAssociative (g₁ f) f- ≡⟨ sym isAssociative
g₁ (f f-) ≡⟨ cong (_∘_ g₁) right-inv g₁ (f f-) ≡⟨ cong (_∘_ g₁) right-inv
g₁ 𝟙 ≡⟨ proj₁ ident g₁ 𝟙 ≡⟨ proj₁ isIdentity
g₁ g₁
iso-is-mono : Isomorphism f Monomorphism {X = X} f iso-is-mono : Isomorphism f Monomorphism {X = X} f
iso-is-mono (f- , (left-inv , right-inv)) g₀ g₁ eq = iso-is-mono (f- , (left-inv , right-inv)) g₀ g₁ eq =
begin begin
g₀ ≡⟨ sym (proj₂ ident) g₀ ≡⟨ sym (proj₂ isIdentity)
𝟙 g₀ ≡⟨ cong (λ φ φ g₀) (sym left-inv) 𝟙 g₀ ≡⟨ cong (λ φ φ g₀) (sym left-inv)
(f- f) g₀ ≡⟨ sym isAssociative (f- f) g₀ ≡⟨ sym isAssociative
f- (f g₀) ≡⟨ cong (_∘_ f-) eq f- (f g₀) ≡⟨ cong (_∘_ f-) eq
f- (f g₁) ≡⟨ isAssociative f- (f g₁) ≡⟨ isAssociative
(f- f) g₁ ≡⟨ cong (λ φ φ g₁) left-inv (f- f) g₁ ≡⟨ cong (λ φ φ g₁) left-inv
𝟙 g₁ ≡⟨ proj₂ ident 𝟙 g₁ ≡⟨ proj₂ isIdentity
g₁ g₁
iso-is-epi-mono : Isomorphism f Epimorphism {X = X} f × Monomorphism {X = X} f iso-is-epi-mono : Isomorphism f Epimorphism {X = X} f × Monomorphism {X = X} f
@ -70,7 +70,7 @@ module _ { : Level} { : Category } (unprovable : IsCategory (RawCat
module _ {c : Category.Object } where module _ {c : Category.Object } where
eqTrans : (λ _ Transformation (prshf c) (prshf c)) eqTrans : (λ _ Transformation (prshf c) (prshf c))
[ (λ _ x [ .𝟙 x ]) identityTrans (prshf c) ] [ (λ _ x [ .𝟙 x ]) identityTrans (prshf c) ]
eqTrans = funExt λ x funExt λ x .ident .proj₂ eqTrans = funExt λ x funExt λ x .isIdentity .proj₂
open import Cubical.NType.Properties open import Cubical.NType.Properties
open import Cat.Categories.Fun open import Cat.Categories.Fun
@ -78,7 +78,7 @@ module _ { : Level} { : Category } (unprovable : IsCategory (RawCat
:ident: = lemSig (naturalIsProp {F = prshf c} {prshf c}) _ _ eq :ident: = lemSig (naturalIsProp {F = prshf c} {prshf c}) _ _ eq
where where
eq : (λ C x [ .𝟙 x ]) identityTrans (prshf c) eq : (λ C x [ .𝟙 x ]) identityTrans (prshf c)
eq = funExt λ A funExt λ B proj₂ .ident eq = funExt λ A funExt λ B proj₂ .isIdentity
yoneda : Functor (Fun { = Opposite } {𝔻 = 𝓢}) yoneda : Functor (Fun { = Opposite } {𝔻 = 𝓢})
yoneda = record yoneda = record
@ -87,7 +87,7 @@ module _ { : Level} { : Category } (unprovable : IsCategory (RawCat
; func→ = :func→: ; func→ = :func→:
} }
; isFunctor = record ; isFunctor = record
{ ident = :ident: { isIdentity = :ident:
; distrib = {!!} ; distrib = {!!}
} }
} }