Skip to content

Commit 1eb993f

Browse files
committed
1 parent 3f22ac3 commit 1eb993f

File tree

2 files changed

+28
-2
lines changed

2 files changed

+28
-2
lines changed

Mathlib/Analysis/NormedSpace/Basic.lean

Lines changed: 22 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Patrick Massot, Johannes Hölzl
55
66
! 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
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -203,6 +203,16 @@ theorem frontier_closedBall [NormedSpace ℝ E] (x : E) {r : ℝ} (hr : r ≠ 0)
203203
rw [frontier, closure_closedBall, interior_closedBall x hr, closedBall_diff_ball]
204204
#align frontier_closed_ball frontier_closedBall
205205

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+
206216
instance {E : Type _} [NormedAddCommGroup E] [NormedSpace ℚ E] (e : E) :
207217
DiscreteTopology <| AddSubgroup.zmultiples e := by
208218
rcases eq_or_ne e 0 with (rfl | he)
@@ -430,6 +440,17 @@ theorem frontier_closedBall' [NormedSpace ℝ E] [Nontrivial E] (x : E) (r : ℝ
430440
rw [frontier, closure_closedBall, interior_closedBall' x r, closedBall_diff_ball]
431441
#align frontier_closed_ball' frontier_closedBall'
432442

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+
433454
theorem rescale_to_shell_zpow {c : α} (hc : 1 < ‖c‖) {ε : ℝ} (εpos : 0 < ε) {x : E} (hx : x ≠ 0) :
434455
∃ n : ℤ, c ^ n ≠ 0 ∧ ‖c ^ n • x‖ < ε ∧ ε / ‖c‖ ≤ ‖c ^ n • x‖ ∧ ‖c ^ n‖⁻¹ ≤ ε⁻¹ * ‖c‖ * ‖x‖ :=
435456
rescale_to_shell_semi_normed_zpow hc εpos (mt norm_eq_zero.1 hx)

Mathlib/Topology/MetricSpace/Basic.lean

Lines changed: 6 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébastien Gouëzel
55
66
! This file was ported from Lean 3 source module topology.metric_space.basic
7-
! leanprover-community/mathlib commit 195fcd60ff2bfe392543bceb0ec2adcdb472db4c
7+
! leanprover-community/mathlib commit 8000bbbe2e9d39b84edb993d88781f536a8a3fa8
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -1872,6 +1872,11 @@ theorem closure_closedBall : closure (closedBall x ε) = closedBall x ε :=
18721872
isClosed_ball.closure_eq
18731873
#align metric.closure_closed_ball Metric.closure_closedBall
18741874

1875+
@[simp]
1876+
theorem closure_sphere : closure (sphere x ε) = sphere x ε :=
1877+
isClosed_sphere.closure_eq
1878+
#align metric.closure_sphere Metric.closure_sphere
1879+
18751880
theorem closure_ball_subset_closedBall : closure (ball x ε) ⊆ closedBall x ε :=
18761881
closure_minimal ball_subset_closedBall isClosed_ball
18771882
#align metric.closure_ball_subset_closed_ball Metric.closure_ball_subset_closedBall

0 commit comments

Comments
 (0)