Add another approach for univalence in Set

This commit is contained in:
Frederik Hanghøj Iversen 2018-03-22 11:50:07 +01:00
parent 807a0f3dcd
commit d12122ce60
2 changed files with 61 additions and 8 deletions

View file

@ -1,11 +1,11 @@
-- | The category of homotopy sets -- | The category of homotopy sets
{-# OPTIONS --allow-unsolved-metas --cubical #-} {-# OPTIONS --allow-unsolved-metas --cubical --caching #-}
module Cat.Categories.Sets where module Cat.Categories.Sets where
open import Cat.Prelude hiding (_≃_) open import Cat.Prelude hiding (_≃_)
import Data.Product import Data.Product
open import Function using (_∘_) open import Function using (_∘_ ; _∘_)
open import Cubical.Univalence using (univalence ; con ; _≃_ ; idtoeqv ; ua) open import Cubical.Univalence using (univalence ; con ; _≃_ ; idtoeqv ; ua)
@ -267,12 +267,12 @@ module _ ( : Level) where
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)
conclusion = trivial? step0 step1 step2 conclusion = trivial? step0 step1 step2
t : (hA hB) (hA hB) thierry : (hA hB) (hA hB)
t = sym≃ conclusion thierry = sym≃ conclusion
-- TODO Is the morphism `(_≃_.eqv conclusion)` the same as -- TODO Is the morphism `(_≃_.eqv conclusion)` the same as
-- `(id-to-iso (λ {A} {B} → isIdentity {A} {B}) hA hB)` ? -- `(id-to-iso (λ {A} {B} → isIdentity {A} {B}) hA hB)` ?
res : isEquiv (hA hB) (hA hB) (_≃_.eqv t) res : isEquiv (hA hB) (hA hB) _
res = _≃_.isEqv t res = _≃_.isEqv thierry
thr : (hA hB) (hA hB) thr : (hA hB) (hA hB)
thr = con _ res thr = con _ res
-- p : _ → (hX : Object) → Path (hA ≅ hB) (hA ≡ hB) -- p : _ → (hX : Object) → Path (hA ≅ hB) (hA ≡ hB)
@ -297,8 +297,51 @@ module _ ( : Level) where
-- --
-- is contractible, which implies univalence. -- is contractible, which implies univalence.
univalent[Contr] : hA isContr (Σ[ hB Object ] hA hB) module _ (hA : Object) where
univalent[Contr] hA = {!!} , {!!} open Σ hA renaming (proj₁ to A)
center : Σ[ hB Object ] (hA hB)
center = hA , idIso hA
open Σ center renaming ({-proj₁ to hA ;-} proj₂ to isoA) using ()
module _ (y : Σ[ hC Object ] (hA hC)) where
open Σ y renaming (proj₁ to hC ; proj₂ to hA≅hC)
open Σ hC renaming (proj₁ to C)
open Σ hA≅hC renaming (proj₁ to obv ; proj₂ to iso)
open Σ iso renaming (proj₁ to inv ; proj₂ to areInv)
-- Idea:
-- Have : hA ≅ hC
-- Can I then construct `A Eqv.≅ C`
-- Cuz then it follows from univalence
A≡C : A C
A≡C = ua s
where
s0 : A Eqv.≅ C
s0 = obv , inv , Eqv.toAreInverses areInv
s : A C
s = fromIsomorphism s0
pObj : hA hC
pObj = lemSig (λ _ isSetIsProp) hA hC A≡C
abstract
isoEq
: (λ i Σ (A proj₁ (pObj i)) (Isomorphism {A = hA} {pObj i}))
[ idIso hA hA≅hC ]
isoEq = {!!}
where
d : iso (λ _ Σ (A A) (Isomorphism {A = hA} {hA}))
[ idIso hA iso ]
d iso = {!!}
isContractible : (hA , idIso hA) (hC , hA≅hC)
isContractible = Σ≡ pObj {!isoEq!}
-- isContractible = lemSig prop≅ center y pObj
univalent[Contr] : isContr (Σ[ hB Object ] hA hB)
univalent[Contr] = center , isContractible
univalent : Univalent univalent : Univalent
univalent = from[Contr] univalent[Contr] univalent = from[Contr] univalent[Contr]

View file

@ -27,6 +27,16 @@ module _ {a b : Level} where
Isomorphism : (f : A B) Set _ Isomorphism : (f : A B) Set _
Isomorphism f = Σ (B A) λ g AreInverses f g Isomorphism f = Σ (B A) λ g AreInverses f g
module _ {f : A B} {g : B A}
(inv : (g f) idFun A
× (f g) idFun B) where
open Σ inv renaming (fst to ve-re ; snd to re-ve)
toAreInverses : AreInverses f g
toAreInverses = record
{ verso-recto = ve-re
; recto-verso = re-ve
}
_≅_ : Set a Set b Set _ _≅_ : Set a Set b Set _
A B = Σ (A B) Isomorphism A B = Σ (A B) Isomorphism