diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index f08a0fafc0a2c..1d353cdafed19 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -1721,9 +1721,8 @@ theorem closure_image_closure (h : Continuous f) : (closure_mono <| image_subset _ subset_closure) theorem closure_subset_preimage_closure_image (h : Continuous f) : - closure s ⊆ f ⁻¹' closure (f '' s) := by - rw [← Set.image_subset_iff] - exact image_closure_subset_closure_image h + closure s ⊆ f ⁻¹' closure (f '' s) := + (mapsTo_image _ _).closure h #align closure_subset_preimage_closure_image closure_subset_preimage_closure_image theorem map_mem_closure {t : Set Y} (hf : Continuous f) @@ -1737,6 +1736,15 @@ theorem Set.MapsTo.closure_left {t : Set Y} (h : MapsTo f s t) ht.closure_eq ▸ h.closure hc #align set.maps_to.closure_left Set.MapsTo.closure_left +theorem Filter.Tendsto.lift'_closure (hf : Continuous f) {l l'} (h : Tendsto f l l') : + Tendsto f (l.lift' closure) (l'.lift' closure) := + tendsto_lift'.2 fun s hs ↦ by + filter_upwards [mem_lift' (h hs)] using (mapsTo_preimage _ _).closure hf + +theorem tendsto_lift'_closure_nhds (hf : Continuous f) (x : X) : + Tendsto f ((𝓝 x).lift' closure) ((𝓝 (f x)).lift' closure) := + (hf.tendsto x).lift'_closure hf + /-! ### Function with dense range -/ diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index a2ec8140b62dd..f4f7a20f636d9 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -1079,10 +1079,13 @@ theorem IsCompact.finite_compact_cover {s : Set X} (hs : IsCompact s) {ι : Type · simp only [Finset.set_biUnion_insert_update _ hx, hK, h3K] #align is_compact.finite_compact_cover IsCompact.finite_compact_cover -theorem Inducing.r1Space [TopologicalSpace Y] {f : Y → X} (hf : Inducing f) : R1Space Y where - specializes_or_disjoint_nhds _ _ := by - simpa only [← hf.specializes_iff, hf.nhds_eq_comap, or_iff_not_imp_left, - ← disjoint_nhds_nhds_iff_not_specializes] using Filter.disjoint_comap +theorem R1Space.of_continuous_specializes_imp [TopologicalSpace Y] {f : Y → X} (hc : Continuous f) + (hspec : ∀ x y, f x ⤳ f y → x ⤳ y) : R1Space Y where + specializes_or_disjoint_nhds x y := (specializes_or_disjoint_nhds (f x) (f y)).imp (hspec x y) <| + ((hc.tendsto _).disjoint · (hc.tendsto _)) + +theorem Inducing.r1Space [TopologicalSpace Y] {f : Y → X} (hf : Inducing f) : R1Space Y := + .of_continuous_specializes_imp hf.continuous fun _ _ ↦ hf.specializes_iff.1 protected theorem R1Space.induced (f : Y → X) : @R1Space Y (.induced f ‹_›) := @Inducing.r1Space _ _ _ _ (.induced f _) f (inducing_induced f) @@ -1396,6 +1399,14 @@ theorem exists_open_nhds_disjoint_closure [T25Space X] {x y : X} (h : x ≠ y) : (disjoint_lift'_closure_nhds.2 h) #align exists_open_nhds_disjoint_closure exists_open_nhds_disjoint_closure +theorem T25Space.of_injective_continuous [TopologicalSpace Y] [T25Space Y] {f : X → Y} + (hinj : Injective f) (hcont : Continuous f) : T25Space X where + t2_5 x y hne := (tendsto_lift'_closure_nhds hcont x).disjoint (t2_5 <| hinj.ne hne) + (tendsto_lift'_closure_nhds hcont y) + +instance [T25Space X] {p : X → Prop} : T25Space {x // p x} := + .of_injective_continuous Subtype.val_injective continuous_subtype_val + section limUnder variable [T2Space X] {f : Filter X}