module Homotopy.Loopspace where
private variable
: Level
ℓ ℓ' : Type∙ ℓ A B C
Loop spaces🔗
Given a pointed type (with basepoint we refer to the type as the loop space of , and write it Since we always have this type is itself naturally pointed.
: Type∙ ℓ → Type∙ ℓ
Ω¹ (A , a₀) = Path A a₀ a₀ , refl Ω¹
If is a pointed map and then we have which, while a loop in the type underlying does not belong to However, conjugation by is an equivalence from to so that lands in the right type. Since we have we see that is itself a pointed map Conjugating by while necessary, introduces some complications in showing that this makes into a functor — if we could “get away with” only using then this would be definitional.
Moreover, it initially looks as though we will have to identify the pointedness proofs as well as the underlying maps, which would require us to state and prove several annoying higher coherences involving conjugation.
opaque: (A →∙ B) → Ω¹ A →∙ Ω¹ B
Ω¹-map (f , pt) .fst α = conj pt (ap f α)
Ω¹-map (f , pt) .snd = conj-of-refl pt Ω¹-map
However, path (and thus loop) spaces are homogeneous, so we can obtain pointed homotopies from unpointed ones. This doesn’t prevent us from having to do path algebra with conjugation, but it does mean we don’t accumulate coherences on our coherences. We can thus show without much pain that preserves identities, composition, and the zero map.
: Ω¹-map {A = A} id∙ ≡ id∙
Ω¹-map-id = homogeneous-funext∙ λ α →
Ω¹-map-id (ap id α) ≡⟨ conj-refl α ⟩
conj refl
α ∎
: (f : B →∙ C) (g : A →∙ B) → Ω¹-map f ∘∙ Ω¹-map g ≡ Ω¹-map (f ∘∙ g)
Ω¹-map-∘ (f , f*) (g , g*) = homogeneous-funext∙ λ α →
Ω¹-map-∘ (ap f (conj g* (ap g α))) ≡⟨ ap (conj f*) (ap-conj f g* (ap g α)) ⟩
conj f* (conj (ap f g*) (ap (f ∘ g) α)) ≡⟨ conj-∙ _ _ _ ⟩
conj f* (ap f g* ∙ f*) (ap (f ∘ g) α) ∎
conj
: Ω¹-map (zero∙ {A = A} {B = B}) ≡ zero∙
Ω¹-map-zero {B = B} = Σ-pathp (funext λ _ → conj-refl _) conj-refl-square Ω¹-map-zero
: Maps∙ A B →∙ Maps∙ (Ω¹ A) (Ω¹ B)
Ω¹-map∙ .fst = Ω¹-map
Ω¹-map∙ .snd = Ω¹-map-zero
Ω¹-map∙
: Ω¹ (Maps∙ A B) ≃∙ Maps∙ A (Ω¹ B)
Ω-Maps∙ .fst = twist , eqv where
Ω-Maps∙ : Ω¹ (Maps∙ _ _) .fst → Maps∙ _ (Ω¹ _) .fst
twist .fst x i = p i .fst x
twist p .snd i j = p j .snd i
twist p
: is-equiv twist
eqv = is-iso→is-equiv λ where
eqv .is-iso.from f i → (λ x → f .fst x i) , (λ j → f .snd j i)
.is-iso.rinv x → refl
.is-iso.linv x → refl
{B = B} .snd i = (λ x j → B .snd) , λ j k → B .snd Ω-Maps∙
In passing, we note that since of a (pointed) equivalence is itself an equivalence, and so is conjugation, the functorial action of defined above carries equivalences to equivalences.
opaque
unfolding Ω¹-map
: A ≃∙ B → Ω¹ A ≃∙ Ω¹ B
Ω¹-ap .fst .fst = Ω¹-map (Equiv∙.to∙ f) .fst
Ω¹-ap f .fst .snd = ∘-is-equiv (conj-is-equiv (f .snd)) (ap-equiv (f .fst) .snd)
Ω¹-ap f .snd = Ω¹-map (Equiv∙.to∙ f) .snd Ω¹-ap f
opaque
unfolding Ω¹-ap conj
: (e : A ≃∙ B) → Equiv∙.inverse (Ω¹-ap e) ≡ Ω¹-ap (Equiv∙.inverse e)
Ω¹-ap-inv (e , pt) = ≃∙-ext $ homogeneous-funext∙ λ l →
Ω¹-ap-inv let e⁻¹ = Equiv.from e in
(Equiv.η e _) ⌜ ap e⁻¹ (conj (sym pt) l) ⌝ ≡⟨ ap! (ap-conj e⁻¹ _ _) ⟩
conj (Equiv.η e _) (conj (ap e⁻¹ (sym pt)) (ap e⁻¹ l)) ≡⟨ conj-∙ _ _ _ ⟩
conj _ ⌝ (ap e⁻¹ l) ≡˘⟨ ap¡ (sym-∙ _ _) ⟩
conj ⌜ (sym (Equiv.adjunctl e pt)) (ap e⁻¹ l) ∎ conj
Finally, since both conjugation and do so, preserves path composition.
opaque
unfolding Ω¹-map
Ω¹-map-∙: ∀ (f : A →∙ B) p q
→ Ω¹-map f · (p ∙ q) ≡ Ω¹-map f · p ∙ Ω¹-map f · q
(f , f*) p q =
Ω¹-map-∙ (ap f (p ∙ q)) ≡⟨ ap (conj f*) (ap-∙ f _ _) ⟩
conj f* (ap f p ∙ ap f q) ≡⟨ conj-of-∙ f* _ _ ⟩
conj f* (ap f p) ∙ conj f* (ap f q) ∎ conj f*
Higher loop spaces🔗
Since is an endofunctor of pointed types, it can be iterated, producing the higher loop spaces of with the convention that is simply We can also equip with a functorial action by iterating that of
: Nat → Type∙ ℓ → Type∙ ℓ
Ωⁿ = A
Ωⁿ zero A (suc n) A = Ω¹ (Ωⁿ n A)
Ωⁿ
: ∀ n (e : A →∙ B) → Ωⁿ n A →∙ Ωⁿ n B
Ωⁿ-map = f
Ωⁿ-map zero f (suc n) f = Ω¹-map (Ωⁿ-map n f) Ωⁿ-map
To show that this is a pointed functorial action preserving path
composition, we simply iterate the proofs that
is functorial. This gives us the following battery of lemmas:
: ∀ n → Ωⁿ-map {A = A} n id∙ ≡ id∙
Ωⁿ-map-id : ∀ n → Ωⁿ-map n (zero∙ {A = A} {B = B}) ≡ zero∙
Ωⁿ-map-zero
Ωⁿ-map-∘: ∀ n → (f : B →∙ C) (g : A →∙ B)
→ Ωⁿ-map n f ∘∙ Ωⁿ-map n g ≡ Ωⁿ-map n (f ∘∙ g)
Ωⁿ-map-∙: ∀ n (f : A →∙ B) p q
→ Ωⁿ-map (suc n) f · (p ∙ q)
(suc n) f · p ∙ Ωⁿ-map (suc n) f · q ≡ Ωⁿ-map
: ∀ n → Ωⁿ-map {A = A} n id∙ ≡ id∙
Ωⁿ-map-id : ∀ n → Ωⁿ-map n (zero∙ {A = A} {B = B}) ≡ zero∙
Ωⁿ-map-zero
Ωⁿ-map-∘: ∀ n → (f : B →∙ C) (g : A →∙ B)
→ Ωⁿ-map n f ∘∙ Ωⁿ-map n g ≡ Ωⁿ-map n (f ∘∙ g)
Ωⁿ-map-∙: ∀ n (f : A →∙ B) p q
→ Ωⁿ-map (suc n) f · (p ∙ q)
(suc n) f · p ∙ Ωⁿ-map (suc n) f · q ≡ Ωⁿ-map
= refl
Ωⁿ-map-id zero (suc n) = ap Ω¹-map (Ωⁿ-map-id n) ∙ Ω¹-map-id
Ωⁿ-map-id
= refl
Ωⁿ-map-∘ zero f g (suc n) f g = Ω¹-map-∘ (Ωⁿ-map n f) (Ωⁿ-map n g) ∙ ap Ω¹-map (Ωⁿ-map-∘ n f g)
Ωⁿ-map-∘
= refl
Ωⁿ-map-zero zero (suc n) = ap Ω¹-map (Ωⁿ-map-zero n) ∙ Ω¹-map-zero
Ωⁿ-map-zero
= Ω¹-map-∙ (Ωⁿ-map n f) p q Ωⁿ-map-∙ n f p q
opaque
unfolding Ω¹-ap
: ∀ n (f : A →∙ B) → is-equiv (f .fst) → is-equiv (Ωⁿ-map n f .fst)
Ωⁿ-map-equiv = e
Ωⁿ-map-equiv zero f e (suc n) f e = Ω¹-ap ((_ , Ωⁿ-map-equiv n f e) , _) .fst .snd
Ωⁿ-map-equiv
: ∀ n → Maps∙ A B →∙ Maps∙ (Ωⁿ n A) (Ωⁿ n B)
Ωⁿ-map∙ .fst = Ωⁿ-map n
Ωⁿ-map∙ n .snd = Ωⁿ-map-zero n
Ωⁿ-map∙ n
Ωⁿ-ap: ∀ {ℓ ℓ'} {A : Type∙ ℓ} {B : Type∙ ℓ'} n (e : A ≃∙ B)
→ Ωⁿ n A ≃∙ Ωⁿ n B
{A = A} {B = B} n e∙@((e , eq) , p) = record
Ωⁿ-ap { fst = e' .fst , Ωⁿ-map-equiv n (e , p) eq
; snd = e' .snd
}
where e' = Ωⁿ-map n (e , p)
: ∀ n → Ωⁿ-ap {A = A} n id≃∙ ≡ id≃∙
Ωⁿ-ap-id with p ← Ωⁿ-map-id n = Σ-pathp (Σ-prop-path! (ap fst p)) (ap snd p) Ωⁿ-ap-id n
Loop spaces at successors🔗
Note that, unlike the HoTT book (2013), we have defined iterated loop spaces so that they satisfy the recurrence rather than This has the benefit of definitionally “exposing” that is a iterated path type as soon as is a literal number, but it does significantly complicate showing that the lifting agrees with they don’t even live in the same type— and this identification turns out to be a key component in Whitehead’s lemma.
We can, however, recursively identify with and this identification naturally generates a pointed equivalence. If desired, we could define the “book” version of and use this identification to prove that it agrees with ours.
: ∀ (A : Type∙ ℓ) n → Ωⁿ (suc n) A ≡ Ωⁿ n (Ω¹ A)
Ωⁿ-sucP = refl
Ωⁿ-sucP A zero (suc n) = ap Ω¹ (Ωⁿ-sucP A n)
Ωⁿ-sucP A
: ∀ (A : Type∙ ℓ) n → Ωⁿ (suc n) A ≃∙ Ωⁿ n (Ω¹ A)
Ωⁿ-suc = path→equiv∙ (Ωⁿ-sucP A n) Ωⁿ-suc A n
Over this identification (on both the domain and codomain) we can then, and once again recursively, identify with
Ω-suc-naturalP: ∀ {A : Type∙ ℓ} {B : Type∙ ℓ'} n (f : A →∙ B)
→ PathP (λ i → Ωⁿ-sucP A n i →∙ Ωⁿ-sucP B n i)
(Ωⁿ-map (suc n) f) (Ωⁿ-map n (Ω¹-map f))
= refl
Ω-suc-naturalP zero f (suc n) f = apd (λ i → Ω¹-map) (Ω-suc-naturalP n f) Ω-suc-naturalP
abstract
Ω-suc-natural: ∀ {A : Type∙ ℓ} {B : Type∙ ℓ'} n (f : A →∙ B)
→ Equiv∙.to∙ (Ωⁿ-suc B n) ∘∙ Ωⁿ-map (suc n) f ∘∙ Equiv∙.from∙ (Ωⁿ-suc A n)
(Ω¹-map f)
≡ Ωⁿ-map n {A = A} {B = B} n f =
Ω-suc-natural let
instance _ : Homogeneous (Ωⁿ n (Ωⁿ 1 B) .fst)
_ = subst Homogeneous (ap fst (Ωⁿ-sucP B n)) auto
in homogeneous-funext∙ λ x → from-pathp {A = λ i → ⌞ Ωⁿ-sucP B n i ⌟} λ i →
.fst (coe1→i (λ i → ⌞ Ωⁿ-sucP A n i ⌟) i x) Ω-suc-naturalP n f i
References
- Univalent Foundations Program, The. 2013. Homotopy Type Theory: Univalent Foundations of Mathematics. Institute for Advanced Study: https://homotopytypetheory.org/book.