Construct the morphism for equivalence 2

I must still show that they are inverses.
This commit is contained in:
Frederik Hanghøj Iversen 2018-04-13 13:24:17 +02:00
parent 0ced448fa6
commit e25ef31907
2 changed files with 77 additions and 66 deletions

View file

@ -327,6 +327,7 @@ module _ {a b : Level} ( : RawCategory a b) where
asTypeIso : TypeIsomorphism (idToIso A B)
asTypeIso = toIso _ _ univalent
-- FIXME Rename
inverse-from-to-iso' : AreInverses (idToIso A B) isoToId
inverse-from-to-iso' = snd iso
@ -377,6 +378,62 @@ module _ {a b : Level} ( : RawCategory a b) where
where
lem : p~ <<< p* identity
lem = fst (snd (snd (idToIso _ _ p)))
module _ {A B X : Object} (iso : A B) where
private
p : A B
p = isoToId iso
p-dom : Arrow A X Arrow B X
p-dom = cong (λ x Arrow x X) p
p-cod : Arrow X A Arrow X B
p-cod = cong (λ x Arrow X x) p
lem : {A B} {x : A B} idToIso A B (isoToId x) x
lem {x = x} i = snd inverse-from-to-iso' i x
open Σ iso renaming (fst to ι) using ()
open Σ (snd iso) renaming (fst to ι~ ; snd to inv)
coe-dom : {f : Arrow A X} coe p-dom f f <<< ι~
coe-dom {f} = begin
coe p-dom f
≡⟨ 9-1-9 p refl f
fst (idToIso _ _ refl) <<< f <<< fst (snd (idToIso _ _ p))
≡⟨ cong (λ φ φ <<< f <<< fst (snd (idToIso _ _ p))) subst-neutral
identity <<< f <<< fst (snd (idToIso _ _ p))
≡⟨ cong (λ φ identity <<< f <<< φ) (cong (λ x (fst (snd x))) lem)
identity <<< f <<< ι~
≡⟨ cong (_<<< ι~) leftIdentity
f <<< ι~
coe-cod : {f : Arrow X A} coe p-cod f ι <<< f
coe-cod {f} = begin
coe p-cod f
≡⟨ 9-1-9 refl p f
fst (idToIso _ _ p) <<< f <<< fst (snd (idToIso _ _ refl))
≡⟨ cong (λ φ fst (idToIso _ _ p) <<< f <<< φ) subst-neutral
fst (idToIso _ _ p) <<< f <<< identity
≡⟨ cong (λ φ φ <<< f <<< identity) (cong fst lem)
ι <<< f <<< identity
≡⟨ sym isAssociative
ι <<< (f <<< identity)
≡⟨ cong (ι <<<_) rightIdentity
ι <<< f
module _ {f : Arrow A X} {g : Arrow B X} (q : PathP (λ i p-dom i) f g) where
domain-twist : g f <<< ι~
domain-twist = begin
g ≡⟨ sym (coe-lem q)
coe p-dom f ≡⟨ coe-dom
f <<< ι~
-- This can probably also just be obtained from the above my taking the
-- symmetric isomorphism.
domain-twist0 : f g <<< ι
domain-twist0 = begin
f ≡⟨ sym rightIdentity
f <<< identity ≡⟨ cong (f <<<_) (sym (fst inv))
f <<< (ι~ <<< ι) ≡⟨ isAssociative
f <<< ι~ <<< ι ≡⟨ cong (_<<< ι) (sym domain-twist)
g <<< ι
-- | All projections are propositions.
module Propositionality where

View file

@ -238,7 +238,7 @@ module Try0 {a b : Level} { : Category a b}
-- I have `φ p` in scope, but surely `p` and `x` are the same - though
-- perhaps not definitonally.
, (λ{ (iso , x) .isoToId iso , x})
, funExt (λ{ (p , q , r) Σ≡ (sym iso-id-inv) (toPathP {A = λ i {!!}} {!!})})
, funExt (λ{ (p , q , r) Σ≡ (sym iso-id-inv) {!!}})
, funExt (λ x Σ≡ (sym id-iso-inv) {!!})
step2
: Σ (X .≅ Y) (λ iso
@ -249,55 +249,9 @@ module Try0 {a b : Level} { : Category a b}
)
((X , xa , xb) (Y , ya , yb))
step2
= ( λ{ ((f , f~ , inv-f) , p , q)
let
r : X Y
r = .isoToId (f , f~ , inv-f)
r-arrow : (.Arrow X A) (.Arrow Y A)
r-arrow i = .Arrow (r i) A
t : coe r-arrow xa .identity .<<< xa .<<< f~
t = {!? ≡ ?!}
lem : {A B} {x : A .≅ B} .idToIso A B (.isoToId x) x
lem = λ{ {x = x} i snd .inverse-from-to-iso' i x}
lem' : {A B} {x : A B} .isoToId (.idToIso _ _ x) x
lem' = λ{ {x = x} i fst .inverse-from-to-iso' i x}
h : ya xa .<<< f~
h = begin
ya ≡⟨ sym (coe-lem p)
coe r-arrow xa
≡⟨ .9-1-9 r refl xa
fst (.idToIso _ _ refl) .<<< xa .<<< fst (snd (.idToIso _ _ r))
≡⟨ cong (λ φ φ .<<< xa .<<< fst (snd (.idToIso _ _ r))) subst-neutral
.identity .<<< xa .<<< fst (snd (.idToIso _ _ r))
≡⟨ cong (λ φ .identity .<<< xa .<<< φ) (cong (λ x (fst (snd x))) lem)
.identity .<<< xa .<<< f~
≡⟨ cong (._<<< f~) .leftIdentity
xa .<<< f~
complicated : Y X
complicated = (λ z sym (.isoToId (f , f~ , inv-f)) z)
in
( f , (begin
[ ya f ] ≡⟨⟩
ya .<<< f ≡⟨ sym .leftIdentity
.identity .<<< (ya .<<< f)
≡⟨ .isAssociative
.identity .<<< ya .<<< f
≡⟨ cong (λ φ φ .<<< ya .<<< f) (sym subst-neutral)
fst (.idToIso _ _ refl) .<<< ya .<<< f
≡⟨ cong (λ φ fst (.idToIso _ _ refl) .<<< ya .<<< φ)
(begin
f ≡⟨ {!cong (λ x → (fst (snd x))) ((lem {x = f~ , f , swap inv-f}))!}
-- f ≡⟨ cong fst (sym (lem {f , f~ , inv-f})) ⟩
-- fst ( .idToIso X Y ( .isoToId (f , f~ , inv-f))) ≡⟨ {!fst inv-f!} ⟩
fst (snd (.idToIso Y X (λ z sym (.isoToId (f , f~ , inv-f)) z))) )
_ .<<< ya .<<< fst (snd (.idToIso _ _ _))
≡⟨ sym (.9-1-9 (sym r) refl ya)
coe (sym r-arrow) ya
≡⟨ coe-lem (λ i p (~ i))
xa ) , {!!})
, ( f~ , sym h , {!!})
= ( λ{ (iso@(f , f~ , inv-f) , p , q)
( f , sym (.domain-twist0 iso p) , sym (.domain-twist0 iso q))
, ( f~ , sym (.domain-twist iso p) , sym (.domain-twist iso q))
, lemA (fst inv-f)
, lemA (snd inv-f)
}
@ -308,26 +262,26 @@ module Try0 {a b : Level} { : Category a b}
iso = fst f , fst f~ , cong fst inv-f , cong fst inv-f~
p : X Y
p = .isoToId iso
lem : xa _
lem = begin
xa ≡⟨ sym (fst (snd f))
ya .<<< fst f ≡⟨ sym .leftIdentity
.identity .<<< (ya .<<< fst f) ≡⟨ .isAssociative
.identity .<<< ya .<<< fst f ≡⟨ {!sym (.9-1-9 ? refl ya)!}
_ .<<< ya .<<< _ ≡⟨ sym (.9-1-9 (sym p) refl ya)
coe (λ i .Arrow (p (~ i)) A) ya
lem0 = begin
xa ≡⟨ {!!}
coe _ xa ≡⟨ {!!}
ya .<<< fst f
helper : PathP (λ i .Arrow (p i) A) xa ya
helper = {!snd f!}
in iso , ? , {!!}})
pA : .Arrow X A .Arrow Y A
pA = cong (λ x .Arrow x A) p
pB : .Arrow X B .Arrow Y B
pB = cong (λ x .Arrow x B) p
k0 = begin
coe pB xb ≡⟨ .coe-dom iso
xb .<<< fst f~ ≡⟨ snd (snd f~)
yb
k1 = begin
coe pA xa ≡⟨ .coe-dom iso
xa .<<< fst f~ ≡⟨ fst (snd f~)
ya
helper : PathP (λ i pA i) xa ya
helper = coe-lem-inv k1
in iso , coe-lem-inv k1 , coe-lem-inv k0})
, record
{ fst = funExt (λ x lemSig
(λ x propSig prop0 (λ _ prop1))
_ _
(Σ≡ {!!} (.propIsomorphism _ _ _)))
(Σ≡ {!!} (.propIsomorphism _ _ _)))
; snd = funExt (λ{ (f , _) lemSig propIsomorphism _ _ {!refl!}})
}
where