Use a single version of \simeq

This commit is contained in:
Frederik Hanghøj Iversen 2018-04-06 18:27:24 +02:00
parent 36d92c7ceb
commit 69689e7b2a
8 changed files with 91 additions and 149 deletions

View file

@ -33,9 +33,6 @@ module _ (a b : Level) where
isIdentity : IsIdentity λ { {A} identity {A} } isIdentity : IsIdentity λ { {A} identity {A} }
isIdentity = (Σ≡ refl refl) , Σ≡ refl refl isIdentity = (Σ≡ refl refl) , Σ≡ refl refl
open import Cubical.NType.Properties
open import Cubical.Sigma
isPreCategory : IsPreCategory RawFam isPreCategory : IsPreCategory RawFam
IsPreCategory.isAssociative isPreCategory IsPreCategory.isAssociative isPreCategory
{A} {B} {C} {D} {f} {g} {h} = isAssociative {A} {B} {C} {D} {f} {g} {h} {A} {B} {C} {D} {f} {g} {h} = isAssociative {A} {B} {C} {D} {f} {g} {h}

View file

@ -78,7 +78,6 @@ module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : C
F[ F G~ ] ≡⟨ prop1 F[ F G~ ] ≡⟨ prop1
idFunctor idFunctor
open import Cubical.Univalence
p0 : F G p0 : F G
p0 = begin p0 = begin
F ≡⟨ sym Functors.rightIdentity F ≡⟨ sym Functors.rightIdentity

View file

@ -2,21 +2,15 @@
{-# OPTIONS --allow-unsolved-metas --cubical --caching #-} {-# OPTIONS --allow-unsolved-metas --cubical --caching #-}
module Cat.Categories.Sets where module Cat.Categories.Sets where
open import Cat.Prelude as P hiding (_≃_) open import Cat.Prelude as P
open import Function using (_∘_ ; _∘_) open import Function using (_∘_ ; _∘_)
open import Cubical.Univalence using (univalence ; con ; _≃_ ; idtoeqv ; ua)
open import Cat.Category open import Cat.Category
open import Cat.Category.Functor open import Cat.Category.Functor
open import Cat.Category.Product open import Cat.Category.Product
open import Cat.Wishlist open import Cat.Wishlist
open import Cat.Equivalence as Eqv using (AreInverses ; module Equiv ; module NoEta) open import Cat.Equivalence renaming (_≅_ to _≈_)
open NoEta
module Equivalence = Equivalence
_⊙_ : {a b c : Level} {A : Set a} {B : Set b} {C : Set c} (A B) (B C) A C _⊙_ : {a b c : Level} {A : Set a} {B : Set b} {C : Set c} (A B) (B C) A C
eqA eqB = Equivalence.compose eqA eqB eqA eqB = Equivalence.compose eqA eqB
@ -52,7 +46,7 @@ module _ ( : Level) where
open IsPreCategory isPreCat hiding (_∘_) open IsPreCategory isPreCat hiding (_∘_)
isIso = Eqv.Isomorphism isIso = TypeIsomorphism
module _ {hA hB : hSet } where module _ {hA hB : hSet } where
open Σ hA renaming (fst to A ; snd to sA) open Σ hA renaming (fst to A ; snd to sA)
open Σ hB renaming (fst to B ; snd to sB) open Σ hB renaming (fst to B ; snd to sB)
@ -95,7 +89,7 @@ module _ ( : Level) where
module _ {a b : Level} {A : Set a} {P : A Set b} where module _ {a b : Level} {A : Set a} {P : A Set b} where
lem2 : ((x : A) isProp (P x)) (p q : Σ A P) lem2 : ((x : A) isProp (P x)) (p q : Σ A P)
(p q) (fst p fst q) (p q) (fst p fst q)
lem2 pA p q = fromIsomorphism iso lem2 pA p q = fromIsomorphism _ _ iso
where where
f : {p q} p q fst p fst q f : {p q} p q fst p fst q
f e i = fst (e i) f e i = fst (e i)
@ -111,7 +105,7 @@ module _ ( : Level) where
{ verso-recto = funExt ve-re { verso-recto = funExt ve-re
; recto-verso = funExt re-ve ; recto-verso = funExt re-ve
} }
iso : (p q) Eqv.≅ (fst p fst q) iso : (p q) (fst p fst q)
iso = f , g , inv iso = f , g , inv
lem3 : {c} {Q : A Set (c b)} lem3 : {c} {Q : A Set (c b)}
@ -119,12 +113,12 @@ module _ ( : Level) where
lem3 {Q = Q} eA = res lem3 {Q = Q} eA = res
where where
f : Σ A P Σ A Q f : Σ A P Σ A Q
f (a , pA) = a , _≃_.eqv (eA a) pA f (a , pA) = a , fst (eA a) pA
g : Σ A Q Σ A P g : Σ A Q Σ A P
g (a , qA) = a , g' qA g (a , qA) = a , g' qA
where where
k : Eqv.Isomorphism _ k : TypeIsomorphism _
k = Equiv≃.toIso _ _ (_≃_.isEqv (eA a)) k = toIso _ _ (snd (eA a))
open Σ k renaming (fst to g') open Σ k renaming (fst to g')
ve-re : (x : Σ A P) (g f) x x ve-re : (x : Σ A P) (g f) x x
ve-re x i = fst x , eq i ve-re x i = fst x , eq i
@ -133,16 +127,16 @@ module _ ( : Level) where
eq = begin eq = begin
snd ((g f) x) ≡⟨⟩ snd ((g f) x) ≡⟨⟩
snd (g (f (a , pA))) ≡⟨⟩ snd (g (f (a , pA))) ≡⟨⟩
g' (_≃_.eqv (eA a) pA) ≡⟨ lem g' (fst (eA a) pA) ≡⟨ lem
pA pA
where where
open Σ x renaming (fst to a ; snd to pA) open Σ x renaming (fst to a ; snd to pA)
k : Eqv.Isomorphism _ k : TypeIsomorphism _
k = Equiv≃.toIso _ _ (_≃_.isEqv (eA a)) k = toIso _ _ (snd (eA a))
open Σ k renaming (fst to g' ; snd to inv) open Σ k renaming (fst to g' ; snd to inv)
module A = AreInverses inv module A = AreInverses inv
-- anti-funExt -- anti-funExt
lem : (g' (_≃_.eqv (eA a))) pA pA lem : (g' (fst (eA a))) pA pA
lem i = A.verso-recto i pA lem i = A.verso-recto i pA
re-ve : (x : Σ A Q) (f g) x x re-ve : (x : Σ A Q) (f g) x x
re-ve x i = fst x , eq i re-ve x i = fst x , eq i
@ -150,11 +144,11 @@ module _ ( : Level) where
open Σ x renaming (fst to a ; snd to qA) open Σ x renaming (fst to a ; snd to qA)
eq = begin eq = begin
snd ((f g) x) ≡⟨⟩ snd ((f g) x) ≡⟨⟩
_≃_.eqv (eA a) (g' qA) ≡⟨ (λ i A.recto-verso i qA) fst (eA a) (g' qA) ≡⟨ (λ i A.recto-verso i qA)
qA qA
where where
k : Eqv.Isomorphism _ k : TypeIsomorphism _
k = Equiv≃.toIso _ _ (_≃_.isEqv (eA a)) k = toIso _ _ (snd (eA a))
open Σ k renaming (fst to g' ; snd to inv) open Σ k renaming (fst to g' ; snd to inv)
module A = AreInverses inv module A = AreInverses inv
inv : AreInverses f g inv : AreInverses f g
@ -162,10 +156,10 @@ module _ ( : Level) where
{ verso-recto = funExt ve-re { verso-recto = funExt ve-re
; recto-verso = funExt re-ve ; recto-verso = funExt re-ve
} }
iso : Σ A P Eqv.≅ Σ A Q iso : Σ A P Σ A Q
iso = f , g , inv iso = f , g , inv
res : Σ A P Σ A Q res : Σ A P Σ A Q
res = fromIsomorphism iso res = fromIsomorphism _ _ iso
module _ {a b : Level} {A : Set a} {B : Set b} where module _ {a b : Level} {A : Set a} {B : Set b} where
lem4 : isSet A isSet B (f : A B) lem4 : isSet A isSet B (f : A B)
@ -173,20 +167,20 @@ module _ ( : Level) where
lem4 sA sB f = lem4 sA sB f =
let let
obv : isEquiv A B f isIso f obv : isEquiv A B f isIso f
obv = Equiv≃.toIso A B obv = toIso A B
inv : isIso f isEquiv A B f inv : isIso f isEquiv A B f
inv = Equiv≃.fromIso A B inv = fromIso A B
re-ve : (x : isEquiv A B f) (inv obv) x x re-ve : (x : isEquiv A B f) (inv obv) x x
re-ve = Equiv≃.inverse-from-to-iso A B re-ve = inverse-from-to-iso A B
ve-re : (x : isIso f) (obv inv) x x ve-re : (x : isIso f) (obv inv) x x
ve-re = Equiv≃.inverse-to-from-iso A B sA sB ve-re = inverse-to-from-iso A B sA sB
iso : isEquiv A B f Eqv.≅ isIso f iso : isEquiv A B f isIso f
iso = obv , inv , iso = obv , inv ,
record record
{ verso-recto = funExt re-ve { verso-recto = funExt re-ve
; recto-verso = funExt ve-re ; recto-verso = funExt ve-re
} }
in fromIsomorphism iso in fromIsomorphism _ _ iso
module _ {hA hB : Object} where module _ {hA hB : Object} where
open Σ hA renaming (fst to A ; snd to sA) open Σ hA renaming (fst to A ; snd to sA)
@ -198,33 +192,15 @@ module _ ( : Level) where
-- univalence -- univalence
step1 : Σ (A B) (isEquiv A B) (A B) step1 : Σ (A B) (isEquiv A B) (A B)
step1 = hh h step1 = sym≃ univalence
where
h : (A B) (A B)
h = sym≃ (univalence {A = A} {B})
obv : Σ (A B) (isEquiv A B) A B
obv = Eqv.deEta
inv : A B Σ (A B) (isEquiv A B)
inv = Eqv.doEta
re-ve : (x : _) (inv obv) x x
re-ve x = refl
-- Because _≃_ does not have eta equality!
ve-re : (x : _) (obv inv) x x
ve-re (con eqv isEqv) i = con eqv isEqv
areInv : AreInverses obv inv
areInv = record { verso-recto = funExt re-ve ; recto-verso = funExt ve-re }
eqv : Σ (A B) (isEquiv A B) Eqv.≅ (A B)
eqv = obv , inv , areInv
hh : Σ (A B) (isEquiv A B) (A B)
hh = fromIsomorphism eqv
-- lem2 with propIsSet -- lem2 with propIsSet
step2 : (A B) (hA hB) step2 : (A B) (hA hB)
step2 = sym≃ (lem2 (λ A isSetIsProp) hA hB) step2 = sym≃ (lem2 (λ A isSetIsProp) hA hB)
-- Go from an isomorphism on sets to an isomorphism on homotopic sets -- Go from an isomorphism on sets to an isomorphism on homotopic sets
trivial? : (hA hB) (A Eqv.≅ B) trivial? : (hA hB) (A B)
trivial? = sym≃ (fromIsomorphism res) trivial? = sym≃ (fromIsomorphism _ _ res)
where where
fwd : Σ (A B) isIso hA hB fwd : Σ (A B) isIso hA hB
fwd (f , g , inv) = f , g , inv.toPair fwd (f , g , inv) = f , g , inv.toPair
@ -232,7 +208,7 @@ module _ ( : Level) where
module inv = AreInverses inv module inv = AreInverses inv
bwd : hA hB Σ (A B) isIso bwd : hA hB Σ (A B) isIso
bwd (f , g , x , y) = f , g , record { verso-recto = x ; recto-verso = y } bwd (f , g , x , y) = f , g , record { verso-recto = x ; recto-verso = y }
res : Σ (A B) isIso Eqv.≅ (hA hB) res : Σ (A B) isIso (hA hB)
res = fwd , bwd , record { verso-recto = refl ; recto-verso = refl } res = fwd , bwd , record { verso-recto = refl ; recto-verso = refl }
conclusion : (hA hB) (hA hB) conclusion : (hA hB) (hA hB)
@ -274,7 +250,6 @@ module _ { : Level} where
private private
𝓢 = 𝓢𝓮𝓽 𝓢 = 𝓢𝓮𝓽
open Category 𝓢 open Category 𝓢
open import Cubical.Sigma
module _ (hA hB : Object) where module _ (hA hB : Object) where
open Σ hA renaming (fst to A ; snd to sA) open Σ hA renaming (fst to A ; snd to sA)

View file

@ -29,6 +29,7 @@
module Cat.Category where module Cat.Category where
open import Cat.Prelude open import Cat.Prelude
open import Cat.Equivalence as Equivalence renaming (_≅_ to _≈_ ; Isomorphism to TypeIsomorphism) hiding (preorder≅)
import Function import Function
@ -122,8 +123,6 @@ record RawCategory (a b : Level) : Set (lsuc (a ⊔ b)) where
import Cat.Equivalence as E import Cat.Equivalence as E
open E public using () renaming (Isomorphism to TypeIsomorphism) open E public using () renaming (Isomorphism to TypeIsomorphism)
open E using (module Equiv)
open Equiv≃ using (fromIso)
univalenceFromIsomorphism : {A B : Object} univalenceFromIsomorphism : {A B : Object}
TypeIsomorphism (idToIso A B) isEquiv (A B) (A B) (idToIso A B) TypeIsomorphism (idToIso A B) isEquiv (A B) (A B) (idToIso A B)
@ -299,10 +298,8 @@ module _ {a b : Level} ( : RawCategory a b) where
univalent≃ = _ , univalent univalent≃ = _ , univalent
module _ {A B : Object} where module _ {A B : Object} where
open import Cat.Equivalence using (module Equiv)
iso-to-id : (A B) (A B) iso-to-id : (A B) (A B)
iso-to-id = fst (Equiv≃.toIso _ _ univalent) iso-to-id = fst (toIso _ _ univalent)
-- | All projections are propositions. -- | All projections are propositions.
module Propositionality where module Propositionality where
@ -321,7 +318,6 @@ module _ {a b : Level} ( : RawCategory a b) where
open Σ Yt renaming (fst to Y ; snd to Yit) open Σ Yt renaming (fst to Y ; snd to Yit)
open Σ (Xit {Y}) renaming (fst to Y→X) using () open Σ (Xit {Y}) renaming (fst to Y→X) using ()
open Σ (Yit {X}) renaming (fst to X→Y) using () open Σ (Yit {X}) renaming (fst to X→Y) using ()
open import Cat.Equivalence hiding (_≅_)
-- Need to show `left` and `right`, what we know is that the arrows are -- Need to show `left` and `right`, what we know is that the arrows are
-- unique. Well, I know that if I compose these two arrows they must give -- unique. Well, I know that if I compose these two arrows they must give
-- the identity, since also the identity is the unique such arrow (by X -- the identity, since also the identity is the unique such arrow (by X
@ -336,10 +332,10 @@ module _ {a b : Level} ( : RawCategory a b) where
right = Yprop _ _ right = Yprop _ _
iso : X Y iso : X Y
iso = X→Y , Y→X , left , right iso = X→Y , Y→X , left , right
fromIso : X Y X Y fromIso' : X Y X Y
fromIso = fst (Equiv≃.toIso (X Y) (X Y) univalent) fromIso' = fst (toIso (X Y) (X Y) univalent)
p0 : X Y p0 : X Y
p0 = fromIso iso p0 = fromIso' iso
p1 : (λ i IsTerminal (p0 i)) [ Xit Yit ] p1 : (λ i IsTerminal (p0 i)) [ Xit Yit ]
p1 = lemPropF propIsTerminal p0 p1 = lemPropF propIsTerminal p0
res : Xt Yt res : Xt Yt
@ -354,7 +350,6 @@ module _ {a b : Level} ( : RawCategory a b) where
open Σ Yi renaming (fst to Y ; snd to Yii) open Σ Yi renaming (fst to Y ; snd to Yii)
open Σ (Xii {Y}) renaming (fst to Y→X) using () open Σ (Xii {Y}) renaming (fst to Y→X) using ()
open Σ (Yii {X}) renaming (fst to X→Y) using () open Σ (Yii {X}) renaming (fst to X→Y) using ()
open import Cat.Equivalence hiding (_≅_)
-- Need to show `left` and `right`, what we know is that the arrows are -- Need to show `left` and `right`, what we know is that the arrows are
-- unique. Well, I know that if I compose these two arrows they must give -- unique. Well, I know that if I compose these two arrows they must give
-- the identity, since also the identity is the unique such arrow (by X -- the identity, since also the identity is the unique such arrow (by X
@ -369,10 +364,10 @@ module _ {a b : Level} ( : RawCategory a b) where
right = Xprop _ _ right = Xprop _ _
iso : X Y iso : X Y
iso = Y→X , X→Y , right , left iso = Y→X , X→Y , right , left
fromIso : X Y X Y fromIso' : X Y X Y
fromIso = fst (Equiv≃.toIso (X Y) (X Y) univalent) fromIso' = fst (toIso (X Y) (X Y) univalent)
p0 : X Y p0 : X Y
p0 = fromIso iso p0 = fromIso' iso
p1 : (λ i IsInitial (p0 i)) [ Xii Yii ] p1 : (λ i IsInitial (p0 i)) [ Xii Yii ]
p1 = lemPropF propIsInitial p0 p1 = lemPropF propIsInitial p0
res : Xi Yi res : Xi Yi
@ -436,9 +431,12 @@ module _ {a b : Level} ( : RawCategory a b) where
propIsCategory : isProp (IsCategory ) propIsCategory : isProp (IsCategory )
propIsCategory = done propIsCategory = done
-- | Univalent categories -- | Univalent categories
-- --
-- Just bundles up the data with witnesses inhabiting the propositions. -- Just bundles up the data with witnesses inhabiting the propositions.
-- Question: Should I remove the type `Category`?
record Category (a b : Level) : Set (lsuc (a b)) where record Category (a b : Level) : Set (lsuc (a b)) where
field field
raw : RawCategory a b raw : RawCategory a b
@ -459,10 +457,8 @@ module _ {a b : Level} { 𝔻 : Category a b} where
isCategoryEq = lemPropF propIsCategory rawEq isCategoryEq = lemPropF propIsCategory rawEq
Category≡ : 𝔻 Category≡ : 𝔻
Category≡ i = record Category.raw (Category≡ i) = rawEq i
{ raw = rawEq i Category.isCategory (Category≡ i) = isCategoryEq i
; isCategory = isCategoryEq i
}
-- | Syntax for arrows- and composition in a given category. -- | Syntax for arrows- and composition in a given category.
module _ {a b : Level} ( : Category a b) where module _ {a b : Level} ( : Category a b) where
@ -501,9 +497,8 @@ module Opposite {a b : Level} where
open IsPreCategory isPreCategory open IsPreCategory isPreCategory
module _ {A B : .Object} where module _ {A B : .Object} where
open import Cat.Equivalence as Equivalence hiding (_≅_)
k : Equivalence.Isomorphism (.idToIso A B) k : Equivalence.Isomorphism (.idToIso A B)
k = Equiv≃.toIso _ _ .univalent k = toIso _ _ .univalent
open Σ k renaming (fst to f ; snd to inv) open Σ k renaming (fst to f ; snd to inv)
open AreInverses inv open AreInverses inv
@ -568,7 +563,7 @@ module Opposite {a b : Level} where
h = ff , invv h = ff , invv
univalent : isEquiv (A B) (A B) univalent : isEquiv (A B) (A B)
(Univalence.idToIso (swap .isIdentity) A B) (Univalence.idToIso (swap .isIdentity) A B)
univalent = Equiv≃.fromIso _ _ h univalent = fromIso _ _ h
isCategory : IsCategory opRaw isCategory : IsCategory opRaw
IsCategory.isPreCategory isCategory = isPreCategory IsCategory.isPreCategory isCategory = isPreCategory

View file

@ -5,7 +5,6 @@ open import Cat.Prelude
open import Function open import Function
open import Cubical open import Cubical
open import Cubical.NType.Properties using (lemPropF)
open import Cat.Category open import Cat.Category

View file

@ -1,8 +1,8 @@
{-# OPTIONS --allow-unsolved-metas --cubical #-} {-# OPTIONS --allow-unsolved-metas --cubical #-}
module Cat.Category.Product where module Cat.Category.Product where
open import Cubical.NType.Properties
open import Cat.Prelude as P hiding (_×_ ; fst ; snd) open import Cat.Prelude as P hiding (_×_ ; fst ; snd)
open import Cat.Equivalence hiding (_≅_)
-- module P = Cat.Prelude -- module P = Cat.Prelude
open import Cat.Category open import Cat.Category
@ -285,10 +285,8 @@ module Try0 {a b : Level} { : Category a b}
open Category cat open Category cat
open import Cat.Equivalence
lemma : Terminal Product A B lemma : Terminal Product A B
lemma = Equiv≃.fromIsomorphism Terminal (Product A B) (f , g , inv) lemma = fromIsomorphism Terminal (Product A B) (f , g , inv)
where where
f : Terminal Product A B f : Terminal Product A B
f ((X , x0 , x1) , uniq) = p f ((X , x0 , x1) , uniq) = p

View file

@ -3,11 +3,26 @@ module Cat.Equivalence where
open import Cubical.Primitives open import Cubical.Primitives
open import Cubical.FromStdLib renaming (-max to _⊔_) open import Cubical.FromStdLib renaming (-max to _⊔_)
-- FIXME: Don't hide ≃
open import Cubical.PathPrelude hiding (inverse ; _≃_) open import Cubical.PathPrelude hiding (inverse ; _≃_)
open import Cubical.PathPrelude using (isEquiv ; isContr ; fiber) public open import Cubical.PathPrelude using (isEquiv ; isContr ; fiber) public
open import Cubical.GradLemma open import Cubical.GradLemma
open import Cat.Prelude using (lemPropF ; setPi ; lemSig ; propSet ; Preorder ; equalityIsEquivalence) open import Cat.Prelude using (lemPropF ; setPi ; lemSig ; propSet ; Preorder ; equalityIsEquivalence ; _≃_)
import Cubical.Univalence as U
module _ {a b : Level} {A : Set a} {B : Set b} where
open Cubical.PathPrelude
deEta : A B A U.≃ B
deEta (a , b) = U.con a b
doEta : A U.≃ B A B
doEta (U.con eqv isEqv) = eqv , isEqv
module _ { : Level} {A B : Set } where
open Cubical.PathPrelude
ua : A B A B
ua (f , isEqv) = U.ua (U.con f isEqv)
module _ {a b : Level} where module _ {a b : Level} where
private private
@ -242,7 +257,6 @@ module _ {a b : Level} (A : Set a) (B : Set b) where
where where
import Cubical.NType.Properties as P import Cubical.NType.Properties as P
module Equiv where
open Equiv ≃isEquiv public open Equiv ≃isEquiv public
module _ {a b : Level} {A : Set a} {B : Set b} where module _ {a b : Level} {A : Set a} {B : Set b} where
@ -273,20 +287,19 @@ module _ {a b : Level} {A : Set a} {B : Set b} where
} }
composeIsEquiv : isEquiv A B f isEquiv B C g isEquiv A C (g f) composeIsEquiv : isEquiv A B f isEquiv B C g isEquiv A C (g f)
composeIsEquiv a b = Equiv≃.fromIso A C (composeIsomorphism a' b') composeIsEquiv a b = fromIso A C (composeIsomorphism a' b')
where where
a' = Equiv≃.toIso A B a a' = toIso A B a
b' = Equiv≃.toIso B C b b' = toIso B C b
composeIso : {c : Level} {C : Set c} (A B) (B C) A C composeIso : {c : Level} {C : Set c} (A B) (B C) A C
composeIso {C = C} (f , iso-f) (g , iso-g) = g f , composeIsomorphism iso-f iso-g composeIso {C = C} (f , iso-f) (g , iso-g) = g f , composeIsomorphism iso-f iso-g
-- Gives the quasi inverse from an equivalence. -- Gives the quasi inverse from an equivalence.
module Equivalence (e : A B) where module Equivalence (e : A B) where
open Equiv≃ A B public
private private
iso : Isomorphism (fst e) iso : Isomorphism (fst e)
iso = snd (toIsomorphism e) iso = snd (toIsomorphism _ _ e)
open AreInverses (snd iso) public open AreInverses (snd iso) public
@ -303,9 +316,7 @@ module _ {a b : Level} {A : Set a} {B : Set b} where
} }
symmetry : B A symmetry : B A
symmetry = B≃A.fromIsomorphism symmetryIso symmetry = fromIsomorphism _ _ symmetryIso
where
module B≃A = Equiv≃ B A
preorder≅ : ( : Level) Preorder _ _ _ preorder≅ : ( : Level) Preorder _ _ _
preorder≅ = record preorder≅ = record
@ -323,54 +334,24 @@ preorder≅ = record
; trans = composeIso ; trans = composeIso
} }
} }
module _ { : Level} {A B : Set } where
module _ {a b : Level} {A : Set a} {B : Set b} where univalence : (A B) (A B)
open import Cubical.PathPrelude renaming (_≃_ to _≃η_) univalence = Equivalence.compose u' aux
open import Cubical.Univalence using (_≃_) where
u : (A B) U.≃ (A U.≃ B)
doEta : A B A ≃η B u = U.univalence
doEta (_≃_.con eqv isEqv) = eqv , isEqv u' : (A B) (A U.≃ B)
u' = doEta u
deEta : A ≃η B A B aux : (A U.≃ B) (A B)
deEta (eqv , isEqv) = _≃_.con eqv isEqv aux = fromIsomorphism _ _ (doEta , deEta , record { verso-recto = funExt (λ{ (U.con _ _) refl}) ; recto-verso = refl })
module NoEta {a b : Level} {A : Set a} {B : Set b} where
open import Cubical.PathPrelude renaming (_≃_ to _≃η_)
open import Cubical.Univalence using (_≃_)
module Equivalence (e : A B) where
open Equivalence (doEta e) hiding
( toIsomorphism ; fromIsomorphism ; _~_
; compose ; symmetryIso ; symmetry ) public
compose : {c : Level} {C : Set c} (B C) A C
compose ee = deEta (Equivalence.compose (doEta e) (doEta ee))
symmetry : B A
symmetry = deEta (Equivalence.symmetry (doEta e))
-- fromIso : {f : A → B} → Isomorphism f → isEquiv f
-- fromIso = ?
-- toIso : {f : A → B} → isEquiv f → Isomorphism f
-- toIso = ?
fromIsomorphism : A B A B
fromIsomorphism (f , iso) = _≃_.con f (Equiv≃.fromIso _ _ iso)
toIsomorphism : A B A B
toIsomorphism (_≃_.con f eqv) = f , Equiv≃.toIso _ _ eqv
-- A few results that I have not generalized to work with both the eta and no-eta variable of ≃ -- A few results that I have not generalized to work with both the eta and no-eta variable of ≃
module _ {a b : Level} {A : Set a} {P : A Set b} where module _ {a b : Level} {A : Set a} {P : A Set b} where
open NoEta
open import Cubical.Univalence using (_≃_)
-- Equality on sigma's whose second component is a proposition is equivalent -- Equality on sigma's whose second component is a proposition is equivalent
-- to equality on their first components. -- to equality on their first components.
equivPropSig : ((x : A) isProp (P x)) (p q : Σ A P) equivPropSig : ((x : A) isProp (P x)) (p q : Σ A P)
(p q) (fst p fst q) (p q) (fst p fst q)
equivPropSig pA p q = fromIsomorphism iso equivPropSig pA p q = fromIsomorphism _ _ iso
where where
f : {p q} p q fst p fst q f : {p q} p q fst p fst q
f e i = fst (e i) f e i = fst (e i)
@ -396,12 +377,12 @@ module _ {a b : Level} {A : Set a} {P : A → Set b} where
equivSigSnd {Q = Q} eA = res equivSigSnd {Q = Q} eA = res
where where
f : Σ A P Σ A Q f : Σ A P Σ A Q
f (a , pA) = a , _≃_.eqv (eA a) pA f (a , pA) = a , fst (eA a) pA
g : Σ A Q Σ A P g : Σ A Q Σ A P
g (a , qA) = a , g' qA g (a , qA) = a , g' qA
where where
k : Isomorphism _ k : Isomorphism _
k = Equiv≃.toIso _ _ (_≃_.isEqv (eA a)) k = toIso _ _ (snd (eA a))
open Σ k renaming (fst to g') open Σ k renaming (fst to g')
ve-re : (x : Σ A P) (g f) x x ve-re : (x : Σ A P) (g f) x x
ve-re x i = fst x , eq i ve-re x i = fst x , eq i
@ -410,16 +391,16 @@ module _ {a b : Level} {A : Set a} {P : A → Set b} where
eq = begin eq = begin
snd ((g f) x) ≡⟨⟩ snd ((g f) x) ≡⟨⟩
snd (g (f (a , pA))) ≡⟨⟩ snd (g (f (a , pA))) ≡⟨⟩
g' (_≃_.eqv (eA a) pA) ≡⟨ lem g' (fst (eA a) pA) ≡⟨ lem
pA pA
where where
open Σ x renaming (fst to a ; snd to pA) open Σ x renaming (fst to a ; snd to pA)
k : Isomorphism _ k : Isomorphism _
k = Equiv≃.toIso _ _ (_≃_.isEqv (eA a)) k = toIso _ _ (snd (eA a))
open Σ k renaming (fst to g' ; snd to inv) open Σ k renaming (fst to g' ; snd to inv)
module A = AreInverses inv module A = AreInverses inv
-- anti-funExt -- anti-funExt
lem : (g' (_≃_.eqv (eA a))) pA pA lem : (g' (fst (eA a))) pA pA
lem i = A.verso-recto i pA lem i = A.verso-recto i pA
re-ve : (x : Σ A Q) (f g) x x re-ve : (x : Σ A Q) (f g) x x
re-ve x i = fst x , eq i re-ve x i = fst x , eq i
@ -427,11 +408,11 @@ module _ {a b : Level} {A : Set a} {P : A → Set b} where
open Σ x renaming (fst to a ; snd to qA) open Σ x renaming (fst to a ; snd to qA)
eq = begin eq = begin
snd ((f g) x) ≡⟨⟩ snd ((f g) x) ≡⟨⟩
_≃_.eqv (eA a) (g' qA) ≡⟨ (λ i A.recto-verso i qA) fst (eA a) (g' qA) ≡⟨ (λ i A.recto-verso i qA)
qA qA
where where
k : Isomorphism _ k : Isomorphism _
k = Equiv≃.toIso _ _ (_≃_.isEqv (eA a)) k = toIso _ _ (snd (eA a))
open Σ k renaming (fst to g' ; snd to inv) open Σ k renaming (fst to g' ; snd to inv)
module A = AreInverses inv module A = AreInverses inv
inv : AreInverses f g inv : AreInverses f g
@ -442,11 +423,9 @@ module _ {a b : Level} {A : Set a} {P : A → Set b} where
iso : Σ A P Σ A Q iso : Σ A P Σ A Q
iso = f , g , inv iso = f , g , inv
res : Σ A P Σ A Q res : Σ A P Σ A Q
res = fromIsomorphism iso res = fromIsomorphism _ _ iso
module _ {a b : Level} {A : Set a} {B : Set b} where module _ {a b : Level} {A : Set a} {B : Set b} where
open NoEta
open import Cubical.Univalence using (_≃_)
-- Equivalence is equivalent to isomorphism when the domain and codomain of -- Equivalence is equivalent to isomorphism when the domain and codomain of
-- the equivalence is a set. -- the equivalence is a set.
equivSetIso : isSet A isSet B (f : A B) equivSetIso : isSet A isSet B (f : A B)
@ -454,17 +433,17 @@ module _ {a b : Level} {A : Set a} {B : Set b} where
equivSetIso sA sB f = equivSetIso sA sB f =
let let
obv : isEquiv A B f Isomorphism f obv : isEquiv A B f Isomorphism f
obv = Equiv≃.toIso A B obv = toIso A B
inv : Isomorphism f isEquiv A B f inv : Isomorphism f isEquiv A B f
inv = Equiv≃.fromIso A B inv = fromIso A B
re-ve : (x : isEquiv A B f) (inv obv) x x re-ve : (x : isEquiv A B f) (inv obv) x x
re-ve = Equiv≃.inverse-from-to-iso A B re-ve = inverse-from-to-iso A B
ve-re : (x : Isomorphism f) (obv inv) x x ve-re : (x : Isomorphism f) (obv inv) x x
ve-re = Equiv≃.inverse-to-from-iso A B sA sB ve-re = inverse-to-from-iso A B sA sB
iso : isEquiv A B f Isomorphism f iso : isEquiv A B f Isomorphism f
iso = obv , inv , iso = obv , inv ,
record record
{ verso-recto = funExt re-ve { verso-recto = funExt re-ve
; recto-verso = funExt ve-re ; recto-verso = funExt ve-re
} }
in fromIsomorphism iso in fromIsomorphism _ _ iso

View file

@ -24,7 +24,7 @@ open import Cubical.NType.Properties
using using
( lemPropF ; lemSig ; lemSigP ; isSetIsProp ( lemPropF ; lemSig ; lemSigP ; isSetIsProp
; propPi ; propPiImpl ; propHasLevel ; setPi ; propSet ; propPi ; propPiImpl ; propHasLevel ; setPi ; propSet
; propSig) ; propSig ; equivPreservesNType)
public public
propIsContr : { : Level} {A : Set } isProp (isContr A) propIsContr : { : Level} {A : Set } isProp (isContr A)