Merge remote-tracking branch 'Saizan/dev' into dev

This commit is contained in:
Frederik Hanghøj Iversen 2018-05-01 19:00:04 +02:00
commit 4b9fe0f5bb
3 changed files with 11 additions and 37 deletions

View file

@ -348,7 +348,7 @@ module _ {a b : Level} ( : RawCategory a b) where
coe refl f ≡⟨ id-coe coe refl f ≡⟨ id-coe
f ≡⟨ sym rightIdentity f ≡⟨ sym rightIdentity
f <<< identity ≡⟨ cong (f <<<_) (sym subst-neutral) f <<< identity ≡⟨ cong (f <<<_) (sym subst-neutral)
f <<< _ ≡⟨ {!!} _ ) a' p f <<< _ ≡⟨⟩ _ ) a' p
module _ {b' : Object} (p : b b') where module _ {b' : Object} (p : b b') where
private private
@ -525,9 +525,9 @@ module _ {a b : Level} ( : RawCategory a b) where
groupoidObject : isGrpd Object groupoidObject : isGrpd Object
groupoidObject A B = res groupoidObject A B = res
where where
open import Data.Nat using (_≤_ ; z≤n ; s≤s) open import Data.Nat using (_≤_ ; ≤′-refl ; ≤′-step)
setIso : x isSet (Isomorphism x) setIso : x isSet (Isomorphism x)
setIso x = ntypeCommulative ((s≤s {n = 1} z≤n)) (propIsomorphism x) setIso x = ntypeCumulative {n = 1} (≤′-step ≤′-refl) (propIsomorphism x)
step : isSet (A B) step : isSet (A B)
step = setSig {sA = arrowsAreSets} {sB = setIso} step = setSig {sA = arrowsAreSets} {sB = setIso}
res : isSet (A B) res : isSet (A B)

View file

@ -20,7 +20,7 @@
{-# OPTIONS --allow-unsolved-metas --cubical #-} {-# OPTIONS --allow-unsolved-metas --cubical #-}
open import Cat.Prelude open import Cat.Prelude
open import Data.Nat using (_≤_ ; z≤n ; s≤s) open import Data.Nat using (_≤_ ; ≤′-refl ; ≤′-step)
module Nat = Data.Nat module Nat = Data.Nat
open import Cat.Category open import Cat.Category
@ -112,8 +112,8 @@ module Properties where
naturalIsSet : (θ : Transformation F G) isSet (Natural F G θ) naturalIsSet : (θ : Transformation F G) isSet (Natural F G θ)
naturalIsSet θ = naturalIsSet θ =
ntypeCommulative ntypeCumulative {n = 1}
(s≤s {n = Nat.suc Nat.zero} z≤n) (Data.Nat.≤′-step Data.Nat.≤′-refl)
(naturalIsProp θ) (naturalIsProp θ)
naturalTransformationIsSet : isSet (NaturalTransformation F G) naturalTransformationIsSet : isSet (NaturalTransformation F G)

View file

@ -1,42 +1,16 @@
{-# OPTIONS --allow-unsolved-metas #-} {-# OPTIONS --allow-unsolved-metas #-}
module Cat.Wishlist where module Cat.Wishlist where
open import Level hiding (suc) open import Level hiding (suc; zero)
open import Cubical open import Cubical
open import Cubical.NType open import Cubical.NType
open import Data.Nat using (_≤_ ; z≤n ; s≤s ; zero ; suc) open import Data.Nat using (_≤_ ; ≤′-refl ; ≤′-step ; zero ; suc)
open import Agda.Builtin.Sigma open import Agda.Builtin.Sigma
open import Cubical.NType.Properties open import Cubical.NType.Properties
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
module _ { : Level} {A : Set } where module _ { : Level} {A : Set } where
ntypeCommulative : {n m} n m HasLevel n ⟩₋₂ A HasLevel m ⟩₋₂ A ntypeCumulative : {n m} n ≤′ m HasLevel n ⟩₋₂ A HasLevel m ⟩₋₂ A
ntypeCommulative {n = zero} {m} z≤n lvl = {!contrInitial ⟨ m ⟩₋₂ lvl!} ntypeCumulative {m} ≤′-refl lvl = lvl
ntypeCommulative {n = .(suc _)} {.(suc _)} (s≤s x) lvl = {!!} ntypeCumulative {n} {suc m} (≤′-step le) lvl = HasLevel+1 m ⟩₋₂ (ntypeCumulative le lvl)