Documentation

Mathlib.Topology.UniformSpace.CompactConvergence

Compact convergence (uniform convergence on compact sets) #

Given a topological space α and a uniform space β (e.g., a metric space or a topological group), the space of continuous maps C(α, β) carries a natural uniform space structure. We define this uniform space structure in this file and also prove the following properties of the topology it induces on C(α, β):

  1. Given a sequence of continuous functions Fₙ : α → β together with some continuous f : α → β, then Fₙ converges to f as a sequence in C(α, β) iff Fₙ converges to f uniformly on each compact subset K of α.
  2. Given Fₙ and f as above and suppose α is locally compact, then Fₙ converges to f iff Fₙ converges to f locally uniformly.
  3. The topology coincides with the compact-open topology.

Property 1 is essentially true by definition, 2 follows from basic results about uniform convergence, but 3 requires a little work and uses the Lebesgue number lemma.

The uniform space structure #

Given subsets K ⊆ α and V ⊆ β × β, let E(K, V) ⊆ C(α, β) × C(α, β) be the set of pairs of continuous functions α → β which are V-close on K: $$ E(K, V) = { (f, g) | ∀ (x ∈ K), (f x, g x) ∈ V }. $$ Fixing some f ∈ C(α, β), let N(K, V, f) ⊆ C(α, β) be the set of continuous functions α → β which are V-close to f on K: $$ N(K, V, f) = { g | ∀ (x ∈ K), (f x, g x) ∈ V }. $$ Using this notation we can describe the uniform space structure and the topology it induces. Specifically:

The topology on C(α, β) thus has a natural subbasis (the compact-open subbasis) and a natural neighbourhood basis (the compact-convergence neighbourhood basis).

Main definitions / results #

Implementation details #

We use the forgetful inheritance pattern (see Note [forgetful inheritance]) to make the topology of the uniform space structure on C(α, β) definitionally equal to the compact-open topology.

TODO #

def ContinuousMap.compactConvNhd {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] (K : Set α) (V : Set (β × β)) (f : C(α, β)) :
Set C(α, β)

Given K ⊆ α, V ⊆ β × β, and f : C(α, β), we define ContinuousMap.compactConvNhd K V f to be the set of g : C(α, β) that are V-close to f on K.

Equations
Instances For
    theorem ContinuousMap.self_mem_compactConvNhd {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] {K : Set α} {V : Set (β × β)} (f : C(α, β)) (hV : V uniformity β) :
    theorem ContinuousMap.compactConvNhd_mono {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] {K : Set α} {V : Set (β × β)} (f : C(α, β)) {V' : Set (β × β)} (hV' : V' V) :
    theorem ContinuousMap.compactConvNhd_mem_comp {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] {K : Set α} {V : Set (β × β)} (f : C(α, β)) {g₁ : C(α, β)} {g₂ : C(α, β)} {V' : Set (β × β)} (hg₁ : g₁ ContinuousMap.compactConvNhd K V f) (hg₂ : g₂ ContinuousMap.compactConvNhd K V' g₁) :
    theorem ContinuousMap.compactConvNhd_nhd_basis {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] {K : Set α} {V : Set (β × β)} (f : C(α, β)) (hV : V uniformity β) :

    A key property of ContinuousMap.compactConvNhd. It allows us to apply TopologicalSpace.nhds_mkOfNhds_filterBasis below.

    theorem ContinuousMap.compactConvNhd_subset_inter {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] (f : C(α, β)) (K₁ : Set α) (K₂ : Set α) (V₁ : Set (β × β)) (V₂ : Set (β × β)) :
    theorem ContinuousMap.compactConvNhd_filter_isBasis {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] (f : C(α, β)) :
    Filter.IsBasis (fun (KV : Set α × Set (β × β)) => IsCompact KV.1 KV.2 uniformity β) fun (KV : Set α × Set (β × β)) => ContinuousMap.compactConvNhd KV.1 KV.2 f

    A filter basis for the neighbourhood filter of a point in the compact-convergence topology.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For

      The compact-convergence topology. In fact, see ContinuousMap.compactOpen_eq_compactConvergence this is the same as the compact-open topology. This definition is thus an auxiliary convenience definition and is unlikely to be of direct use.

      Equations
      Instances For
        theorem ContinuousMap.hasBasis_nhds_compactConvergence {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] (f : C(α, β)) :
        Filter.HasBasis (nhds f) (fun (p : Set α × Set (β × β)) => IsCompact p.1 p.2 uniformity β) fun (p : Set α × Set (β × β)) => ContinuousMap.compactConvNhd p.1 p.2 f
        theorem ContinuousMap.tendsto_iff_forall_compact_tendstoUniformlyOn' {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] (f : C(α, β)) {ι : Type u₃} {p : Filter ι} {F : ιC(α, β)} :
        Filter.Tendsto F p (nhds f) ∀ (K : Set α), IsCompact KTendstoUniformlyOn (fun (i : ι) (a : α) => (F i) a) (f) p K

        This is an auxiliary lemma and is unlikely to be of direct use outside of this file. See ContinuousMap.tendsto_iff_forall_compact_tendstoUniformlyOn below for the useful version where the topology is picked up via typeclass inference.

        theorem ContinuousMap.compactConvNhd_subset_compactOpen {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] {K : Set α} (f : C(α, β)) (hK : IsCompact K) {U : Set β} (hU : IsOpen U) (hf : f ContinuousMap.CompactOpen.gen K U) :

        Any point of ContinuousMap.CompactOpen.gen K U is also an interior point wrt the topology of compact convergence.

        The topology of compact convergence is thus at least as fine as the compact-open topology.

        theorem ContinuousMap.iInter_compactOpen_gen_subset_compactConvNhd {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] {K : Set α} {V : Set (β × β)} (f : C(α, β)) (hK : IsCompact K) (hV : V uniformity β) :
        ∃ (ι : Type u₁) (x : Fintype ι) (C : ιSet α) (_ : ∀ (i : ι), IsCompact (C i)) (U : ιSet β) (_ : ∀ (i : ι), IsOpen (U i)), f ⋂ (i : ι), ContinuousMap.CompactOpen.gen (C i) (U i) ⋂ (i : ι), ContinuousMap.CompactOpen.gen (C i) (U i) ContinuousMap.compactConvNhd K V f

        The point f in ContinuousMap.compactConvNhd K V f is also an interior point wrt the compact-open topology.

        Since ContinuousMap.compactConvNhd K V f are a neighbourhood basis at f for each f, it follows that the compact-open topology is at least as fine as the topology of compact convergence.

        theorem ContinuousMap.compactOpen_eq_compactConvergence {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] :
        ContinuousMap.compactOpen = ContinuousMap.compactConvergenceTopology

        The compact-open topology is equal to the compact-convergence topology.

        The filter on C(α, β) × C(α, β) which underlies the uniform space structure on C(α, β).

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          theorem ContinuousMap.hasBasis_compactConvergenceUniformity_aux {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] :
          Filter.HasBasis ContinuousMap.compactConvergenceUniformity (fun (p : Set α × Set (β × β)) => IsCompact p.1 p.2 uniformity β) fun (p : Set α × Set (β × β)) => {fg : C(α, β) × C(α, β) | xp.1, (fg.1 x, fg.2 x) p.2}
          theorem ContinuousMap.mem_compactConvergenceUniformity {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] (X : Set (C(α, β) × C(α, β))) :
          X ContinuousMap.compactConvergenceUniformity ∃ (K : Set α) (V : Set (β × β)), IsCompact K V uniformity β {fg : C(α, β) × C(α, β) | xK, (fg.1 x, fg.2 x) V} X

          An intermediate lemma. Usually ContinuousMap.mem_compactConvergence_entourage_iff is more useful.

          Note that we ensure the induced topology is definitionally the compact-open topology.

          Equations
          • One or more equations did not get rendered due to their size.
          theorem ContinuousMap.mem_compactConvergence_entourage_iff {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] (X : Set (C(α, β) × C(α, β))) :
          X uniformity C(α, β) ∃ (K : Set α) (V : Set (β × β)), IsCompact K V uniformity β {fg : C(α, β) × C(α, β) | xK, (fg.1 x, fg.2 x) V} X
          theorem ContinuousMap.hasBasis_compactConvergenceUniformity {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] :
          Filter.HasBasis (uniformity C(α, β)) (fun (p : Set α × Set (β × β)) => IsCompact p.1 p.2 uniformity β) fun (p : Set α × Set (β × β)) => {fg : C(α, β) × C(α, β) | xp.1, (fg.1 x, fg.2 x) p.2}
          theorem Filter.HasBasis.compactConvergenceUniformity {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] {ι : Type u_1} {pi : ιProp} {s : ιSet (β × β)} (h : Filter.HasBasis (uniformity β) pi s) :
          Filter.HasBasis (uniformity C(α, β)) (fun (p : Set α × ι) => IsCompact p.1 pi p.2) fun (p : Set α × ι) => {fg : C(α, β) × C(α, β) | xp.1, (fg.1 x, fg.2 x) s p.2}
          theorem ContinuousMap.tendsto_iff_forall_compact_tendstoUniformlyOn {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] {f : C(α, β)} {ι : Type u₃} {p : Filter ι} {F : ιC(α, β)} :
          Filter.Tendsto F p (nhds f) ∀ (K : Set α), IsCompact KTendstoUniformlyOn (fun (i : ι) (a : α) => (F i) a) (f) p K
          theorem ContinuousMap.tendsto_of_tendstoLocallyUniformly {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] {f : C(α, β)} {ι : Type u₃} {p : Filter ι} {F : ιC(α, β)} (h : TendstoLocallyUniformly (fun (i : ι) (a : α) => (F i) a) (f) p) :

          Locally uniform convergence implies convergence in the compact-open topology.

          theorem ContinuousMap.tendsto_iff_tendstoLocallyUniformly {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] {f : C(α, β)} {ι : Type u₃} {p : Filter ι} {F : ιC(α, β)} [WeaklyLocallyCompactSpace α] :
          Filter.Tendsto F p (nhds f) TendstoLocallyUniformly (fun (i : ι) (a : α) => (F i) a) (f) p

          In a weakly locally compact space, convergence in the compact-open topology is the same as locally uniform convergence.

          The right-to-left implication holds in any topological space, see ContinuousMap.tendsto_of_tendstoLocallyUniformly.

          @[deprecated ContinuousMap.tendsto_iff_tendstoLocallyUniformly]
          theorem ContinuousMap.tendstoLocallyUniformly_of_tendsto {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] {f : C(α, β)} {ι : Type u₃} {p : Filter ι} {F : ιC(α, β)} [WeaklyLocallyCompactSpace α] (h : Filter.Tendsto F p (nhds f)) :
          TendstoLocallyUniformly (fun (i : ι) (a : α) => (F i) a) (f) p
          theorem ContinuousMap.hasBasis_compactConvergenceUniformity_of_compact {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] [CompactSpace α] :
          Filter.HasBasis (uniformity C(α, β)) (fun (V : Set (β × β)) => V uniformity β) fun (V : Set (β × β)) => {fg : C(α, β) × C(α, β) | ∀ (x : α), (fg.1 x, fg.2 x) V}
          theorem ContinuousMap.tendsto_iff_tendstoUniformly {α : Type u₁} {β : Type u₂} [TopologicalSpace α] [UniformSpace β] {f : C(α, β)} {ι : Type u₃} {p : Filter ι} {F : ιC(α, β)} [CompactSpace α] :
          Filter.Tendsto F p (nhds f) TendstoUniformly (fun (i : ι) (a : α) => (F i) a) (f) p

          Convergence in the compact-open topology is the same as uniform convergence for sequences of continuous functions on a compact space.