added Z=NuN and Groupoid laws exercises in quest2
This commit is contained in:
parent
ccb193a881
commit
ee0f734da2
@ -9,3 +9,9 @@ open import Cubical.Foundations.Prelude
|
|||||||
) public
|
) public
|
||||||
open import Cubical.HITs.S1 using (S¹ ; base ; loop) public
|
open import Cubical.HITs.S1 using (S¹ ; base ; loop) public
|
||||||
open import 1FundamentalGroup.Quest1Solutions public
|
open import 1FundamentalGroup.Quest1Solutions public
|
||||||
|
|
||||||
|
refl∙refl : {A : Type} {a : A} → refl ∙ refl ≡ refl {x = a}
|
||||||
|
refl∙refl {a = a} = sym (λ i j → compPath-filler (refl {x = a}) refl i j)
|
||||||
|
|
||||||
|
symRefl : {A : Type} {a : A} → sym refl ≡ refl {x = a}
|
||||||
|
symRefl = refl
|
||||||
|
@ -62,3 +62,4 @@ true≢false' : true ≡ false → ⊥
|
|||||||
true≢false' = {!!}
|
true≢false' = {!!}
|
||||||
|
|
||||||
-}
|
-}
|
||||||
|
|
||||||
|
@ -2,40 +2,63 @@
|
|||||||
module 1FundamentalGroup.Quest2Solutions where
|
module 1FundamentalGroup.Quest2Solutions where
|
||||||
open import 1FundamentalGroup.Preambles.P2
|
open import 1FundamentalGroup.Preambles.P2
|
||||||
|
|
||||||
-- sucℤ : ℤ → ℤ
|
|
||||||
-- sucℤ (pos n) = pos (suc n)
|
|
||||||
-- sucℤ (negsuc zero) = pos zero
|
|
||||||
-- sucℤ (negsuc (suc n)) = negsuc n
|
|
||||||
|
|
||||||
-- predℤ : ℤ → ℤ
|
data _⊔_ (A B : Type) : Type where
|
||||||
-- predℤ (pos zero) = negsuc zero
|
|
||||||
-- predℤ (pos (suc n)) = pos n
|
|
||||||
-- predℤ (negsuc n) = negsuc (suc n)
|
|
||||||
|
|
||||||
-- sucℤIso : ℤ ≅ ℤ
|
inl : A → A ⊔ B
|
||||||
-- sucℤIso = iso sucℤ predℤ s r where
|
inr : B → A ⊔ B
|
||||||
|
|
||||||
-- s : section sucℤ predℤ
|
ℤ≡ℕ⊔ℕ : ℤ ≡ ℕ ⊔ ℕ
|
||||||
-- s (pos zero) = refl
|
ℤ≡ℕ⊔ℕ = isoToPath (iso fun inv rightInv leftInv) where
|
||||||
-- s (pos (suc n)) = refl
|
|
||||||
-- s (negsuc zero) = refl
|
|
||||||
-- s (negsuc (suc n)) = refl
|
|
||||||
|
|
||||||
-- r : retract sucℤ predℤ
|
fun : ℤ → ℕ ⊔ ℕ
|
||||||
-- r (pos zero) = refl
|
fun (pos n) = inl n
|
||||||
-- r (pos (suc n)) = refl
|
fun (negsuc n) = inr n
|
||||||
-- r (negsuc zero) = refl
|
|
||||||
-- r (negsuc (suc n)) = refl
|
|
||||||
|
|
||||||
-- sucℤPath : ℤ ≡ ℤ
|
inv : ℕ ⊔ ℕ → ℤ
|
||||||
-- sucℤPath = isoToPath sucℤIso
|
inv (inl n) = pos n
|
||||||
|
inv (inr n) = negsuc n
|
||||||
|
|
||||||
-- helix : S¹ → Type
|
rightInv : section fun inv
|
||||||
-- helix base = ℤ
|
rightInv (inl n) = refl
|
||||||
-- helix (loop i) = sucℤPath i
|
rightInv (inr n) = refl
|
||||||
|
|
||||||
-- windingNumberBase : base ≡ base → ℤ
|
leftInv : retract fun inv
|
||||||
-- windingNumberBase p = endPt helix p (pos zero)
|
leftInv (pos n) = refl
|
||||||
|
leftInv (negsuc n) = refl
|
||||||
|
|
||||||
-- windingNumber : (x : S¹) → base ≡ x → helix x
|
∙refl : {A : Type} {x y : A} (p : x ≡ y) → p ∙ refl ≡ p
|
||||||
-- windingNumber x p = endPt helix p (pos zero)
|
∙refl = J (λ y p → p ∙ refl ≡ p) refl∙refl
|
||||||
|
|
||||||
|
refl∙ : {A : Type} {x y : A} (p : x ≡ y) → refl ∙ p ≡ p
|
||||||
|
refl∙ = J (λ y p → refl ∙ p ≡ p) refl∙refl
|
||||||
|
|
||||||
|
∙sym : {A : Type} {x y : A} (p : x ≡ y) → p ∙ sym p ≡ refl
|
||||||
|
∙sym = J (λ y p → p ∙ sym p ≡ refl)
|
||||||
|
(
|
||||||
|
refl ∙ sym refl
|
||||||
|
≡⟨ cong (λ p → refl ∙ p) symRefl ⟩
|
||||||
|
refl ∙ refl
|
||||||
|
≡⟨ refl∙refl ⟩
|
||||||
|
refl ∎)
|
||||||
|
|
||||||
|
sym∙ : {A : Type} {x y : A} (p : x ≡ y) → (sym p) ∙ p ≡ refl
|
||||||
|
sym∙ = J (λ y p → (sym p) ∙ p ≡ refl)
|
||||||
|
(
|
||||||
|
(sym refl) ∙ refl
|
||||||
|
≡⟨ cong (λ p → p ∙ refl) symRefl ⟩
|
||||||
|
refl ∙ refl
|
||||||
|
≡⟨ refl∙refl ⟩
|
||||||
|
refl ∎)
|
||||||
|
|
||||||
|
assoc : {A : Type} {w x : A} (p : w ≡ x) {y : A} (q : x ≡ y) {z : A} (r : y ≡ z)
|
||||||
|
→ (p ∙ q) ∙ r ≡ p ∙ (q ∙ r)
|
||||||
|
assoc {A} {w} = J
|
||||||
|
-- casing on p
|
||||||
|
(λ x p → {y : A} (q : x ≡ y) {z : A} (r : y ≡ z) → (p ∙ q) ∙ r ≡ p ∙ (q ∙ r))
|
||||||
|
-- reduce to showing when p = refl
|
||||||
|
λ q r → (refl ∙ q) ∙ r
|
||||||
|
≡⟨ cong (λ p' → p' ∙ r) (refl∙ q) ⟩
|
||||||
|
q ∙ r
|
||||||
|
≡⟨ sym (refl∙ (q ∙ r)) ⟩
|
||||||
|
refl ∙ q ∙ r ∎
|
||||||
|
Binary file not shown.
Binary file not shown.
Loading…
Reference in New Issue
Block a user