@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Patrick Massot, Johannes Hölzl
5
5
6
6
! This file was ported from Lean 3 source module analysis.normed_space.basic
7
- ! leanprover-community/mathlib commit d3af0609f6db8691dffdc3e1fb7feb7da72698f2
7
+ ! leanprover-community/mathlib commit 8000bbbe2e9d39b84edb993d88781f536a8a3fa8
8
8
! Please do not edit these lines, except to modify the commit id
9
9
! if you have ported upstream changes.
10
10
-/
@@ -203,6 +203,16 @@ theorem frontier_closedBall [NormedSpace ℝ E] (x : E) {r : ℝ} (hr : r ≠ 0)
203
203
rw [frontier, closure_closedBall, interior_closedBall x hr, closedBall_diff_ball]
204
204
#align frontier_closed_ball frontier_closedBall
205
205
206
+ theorem interior_sphere [NormedSpace ℝ E] (x : E) {r : ℝ} (hr : r ≠ 0 ) :
207
+ interior (sphere x r) = ∅ := by
208
+ rw [← frontier_closedBall x hr, interior_frontier isClosed_ball]
209
+ #align interior_sphere interior_sphere
210
+
211
+ theorem frontier_sphere [NormedSpace ℝ E] (x : E) {r : ℝ} (hr : r ≠ 0 ) :
212
+ frontier (sphere x r) = sphere x r := by
213
+ rw [isClosed_sphere.frontier_eq, interior_sphere x hr, diff_empty]
214
+ #align frontier_sphere frontier_sphere
215
+
206
216
instance {E : Type _} [NormedAddCommGroup E] [NormedSpace ℚ E] (e : E) :
207
217
DiscreteTopology <| AddSubgroup.zmultiples e := by
208
218
rcases eq_or_ne e 0 with (rfl | he)
@@ -430,6 +440,17 @@ theorem frontier_closedBall' [NormedSpace ℝ E] [Nontrivial E] (x : E) (r : ℝ
430
440
rw [frontier, closure_closedBall, interior_closedBall' x r, closedBall_diff_ball]
431
441
#align frontier_closed_ball' frontier_closedBall'
432
442
443
+ @[simp]
444
+ theorem interior_sphere' [NormedSpace ℝ E] [Nontrivial E] (x : E) (r : ℝ) :
445
+ interior (sphere x r) = ∅ := by rw [← frontier_closedBall' x, interior_frontier isClosed_ball]
446
+ #align interior_sphere' interior_sphere'
447
+
448
+ @[simp]
449
+ theorem frontier_sphere' [NormedSpace ℝ E] [Nontrivial E] (x : E) (r : ℝ) :
450
+ frontier (sphere x r) = sphere x r := by
451
+ rw [isClosed_sphere.frontier_eq, interior_sphere' x, diff_empty]
452
+ #align frontier_sphere' frontier_sphere'
453
+
433
454
theorem rescale_to_shell_zpow {c : α} (hc : 1 < ‖c‖) {ε : ℝ} (εpos : 0 < ε) {x : E} (hx : x ≠ 0 ) :
434
455
∃ n : ℤ, c ^ n ≠ 0 ∧ ‖c ^ n • x‖ < ε ∧ ε / ‖c‖ ≤ ‖c ^ n • x‖ ∧ ‖c ^ n‖⁻¹ ≤ ε⁻¹ * ‖c‖ * ‖x‖ :=
435
456
rescale_to_shell_semi_normed_zpow hc εpos (mt norm_eq_zero.1 hx)
0 commit comments