nth homotopy group #
We define the nth homotopy group at x : X, π_n X x, as the equivalence classes
of functions from the n-dimensional cube to the topological space X
that send the boundary to the base point x, up to homotopic equivalence.
Note that such functions are generalized loops GenLoop (Fin n) x; in particular
GenLoop (Fin 1) x ≃ Path x x.
We show that π_0 X x is equivalent to the path-connected components, and
that π_1 X x is equivalent to the fundamental group at x.
We provide a group instance using path composition and show commutativity when n > 1.
definitions #
GenLoop N xis the type of continuous functionsI^N → Xthat send the boundary tox,HomotopyGroup.Pi n X xdenotedπ_ n X xis the quotient ofGenLoop (Fin n) xby homotopy relative to the boundary,- group instance
Group (π_(n+1) X x), - commutative group instance
CommGroup (π_(n+2) X x).
TODO:
Ω^M (Ω^N X) ≃ₜ Ω^(M⊕N) X, andΩ^M X ≃ₜ Ω^N XwhenM ≃ N. Similarly forπ_.- Path-induced homomorphisms. Show that
HomotopyGroup.pi1EquivFundamentalGroupis a group isomorphism. - Examples with
𝕊^n:π_n (𝕊^n) = ℤ,π_m (𝕊^n)trivial form < n. - Actions of π_1 on π_n.
- Lie algebra:
⁅π_(n+1), π_(m+1)⁆contained inπ_(n+m+1).
I^N is notation (in the Topology namespace) for N → I,
i.e. the unit cube indexed by a type N.
Equations
- Topology.«termI^_» = Lean.ParserDescr.node `Topology.«termI^_» 1022 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol "I^") (Lean.ParserDescr.cat `term 0))
Instances For
The points in a cube with at least one projection equal to 0 or 1.
Equations
- Cube.boundary N = {y : N → ↑unitInterval | ∃ (i : N), y i = 0 ∨ y i = 1}
Instances For
The forward direction of the homeomorphism between the cube $I^N$ and $I × I^{N\setminus\{j\}}$.
Equations
Instances For
The backward direction of the homeomorphism between the cube $I^N$ and $I × I^{N\setminus\{j\}}$.
Equations
- Cube.insertAt i = (Homeomorph.funSplitAt (↑unitInterval) i).symm
Instances For
The space of paths with both endpoints equal to a specified point x : X.
Denoted as Ω, within the Topology.Homotopy namespace.
Equations
- Topology.Homotopy.termΩ = Lean.ParserDescr.node `Topology.Homotopy.termΩ 1024 (Lean.ParserDescr.symbol "Ω")
Instances For
Equations
- LoopSpace.inhabited X x = { default := Path.refl x }
The n-dimensional generalized loops based at x in a space X are
continuous functions I^n → X that sends the boundary to x.
We allow an arbitrary indexing type N in place of Fin n here.
Instances For
The n-dimensional generalized loops based at x in a space X are
continuous functions I^n → X that sends the boundary to x.
We allow an arbitrary indexing type N in place of Fin n here.
Equations
- Topology.Homotopy.«termΩ^» = Lean.ParserDescr.node `Topology.Homotopy.«termΩ^» 1024 (Lean.ParserDescr.symbol "Ω^")
Instances For
Equations
- GenLoop.instFunLike = { coe := fun (f : ↑(GenLoop N X x)) => ⇑↑f, coe_injective' := ⋯ }
Copy of a GenLoop with a new map from the unit cube equal to the old one.
Useful to fix definitional equalities.
Instances For
The constant GenLoop at x.
Equations
- GenLoop.const = ⟨ContinuousMap.const (N → ↑unitInterval) x, ⋯⟩
Instances For
Equations
- GenLoop.inhabited = { default := GenLoop.const }
The "homotopic relative to boundary" relation between GenLoops.
Equations
- GenLoop.Homotopic f g = (↑f).HomotopicRel (↑g) (Cube.boundary N)
Instances For
Equations
- GenLoop.Homotopic.setoid N x = { r := GenLoop.Homotopic, iseqv := ⋯ }
Loop from a generalized loop by currying $I^N → X$ into $I → (I^{N\setminus\{j\}} → X)$.
Equations
- GenLoop.toLoop i p = { toFun := fun (t : ↑unitInterval) => ⟨((↑p).comp ↑(Cube.insertAt i)).curry t, ⋯⟩, continuous_toFun := ⋯, source' := ⋯, target' := ⋯ }
Instances For
Generalized loop from a loop by uncurrying $I → (I^{N\setminus\{j\}} → X)$ into $I^N → X$.
Equations
- GenLoop.fromLoop i p = ⟨({ toFun := Subtype.val, continuous_toFun := ⋯ }.comp p.toContinuousMap).uncurry.comp ↑(Cube.splitAt i), ⋯⟩
Instances For
The n+1-dimensional loops are in bijection with the loops in the space of
n-dimensional loops with base point const.
We allow an arbitrary indexing type N in place of Fin n here.
Equations
- GenLoop.loopHomeo i = { toFun := GenLoop.toLoop i, invFun := GenLoop.fromLoop i, left_inv := ⋯, right_inv := ⋯, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Composition with Cube.insertAt as a continuous map.
Equations
- GenLoop.cCompInsert i = { toFun := fun (f : C(N → ↑unitInterval, X)) => f.comp ↑(Cube.insertAt i), continuous_toFun := ⋯ }
Instances For
A homotopy between n+1-dimensional loops p and q constant on the boundary
seen as a homotopy between two paths in the space of n-dimensional paths.
Equations
- GenLoop.homotopyTo i H = ({ toFun := ContinuousMap.curry, continuous_toFun := ⋯ }.comp ((GenLoop.cCompInsert i).comp H.curry)).uncurry
Instances For
The converse to GenLoop.homotopyTo: a homotopy between two loops in the space of
n-dimensional loops can be seen as a homotopy between two n+1-dimensional paths.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Concatenation of two GenLoops along the ith coordinate.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Reversal of a GenLoop along the ith coordinate.
Equations
- GenLoop.symmAt i f = GenLoop.copy (GenLoop.fromLoop i (Path.symm (GenLoop.toLoop i f))) (fun (t : N → ↑unitInterval) => f fun (j : N) => if j = i then unitInterval.symm (t i) else t j) ⋯
Instances For
The nth homotopy group at x defined as the quotient of Ω^n x by the
GenLoop.Homotopic relation.
Equations
- HomotopyGroup N X x = Quotient (GenLoop.Homotopic.setoid N x)
Instances For
Equations
Equivalence between the homotopy group of X and the fundamental group of
Ω^{j // j ≠ i} x.
Equations
- homotopyGroupEquivFundamentalGroup i = (Quotient.congr (GenLoop.loopHomeo i).toEquiv ⋯).trans (CategoryTheory.Groupoid.isoEquivHom { as := GenLoop.const } { as := GenLoop.const }).symm
Instances For
Homotopy group of finite index, denoted as π_n within the Topology namespace.
Equations
- HomotopyGroup.Pi n X x = HomotopyGroup (Fin n) X x
Instances For
Homotopy group of finite index, denoted as π_n within the Topology namespace.
Equations
- Topology.termπ_ = Lean.ParserDescr.node `Topology.termπ_ 1024 (Lean.ParserDescr.symbol "π_")
Instances For
The 0-dimensional generalized loops based at x are in bijection with X.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The homotopy "group" indexed by an empty type is in bijection with
the path components of X, aka the ZerothHomotopy.
Equations
Instances For
The 0th homotopy "group" is in bijection with ZerothHomotopy.
Instances For
The 1-dimensional generalized loops based at x are in bijection with loops at x.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The homotopy group at x indexed by a singleton is in bijection with the fundamental group,
i.e. the loops based at x up to homotopy.
Equations
- homotopyGroupEquivFundamentalGroupOfUnique N = (Quotient.congr (genLoopEquivOfUnique N) ⋯).trans (CategoryTheory.Groupoid.isoEquivHom { as := x } { as := x }).symm
Instances For
The first homotopy group at x is in bijection with the fundamental group.
Equations
Instances For
Group structure on HomotopyGroup N X x for nonempty N (in particular π_(n+1) X x).
Equations
Group structure on HomotopyGroup obtained by pulling back path composition along the
ith direction. The group structures for two different i j : N distribute over each
other, and therefore are equal by the Eckmann-Hilton argument.
Equations
Instances For
Characterization of multiplicative identity
Characterization of multiplication
Characterization of multiplicative inverse
Multiplication on HomotopyGroup N X x is commutative for nontrivial N.
In particular, multiplication on π_(n+2) is commutative.