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.
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 south
Σ-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 north = b₀
loops→Σ-map∙ pair .fst south = 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' north = refl
f' south (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∙ →∙ Ω¹ B∙)
Σ-map∙≃map∙-Ω = Σ-map∙≃loops ∙e loops≃map∙-Ω
Σ-map∙≃map∙-Ω
: Maps∙ (Σ¹ A∙) B∙ ≃∙ Maps∙ A∙ (Ω¹ B∙)
Σ-map∙≃∙map∙-Ω .fst = Σ-map∙≃map∙-Ω
Σ-map∙≃∙map∙-Ω .snd = homogeneous-funext∙ λ a →
Σ-map∙≃∙map∙-Ω (Σ-map∙→loops zero∙) · a ≡⟨⟩
loops→map∙-Ω (refl ∙ refl) ∙ sym (refl ∙ refl) ≡⟨ ap (λ x → x ∙ sym x) (∙-idl _) ⟩
_ ⟩
refl ∙ refl ≡⟨ ∙-idl refl ∎
private
loops-map: ∀ {ℓa ℓb ℓc} {A∙@(A , a₀) : Type∙ ℓa} {B∙@(B , b₀) : Type∙ ℓb} {C∙@(C , c₀) : Type∙ ℓc}
→ (B∙ →∙ C∙)
→ (Σ _ λ b → A → b₀ ≡ b) → (Σ _ λ c → A → c₀ ≡ c)
(f , pt) (b , l) = f b , λ a → sym pt ∙ ap f (l a)
loops-map
Σ-map∙≃loops-naturalr: ∀ {ℓa ℓb ℓc} {A∙@(A , a₀) : Type∙ ℓa} {B∙@(B , b₀) : Type∙ ℓb} {C∙@(C , c₀) : Type∙ ℓc}
→ (f∙ : B∙ →∙ C∙) (l∙ : Σ¹ A∙ →∙ B∙)
→ loops-map {A∙ = A∙} f∙ (Σ-map∙→loops {A∙ = A∙} l∙)
{A∙ = A∙} (f∙ ∘∙ l∙)
≡ Σ-map∙→loops {A∙ = A , a₀} (f , pt) (l , pt') = refl ,ₚ ext λ a →
Σ-map∙≃loops-naturalr (sym pt' ∙ ap l (merid a)) ≡⟨ ap (sym pt ∙_) (ap-∙ f _ _) ⟩
sym pt ∙ ap f (sym pt') ∙ ap (f ∘ l) (merid a) ≡⟨ ∙-assoc _ _ _ ⟩
sym pt ∙ ap f (sym pt') ⌝ ∙ ap (f ∘ l) (merid a) ≡˘⟨ ap¡ (sym-∙ _ _) ⟩
⌜ sym pt ∙ ap f (ap f pt' ∙ pt) ∙ ap (f ∘ l) (merid a) ∎
sym
opaque
unfolding Ω¹-map
loops≃map∙-Ω-naturalr: ∀ {ℓa ℓb ℓc} {A∙@(A , a₀) : Type∙ ℓa} {B∙@(B , b₀) : Type∙ ℓb} {C∙@(C , c₀) : Type∙ ℓc}
→ (f∙ : B∙ →∙ C∙) (l∙ : Σ _ λ b → A → b₀ ≡ b)
→ Ω¹-map f∙ ∘∙ loops→map∙-Ω {A∙ = A∙} l∙
(loops-map {A∙ = A∙} f∙ l∙)
≡ loops→map∙-Ω {A∙ = A , a₀} f∙@(f , pt) (b , l) = homogeneous-funext∙ λ a →
loops≃map∙-Ω-naturalr .fst (loops→map∙-Ω (b , l) .fst a) ≡⟨⟩
Ω¹-map f∙ (ap f (l a ∙ sym (l a₀))) ≡⟨ conj-defn _ _ ⟩
conj pt (l a ∙ sym (l a₀)) ⌝ ∙ pt ≡⟨ ap! (ap-∙ f _ _) ⟩
sym pt ∙ ⌜ ap f ((ap f (l a) ∙ ap f (sym (l a₀))) ∙ pt) ≡⟨ ap (sym pt ∙_) (sym (∙-assoc _ _ _)) ⟩
sym pt ∙ (ap f (l a) ∙ (ap f (sym (l a₀)) ∙ pt)) ≡⟨ ∙-assoc _ _ _ ⟩
sym pt ∙ (sym pt ∙ ap f (l a)) ∙ ⌜ ap f (sym (l a₀)) ∙ pt ⌝ ≡˘⟨ ap¡ (sym-∙ _ _) ⟩
(sym pt ∙ ap f (l a)) ∙ sym (sym pt ∙ ap f (l a₀)) ∎
Σ-map∙≃map∙-Ω-naturalr: ∀ {ℓa ℓb ℓc} {A∙ : Type∙ ℓa} {B∙ : Type∙ ℓb} {C∙ : Type∙ ℓc}
→ (f : B∙ →∙ C∙) (l : Σ¹ A∙ →∙ B∙)
→ Ω¹-map f ∘∙ Σ-map∙≃map∙-Ω {A∙ = A∙} .fst l ≡ Σ-map∙≃map∙-Ω .fst (f ∘∙ l)
{A∙ = A∙} f∙ l∙ =
Σ-map∙≃map∙-Ω-naturalr .fst l∙ ≡⟨ loops≃map∙-Ω-naturalr f∙ _ ⟩
Ω¹-map f∙ ∘∙ Σ-map∙≃map∙-Ω (loops-map {A∙ = A∙} f∙ (Σ-map∙→loops {A∙ = A∙} l∙)) ≡⟨ ap loops→map∙-Ω (Σ-map∙≃loops-naturalr {A∙ = A∙} f∙ l∙) ⟩
loops→map∙-Ω .fst (f∙ ∘∙ l∙) ∎ Σ-map∙≃map∙-Ω
We give an explicit description of the unit of this adjunction, a map
which will be useful in the proof of the Freudenthal suspension
theorem. The merid
constructor
of the suspension
generates a path between the north
and south
poles, rather than a loop
on either pole; but, since
is pointed, we can “correct” the meridian generated by an element
by composing with the inverse of the meridian at the basepoint, to
properly get an element in the loop space
Moreover, since
is
this is a pointed map.
: ∀ {ℓ} (A : Type∙ ℓ) → A →∙ Ω¹ (Σ¹ A)
suspend∙ (A , a₀) .fst x = merid x ∙ sym (merid a₀)
suspend∙ (A , a₀) .snd = ∙-invr (merid a₀) suspend∙
: ∀ {ℓ} (A : Type∙ ℓ) → ⌞ A ⌟ → Path ⌞ Σ¹ A ⌟ north north
suspend = suspend∙ A∙ .fst x suspend A∙ x
In order to connect this map with the adjunction we built above, we
prove that, for any pointed map
the adjunct of the composite
is the map
In particular, when
is the identity, this uniquely characterises suspend∙
as the unit of the
adjunction.
suspend∙-is-unit: ∀ {ℓa ℓb} {A : Type∙ ℓa} {B : Type∙ ℓb} (f : A →∙ B)
→ Equiv.from Σ-map∙≃map∙-Ω (suspend∙ B ∘∙ f) ≡ Susp-map∙ f
{B = B , b₀} f = funext∙
suspend∙-is-unit (Susp-elim _ refl (merid b₀) λ a →
(flip₁ (∙-filler (merid (f .fst a)) (sym (merid b₀)))))
transpose refl
opaque
unfolding Ω¹-map
suspend∙-natural: ∀ {ℓa ℓb} {A : Type∙ ℓa} {B : Type∙ ℓb} (f : A →∙ B)
→ Ω¹-map (Susp-map∙ f) ∘∙ suspend∙ A ≡ suspend∙ B ∘∙ f
{A = A∙@(A , a₀)} {B = B∙@(B , b₀)} f∙@(f , pt) =
suspend∙-natural λ a →
homogeneous-funext∙ (ap (Susp-map f) (merid a ∙ sym (merid a₀))) ≡⟨ conj-refl _ ⟩
conj refl (Susp-map f) (merid a ∙ sym (merid a₀)) ≡⟨ ap-∙ (Susp-map f) (merid a) (sym (merid a₀)) ⟩
ap (f a) ∙ sym (merid ⌜ f a₀ ⌝) ≡⟨ ap! pt ⟩
merid (f a) ∙ sym (merid b₀) ∎ merid
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. 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
opaque
: ∀ {ℓ} {A : Type∙ ℓ} n → Maps∙ (Sⁿ n) A ≃∙ Ωⁿ n A
Ωⁿ≃∙Sⁿ-map {A = A} zero = Iso→Equiv (from , iso to (λ _ → refl) invr) , refl where
Ωⁿ≃∙Sⁿ-map to : A .fst → ((Susp ⊥ , north) →∙ A)
to x .fst north = A .snd
to x .fst south = x
to x .snd = refl
: ((Susp ⊥ , north) →∙ A) → A .fst
from = f .fst south
from f
: is-right-inverse from to
invr (x , p) = Σ-pathp (funext (λ { north → sym p ; south → refl }))
invr λ i j → p (~ i ∨ j)
{A = A} (suc n) =
Ωⁿ≃∙Sⁿ-map (Σ¹ (Sⁿ n)) A ≃∙⟨ Σ-map∙≃∙map∙-Ω ⟩
Maps∙ (Sⁿ n) (Ωⁿ 1 A) ≃∙⟨ Ωⁿ≃∙Sⁿ-map n ⟩
Maps∙ (Ωⁿ 1 A) ≃∙˘⟨ Ωⁿ-suc _ n ⟩
Ωⁿ n (suc n) A ≃∙∎
Ωⁿ
: ∀ {ℓ} {A : Type∙ ℓ} n → (Sⁿ n →∙ A) ≃ Ωⁿ n A .fst
Ωⁿ≃Sⁿ-map = Ωⁿ≃∙Sⁿ-map n .fst Ωⁿ≃Sⁿ-map n
opaque unfolding Ωⁿ≃∙Sⁿ-map
We also show that this equivalence is natural in
Ωⁿ≃Sⁿ-map-natural: ∀ {ℓ ℓ'} {A : Type∙ ℓ} {B : Type∙ ℓ'} n
→ (f : A →∙ B) (l : Sⁿ n →∙ A)
→ Ωⁿ-map n f · (Ωⁿ≃Sⁿ-map n · l) ≡ Ωⁿ≃Sⁿ-map n · (f ∘∙ l)
Ωⁿ≃Sⁿ-map-natural: ∀ {ℓ ℓ'} {A : Type∙ ℓ} {B : Type∙ ℓ'} n
→ (f : A →∙ B) (l : Sⁿ n →∙ A)
→ Ωⁿ-map n f · (Ωⁿ≃Sⁿ-map n · l) ≡ Ωⁿ≃Sⁿ-map n · (f ∘∙ l)
= refl
Ωⁿ≃Sⁿ-map-natural zero f l (suc n) f l =
Ωⁿ≃Sⁿ-map-natural (suc n) f · (Equiv.from (Ωⁿ-suc _ n .fst) (Ωⁿ≃Sⁿ-map n · (Σ-map∙≃map∙-Ω · l)))
Ωⁿ-map .adjunctl (Ωⁿ-suc _ n .fst) (Ω-suc-natural n f ·ₚ _) ⟩
≡⟨ Equiv.from (Ωⁿ-suc _ n .fst) ⌜ Ωⁿ-map n (Ω¹-map f) .fst (Ωⁿ≃Sⁿ-map n · (Σ-map∙≃map∙-Ω · l)) ⌝
Equiv(Ωⁿ≃Sⁿ-map-natural n (Ω¹-map f) _) ⟩
≡⟨ ap! .from (Ωⁿ-suc _ n .fst) (Ωⁿ≃Sⁿ-map n · ⌜ Ω¹-map f ∘∙ Σ-map∙≃map∙-Ω · l ⌝)
Equiv(Σ-map∙≃map∙-Ω-naturalr f l) ⟩
≡⟨ ap! .from (Ωⁿ-suc _ n .fst) (Ωⁿ≃Sⁿ-map n · (Σ-map∙≃map∙-Ω · (f ∘∙ l)))
Equiv ∎
Ωⁿ≃Sⁿ-map-inv-natural: ∀ {ℓ ℓ'} {A : Type∙ ℓ} {B : Type∙ ℓ'} n
→ (f : A →∙ B) (l : ⌞ Ωⁿ n A ⌟)
→ f ∘∙ Equiv.from (Ωⁿ≃Sⁿ-map n) l ≡ Equiv.from (Ωⁿ≃Sⁿ-map n) (Ωⁿ-map n f .fst l)
= Equiv.adjunctl (Ωⁿ≃Sⁿ-map n) $
Ωⁿ≃Sⁿ-map-inv-natural n f l (Ωⁿ≃Sⁿ-map-natural n f (Equiv.from (Ωⁿ≃Sⁿ-map n) l))
sym (Ωⁿ-map n f .fst) (Equiv.ε (Ωⁿ≃Sⁿ-map n) l) ∙ ap