Documentation

Mathlib.Topology.Homotopy.LocallyContractible

Strongly locally contractible spaces #

This file defines LocallyContractibleSpace and StronglyLocallyContractibleSpace.

Main definitions #

Main results #

TODO #

Notes #

Terminology: The classical definition of locally contractible (LC) requires that for every point x and neighborhood U ∋ x, there exists a neighborhood V ∋ x with V ⊆ U such that the inclusion V ↪ U is null-homotopic. The definition here is strictly stronger: we require contractible neighborhoods to form a neighborhood basis. This is often called strongly locally contractible (SLC).

Hierarchy of notions:

This naming is not used uniformly: according to https://ncatlab.org/nlab/show/locally+contractible+space the second and third notion here could also be called "locally contractible" and "semilocally contractible" respectively. We've enquired at https://math.stackexchange.com/questions/5109428/terminology-for-local-contractibility-locally-contractible-vs-strongly-local in the hope of gettting definitive naming advice.

The Borsuk-Mazurkiewicz counterexample [BM34] shows that classical LC does not imply SLC. Moreover, from a contractible neighborhood S one generally cannot shrink to an open V ⊆ S that remains contractible, so requiring neighborhoods to be open is potentially strictly stronger than SLC.

Classical local contractibility: for every point and every neighborhood U, there exists a neighborhood V ⊆ U such that the inclusion V ↪ U is null-homotopic.

This is weaker than StronglyLocallyContractibleSpace.

Equations
Instances For

    A topological space is strongly locally contractible if, at every point, contractible neighborhoods form a neighborhood basis. Here "contractible" means contractible as a subspace.

    This is strictly stronger than the classical notion of locally contractible, which only requires null-homotopic inclusions.

    Instances
      theorem StronglyLocallyContractibleSpace.of_bases {X : Type u_1} [TopologicalSpace X] {ι : Type u_3} {p : XιProp} {s : XιSet X} (h : ∀ (x : X), (nhds x).HasBasis (p x) (s x)) (h' : ∀ (x : X) (i : ι), p x iContractibleSpace (s x i)) :

      A helper to construct a strongly locally contractible space from a neighborhood basis where each basis element is contractible as a subspace.

      theorem contractible_subset_basis {X : Type u_1} [TopologicalSpace X] {x : X} [StronglyLocallyContractibleSpace X] {U : Set X} (h : IsOpen U) (hx : x U) :
      (nhds x).HasBasis (fun (s : Set X) => s nhds x ContractibleSpace s s U) id

      In a strongly locally contractible space, for any open set U containing x, there is a basis of contractible neighborhoods of x contained in U.

      @[instance 100]

      Strongly locally contractible spaces are locally path-connected.

      Open embeddings preserve strong local contractibility: if X is strongly locally contractible and e : Y → X is an open embedding, then Y is strongly locally contractible.

      Open subsets of strongly locally contractible spaces are strongly locally contractible.

      The strong notion (contractible neighborhood basis) implies the classical notion (null-homotopic inclusions). The converse is false by the Borsuk-Mazurkiewicz counterexample [BM34]; see also [Mel12] for discussion and the Whitehead manifold example.