Show that objects are groupoids

This commit is contained in:
Frederik Hanghøj Iversen 2018-04-13 15:22:13 +02:00
parent 5afa835787
commit 7b45b5cdc3
4 changed files with 50 additions and 27 deletions

View File

@ -496,6 +496,20 @@ module _ {a b : Level} ( : RawCategory a b) where
res : Xi Yi
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
open RawCategory
open Univalence

View File

@ -216,8 +216,14 @@ module Try0 {a b : Level} { : Category a b}
)
step1
= 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
(symIso (_ , .asTypeIso {X} {Y}) .snd))
(isoSigFst
{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
: Σ (X .≅ Y) (λ iso

View File

@ -21,7 +21,7 @@ open import Cubical.GradLemma
using (gradLemma)
public
open import Cubical.NType
using (⟨-2⟩ ; ⟨-1⟩ ; ⟨0⟩ ; TLevel ; HasLevel)
using (⟨-2⟩ ; ⟨-1⟩ ; ⟨0⟩ ; TLevel ; HasLevel ; isGrpd)
public
open import Cubical.NType.Properties
using
@ -96,3 +96,5 @@ module _ { : Level} {A : Set } {a : A} where
pathJ (λ y x A) _ A refl ≡⟨ pathJprop {x = a} (λ y x A) _
_ ≡⟨ pathJprop {x = a} (λ y x A) _
a
open import Cat.Wishlist public

View File

@ -9,31 +9,32 @@ open import Agda.Builtin.Sigma
open import Cubical.NType.Properties
step : {} {A : Set } isContr A (x y : A) isContr (x y)
step (a , contr) x y = {!p , c!}
-- where
-- p : x ≡ y
-- p = begin
-- x ≡⟨ sym (contr x) ⟩
-- a ≡⟨ contr y ⟩
-- y ∎
-- c : (q : x ≡ y) → p ≡ q
-- c q i j = contr (p {!!}) {!!}
private
step : {} {A : Set } isContr A (x y : A) isContr (x y)
step (a , contr) x y = {!p , c!}
-- where
-- p : x ≡ y
-- p = begin
-- x ≡⟨ sym (contr x) ⟩
-- a ≡⟨ contr y ⟩
-- y ∎
-- c : (q : x ≡ y) → p ≡ q
-- c q i j = contr (p {!!}) {!!}
-- Contractible types have any given homotopy level.
contrInitial : { : Level} {A : Set } n isContr A HasLevel n A
contrInitial ⟨-2⟩ contr = contr
-- lem' (S ⟨-2⟩) (a , contr) = {!step!}
contrInitial (S ⟨-2⟩) (a , contr) x y = begin
x ≡⟨ sym (contr x)
a ≡⟨ contr y
y
contrInitial (S (S n)) contr x y = {!lvl!} -- Why is this not well-founded?
where
c : isContr (x y)
c = step contr x y
lvl : HasLevel (S n) (x y)
lvl = contrInitial {A = x y} (S n) c
-- Contractible types have any given homotopy level.
contrInitial : { : Level} {A : Set } n isContr A HasLevel n A
contrInitial ⟨-2⟩ contr = contr
-- lem' (S ⟨-2⟩) (a , contr) = {!step!}
contrInitial (S ⟨-2⟩) (a , contr) x y = begin
x ≡⟨ sym (contr x)
a ≡⟨ contr y
y
contrInitial (S (S n)) contr x y = {!lvl!} -- Why is this not well-founded?
where
c : isContr (x y)
c = step contr x y
lvl : HasLevel (S n) (x y)
lvl = contrInitial {A = x y} (S n) c
module _ { : Level} {A : Set } where
ntypeCommulative : {n m} n m HasLevel n ⟩₋₂ A HasLevel m ⟩₋₂ A