Documentation

Mathlib.Topology.ContinuousFunction.CocompactMap

Cocompact continuous maps #

The type of cocompact continuous maps are those which tend to the cocompact filter on the codomain along the cocompact filter on the domain. When the domain and codomain are Hausdorff, this is equivalent to many other conditions, including that preimages of compact sets are compact.

Cocompact continuous maps #

structure CocompactMap (α : Type u) (β : Type v) [TopologicalSpace α] [TopologicalSpace β] extends ContinuousMap :
Type (max u v)

A cocompact continuous map is a continuous function between topological spaces which tends to the cocompact filter along the cocompact filter. Functions for which preimages of compact sets are compact always satisfy this property, and the converse holds for cocompact continuous maps when the codomain is Hausdorff (see CocompactMap.tendsto_of_forall_preimage and CocompactMap.isCompact_preimage).

Cocompact maps thus generalise proper maps, with which they correspond when the codomain is Hausdorff.

Instances For
    class CocompactMapClass (F : Type u_1) (α : outParam (Type u_2)) (β : outParam (Type u_3)) [TopologicalSpace α] [TopologicalSpace β] [FunLike F α β] extends ContinuousMapClass :

    CocompactMapClass F α β states that F is a type of cocompact continuous maps.

    You should also extend this typeclass when you extend CocompactMap.

    Instances
      def CocompactMapClass.toCocompactMap {F : Type u_1} {α : Type u_2} {β : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [FunLike F α β] [CocompactMapClass F α β] (f : F) :

      Turn an element of a type F satisfying CocompactMapClass F α β into an actual CocompactMap. This is declared as the default coercion from F to CocompactMap α β.

      Equations
      • f = let __src := f; { toContinuousMap := __src, cocompact_tendsto' := }
      Instances For
        instance CocompactMapClass.instCoeTCCocompactMap {F : Type u_1} {α : Type u_2} {β : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [FunLike F α β] [CocompactMapClass F α β] :
        Equations
        • CocompactMapClass.instCoeTCCocompactMap = { coe := CocompactMapClass.toCocompactMap }
        Equations
        • CocompactMap.instFunLikeCocompactMap = { coe := fun (f : CocompactMap α β) => f.toFun, coe_injective' := }
        @[simp]
        theorem CocompactMap.coe_toContinuousMap {α : Type u_1} {β : Type u_2} [TopologicalSpace α] [TopologicalSpace β] {f : CocompactMap α β} :
        f.toContinuousMap = f
        theorem CocompactMap.ext {α : Type u_1} {β : Type u_2} [TopologicalSpace α] [TopologicalSpace β] {f : CocompactMap α β} {g : CocompactMap α β} (h : ∀ (x : α), f x = g x) :
        f = g
        def CocompactMap.copy {α : Type u_1} {β : Type u_2} [TopologicalSpace α] [TopologicalSpace β] (f : CocompactMap α β) (f' : αβ) (h : f' = f) :

        Copy of a CocompactMap with a new toFun equal to the old one. Useful to fix definitional equalities.

        Equations
        • CocompactMap.copy f f' h = { toContinuousMap := { toFun := f', continuous_toFun := }, cocompact_tendsto' := }
        Instances For
          @[simp]
          theorem CocompactMap.coe_copy {α : Type u_1} {β : Type u_2} [TopologicalSpace α] [TopologicalSpace β] (f : CocompactMap α β) (f' : αβ) (h : f' = f) :
          (CocompactMap.copy f f' h) = f'
          theorem CocompactMap.copy_eq {α : Type u_1} {β : Type u_2} [TopologicalSpace α] [TopologicalSpace β] (f : CocompactMap α β) (f' : αβ) (h : f' = f) :
          @[simp]
          theorem CocompactMap.coe_mk {α : Type u_1} {β : Type u_2} [TopologicalSpace α] [TopologicalSpace β] (f : C(α, β)) (h : Filter.Tendsto (f) (Filter.cocompact α) (Filter.cocompact β)) :
          { toContinuousMap := f, cocompact_tendsto' := h } = f

          The identity as a cocompact continuous map.

          Equations
          Instances For
            @[simp]
            theorem CocompactMap.coe_id (α : Type u_1) [TopologicalSpace α] :
            (CocompactMap.id α) = id
            Equations
            def CocompactMap.comp {α : Type u_1} {β : Type u_2} {γ : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [TopologicalSpace γ] (f : CocompactMap β γ) (g : CocompactMap α β) :

            The composition of cocompact continuous maps, as a cocompact continuous map.

            Equations
            Instances For
              @[simp]
              theorem CocompactMap.coe_comp {α : Type u_1} {β : Type u_2} {γ : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [TopologicalSpace γ] (f : CocompactMap β γ) (g : CocompactMap α β) :
              (CocompactMap.comp f g) = f g
              @[simp]
              theorem CocompactMap.comp_apply {α : Type u_1} {β : Type u_2} {γ : Type u_3} [TopologicalSpace α] [TopologicalSpace β] [TopologicalSpace γ] (f : CocompactMap β γ) (g : CocompactMap α β) (a : α) :
              (CocompactMap.comp f g) a = f (g a)
              @[simp]
              theorem CocompactMap.comp_assoc {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} [TopologicalSpace α] [TopologicalSpace β] [TopologicalSpace γ] [TopologicalSpace δ] (f : CocompactMap γ δ) (g : CocompactMap β γ) (h : CocompactMap α β) :
              @[simp]
              theorem CocompactMap.id_comp {α : Type u_1} {β : Type u_2} [TopologicalSpace α] [TopologicalSpace β] (f : CocompactMap α β) :
              @[simp]
              theorem CocompactMap.comp_id {α : Type u_1} {β : Type u_2} [TopologicalSpace α] [TopologicalSpace β] (f : CocompactMap α β) :
              theorem CocompactMap.tendsto_of_forall_preimage {α : Type u_1} {β : Type u_2} [TopologicalSpace α] [TopologicalSpace β] {f : αβ} (h : ∀ (s : Set β), IsCompact sIsCompact (f ⁻¹' s)) :
              theorem CocompactMap.isCompact_preimage {α : Type u_1} {β : Type u_2} [TopologicalSpace α] [TopologicalSpace β] [T2Space β] (f : CocompactMap α β) ⦃s : Set β (hs : IsCompact s) :
              IsCompact (f ⁻¹' s)

              If the codomain is Hausdorff, preimages of compact sets are compact under a cocompact continuous map.

              @[simp]
              theorem Homeomorph.toCocompactMap_toFun {α : Type u_1} {β : Type u_2} [TopologicalSpace α] [TopologicalSpace β] (f : α ≃ₜ β) (a : α) :
              def Homeomorph.toCocompactMap {α : Type u_1} {β : Type u_2} [TopologicalSpace α] [TopologicalSpace β] (f : α ≃ₜ β) :

              A homeomorphism is a cocompact map.

              Equations
              Instances For