Documentation

Mathlib.Topology.Continuous

Continuity in topological spaces #

For topological spaces X and Y, a function f : X → Y and a point x : X, ContinuousAt f x means f is continuous at x, and global continuity is Continuous f. There is also a version of continuity PContinuous for partially defined functions.

Tags #

continuity, continuous function

theorem continuous_def {X : Type u_1} {Y : Type u_2} {x✝ : TopologicalSpace X} {x✝¹ : TopologicalSpace Y} {f : XY} :
Continuous f ∀ (s : Set Y), IsOpen sIsOpen (f ⁻¹' s)
theorem IsOpen.preimage {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} (hf : Continuous f) {t : Set Y} (h : IsOpen t) :
theorem Equiv.isOpenMap_symm_iff {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] (e : X Y) :
theorem continuous_congr {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f g : XY} (h : ∀ (x : X), f x = g x) :
theorem Continuous.congr {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f g : XY} (h : Continuous f) (h' : ∀ (x : X), f x = g x) :
theorem ContinuousAt.tendsto {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {x : X} (h : ContinuousAt f x) :
Filter.Tendsto f (nhds x) (nhds (f x))
theorem continuousAt_def {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {x : X} :
ContinuousAt f x Anhds (f x), f ⁻¹' A nhds x
theorem continuousAt_congr {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {x : X} {g : XY} (h : f =ᶠ[nhds x] g) :
theorem ContinuousAt.congr {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {x : X} {g : XY} (hf : ContinuousAt f x) (h : f =ᶠ[nhds x] g) :
theorem ContinuousAt.preimage_mem_nhds {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {x : X} {t : Set Y} (h : ContinuousAt f x) (ht : t nhds (f x)) :
theorem ContinuousAt.eventually_mem {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {x : X} (hf : ContinuousAt f x) {s : Set Y} (hs : s nhds (f x)) :
∀ᶠ (y : X) in nhds x, f y s

If f x ∈ s ∈ 𝓝 (f x) for continuous f, then f y ∈ s near x.

This is essentially Filter.Tendsto.eventually_mem, but infers in more cases when applied.

theorem not_continuousAt_of_tendsto {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {l₁ : Filter X} {l₂ : Filter Y} {x : X} (hf : Filter.Tendsto f l₁ l₂) [l₁.NeBot] (hl₁ : l₁ nhds x) (hl₂ : Disjoint (nhds (f x)) l₂) :

If a function f tends to somewhere other than 𝓝 (f x) at x, then f is not continuous at x

theorem ClusterPt.map {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {x : X} {lx : Filter X} {ly : Filter Y} (H : ClusterPt x lx) (hfc : ContinuousAt f x) (hf : Filter.Tendsto f lx ly) :
ClusterPt (f x) ly
theorem preimage_interior_subset_interior_preimage {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {t : Set Y} (hf : Continuous f) :

See also interior_preimage_subset_preimage_interior.

theorem continuous_id' {X : Type u_1} [TopologicalSpace X] :
Continuous fun (x : X) => x
theorem Continuous.comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] {f : XY} {g : YZ} (hg : Continuous g) (hf : Continuous f) :
theorem Continuous.comp' {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] {f : XY} {g : YZ} (hg : Continuous g) (hf : Continuous f) :
Continuous fun (x : X) => g (f x)
theorem Continuous.iterate {X : Type u_1} [TopologicalSpace X] {f : XX} (h : Continuous f) (n : ) :
theorem ContinuousAt.comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] {f : XY} {x : X} {g : YZ} (hg : ContinuousAt g (f x)) (hf : ContinuousAt f x) :
theorem ContinuousAt.comp' {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] {f : XY} {g : YZ} {x : X} (hg : ContinuousAt g (f x)) (hf : ContinuousAt f x) :
ContinuousAt (fun (x : X) => g (f x)) x
theorem ContinuousAt.comp_of_eq {X : Type u_1} {Y : Type u_2} {Z : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] {f : XY} {x : X} {y : Y} {g : YZ} (hg : ContinuousAt g y) (hf : ContinuousAt f x) (hy : f x = y) :

See note [comp_of_eq lemmas]

theorem Continuous.tendsto {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} (hf : Continuous f) (x : X) :
Filter.Tendsto f (nhds x) (nhds (f x))
theorem Continuous.tendsto' {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} (hf : Continuous f) (x : X) (y : Y) (h : f x = y) :

A version of Continuous.tendsto that allows one to specify a simpler form of the limit. E.g., one can write continuous_exp.tendsto' 0 1 exp_zero.

theorem Continuous.continuousAt {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {x : X} (h : Continuous f) :
theorem continuous_iff_continuousAt {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} :
Continuous f ∀ (x : X), ContinuousAt f x
theorem continuousAt_const {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {x : X} {y : Y} :
ContinuousAt (fun (x : X) => y) x
theorem continuous_const {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {y : Y} :
Continuous fun (x : X) => y
theorem Filter.EventuallyEq.continuousAt {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {x : X} {y : Y} (h : f =ᶠ[nhds x] fun (x : X) => y) :
theorem continuous_of_const {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} (h : ∀ (x y : X), f x = f y) :
theorem continuousAt_id {X : Type u_1} [TopologicalSpace X] {x : X} :
theorem continuousAt_id' {X : Type u_1} [TopologicalSpace X] (y : X) :
ContinuousAt (fun (x : X) => x) y
theorem ContinuousAt.iterate {X : Type u_1} [TopologicalSpace X] {x : X} {f : XX} (hf : ContinuousAt f x) (hx : f x = x) (n : ) :
theorem continuous_iff_isClosed {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} :
Continuous f ∀ (s : Set Y), IsClosed sIsClosed (f ⁻¹' s)
theorem IsClosed.preimage {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} (hf : Continuous f) {t : Set Y} (h : IsClosed t) :
theorem mem_closure_image {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {s : Set X} {x : X} (hf : ContinuousAt f x) (hx : x closure s) :
f x closure (f '' s)
theorem Continuous.closure_preimage_subset {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} (hf : Continuous f) (t : Set Y) :
theorem Continuous.frontier_preimage_subset {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} (hf : Continuous f) (t : Set Y) :
theorem Set.MapsTo.closure {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {s : Set X} {t : Set Y} (h : MapsTo f s t) (hc : Continuous f) :

If a continuous map f maps s to t, then it maps closure s to closure t.

theorem image_closure_subset_closure_image {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {s : Set X} (h : Continuous f) :
f '' closure s closure (f '' s)

See also IsClosedMap.closure_image_eq_of_continuous.

theorem closure_image_closure {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {s : Set X} (h : Continuous f) :
theorem closure_subset_preimage_closure_image {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {s : Set X} (h : Continuous f) :
theorem map_mem_closure {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {s : Set X} {x : X} {t : Set Y} (hf : Continuous f) (hx : x closure s) (ht : Set.MapsTo f s t) :
f x closure t
theorem Set.MapsTo.closure_left {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {s : Set X} {t : Set Y} (h : MapsTo f s t) (hc : Continuous f) (ht : IsClosed t) :
MapsTo f (closure s) t

If a continuous map f maps s to a closed set t, then it maps closure s to t.

theorem Filter.Tendsto.lift'_closure {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} (hf : Continuous f) {l : Filter X} {l' : Filter Y} (h : Tendsto f l l') :
theorem tendsto_lift'_closure_nhds {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} (hf : Continuous f) (x : X) :

Function with dense range #

theorem Function.Surjective.denseRange {X : Type u_1} [TopologicalSpace X] {α : Type u_4} {f : αX} (hf : Surjective f) :

A surjective map has dense range.

theorem denseRange_iff_closure_range {X : Type u_1} [TopologicalSpace X] {α : Type u_4} {f : αX} :
theorem DenseRange.closure_range {X : Type u_1} [TopologicalSpace X] {α : Type u_4} {f : αX} (h : DenseRange f) :
@[simp]
theorem Continuous.range_subset_closure_image_dense {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {s : Set X} {f : XY} (hf : Continuous f) (hs : Dense s) :
theorem DenseRange.dense_image {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {s : Set X} {f : XY} (hf' : DenseRange f) (hf : Continuous f) (hs : Dense s) :
Dense (f '' s)

The image of a dense set under a continuous map with dense range is a dense set.

theorem DenseRange.subset_closure_image_preimage_of_isOpen {X : Type u_1} [TopologicalSpace X] {α : Type u_4} {f : αX} {s : Set X} (hf : DenseRange f) (hs : IsOpen s) :
s closure (f '' (f ⁻¹' s))

If f has dense range and s is an open set in the codomain of f, then the image of the preimage of s under f is dense in s.

theorem DenseRange.dense_of_mapsTo {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {s : Set X} {f : XY} (hf' : DenseRange f) (hf : Continuous f) (hs : Dense s) {t : Set Y} (ht : Set.MapsTo f s t) :

If a continuous map with dense range maps a dense set to a subset of t, then t is a dense set.

theorem DenseRange.comp {Y : Type u_2} {Z : Type u_3} [TopologicalSpace Y] [TopologicalSpace Z] {α : Type u_4} {g : YZ} {f : αY} (hg : DenseRange g) (hf : DenseRange f) (cg : Continuous g) :

Composition of a continuous map with dense range and a function with dense range has dense range.

theorem DenseRange.nonempty_iff {X : Type u_1} [TopologicalSpace X] {α : Type u_4} {f : αX} (hf : DenseRange f) :
theorem DenseRange.nonempty {X : Type u_1} [TopologicalSpace X] {α : Type u_4} {f : αX} [h : Nonempty X] (hf : DenseRange f) :
noncomputable def DenseRange.some {X : Type u_1} [TopologicalSpace X] {α : Type u_4} {f : αX} (hf : DenseRange f) (x : X) :
α

Given a function f : X → Y with dense range and y : Y, returns some x : X.

Equations
Instances For
    theorem DenseRange.exists_mem_open {X : Type u_1} [TopologicalSpace X] {α : Type u_4} {f : αX} {s : Set X} (hf : DenseRange f) (ho : IsOpen s) (hs : s.Nonempty) :
    ∃ (a : α), f a s
    theorem DenseRange.mem_nhds {X : Type u_1} [TopologicalSpace X] {x : X} {α : Type u_4} {f : αX} {s : Set X} (h : DenseRange f) (hs : s nhds x) :
    ∃ (a : α), f a s

    The library contains many lemmas stating that functions/operations are continuous. There are many ways to formulate the continuity of operations. Some are more convenient than others. Note: for the most part this note also applies to other properties (Measurable, Differentiable, ContinuousOn, ...).

    The traditional way #

    As an example, let's look at addition (+) : M → M → M. We can state that this is continuous in different definitionally equal ways (omitting some typing information)

    However, lemmas with this conclusion are not nice to use in practice because

    1. They confuse the elaborator. The following two examples fail, because of limitations in the elaboration process.
    variable {M : Type*} [Add M] [TopologicalSpace M] [ContinuousAdd M]
    example : Continuous (fun x : M ↦ x + x) :=
      continuous_add.comp _
    
    example : Continuous (fun x : M ↦ x + x) :=
      continuous_add.comp (continuous_id.prodMk continuous_id)
    

    The second is a valid proof, which is accepted if you write it as continuous_add.comp (continuous_id.prodMk continuous_id :)

    1. If the operation has more than 2 arguments, they are impractical to use, because in your application the arguments in the domain might be in a different order or associated differently.

    The convenient way #

    A much more convenient way to write continuity lemmas is like Continuous.add:

    Continuous.add {f g : X → M} (hf : Continuous f) (hg : Continuous g) :
      Continuous (fun x ↦ f x + g x)
    

    The conclusion can be Continuous (f + g), which is definitionally equal. This has the following advantages

    • It supports projection notation, so is shorter to write.
    • Continuous.add _ _ is recognized correctly by the elaborator and gives useful new goals.
    • It works generally, since the domain is a variable.

    As an example for a unary operation, we have Continuous.neg.

    Continuous.neg {f : X → G} (hf : Continuous f) : Continuous (fun x ↦ -f x)
    

    For unary functions, the elaborator is not confused when applying the traditional lemma (like continuous_neg), but it's still convenient to have the short version available (compare hf.neg.neg.neg with continuous_neg.comp <| continuous_neg.comp <| continuous_neg.comp hf).

    As a harder example, consider an operation of the following type:

    def strans {x : F} (γ γ' : Path x x) (t₀ : I) : Path x x
    

    The precise definition is not important, only its type. The correct continuity principle for this operation is something like this:

    {f : X → F} {γ γ' : ∀ x, Path (f x) (f x)} {t₀ s : X → I}
      (hγ : Continuous ↿γ) (hγ' : Continuous ↿γ')
      (ht : Continuous t₀) (hs : Continuous s) :
      Continuous (fun x ↦ strans (γ x) (γ' x) (t x) (s x))
    

    Note that all arguments of strans are indexed over X, even the basepoint x, and the last argument s that arises since Path x x has a coercion to I → F. The paths γ and γ' (which are unary functions from I) become binary functions in the continuity lemma.

    Summary #

    • Make sure that your continuity lemmas are stated in the most general way, and in a convenient form. That means that:
      • The conclusion has a variable X as domain (not something like Y × Z);
      • Wherever possible, all point arguments c : Y are replaced by functions c : X → Y;
      • All n-ary function arguments are replaced by n+1-ary functions (f : Y → Z becomes f : X → Y → Z);
      • All (relevant) arguments have continuity assumptions, and perhaps there are additional assumptions needed to make the operation continuous;
      • The function in the conclusion is fully applied.
    • These remarks are mostly about the format of the conclusion of a continuity lemma. In assumptions it's fine to state that a function with more than 1 argument is continuous using or Function.uncurry.

    Functions with discontinuities #

    In some cases, you want to work with discontinuous functions, and in certain expressions they are still continuous. For example, consider the fractional part of a number, Int.fract : ℝ → ℝ. In this case, you want to add conditions to when a function involving fract is continuous, so you get something like this: (assumption hf could be weakened, but the important thing is the shape of the conclusion)

    lemma ContinuousOn.comp_fract {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y]
        {f : X → ℝ → Y} {g : X → ℝ} (hf : Continuous ↿f) (hg : Continuous g) (h : ∀ s, f s 0 = f s 1) :
        Continuous (fun x ↦ f x (fract (g x)))
    

    With ContinuousAt you can be even more precise about what to prove in case of discontinuities, see e.g. ContinuousAt.comp_div_cases.

    Equations
    Instances For

      Lean's elaborator has trouble elaborating applications of lemmas that state that the composition of two functions satisfy some property at a point, like ContinuousAt.comp / ContDiffAt.comp and ContMDiffWithinAt.comp. The reason is that a lemma like this looks like ContinuousAt g (f x) → ContinuousAt f x → ContinuousAt (g ∘ f) x. Since Lean's elaborator elaborates the arguments from left-to-right, when you write hg.comp hf, the elaborator will try to figure out both f and g from the type of hg. It tries to figure out f just from the point where g is continuous. For example, if hg : ContinuousAt g (a, x) then the elaborator will assign f to the function Prod.mk a, since in that case f x = (a, x). This is undesirable in most cases where f is not a variable. There are some ways to work around this, for example by giving f explicitly, or to force Lean to elaborate hf before elaborating hg, but this is annoying. Another better solution is to reformulate composition lemmas to have the following shape ContinuousAt g y → ContinuousAt f x → f x = y → ContinuousAt (g ∘ f) x. This is even useful if the proof of f x = y is rfl. The reason that this works better is because the type of hg doesn't mention f. Only after elaborating the two ContinuousAt arguments, Lean will try to unify f x with y, which is often easy after having chosen the correct functions for f and g. Here is an example that shows the difference:

      example [TopologicalSpace X] [TopologicalSpace Y] {x₀ : X} (f : X → X → Y)
          (hf : ContinuousAt (Function.uncurry f) (x₀, x₀)) :
          ContinuousAt (fun x ↦ f x x) x₀ :=
        -- hf.comp (continuousAt_id.prod continuousAt_id) -- type mismatch
        -- hf.comp_of_eq (continuousAt_id.prod continuousAt_id) rfl -- works
      
      Equations
      Instances For