Show that objects are groupoids
This commit is contained in:
parent
5afa835787
commit
7b45b5cdc3
|
@ -496,6 +496,20 @@ module _ {ℓa ℓb : Level} (ℂ : RawCategory ℓa ℓb) where
|
||||||
res : Xi ≡ Yi
|
res : Xi ≡ Yi
|
||||||
res = lemSig propIsInitial _ _ (isoToId iso)
|
res = lemSig propIsInitial _ _ (isoToId iso)
|
||||||
|
|
||||||
|
groupoidObject : isGrpd Object
|
||||||
|
groupoidObject A B = res
|
||||||
|
where
|
||||||
|
open import Data.Nat using (_≤_ ; z≤n ; s≤s)
|
||||||
|
setIso : ∀ x → isSet (Isomorphism x)
|
||||||
|
setIso x = ntypeCommulative ((s≤s {n = 1} z≤n)) (propIsomorphism x)
|
||||||
|
step : isSet (A ≅ B)
|
||||||
|
step = setSig {sA = arrowsAreSets} {sB = setIso}
|
||||||
|
res : isSet (A ≡ B)
|
||||||
|
res = equivPreservesNType
|
||||||
|
{A = A ≅ B} {B = A ≡ B} {n = ⟨0⟩}
|
||||||
|
(Equivalence.symmetry (univalent≃ {A = A} {B}))
|
||||||
|
step
|
||||||
|
|
||||||
module _ {ℓa ℓb : Level} (ℂ : RawCategory ℓa ℓb) where
|
module _ {ℓa ℓb : Level} (ℂ : RawCategory ℓa ℓb) where
|
||||||
open RawCategory ℂ
|
open RawCategory ℂ
|
||||||
open Univalence
|
open Univalence
|
||||||
|
|
|
@ -216,8 +216,14 @@ module Try0 {ℓa ℓb : Level} {ℂ : Category ℓa ℓb}
|
||||||
)
|
)
|
||||||
step1
|
step1
|
||||||
= symIso
|
= symIso
|
||||||
(isoSigFst {A = (X ℂ.≅ Y)} {B = (X ≡ Y)} {!!} {Q = \ p → (PathP (λ i → ℂ.Arrow (p i) A) xa ya) × (PathP (λ i → ℂ.Arrow (p i) B) xb yb)} ℂ.isoToId
|
(isoSigFst
|
||||||
(symIso (_ , ℂ.asTypeIso {X} {Y}) .snd))
|
{A = (X ℂ.≅ Y)}
|
||||||
|
{B = (X ≡ Y)}
|
||||||
|
(ℂ.groupoidObject _ _)
|
||||||
|
{Q = \ p → (PathP (λ i → ℂ.Arrow (p i) A) xa ya) × (PathP (λ i → ℂ.Arrow (p i) B) xb yb)}
|
||||||
|
ℂ.isoToId
|
||||||
|
(symIso (_ , ℂ.asTypeIso {X} {Y}) .snd)
|
||||||
|
)
|
||||||
|
|
||||||
step2
|
step2
|
||||||
: Σ (X ℂ.≅ Y) (λ iso
|
: Σ (X ℂ.≅ Y) (λ iso
|
||||||
|
|
|
@ -21,7 +21,7 @@ open import Cubical.GradLemma
|
||||||
using (gradLemma)
|
using (gradLemma)
|
||||||
public
|
public
|
||||||
open import Cubical.NType
|
open import Cubical.NType
|
||||||
using (⟨-2⟩ ; ⟨-1⟩ ; ⟨0⟩ ; TLevel ; HasLevel)
|
using (⟨-2⟩ ; ⟨-1⟩ ; ⟨0⟩ ; TLevel ; HasLevel ; isGrpd)
|
||||||
public
|
public
|
||||||
open import Cubical.NType.Properties
|
open import Cubical.NType.Properties
|
||||||
using
|
using
|
||||||
|
@ -96,3 +96,5 @@ module _ {ℓ : Level} {A : Set ℓ} {a : A} where
|
||||||
pathJ (λ y x → A) _ A refl ≡⟨ pathJprop {x = a} (λ y x → A) _ ⟩
|
pathJ (λ y x → A) _ A refl ≡⟨ pathJprop {x = a} (λ y x → A) _ ⟩
|
||||||
_ ≡⟨ pathJprop {x = a} (λ y x → A) _ ⟩
|
_ ≡⟨ pathJprop {x = a} (λ y x → A) _ ⟩
|
||||||
a ∎
|
a ∎
|
||||||
|
|
||||||
|
open import Cat.Wishlist public
|
||||||
|
|
|
@ -9,31 +9,32 @@ open import Agda.Builtin.Sigma
|
||||||
|
|
||||||
open import Cubical.NType.Properties
|
open import Cubical.NType.Properties
|
||||||
|
|
||||||
step : ∀ {ℓ} {A : Set ℓ} → isContr A → (x y : A) → isContr (x ≡ y)
|
private
|
||||||
step (a , contr) x y = {!p , c!}
|
step : ∀ {ℓ} {A : Set ℓ} → isContr A → (x y : A) → isContr (x ≡ y)
|
||||||
-- where
|
step (a , contr) x y = {!p , c!}
|
||||||
-- p : x ≡ y
|
-- where
|
||||||
-- p = begin
|
-- p : x ≡ y
|
||||||
-- x ≡⟨ sym (contr x) ⟩
|
-- p = begin
|
||||||
-- a ≡⟨ contr y ⟩
|
-- x ≡⟨ sym (contr x) ⟩
|
||||||
-- y ∎
|
-- a ≡⟨ contr y ⟩
|
||||||
-- c : (q : x ≡ y) → p ≡ q
|
-- y ∎
|
||||||
-- c q i j = contr (p {!!}) {!!}
|
-- c : (q : x ≡ y) → p ≡ q
|
||||||
|
-- c q i j = contr (p {!!}) {!!}
|
||||||
|
|
||||||
-- Contractible types have any given homotopy level.
|
-- Contractible types have any given homotopy level.
|
||||||
contrInitial : {ℓ : Level} {A : Set ℓ} → ∀ n → isContr A → HasLevel n A
|
contrInitial : {ℓ : Level} {A : Set ℓ} → ∀ n → isContr A → HasLevel n A
|
||||||
contrInitial ⟨-2⟩ contr = contr
|
contrInitial ⟨-2⟩ contr = contr
|
||||||
-- lem' (S ⟨-2⟩) (a , contr) = {!step!}
|
-- lem' (S ⟨-2⟩) (a , contr) = {!step!}
|
||||||
contrInitial (S ⟨-2⟩) (a , contr) x y = begin
|
contrInitial (S ⟨-2⟩) (a , contr) x y = begin
|
||||||
x ≡⟨ sym (contr x) ⟩
|
x ≡⟨ sym (contr x) ⟩
|
||||||
a ≡⟨ contr y ⟩
|
a ≡⟨ contr y ⟩
|
||||||
y ∎
|
y ∎
|
||||||
contrInitial (S (S n)) contr x y = {!lvl!} -- Why is this not well-founded?
|
contrInitial (S (S n)) contr x y = {!lvl!} -- Why is this not well-founded?
|
||||||
where
|
where
|
||||||
c : isContr (x ≡ y)
|
c : isContr (x ≡ y)
|
||||||
c = step contr x y
|
c = step contr x y
|
||||||
lvl : HasLevel (S n) (x ≡ y)
|
lvl : HasLevel (S n) (x ≡ y)
|
||||||
lvl = contrInitial {A = x ≡ y} (S n) c
|
lvl = contrInitial {A = x ≡ y} (S n) c
|
||||||
|
|
||||||
module _ {ℓ : Level} {A : Set ℓ} where
|
module _ {ℓ : Level} {A : Set ℓ} where
|
||||||
ntypeCommulative : ∀ {n m} → n ≤ m → HasLevel ⟨ n ⟩₋₂ A → HasLevel ⟨ m ⟩₋₂ A
|
ntypeCommulative : ∀ {n m} → n ≤ m → HasLevel ⟨ n ⟩₋₂ A → HasLevel ⟨ m ⟩₋₂ A
|
||||||
|
|
Loading…
Reference in a new issue