module Homotopy.Base where
Synthetic homotopy theory🔗
This module contains the basic definitions for the study of synthetic homotopy theory. Synthetic homotopy theory is the name given to studying in their own terms, i.e., the application of homotopy type theory to computing homotopy invariants of spaces. Central to the theory is the concept of pointed type and pointed map. After all, homotopy groups are no more than the set-truncations of n-fold iterated loop spaces, and loop spaces are always relative to a basepoint.
A helper that will come in handy is Σ∙
, which attaches the north pole as the
basepoint of the suspended space.
: ∀ {ℓ} → Type∙ ℓ → Type∙ ℓ
Σ∙ = Susp∙ (A .fst)
Σ∙ A
: ∀ {ℓ} → Type∙ ℓ → Type∙ ℓ
Ω∙ (A , a) = Path A a a , refl Ω∙
The suspension-loop space adjunction🔗
An important stepping stone in calculating loop spaces of higher types is the suspension-loop space adjunction: basepoint-preserving maps from a suspension are the same thing as basepoint-preserving maps into a loop space. We construct the equivalence in two steps, but both halves are constructed in elementary terms.
First, we’ll prove that
which is slightly involved, but not too much. The actual equivalence
is very straightforward to construct, but proving that the two maps
Σ-map→loops
and loops→Σ-map
are inverses
involves nontrivial path algebra.
module _ {ℓ ℓ'} {A∙@(A , a₀) : Type∙ ℓ} {B∙@(B , b₀) : Type∙ ℓ'} where
: (Σ∙ A∙ →∙ B∙) → (Σ _ λ bs → A → b₀ ≡ bs)
Σ-map∙→loops .fst = f .fst S
Σ-map∙→loops f .snd a = sym (f .snd) ∙ ap (f .fst) (merid a)
Σ-map∙→loops f
: (Σ _ λ bs → A → b₀ ≡ bs) → (Σ∙ A∙ →∙ B∙)
loops→Σ-map∙ .fst N = b₀
loops→Σ-map∙ pair .fst S = pair .fst
loops→Σ-map∙ pair .fst (merid x i) = pair .snd x i
loops→Σ-map∙ pair .snd = refl loops→Σ-map∙ pair
The construction for turning a family of loops into a basepoint-preserving map into is even simpler, perhaps because these are almost definitionally the same thing.
: (Σ _ λ bs → A → b₀ ≡ bs) → (A∙ →∙ Ω∙ B∙)
loops→map∙-Ω (b , l) .fst a = l a ∙ sym (l a₀)
loops→map∙-Ω (b , l) .snd = ∙-invr (l a₀)
loops→map∙-Ω
: (A∙ →∙ Ω∙ B∙) → (Σ _ λ bs → A → b₀ ≡ bs)
map∙-Ω→loops (f , _) .fst = b₀
map∙-Ω→loops (f , _) .snd = f map∙-Ω→loops
The path algebra for showing these are both pairs of inverse equivalences is not very interesting, so I’ve kept it hidden.
: (Σ∙ A∙ →∙ B∙) ≃ (Σ _ λ b → A → b₀ ≡ b)
Σ-map∙≃loops = Iso→Equiv (Σ-map∙→loops , iso loops→Σ-map∙ invr invl) where
Σ-map∙≃loops : is-right-inverse loops→Σ-map∙ Σ-map∙→loops
invr (p , q) = Σ-pathp refl $ funext λ a → ∙-idl (q a)
invr
: is-left-inverse loops→Σ-map∙ Σ-map∙→loops
invl (f , pres) i = funext f' i , λ j → pres (~ i ∨ j) where
invl : (a : Susp A) → loops→Σ-map∙ (Σ-map∙→loops (f , pres)) .fst a ≡ f a
f' = sym pres
f' N = refl
f' S (merid x i) j = ∙-filler₂ (sym pres) (ap f (merid x)) j i
f'
: (Σ _ λ bs → A → b₀ ≡ bs) ≃ (A∙ →∙ Ω∙ B∙)
loops≃map∙-Ω = Iso→Equiv (loops→map∙-Ω , iso map∙-Ω→loops invr invl) where
loops≃map∙-Ω : ∀ {ℓ} {A : Type ℓ} {x : A} (q : x ≡ x) (r : refl ≡ q)
lemma' → ap (λ p → q ∙ sym p) r ∙ ∙-invr q ≡ ∙-idr q ∙ sym r
=
lemma' q r (λ q' r → ap (λ p → q' ∙ sym p) r ∙ ∙-invr q' ≡ ∙-idr q' ∙ sym r)
J (∙-idl _ ∙ sym (∙-idr _))
r
: is-right-inverse map∙-Ω→loops loops→map∙-Ω
invr (b , x) = Σ-pathp (funext (λ a → ap₂ _∙_ refl (ap sym x) ∙ ∙-idr _)) (to-pathp (subst-path-left _ _ ∙ lemma)) where
invr =
lemma (ap₂ _∙_ refl (ap sym x) ∙ ∙-idr (b a₀)) ⌝ ∙ ∙-invr (b a₀) ≡⟨ ap! (sym-∙ (sym _) _) ⟩
⌜ sym (sym (∙-idr (b a₀)) ∙ ap (b a₀ ∙_) (ap sym (sym x))) ∙ ∙-invr (b a₀) ≡⟨ sym (∙-assoc _ _ _) ⟩
(∙-idr (b a₀)) ∙ ⌜ ap (λ p → b a₀ ∙ sym p) (sym x) ∙ ∙-invr (b a₀) ⌝ ≡⟨ ap! (lemma' (b a₀) (sym x)) ⟩
sym (∙-idr (b a₀)) ∙ ∙-idr (b a₀) ∙ x ≡⟨ ∙-cancell _ _ ⟩
sym
x ∎
: is-left-inverse map∙-Ω→loops loops→map∙-Ω
invl (f , p) = Σ-pathp (p a₀) $ to-pathp $ funext $ λ x →
invl _ _ ∙ sym (∙-assoc _ _ _)
subst-path-right _∙_ refl (∙-invl (p a₀)) ∙ ∙-idr _
∙ ap₂ (transport-refl x) ∙ ap p
Composing these equivalences, we get the adjunction:
: (Σ∙ A∙ →∙ B∙) ≃ (A∙ →∙ Ωⁿ 1 B∙)
Σ-map∙≃map∙-Ω = Σ-map∙≃loops ∙e loops≃map∙-Ω Σ-map∙≃map∙-Ω
Loop spaces are equivalently based maps out of spheres🔗
Repeatedly applying the equivalence we just built, we can build an equivalence
thus characterising loop spaces as basepoint-preserving maps out of the sphere.
: ∀ {ℓ} {A : Type∙ ℓ} n → Ωⁿ n (Ω∙ A) ≡ Ωⁿ (suc n) A
reassoc-Ω = refl
reassoc-Ω zero {A = A} (suc n) =
reassoc-Ω 1 (Ωⁿ n (Ωⁿ 1 A)) ≡⟨ ap (Ωⁿ 1) (reassoc-Ω n) ⟩
Ωⁿ 1 (Ωⁿ (suc n) A) ∎ Ωⁿ
As a special case, in the zeroth dimension, we conclude that i.e., basepoint-preserving maps from the booleans (based at either point) are the same thing as points of
: ∀ {ℓ} {A : Type∙ ℓ} n → (Sⁿ n →∙ A) ≃ Ωⁿ n A .fst
Ωⁿ≃Sⁿ-map {A = A} zero = Iso→Equiv (from , iso to (λ _ → refl) invr) where
Ωⁿ≃Sⁿ-map to : A .fst → ((Susp ⊥ , N) →∙ A)
to x .fst N = A .snd
to x .fst S = x
to x .snd = refl
: ((Susp ⊥ , N) →∙ A) → A .fst
from = f .fst S
from f
: is-right-inverse from to
invr (x , p) = Σ-pathp (funext (λ { N → sym p ; S → refl })) λ i j → p (~ i ∨ j)
invr
{A = A} (suc n) =
Ωⁿ≃Sⁿ-map (Σ∙ (Susp _ , N) →∙ A) ≃⟨ Σ-map∙≃map∙-Ω ⟩
((Susp (Sⁿ⁻¹ n) , N) →∙ Ωⁿ 1 A) ≃⟨ Ωⁿ≃Sⁿ-map n ⟩
(Ωⁿ 1 A) .fst ≃⟨ path→equiv (ap fst (reassoc-Ω n)) ⟩
Ωⁿ n (suc n) A .fst ≃∎ Ωⁿ