Documentation

Mathlib.Topology.NhdsWithin

Neighborhoods relative to a subset #

This file develops API on the relative version nhdsWithin of nhds, which is defined in previous definition files.

Their basic properties studied in this file include the relationship between neighborhood filters relative to a set and neighborhood filters in the corresponding subtype, and are in later files used to develop relativ versions ContinuousOn and ContinuousWithinAt of Continuous and ContinuousAt.

Notation #

Properties of the neighborhood-within filter #

@[simp]
theorem nhds_bind_nhdsWithin {α : Type u_1} [TopologicalSpace α] {a : α} {s : Set α} :
((nhds a).bind fun (x : α) => nhdsWithin x s) = nhdsWithin a s
@[simp]
theorem eventually_nhds_nhdsWithin {α : Type u_1} [TopologicalSpace α] {a : α} {s : Set α} {p : αProp} :
(∀ᶠ (y : α) in nhds a, ∀ᶠ (x : α) in nhdsWithin y s, p x) ∀ᶠ (x : α) in nhdsWithin a s, p x
theorem eventually_nhdsWithin_iff {α : Type u_1} [TopologicalSpace α] {a : α} {s : Set α} {p : αProp} :
(∀ᶠ (x : α) in nhdsWithin a s, p x) ∀ᶠ (x : α) in nhds a, x sp x
theorem frequently_nhdsWithin_iff {α : Type u_1} [TopologicalSpace α] {z : α} {s : Set α} {p : αProp} :
(∃ᶠ (x : α) in nhdsWithin z s, p x) ∃ᶠ (x : α) in nhds z, p x x s
theorem mem_closure_ne_iff_frequently_within {α : Type u_1} [TopologicalSpace α] {z : α} {s : Set α} :
z closure (s \ {z}) ∃ᶠ (x : α) in nhdsWithin z {z}, x s
@[simp]
theorem eventually_eventually_nhdsWithin {α : Type u_1} [TopologicalSpace α] {a : α} {s : Set α} {p : αProp} :
(∀ᶠ (y : α) in nhdsWithin a s, ∀ᶠ (x : α) in nhdsWithin y s, p x) ∀ᶠ (x : α) in nhdsWithin a s, p x
@[simp]
theorem eventually_mem_nhdsWithin_iff {α : Type u_1} [TopologicalSpace α] {x : α} {s t : Set α} :
(∀ᶠ (x' : α) in nhdsWithin x s, t nhdsWithin x' s) t nhdsWithin x s
theorem nhdsWithin_eq {α : Type u_1} [TopologicalSpace α] (a : α) (s : Set α) :
nhdsWithin a s = t{t : Set α | a t IsOpen t}, Filter.principal (t s)
@[simp]
theorem nhdsWithin_univ {α : Type u_1} [TopologicalSpace α] (a : α) :
theorem nhdsWithin_hasBasis {α : Type u_1} [TopologicalSpace α] {ι : Sort u_5} {p : ιProp} {s : ιSet α} {a : α} (h : (nhds a).HasBasis p s) (t : Set α) :
(nhdsWithin a t).HasBasis p fun (i : ι) => s i t
theorem nhdsWithin_basis_open {α : Type u_1} [TopologicalSpace α] (a : α) (t : Set α) :
(nhdsWithin a t).HasBasis (fun (u : Set α) => a u IsOpen u) fun (u : Set α) => u t
theorem mem_nhdsWithin {α : Type u_1} [TopologicalSpace α] {t : Set α} {a : α} {s : Set α} :
t nhdsWithin a s ∃ (u : Set α), IsOpen u a u u s t
theorem mem_nhdsWithin_iff_exists_mem_nhds_inter {α : Type u_1} [TopologicalSpace α] {t : Set α} {a : α} {s : Set α} :
t nhdsWithin a s unhds a, u s t
theorem diff_mem_nhdsWithin_compl {α : Type u_1} [TopologicalSpace α] {x : α} {s : Set α} (hs : s nhds x) (t : Set α) :
theorem diff_mem_nhdsWithin_diff {α : Type u_1} [TopologicalSpace α] {x : α} {s t : Set α} (hs : s nhdsWithin x t) (t' : Set α) :
s \ t' nhdsWithin x (t \ t')
theorem nhds_of_nhdsWithin_of_nhds {α : Type u_1} [TopologicalSpace α] {s t : Set α} {a : α} (h1 : s nhds a) (h2 : t nhdsWithin a s) :
t nhds a
theorem mem_nhdsWithin_iff_eventually {α : Type u_1} [TopologicalSpace α] {s t : Set α} {x : α} :
t nhdsWithin x s ∀ᶠ (y : α) in nhds x, y sy t
theorem mem_nhdsWithin_iff_eventuallyEq {α : Type u_1} [TopologicalSpace α] {s t : Set α} {x : α} :
theorem nhdsWithin_eq_iff_eventuallyEq {α : Type u_1} [TopologicalSpace α] {s t : Set α} {x : α} :
theorem nhdsWithin_le_iff {α : Type u_1} [TopologicalSpace α] {s t : Set α} {x : α} :
theorem preimage_nhdsWithin_coinduced' {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {X : αβ} {s : Set β} {t : Set α} {a : α} (h : a t) (hs : s nhds (X a)) :
theorem mem_nhdsWithin_of_mem_nhds {α : Type u_1} [TopologicalSpace α] {s t : Set α} {a : α} (h : s nhds a) :
theorem self_mem_nhdsWithin {α : Type u_1} [TopologicalSpace α] {a : α} {s : Set α} :
theorem eventually_mem_nhdsWithin {α : Type u_1} [TopologicalSpace α] {a : α} {s : Set α} :
∀ᶠ (x : α) in nhdsWithin a s, x s
theorem inter_mem_nhdsWithin {α : Type u_1} [TopologicalSpace α] (s : Set α) {t : Set α} {a : α} (h : t nhds a) :
theorem pure_le_nhdsWithin {α : Type u_1} [TopologicalSpace α] {a : α} {s : Set α} (ha : a s) :
theorem mem_of_mem_nhdsWithin {α : Type u_1} [TopologicalSpace α] {a : α} {s t : Set α} (ha : a s) (ht : t nhdsWithin a s) :
a t
theorem Filter.Eventually.self_of_nhdsWithin {α : Type u_1} [TopologicalSpace α] {p : αProp} {s : Set α} {x : α} (h : ∀ᶠ (y : α) in nhdsWithin x s, p y) (hx : x s) :
p x
theorem tendsto_const_nhdsWithin {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {l : Filter β} {s : Set α} {a : α} (ha : a s) :
Filter.Tendsto (fun (x : β) => a) l (nhdsWithin a s)
theorem nhdsWithin_restrict'' {α : Type u_1} [TopologicalSpace α] {a : α} (s : Set α) {t : Set α} (h : t nhdsWithin a s) :
theorem nhdsWithin_restrict' {α : Type u_1} [TopologicalSpace α] {a : α} (s : Set α) {t : Set α} (h : t nhds a) :
theorem nhdsWithin_restrict {α : Type u_1} [TopologicalSpace α] {a : α} (s : Set α) {t : Set α} (h₀ : a t) (h₁ : IsOpen t) :
theorem nhdsWithin_le_of_mem {α : Type u_1} [TopologicalSpace α] {a : α} {s t : Set α} (h : s nhdsWithin a t) :
theorem nhdsWithin_le_nhds {α : Type u_1} [TopologicalSpace α] {a : α} {s : Set α} :
theorem nhdsWithin_eq_nhdsWithin' {α : Type u_1} [TopologicalSpace α] {a : α} {s t u : Set α} (hs : s nhds a) (h₂ : t s = u s) :
theorem nhdsWithin_eq_nhdsWithin {α : Type u_1} [TopologicalSpace α] {a : α} {s t u : Set α} (h₀ : a s) (h₁ : IsOpen s) (h₂ : t s = u s) :
@[simp]
theorem nhdsWithin_eq_nhds {α : Type u_1} [TopologicalSpace α] {a : α} {s : Set α} :
theorem IsOpen.nhdsWithin_eq {α : Type u_1} [TopologicalSpace α] {a : α} {s : Set α} (h : IsOpen s) (ha : a s) :
theorem preimage_nhds_within_coinduced {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {X : αβ} {s : Set β} {t : Set α} {a : α} (h : a t) (ht : IsOpen t) (hs : s nhds (X a)) :
@[simp]
theorem nhdsWithin_empty {α : Type u_1} [TopologicalSpace α] (a : α) :
theorem nhdsWithin_union {α : Type u_1} [TopologicalSpace α] (a : α) (s t : Set α) :
nhdsWithin a (s t) = nhdsWithin a snhdsWithin a t
theorem nhds_eq_nhdsWithin_sup_nhdsWithin {α : Type u_1} [TopologicalSpace α] (b : α) {I₁ I₂ : Set α} (hI : Set.univ = I₁ I₂) :
nhds b = nhdsWithin b I₁nhdsWithin b I₂
theorem union_mem_nhds_of_mem_nhdsWithin {α : Type u_1} [TopologicalSpace α] {b : α} {I₁ I₂ : Set α} (h : Set.univ = I₁ I₂) {L : Set α} (hL : L nhdsWithin b I₁) {R : Set α} (hR : R nhdsWithin b I₂) :
L R nhds b

If L and R are neighborhoods of b within sets whose union is Set.univ, then L ∪ R is a neighborhood of b.

Writing a punctured neighborhood filter as a sup of left and right filters.

theorem nhds_of_Ici_Iic {α : Type u_1} [TopologicalSpace α] [LinearOrder α] {b : α} {L : Set α} (hL : L nhdsWithin b (Set.Iic b)) {R : Set α} (hR : R nhdsWithin b (Set.Ici b)) :

Obtain a "predictably-sided" neighborhood of b from two one-sided neighborhoods.

theorem nhdsWithin_biUnion {α : Type u_1} [TopologicalSpace α] {ι : Type u_5} {I : Set ι} (hI : I.Finite) (s : ιSet α) (a : α) :
nhdsWithin a (⋃ iI, s i) = iI, nhdsWithin a (s i)
theorem nhdsWithin_sUnion {α : Type u_1} [TopologicalSpace α] {S : Set (Set α)} (hS : S.Finite) (a : α) :
nhdsWithin a (⋃₀ S) = sS, nhdsWithin a s
theorem nhdsWithin_iUnion {α : Type u_1} [TopologicalSpace α] {ι : Sort u_5} [Finite ι] (s : ιSet α) (a : α) :
nhdsWithin a (⋃ (i : ι), s i) = ⨆ (i : ι), nhdsWithin a (s i)
theorem nhdsWithin_inter {α : Type u_1} [TopologicalSpace α] (a : α) (s t : Set α) :
nhdsWithin a (s t) = nhdsWithin a snhdsWithin a t
theorem nhdsWithin_inter' {α : Type u_1} [TopologicalSpace α] (a : α) (s t : Set α) :
theorem nhdsWithin_inter_of_mem {α : Type u_1} [TopologicalSpace α] {a : α} {s t : Set α} (h : s nhdsWithin a t) :
theorem nhdsWithin_inter_of_mem' {α : Type u_1} [TopologicalSpace α] {a : α} {s t : Set α} (h : t nhdsWithin a s) :
@[simp]
theorem nhdsWithin_singleton {α : Type u_1} [TopologicalSpace α] (a : α) :
@[simp]
theorem nhdsWithin_insert {α : Type u_1} [TopologicalSpace α] (a : α) (s : Set α) :
nhdsWithin a (insert a s) = pure anhdsWithin a s
theorem mem_nhdsWithin_insert {α : Type u_1} [TopologicalSpace α] {a : α} {s t : Set α} :
theorem insert_mem_nhdsWithin_insert {α : Type u_1} [TopologicalSpace α] {a : α} {s t : Set α} (h : t nhdsWithin a s) :
theorem insert_mem_nhds_iff {α : Type u_1} [TopologicalSpace α] {a : α} {s : Set α} :
@[simp]
theorem nhdsNE_sup_pure {α : Type u_1} [TopologicalSpace α] (a : α) :
@[simp]
theorem pure_sup_nhdsNE {α : Type u_1} [TopologicalSpace α] (a : α) :
theorem continuousAt_iff_punctured_nhds {α : Type u_1} {β : Type u_2} [TopologicalSpace α] [TopologicalSpace β] {f : αβ} {a : α} :
theorem nhdsWithin_prod {α : Type u_1} {β : Type u_2} [TopologicalSpace α] [TopologicalSpace β] {s u : Set α} {t v : Set β} {a : α} {b : β} (hu : u nhdsWithin a s) (hv : v nhdsWithin b t) :
u ×ˢ v nhdsWithin (a, b) (s ×ˢ t)
theorem Filter.EventuallyEq.mem_interior {α : Type u_1} [TopologicalSpace α] {x : α} {s t : Set α} (hst : s =ᶠ[nhds x] t) (h : x interior s) :
theorem Filter.EventuallyEq.mem_interior_iff {α : Type u_1} [TopologicalSpace α] {x : α} {s t : Set α} (hst : s =ᶠ[nhds x] t) :
theorem nhdsWithin_pi_eq' {ι : Type u_5} {X : ιType u_6} [(i : ι) → TopologicalSpace (X i)] {I : Set ι} (hI : I.Finite) (s : (i : ι) → Set (X i)) (x : (i : ι) → X i) :
nhdsWithin x (I.pi s) = ⨅ (i : ι), Filter.comap (fun (x : (i : ι) → X i) => x i) (nhds (x i)⨅ (_ : i I), Filter.principal (s i))
theorem nhdsWithin_pi_eq {ι : Type u_5} {X : ιType u_6} [(i : ι) → TopologicalSpace (X i)] {I : Set ι} (hI : I.Finite) (s : (i : ι) → Set (X i)) (x : (i : ι) → X i) :
nhdsWithin x (I.pi s) = (⨅ iI, Filter.comap (fun (x : (i : ι) → X i) => x i) (nhdsWithin (x i) (s i)))⨅ (i : ι), ⨅ (_ : iI), Filter.comap (fun (x : (i : ι) → X i) => x i) (nhds (x i))
theorem nhdsWithin_pi_univ_eq {ι : Type u_5} {X : ιType u_6} [(i : ι) → TopologicalSpace (X i)] [Finite ι] (s : (i : ι) → Set (X i)) (x : (i : ι) → X i) :
nhdsWithin x (Set.univ.pi s) = ⨅ (i : ι), Filter.comap (fun (x : (i : ι) → X i) => x i) (nhdsWithin (x i) (s i))
theorem nhdsWithin_pi_eq_bot {ι : Type u_5} {X : ιType u_6} [(i : ι) → TopologicalSpace (X i)] {I : Set ι} {s : (i : ι) → Set (X i)} {x : (i : ι) → X i} :
nhdsWithin x (I.pi s) = iI, nhdsWithin (x i) (s i) =
theorem nhdsWithin_pi_neBot {ι : Type u_5} {X : ιType u_6} [(i : ι) → TopologicalSpace (X i)] {I : Set ι} {s : (i : ι) → Set (X i)} {x : (i : ι) → X i} :
(nhdsWithin x (I.pi s)).NeBot iI, (nhdsWithin (x i) (s i)).NeBot
instance instNeBotNhdsWithinUnivPi {ι : Type u_5} {X : ιType u_6} [(i : ι) → TopologicalSpace (X i)] {s : (i : ι) → Set (X i)} {x : (i : ι) → X i} [∀ (i : ι), (nhdsWithin (x i) (s i)).NeBot] :
instance Pi.instNeBotNhdsWithinIio {ι : Type u_5} {X : ιType u_6} [(i : ι) → TopologicalSpace (X i)] [Nonempty ι] [(i : ι) → Preorder (X i)] {x : (i : ι) → X i} [∀ (i : ι), (nhdsWithin (x i) (Set.Iio (x i))).NeBot] :
instance Pi.instNeBotNhdsWithinIoi {ι : Type u_5} {X : ιType u_6} [(i : ι) → TopologicalSpace (X i)] [Nonempty ι] [(i : ι) → Preorder (X i)] {x : (i : ι) → X i} [∀ (i : ι), (nhdsWithin (x i) (Set.Ioi (x i))).NeBot] :
theorem Filter.Tendsto.piecewise_nhdsWithin {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {f g : αβ} {t : Set α} [(x : α) → Decidable (x t)] {a : α} {s : Set α} {l : Filter β} (h₀ : Tendsto f (nhdsWithin a (s t)) l) (h₁ : Tendsto g (nhdsWithin a (s t)) l) :
Tendsto (t.piecewise f g) (nhdsWithin a s) l
theorem Filter.Tendsto.if_nhdsWithin {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {f g : αβ} {p : αProp} [DecidablePred p] {a : α} {s : Set α} {l : Filter β} (h₀ : Tendsto f (nhdsWithin a (s {x : α | p x})) l) (h₁ : Tendsto g (nhdsWithin a (s {x : α | ¬p x})) l) :
Tendsto (fun (x : α) => if p x then f x else g x) (nhdsWithin a s) l
theorem map_nhdsWithin {α : Type u_1} {β : Type u_2} [TopologicalSpace α] (f : αβ) (a : α) (s : Set α) :
Filter.map f (nhdsWithin a s) = t{t : Set α | a t IsOpen t}, Filter.principal (f '' (t s))
theorem tendsto_nhdsWithin_mono_left {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {f : αβ} {a : α} {s t : Set α} {l : Filter β} (hst : s t) (h : Filter.Tendsto f (nhdsWithin a t) l) :
theorem tendsto_nhdsWithin_mono_right {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {f : βα} {l : Filter β} {a : α} {s t : Set α} (hst : s t) (h : Filter.Tendsto f l (nhdsWithin a s)) :
theorem tendsto_nhdsWithin_of_tendsto_nhds {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {f : αβ} {a : α} {s : Set α} {l : Filter β} (h : Filter.Tendsto f (nhds a) l) :
theorem eventually_mem_of_tendsto_nhdsWithin {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {f : βα} {a : α} {s : Set α} {l : Filter β} (h : Filter.Tendsto f l (nhdsWithin a s)) :
∀ᶠ (i : β) in l, f i s
theorem tendsto_nhds_of_tendsto_nhdsWithin {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {f : βα} {a : α} {s : Set α} {l : Filter β} (h : Filter.Tendsto f l (nhdsWithin a s)) :
theorem nhdsWithin_neBot_of_mem {α : Type u_1} [TopologicalSpace α] {s : Set α} {x : α} (hx : x s) :
theorem IsClosed.mem_of_nhdsWithin_neBot {α : Type u_1} [TopologicalSpace α] {s : Set α} (hs : IsClosed s) {x : α} (hx : (nhdsWithin x s).NeBot) :
x s
theorem DenseRange.nhdsWithin_neBot {α : Type u_1} [TopologicalSpace α] {ι : Type u_5} {f : ια} (h : DenseRange f) (x : α) :
theorem mem_closure_pi {ι : Type u_5} {α : ιType u_6} [(i : ι) → TopologicalSpace (α i)] {I : Set ι} {s : (i : ι) → Set (α i)} {x : (i : ι) → α i} :
x closure (I.pi s) iI, x i closure (s i)
theorem closure_pi_set {ι : Type u_5} {α : ιType u_6} [(i : ι) → TopologicalSpace (α i)] (I : Set ι) (s : (i : ι) → Set (α i)) :
closure (I.pi s) = I.pi fun (i : ι) => closure (s i)
theorem dense_pi {ι : Type u_5} {α : ιType u_6} [(i : ι) → TopologicalSpace (α i)] {s : (i : ι) → Set (α i)} (I : Set ι) (hs : iI, Dense (s i)) :
Dense (I.pi s)
theorem DenseRange.piMap {ι : Type u_5} {X : ιType u_6} {Y : ιType u_7} [(i : ι) → TopologicalSpace (Y i)] {f : (i : ι) → X iY i} (hf : ∀ (i : ι), DenseRange (f i)) :
theorem eventuallyEq_nhdsWithin_iff {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {f g : αβ} {s : Set α} {a : α} :
f =ᶠ[nhdsWithin a s] g ∀ᶠ (x : α) in nhds a, x sf x = g x
theorem eventuallyEq_nhds_of_eventuallyEq_nhdsNE {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {f g : αβ} {a : α} (h₁ : f =ᶠ[nhdsWithin a {a}] g) (h₂ : f a = g a) :

Two functions agree on a neighborhood of x if they agree at x and in a punctured neighborhood.

theorem eventuallyEq_nhdsWithin_of_eqOn {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {f g : αβ} {s : Set α} {a : α} (h : Set.EqOn f g s) :
theorem Set.EqOn.eventuallyEq_nhdsWithin {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {f g : αβ} {s : Set α} {a : α} (h : EqOn f g s) :
theorem tendsto_nhdsWithin_congr {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {f g : αβ} {s : Set α} {a : α} {l : Filter β} (hfg : xs, f x = g x) (hf : Filter.Tendsto f (nhdsWithin a s) l) :
theorem eventually_nhdsWithin_of_forall {α : Type u_1} [TopologicalSpace α] {s : Set α} {a : α} {p : αProp} (h : xs, p x) :
∀ᶠ (x : α) in nhdsWithin a s, p x
theorem tendsto_nhdsWithin_of_tendsto_nhds_of_eventually_within {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {a : α} {l : Filter β} {s : Set α} (f : βα) (h1 : Filter.Tendsto f l (nhds a)) (h2 : ∀ᶠ (x : β) in l, f x s) :
theorem tendsto_nhdsWithin_iff {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {a : α} {l : Filter β} {s : Set α} {f : βα} :
Filter.Tendsto f l (nhdsWithin a s) Filter.Tendsto f l (nhds a) ∀ᶠ (n : β) in l, f n s
@[simp]
theorem tendsto_nhdsWithin_range {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {a : α} {l : Filter β} {f : βα} :
theorem Filter.EventuallyEq.eq_of_nhdsWithin {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {s : Set α} {f g : αβ} {a : α} (h : f =ᶠ[nhdsWithin a s] g) (hmem : a s) :
f a = g a
theorem eventually_nhdsWithin_of_eventually_nhds {α : Type u_1} [TopologicalSpace α] {s : Set α} {a : α} {p : αProp} (h : ∀ᶠ (x : α) in nhds a, p x) :
∀ᶠ (x : α) in nhdsWithin a s, p x
theorem Set.MapsTo.preimage_mem_nhdsWithin {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {f : αβ} {s : Set α} {t : Set β} {x : α} (hst : MapsTo f s t) :

nhdsWithin and subtypes #

theorem mem_nhdsWithin_subtype {α : Type u_1} [TopologicalSpace α] {s : Set α} {a : { x : α // x s }} {t u : Set { x : α // x s }} :
theorem nhdsWithin_subtype {α : Type u_1} [TopologicalSpace α] (s : Set α) (a : { x : α // x s }) (t : Set { x : α // x s }) :
theorem nhdsWithin_eq_map_subtype_coe {α : Type u_1} [TopologicalSpace α] {s : Set α} {a : α} (h : a s) :
theorem mem_nhds_subtype_iff_nhdsWithin {α : Type u_1} [TopologicalSpace α] {s : Set α} {a : s} {t : Set s} :
theorem preimage_coe_mem_nhds_subtype {α : Type u_1} [TopologicalSpace α] {s t : Set α} {a : s} :
theorem eventually_nhds_subtype_iff {α : Type u_1} [TopologicalSpace α] (s : Set α) (a : s) (P : αProp) :
(∀ᶠ (x : s) in nhds a, P x) ∀ᶠ (x : α) in nhdsWithin (↑a) s, P x
theorem frequently_nhds_subtype_iff {α : Type u_1} [TopologicalSpace α] (s : Set α) (a : s) (P : αProp) :
(∃ᶠ (x : s) in nhds a, P x) ∃ᶠ (x : α) in nhdsWithin (↑a) s, P x
theorem tendsto_nhdsWithin_iff_subtype {α : Type u_1} {β : Type u_2} [TopologicalSpace α] {s : Set α} {a : α} (h : a s) (f : αβ) (l : Filter β) :