added arc0/quest5 and arc1/quest3

This commit is contained in:
Jlh18 2021-11-01 00:07:28 +00:00
parent 334bb6eace
commit eb5587b25f
6 changed files with 66 additions and 3 deletions

View File

@ -104,10 +104,10 @@ Trans {x = x} {z = z} = J (λ y1 p → y1 ≡ z → x ≡ z) λ q → q
_∙_ = Trans -- input \.
_≡⟨_⟩_ : (x : A) x y y z x z
_≡⟨_⟩_ : (x : A) x y y z x z -- input \< and \>
_ ≡⟨ x≡y y≡z = x≡y y≡z
_∎ : (x : A) x x
_∎ : (x : A) x x -- input \qed
_ = refl
infixr 30 _∙_

View File

@ -0,0 +1,50 @@
module 0Trinitarianism.Quest5Solutions where
open import Cubical.Foundations.Prelude renaming (subst to endPt ; transport to pathToFun)
open import Cubical.HITs.S1 using ( ; base ; loop )
open import Cubical.Foundations.Isomorphism renaming (Iso to _≅_)
open import Cubical.Foundations.Path
open import 1FundamentalGroup.Quest0Solutions
open import Cubical.Data.Bool
PathD : {A0 A1 : Type} (A : A0 A1) (x : A0) (y : A1) Type
PathD A x y = pathToFun A x y
syntax PathD A x y = x y along A
outOfS¹P : (B : Type) (b : B base) PathP (λ i B (loop i)) b b (x : ) B x
outOfS¹P B b p base = b
outOfS¹P B b p (loop i) = p i
outOfS¹D : (B : Type) (b : B base) b b along (λ i B (loop i)) (x : ) B x
outOfS¹D B b p x = outOfS¹P B b (_≅_.inv (PathPIsoPath (λ i B (loop i)) b b) p) x
-- example : (x : S¹) → doubleCover x → Bool
-- example = outOfS¹D (λ x → doubleCover x → Bool) (λ x → true)
-- (pathToFun (λ i → (doubleCover (loop i)) → Bool) (λ x → true)
-- ≡⟨ refl ⟩
-- pathToFun (λ i → flipPath i → Bool) (λ x → true)
-- ≡⟨ refl ⟩
-- (λ x → true)
-- ∎)
example : (x : ) doubleCover x doubleCover x
example base = Flip
example (loop i) = p i where
p : PathP (λ i doubleCover (loop i) doubleCover (loop i)) Flip Flip
p = {!!}
example' : (x : ) doubleCover x doubleCover x
example' = outOfS¹D (λ x doubleCover x doubleCover x) Flip
(
pathToFun (λ i doubleCover (loop i) doubleCover (loop i)) Flip
≡⟨ refl
pathToFun (λ i flipPath i flipPath i) Flip
≡⟨ funExt lem
Flip
) where
lem : (x : Bool) pathToFun (λ i flipPath i flipPath i) Flip x Flip x
lem false = refl
lem true = refl

View File

@ -0,0 +1 @@
module 1FundamentalGroup.Quest3 where

View File

@ -1,9 +1,21 @@
module 1FundamentalGroup.Quest3Solutions where
open import Cubical.Foundations.Prelude
open import Cubical.HITs.S1
open import Cubical.HITs.S1 using ( ; base ; loop )
open import 1FundamentalGroup.Quest1Solutions
open import Cubical.Foundations.Path
outOfS¹ : (B : Type) (b : B base) PathP (λ i B (loop i)) b b (x : ) B x
outOfS¹ B b p base = b
outOfS¹ B b p (loop i) = p i
rewind : (x : ) helix x base x
rewind base = loop_times
rewind (loop i) = path i where
path : PathP (λ i helix (loop i) base loop i) loop_times loop_times
path = {!!}
thing = PathPIsoPath

Binary file not shown.