@@ -18,7 +18,7 @@ about these definitions.
18
18
variables {α : Type *} {β : Type *} {γ : Type *} {ι : Type *}
19
19
20
20
noncomputable theory
21
- open filter metric
21
+ open filter metric function set
22
22
open_locale topological_space big_operators nnreal ennreal uniformity pointwise
23
23
24
24
section semi_normed_group
@@ -109,8 +109,8 @@ hg.op_zero_is_bounded_under_le hf (flip (•)) (λ x y, ((norm_smul y x).trans (
109
109
theorem closure_ball [normed_space ℝ E] (x : E) {r : ℝ} (hr : r ≠ 0 ) :
110
110
closure (ball x r) = closed_ball x r :=
111
111
begin
112
- refine set. subset.antisymm closure_ball_subset_closed_ball (λ y hy, _),
113
- have : continuous_within_at (λ c : ℝ, c • (y - x) + x) (set. Ico 0 1 ) 1 :=
112
+ refine subset.antisymm closure_ball_subset_closed_ball (λ y hy, _),
113
+ have : continuous_within_at (λ c : ℝ, c • (y - x) + x) (Ico 0 1 ) 1 :=
114
114
((continuous_id.smul continuous_const).add continuous_const).continuous_within_at,
115
115
convert this.mem_closure _ _,
116
116
{ rw [one_smul, sub_add_cancel] },
@@ -135,19 +135,19 @@ theorem interior_closed_ball [normed_space ℝ E] (x : E) {r : ℝ} (hr : r ≠
135
135
begin
136
136
cases hr.lt_or_lt with hr hr,
137
137
{ rw [closed_ball_eq_empty.2 hr, ball_eq_empty.2 hr.le, interior_empty] },
138
- refine set. subset.antisymm _ ball_subset_interior_closed_ball,
138
+ refine subset.antisymm _ ball_subset_interior_closed_ball,
139
139
intros y hy,
140
140
rcases (mem_closed_ball.1 $ interior_subset hy).lt_or_eq with hr|rfl, { exact hr },
141
141
set f : ℝ → E := λ c : ℝ, c • (y - x) + x,
142
- suffices : f ⁻¹' closed_ball x (dist y x) ⊆ set. Icc (-1 ) 1 ,
142
+ suffices : f ⁻¹' closed_ball x (dist y x) ⊆ Icc (-1 ) 1 ,
143
143
{ have hfc : continuous f := (continuous_id.smul continuous_const).add continuous_const,
144
144
have hf1 : (1 :ℝ) ∈ f ⁻¹' (interior (closed_ball x $ dist y x)), by simpa [f],
145
- have h1 : (1 :ℝ) ∈ interior (set. Icc (-1 :ℝ) 1 ) :=
145
+ have h1 : (1 :ℝ) ∈ interior (Icc (-1 :ℝ) 1 ) :=
146
146
interior_mono this (preimage_interior_subset_interior_preimage hfc hf1),
147
147
contrapose h1,
148
148
simp },
149
149
intros c hc,
150
- rw [set. mem_Icc, ← abs_le, ← real.norm_eq_abs, ← mul_le_mul_right hr],
150
+ rw [mem_Icc, ← abs_le, ← real.norm_eq_abs, ← mul_le_mul_right hr],
151
151
simpa [f, dist_eq_norm, norm_smul] using hc
152
152
end
153
153
@@ -278,6 +278,29 @@ gives some more context. -/
278
278
@[priority 100 ]
279
279
instance normed_space.to_module' : module α F := normed_space.to_module
280
280
281
+ section surj
282
+
283
+ variables (E) [normed_space ℝ E] [nontrivial E]
284
+
285
+ lemma exists_norm_eq {c : ℝ} (hc : 0 ≤ c) : ∃ x : E, ∥x∥ = c :=
286
+ begin
287
+ rcases exists_ne (0 : E) with ⟨x, hx⟩,
288
+ rw ← norm_ne_zero_iff at hx,
289
+ use c • ∥x∥⁻¹ • x,
290
+ simp [norm_smul, real.norm_of_nonneg hc, hx]
291
+ end
292
+
293
+ @[simp] lemma range_norm : range (norm : E → ℝ) = Ici 0 :=
294
+ subset.antisymm (range_subset_iff.2 norm_nonneg) (λ _, exists_norm_eq E)
295
+
296
+ lemma nnnorm_surjective : surjective (nnnorm : E → ℝ≥0 ) :=
297
+ λ c, (exists_norm_eq E c.coe_nonneg).imp $ λ x h, nnreal.eq h
298
+
299
+ @[simp] lemma range_nnnorm : range (nnnorm : E → ℝ≥0 ) = univ :=
300
+ (nnnorm_surjective E).range_eq
301
+
302
+ end surj
303
+
281
304
theorem interior_closed_ball' [normed_space ℝ E] [nontrivial E] (x : E) (r : ℝ) :
282
305
interior (closed_ball x r) = ball x r :=
283
306
begin
@@ -319,7 +342,7 @@ begin
319
342
rwa norm_pos_iff
320
343
end
321
344
322
- protected lemma normed_space.unbounded_univ : ¬bounded (set. univ : set E) :=
345
+ protected lemma normed_space.unbounded_univ : ¬bounded (univ : set E) :=
323
346
λ h, let ⟨R, hR⟩ := bounded_iff_forall_norm_le.1 h, ⟨x, hx⟩ := normed_space.exists_lt_norm 𝕜 E R
324
347
in hx.not_le (hR x trivial)
325
348
0 commit comments