Use explicit parameter for hSet
This commit is contained in:
parent
ae0ff092f8
commit
8f67ff9f36
|
@ -59,10 +59,8 @@ module _ {ℓ : Level} {A B : Set ℓ} {a : A} where
|
||||||
|
|
||||||
module _ (ℓ : Level) where
|
module _ (ℓ : Level) where
|
||||||
private
|
private
|
||||||
open import Cubical.Universe using (hSet) public
|
|
||||||
|
|
||||||
SetsRaw : RawCategory (lsuc ℓ) ℓ
|
SetsRaw : RawCategory (lsuc ℓ) ℓ
|
||||||
RawCategory.Object SetsRaw = hSet {ℓ}
|
RawCategory.Object SetsRaw = hSet ℓ
|
||||||
RawCategory.Arrow SetsRaw (T , _) (U , _) = T → U
|
RawCategory.Arrow SetsRaw (T , _) (U , _) = T → U
|
||||||
RawCategory.𝟙 SetsRaw = Function.id
|
RawCategory.𝟙 SetsRaw = Function.id
|
||||||
RawCategory._∘_ SetsRaw = Function._∘′_
|
RawCategory._∘_ SetsRaw = Function._∘′_
|
||||||
|
@ -79,7 +77,7 @@ module _ (ℓ : Level) where
|
||||||
arrowsAreSets {B = (_ , s)} = setPi λ _ → s
|
arrowsAreSets {B = (_ , s)} = setPi λ _ → s
|
||||||
|
|
||||||
isIso = Eqv.Isomorphism
|
isIso = Eqv.Isomorphism
|
||||||
module _ {hA hB : hSet {ℓ}} where
|
module _ {hA hB : hSet ℓ} where
|
||||||
open Σ hA renaming (proj₁ to A ; proj₂ to sA)
|
open Σ hA renaming (proj₁ to A ; proj₂ to sA)
|
||||||
open Σ hB renaming (proj₁ to B ; proj₂ to sB)
|
open Σ hB renaming (proj₁ to B ; proj₂ to sB)
|
||||||
lem1 : (f : A → B) → isSet A → isSet B → isProp (isIso f)
|
lem1 : (f : A → B) → isSet A → isSet B → isProp (isIso f)
|
||||||
|
@ -284,7 +282,7 @@ module _ (ℓ : Level) where
|
||||||
univalent' : ∀ hA → isContr (Σ[ hB ∈ Object ] hA ≅ hB)
|
univalent' : ∀ hA → isContr (Σ[ hB ∈ Object ] hA ≅ hB)
|
||||||
univalent' hA = {!!} , {!!}
|
univalent' hA = {!!} , {!!}
|
||||||
|
|
||||||
module _ {hA hB : hSet {ℓ}} where
|
module _ {hA hB : hSet ℓ} where
|
||||||
|
|
||||||
-- Thierry: `thr0` implies univalence.
|
-- Thierry: `thr0` implies univalence.
|
||||||
univalent : isEquiv (hA ≡ hB) (hA ≅ hB) (Univalence.id-to-iso (λ {A} {B} → isIdentity {A} {B}) hA hB)
|
univalent : isEquiv (hA ≡ hB) (hA ≅ hB) (Univalence.id-to-iso (λ {A} {B} → isIdentity {A} {B}) hA hB)
|
||||||
|
|
|
@ -17,7 +17,7 @@ open import Cubical.GradLemma
|
||||||
using (gradLemma)
|
using (gradLemma)
|
||||||
public
|
public
|
||||||
open import Cubical.NType
|
open import Cubical.NType
|
||||||
using (⟨-2⟩)
|
using (⟨-2⟩ ; ⟨-1⟩ ; ⟨0⟩ ; TLevel ; HasLevel)
|
||||||
public
|
public
|
||||||
open import Cubical.NType.Properties
|
open import Cubical.NType.Properties
|
||||||
using
|
using
|
||||||
|
@ -25,7 +25,18 @@ open import Cubical.NType.Properties
|
||||||
; propPi ; propHasLevel ; setPi ; propSet)
|
; propPi ; propHasLevel ; setPi ; propSet)
|
||||||
public
|
public
|
||||||
open import Cubical.Sigma using (setSig ; sigPresSet) public
|
open import Cubical.Sigma using (setSig ; sigPresSet) public
|
||||||
open import Cubical.Universe using (hSet) public
|
|
||||||
|
module _ (ℓ : Level) where
|
||||||
|
-- FIXME Ask if we can push upstream.
|
||||||
|
-- A redefinition of `Cubical.Universe` with an explicit parameter
|
||||||
|
_-type : TLevel → Set (lsuc ℓ)
|
||||||
|
n -type = Σ (Set ℓ) (HasLevel n)
|
||||||
|
|
||||||
|
hSet : Set (lsuc ℓ)
|
||||||
|
hSet = ⟨0⟩ -type
|
||||||
|
|
||||||
|
Prop : Set (lsuc ℓ)
|
||||||
|
Prop = ⟨-1⟩ -type
|
||||||
|
|
||||||
-----------------
|
-----------------
|
||||||
-- * Utilities --
|
-- * Utilities --
|
||||||
|
|
Loading…
Reference in a new issue