Expose naturalIsProp
This commit is contained in:
parent
bc2129b8fc
commit
954a89f8d1
|
@ -38,9 +38,6 @@ module _ {ℓc ℓc' ℓd ℓd' : Level} {ℂ : Category ℓc ℓc'} {𝔻 : Cat
|
||||||
→ (f : ℂ [ A , B ])
|
→ (f : ℂ [ A , B ])
|
||||||
→ 𝔻 [ θ B ∘ F.func→ f ] ≡ 𝔻 [ G.func→ f ∘ θ A ]
|
→ 𝔻 [ θ B ∘ F.func→ f ] ≡ 𝔻 [ G.func→ f ∘ θ A ]
|
||||||
|
|
||||||
-- naturalIsProp : ∀ θ → isProp (Natural θ)
|
|
||||||
-- naturalIsProp θ x y = {!funExt!}
|
|
||||||
|
|
||||||
NaturalTransformation : Set (ℓc ⊔ ℓc' ⊔ ℓd')
|
NaturalTransformation : Set (ℓc ⊔ ℓc' ⊔ ℓd')
|
||||||
NaturalTransformation = Σ Transformation Natural
|
NaturalTransformation = Σ Transformation Natural
|
||||||
|
|
||||||
|
@ -100,13 +97,11 @@ module _ {ℓc ℓc' ℓd ℓd' : Level} {ℂ : Category ℓc ℓc'} {𝔻 : Cat
|
||||||
NatComp = _:⊕:_
|
NatComp = _:⊕:_
|
||||||
|
|
||||||
private
|
private
|
||||||
module _ {F G : Functor ℂ 𝔻} where
|
|
||||||
module 𝔻 = Category 𝔻
|
module 𝔻 = Category 𝔻
|
||||||
|
|
||||||
|
module _ {F G : Functor ℂ 𝔻} where
|
||||||
transformationIsSet : isSet (Transformation F G)
|
transformationIsSet : isSet (Transformation F G)
|
||||||
transformationIsSet _ _ p q i j C = 𝔻.arrowIsSet _ _ (λ l → p l C) (λ l → q l C) i j
|
transformationIsSet _ _ p q i j C = 𝔻.arrowIsSet _ _ (λ l → p l C) (λ l → q l C) i j
|
||||||
IsSet' : {ℓ : Level} (A : Set ℓ) → Set ℓ
|
|
||||||
IsSet' A = {x y : A} → (p q : (λ _ → A) [ x ≡ y ]) → p ≡ q
|
|
||||||
|
|
||||||
naturalIsProp : (θ : Transformation F G) → isProp (Natural F G θ)
|
naturalIsProp : (θ : Transformation F G) → isProp (Natural F G θ)
|
||||||
naturalIsProp θ θNat θNat' = lem
|
naturalIsProp θ θNat θNat' = lem
|
||||||
|
@ -141,18 +136,17 @@ module _ {ℓc ℓc' ℓd ℓd' : Level} {ℂ : Category ℓc ℓc'} {𝔻 : Cat
|
||||||
where
|
where
|
||||||
open Category 𝔻
|
open Category 𝔻
|
||||||
|
|
||||||
module _ {A B : Functor ℂ 𝔻} {f : NaturalTransformation A B} where
|
|
||||||
private
|
private
|
||||||
|
module _ {A B : Functor ℂ 𝔻} {f : NaturalTransformation A B} where
|
||||||
allNatural = naturalIsProp {F = A} {B}
|
allNatural = naturalIsProp {F = A} {B}
|
||||||
f' = proj₁ f
|
f' = proj₁ f
|
||||||
module 𝔻Data = Category 𝔻
|
|
||||||
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 ∘ 𝔻Data.𝟙 ] ≡⟨ proj₁ (𝔻.ident {A} {B}) ⟩
|
𝔻 [ f' C ∘ 𝔻.𝟙 ] ≡⟨ proj₁ 𝔻.ident ⟩
|
||||||
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 {A} {B})
|
eq-l C = proj₂ 𝔻.ident
|
||||||
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
|
||||||
|
|
Loading…
Reference in a new issue