Skip to content

Commit

Permalink
chore(analysis/normed_space): remove 2 norm_zero lemmas (#3558)
Browse files Browse the repository at this point in the history
We have a general `norm_zero` lemma and these lemmas are not used
before we introduce `normed_group` instances.
  • Loading branch information
urkud committed Jul 26, 2020
1 parent 11179d5 commit d7fcd8c
Show file tree
Hide file tree
Showing 4 changed files with 5 additions and 11 deletions.
4 changes: 2 additions & 2 deletions src/analysis/calculus/mean_value.lean
Expand Up @@ -367,7 +367,7 @@ begin
simpa using (has_deriv_at_const x C).mul ((has_deriv_at_id x).sub (has_deriv_at_const x a)) },
convert image_norm_le_of_norm_deriv_right_le_deriv_boundary hg hg' _ hB bound,
{ simp only [g, B] },
{ simp only [g, B], rw [sub_self, _root_.norm_zero, sub_self, mul_zero] }
{ simp only [g, B], rw [sub_self, norm_zero, sub_self, mul_zero] }
end

/-- A function on `[a, b]` with the norm of the derivative within `[a, b]`
Expand Down Expand Up @@ -509,7 +509,7 @@ theorem convex.is_const_of_fderiv_within_eq_zero {s : set E} (hs : convex s)
{x y : E} (hx : x ∈ s) (hy : y ∈ s) :
f x = f y :=
have bound : ∀ x ∈ s, ∥fderiv_within ℝ f s x∥ ≤ 0,
from λ x hx, by simp only [hf' x hx, _root_.norm_zero],
from λ x hx, by simp only [hf' x hx, norm_zero],
by simpa only [(dist_eq_norm _ _).symm, zero_mul, dist_le_zero, eq_comm]
using hs.norm_image_sub_le_of_norm_fderiv_within_le hf bound hx hy

Expand Down
5 changes: 1 addition & 4 deletions src/analysis/normed_space/multilinear.lean
Expand Up @@ -111,7 +111,7 @@ begin
by a power of a scalar `c` with norm `∥c∥ > 1`.-/
by_cases h : ∃i, m i = 0,
{ rcases h with ⟨i, hi⟩,
rw [f.map_coord_zero i hi, _root_.norm_zero],
rw [f.map_coord_zero i hi, norm_zero],
exact mul_nonneg (le_of_lt C_pos) (prod_nonneg (λi hi, norm_nonneg _)) },
{ push_neg at h,
have : ∀i, ∃d:𝕜, d ≠ 0 ∧ ∥d • m i∥ ≤ δ ∧ (δ/∥c∥ ≤ ∥d • m i∥) ∧ (∥d∥⁻¹ ≤ δ⁻¹ * ∥c∥ * ∥m i∥) :=
Expand Down Expand Up @@ -374,9 +374,6 @@ begin
simp }
end

@[simp] lemma norm_zero : ∥(0 : continuous_multilinear_map 𝕜 E₁ E₂)∥ = 0 :=
by rw op_norm_zero_iff

lemma op_norm_smul_le : ∥c • f∥ ≤ ∥c∥ * ∥f∥ :=
(Inf_le _ bounds_bdd_below
⟨mul_nonneg (norm_nonneg _) (op_norm_nonneg _), λ _,
Expand Down
3 changes: 0 additions & 3 deletions src/analysis/normed_space/operator_norm.lean
Expand Up @@ -296,9 +296,6 @@ iff.intro
⟨ge_of_eq rfl, λ _, le_of_eq (by { rw [zero_mul, hf], exact norm_zero })⟩)
(op_norm_nonneg _))

@[simp] lemma norm_zero : ∥(0 : E →L[𝕜] F)∥ = 0 :=
by rw op_norm_zero_iff

/-- The norm of the identity is at most `1`. It is in fact `1`, except when the space is trivial
where it is `0`. It means that one can not do better than an inequality in general. -/
lemma norm_id_le : ∥id 𝕜 E∥ ≤ 1 :=
Expand Down
4 changes: 2 additions & 2 deletions src/measure_theory/bochner_integration.lean
Expand Up @@ -1022,7 +1022,7 @@ begin
by_cases hf : measurable f ∧ integrable f μ,
{ rw [integral_eq f hf.1 hf.2, ← l1.norm_of_fun_eq_lintegral_norm f hf.1 hf.2],
exact l1.norm_integral_le _ },
{ rw [integral_undef hf, _root_.norm_zero],
{ rw [integral_undef hf, norm_zero],
exact to_real_nonneg }
end

Expand Down Expand Up @@ -1197,7 +1197,7 @@ classical.by_cases
... = ∫ a, ∥f a∥ ∂μ : (integral_eq_lintegral_of_nonneg_ae le_ae $ measurable.norm h).symm )
( λh : ¬measurable f,
begin
rw [integral_non_measurable h, _root_.norm_zero],
rw [integral_non_measurable h, norm_zero],
exact integral_nonneg_of_ae le_ae
end )

Expand Down

0 comments on commit d7fcd8c

Please sign in to comment.