Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit a32d58b

Browse files
committed
feat(analysis/*): generalize set_smul_mem_nhds_zero to topological vector spaces (#12779)
The lemma holds for arbitrary topological vector spaces and has nothing to do with normed spaces. Co-authored-by: Eric Wieser <wieser.eric@gmail.com>
1 parent adcfc58 commit a32d58b

File tree

2 files changed

+42
-18
lines changed

2 files changed

+42
-18
lines changed

src/analysis/normed_space/pointwise.lean

Lines changed: 0 additions & 18 deletions
Original file line numberDiff line numberDiff line change
@@ -87,24 +87,6 @@ begin
8787
simpa only [this, dist_eq_norm, add_sub_cancel', mem_closed_ball] using I,
8888
end
8989

90-
lemma set_smul_mem_nhds_zero {s : set E} (hs : s ∈ 𝓝 (0 : E)) {c : 𝕜} (hc : c ≠ 0) :
91-
c • s ∈ 𝓝 (0 : E) :=
92-
begin
93-
obtain ⟨ε, εpos, hε⟩ : ∃ (ε : ℝ) (H : 0 < ε), ball 0 ε ⊆ s := metric.mem_nhds_iff.1 hs,
94-
have : c • ball (0 : E) ε ∈ 𝓝 (0 : E),
95-
{ rw [smul_ball hc, smul_zero],
96-
exact ball_mem_nhds _ (mul_pos (by simpa using hc) εpos) },
97-
exact filter.mem_of_superset this ((set_smul_subset_set_smul_iff₀ hc).2 hε)
98-
end
99-
100-
lemma set_smul_mem_nhds_zero_iff (s : set E) {c : 𝕜} (hc : c ≠ 0) :
101-
c • s ∈ 𝓝 (0 : E) ↔ s ∈ 𝓝(0 : E) :=
102-
begin
103-
refine ⟨λ h, _, λ h, set_smul_mem_nhds_zero h hc⟩,
104-
convert set_smul_mem_nhds_zero h (inv_ne_zero hc),
105-
rw [smul_smul, inv_mul_cancel hc, one_smul],
106-
end
107-
10890
/-- Any ball is the image of a ball centered at the origin under a shift. -/
10991
lemma vadd_ball_zero (x : E) (r : ℝ) : x +ᵥ ball 0 r = ball x r :=
11092
by rw [vadd_ball, vadd_eq_add, add_zero]

src/topology/algebra/const_mul_action.lean

Lines changed: 42 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -375,3 +375,45 @@ begin
375375
simp only [image_smul, not_not, mem_set_of_eq, ne.def] at H,
376376
exact eq_empty_iff_forall_not_mem.mp H (γ • x) ⟨mem_image_of_mem _ x_in_K₀, h'⟩ },
377377
end
378+
379+
section nhds
380+
381+
section mul_action
382+
383+
variables {G₀ : Type*} [group_with_zero G₀] [mul_action G₀ α]
384+
[topological_space α] [has_continuous_const_smul G₀ α]
385+
386+
/-- Scalar multiplication preserves neighborhoods. -/
387+
lemma set_smul_mem_nhds_smul {c : G₀} {s : set α} {x : α} (hs : s ∈ 𝓝 x) (hc : c ≠ 0) :
388+
c • s ∈ 𝓝 (c • x : α) :=
389+
begin
390+
rw mem_nhds_iff at hs ⊢,
391+
obtain ⟨U, hs', hU, hU'⟩ := hs,
392+
exact ⟨c • U, set.smul_set_mono hs', hU.smul₀ hc, set.smul_mem_smul_set hU'⟩,
393+
end
394+
395+
lemma set_smul_mem_nhds_smul_iff {c : G₀} {s : set α} {x : α} (hc : c ≠ 0) :
396+
c • s ∈ 𝓝 (c • x : α) ↔ s ∈ 𝓝 x :=
397+
begin
398+
refine ⟨λ h, _, λ h, set_smul_mem_nhds_smul h hc⟩,
399+
rw [←inv_smul_smul₀ hc x, ←inv_smul_smul₀ hc s],
400+
exact set_smul_mem_nhds_smul h (inv_ne_zero hc),
401+
end
402+
403+
end mul_action
404+
405+
section distrib_mul_action
406+
407+
variables {G₀ : Type*} [group_with_zero G₀] [add_monoid α] [distrib_mul_action G₀ α]
408+
[topological_space α] [has_continuous_const_smul G₀ α]
409+
410+
lemma set_smul_mem_nhds_zero_iff {s : set α} {c : G₀} (hc : c ≠ 0) :
411+
c • s ∈ 𝓝 (0 : α) ↔ s ∈ 𝓝 (0 : α) :=
412+
begin
413+
refine iff.trans _ (set_smul_mem_nhds_smul_iff hc),
414+
rw smul_zero,
415+
end
416+
417+
end distrib_mul_action
418+
419+
end nhds

0 commit comments

Comments
 (0)