@@ -14,7 +14,7 @@ noncomputable theory
14
14
The main definition is the type class `topological space α` which endows a type `α` with a topology.
15
15
Then `set α` gets predicates `is_open`, `is_closed` and functions `interior`, `closure` and
16
16
`frontier`. Each point `x` of `α` gets a neighborhood filter `𝓝 x`. A filter `F` on `α` has
17
- `x` as a cluster point if `is_cluster_pt x F : 𝓝 x ⊓ F ≠ ⊥`. A map `f : ι → α` clusters at `x`
17
+ `x` as a cluster point if `cluster_pt x F : 𝓝 x ⊓ F ≠ ⊥`. A map `f : ι → α` clusters at `x`
18
18
along `F : filter ι` if `map_cluster_pt x F f : cluster_pt x (map f F)`. In particular
19
19
the notion of cluster point of a sequence `u` is `map_cluster_pt x at_top u`.
20
20
@@ -969,6 +969,11 @@ lemma continuous_at.preimage_mem_nhds {f : α → β} {x : α} {t : set β} (h :
969
969
(ht : t ∈ 𝓝 (f x)) : f ⁻¹' t ∈ 𝓝 x :=
970
970
h ht
971
971
972
+ lemma cluster_pt.map {x : α} {la : filter α} {lb : filter β} (H : cluster_pt x la)
973
+ {f : α → β} (hfc : continuous_at f x) (hf : tendsto f la lb) :
974
+ cluster_pt (f x) lb :=
975
+ ne_bot_of_le_ne_bot ((map_ne_bot_iff f).2 H) $ hfc.tendsto.inf hf
976
+
972
977
lemma preimage_interior_subset_interior_preimage {f : α → β} {s : set β}
973
978
(hf : continuous f) : f⁻¹' (interior s) ⊆ interior (f⁻¹' s) :=
974
979
interior_maximal (preimage_mono interior_subset) (is_open_interior.preimage hf)
@@ -1110,23 +1115,21 @@ begin
1110
1115
apply h', rw mem_nhds_sets_iff, exact ⟨s, set.subset.refl _, os, ys⟩
1111
1116
end
1112
1117
1118
+ /-- If a continuous map `f` maps `s` to `t`, then it maps `closure s` to `closure t`. -/
1119
+ lemma set.maps_to.closure {s : set α} {t : set β} {f : α → β} (h : maps_to f s t)
1120
+ (hc : continuous f) : maps_to f (closure s) (closure t) :=
1121
+ begin
1122
+ simp only [maps_to, mem_closure_iff_cluster_pt],
1123
+ exact λ x hx, hx.map hc.continuous_at (tendsto_principal_principal.2 h)
1124
+ end
1125
+
1113
1126
lemma image_closure_subset_closure_image {f : α → β} {s : set α} (h : continuous f) :
1114
1127
f '' closure s ⊆ closure (f '' s) :=
1115
- have ∀ (a : α), cluster_pt a (𝓟 s) → cluster_pt (f a) (𝓟 (f '' s)),
1116
- from assume a ha,
1117
- have h₁ : ¬ map f (𝓝 a ⊓ 𝓟 s) = ⊥,
1118
- by rwa[map_eq_bot_iff],
1119
- have h₂ : map f (𝓝 a ⊓ 𝓟 s) ≤ 𝓝 (f a) ⊓ 𝓟 (f '' s),
1120
- from le_inf
1121
- (le_trans (map_mono inf_le_left) $ by rw [continuous_iff_continuous_at] at h; exact h a)
1122
- (le_trans (map_mono inf_le_right) $ by simp [subset_preimage_image] ),
1123
- ne_bot_of_le_ne_bot h₁ h₂,
1124
- by simp [image_subset_iff, closure_eq_cluster_pts]; assumption
1125
-
1126
- lemma mem_closure {s : set α} {t : set β} {f : α → β} {a : α}
1128
+ ((maps_to_image f s).closure h).image_subset
1129
+
1130
+ lemma map_mem_closure {s : set α} {t : set β} {f : α → β} {a : α}
1127
1131
(hf : continuous f) (ha : a ∈ closure s) (ht : ∀a∈s, f a ∈ t) : f a ∈ closure t :=
1128
- subset.trans (image_closure_subset_closure_image hf) (closure_mono $ image_subset_iff.2 ht) $
1129
- (mem_image_of_mem f ha)
1132
+ set.maps_to.closure ht hf ha
1130
1133
1131
1134
/-!
1132
1135
### Function with dense range
0 commit comments