From 17ef379e997badd73e5eabb4d38f11919ab3c4b3 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 17 Nov 2022 13:13:12 +0000 Subject: [PATCH] refactor(analysis): change the symbol for norm to align with the unicode spec (#17575) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The characters in question are: * U+2016 DOUBLE VERTICAL LINE `‖`. The Unicode Character Database says "used in pairs to indicate norm of a matrix", for instance [here](https://www.fileformat.info/info/unicode/char/2016/index.htm) and [here](https://unicode-explorer.com/c/2016) * U+2225 PARALLEL TO `∥`. On some platforms this renders with a forward slant! Previously `norm` was the latter and `parallel` was the former. This change swaps them around: * `∥x∥`, `S ‖ T`: before * `‖x‖`, `S ∥ T`: after Code using U+2016 `‖` for fintype cardinality has been left unchanged. --- archive/imo/imo1972_q5.lean | 60 +- counterexamples/phillips.lean | 10 +- src/analysis/ODE/gronwall.lean | 16 +- src/analysis/ODE/picard_lindelof.lean | 10 +- src/analysis/analytic/basic.lean | 238 ++++---- src/analysis/analytic/composition.lean | 64 +-- src/analysis/analytic/inverse.lean | 46 +- src/analysis/analytic/radius_liminf.lean | 10 +- src/analysis/analytic/uniqueness.lean | 2 +- .../asymptotics/asymptotic_equivalent.lean | 24 +- src/analysis/asymptotics/asymptotics.lean | 230 ++++---- .../asymptotics/specific_asymptotics.lean | 22 +- .../asymptotics/superpolynomial_decay.lean | 4 +- src/analysis/asymptotics/theta.lean | 8 +- src/analysis/box_integral/basic.lean | 14 +- .../box_integral/divergence_theorem.lean | 30 +- src/analysis/box_integral/integrability.lean | 20 +- src/analysis/calculus/cont_diff.lean | 26 +- src/analysis/calculus/deriv.lean | 58 +- src/analysis/calculus/extend_deriv.lean | 16 +- src/analysis/calculus/fderiv.lean | 48 +- src/analysis/calculus/fderiv_analytic.lean | 6 +- src/analysis/calculus/fderiv_measurable.lean | 158 +++--- src/analysis/calculus/fderiv_symmetric.lean | 34 +- .../calculus/formal_multilinear_series.lean | 2 +- src/analysis/calculus/inverse.lean | 38 +- src/analysis/calculus/local_extr.lean | 4 +- src/analysis/calculus/mean_value.lean | 172 +++--- .../calculus/parametric_integral.lean | 48 +- .../parametric_interval_integral.lean | 4 +- src/analysis/calculus/tangent_cone.lean | 24 +- src/analysis/calculus/taylor.lean | 10 +- .../calculus/uniform_limits_deriv.lean | 28 +- src/analysis/complex/abs_max.lean | 112 ++-- src/analysis/complex/arg.lean | 2 +- src/analysis/complex/basic.lean | 52 +- src/analysis/complex/cauchy_integral.lean | 18 +- src/analysis/complex/conformal.lean | 4 +- src/analysis/complex/isometry.lean | 2 +- src/analysis/complex/liouville.lean | 18 +- src/analysis/complex/open_mapping.lean | 32 +- src/analysis/complex/phragmen_lindelof.lean | 194 +++---- src/analysis/complex/polynomial.lean | 2 +- .../complex/removable_singularity.lean | 4 +- src/analysis/complex/roots_of_unity.lean | 4 +- src/analysis/complex/schwarz.lean | 24 +- src/analysis/convex/gauge.lean | 14 +- src/analysis/convex/integral.lean | 18 +- src/analysis/convex/strict_convex_space.lean | 46 +- src/analysis/convex/topology.lean | 4 +- src/analysis/convex/uniform.lean | 36 +- src/analysis/convolution.lean | 58 +- src/analysis/fourier.lean | 6 +- src/analysis/inner_product_space/adjoint.lean | 24 +- src/analysis/inner_product_space/basic.lean | 280 ++++----- .../inner_product_space/calculus.lean | 48 +- src/analysis/inner_product_space/dual.lean | 4 +- .../gram_schmidt_ortho.lean | 16 +- .../inner_product_space/l2_space.lean | 14 +- .../inner_product_space/lax_milgram.lean | 12 +- .../inner_product_space/orientation.lean | 10 +- src/analysis/inner_product_space/pi_L2.lean | 8 +- .../inner_product_space/projection.lean | 142 ++--- .../inner_product_space/rayleigh.lean | 82 +-- .../inner_product_space/spectrum.lean | 10 +- src/analysis/inner_product_space/two_dim.lean | 48 +- src/analysis/locally_convex/abs_convex.lean | 2 +- .../locally_convex/balanced_core_hull.lean | 20 +- src/analysis/locally_convex/basic.lean | 24 +- src/analysis/locally_convex/bounded.lean | 6 +- .../locally_convex/continuous_of_bounded.lean | 4 +- src/analysis/locally_convex/polar.lean | 12 +- src/analysis/locally_convex/weak_dual.lean | 14 +- .../locally_convex/with_seminorms.lean | 4 +- src/analysis/matrix.lean | 112 ++-- src/analysis/normed/field/basic.lean | 240 ++++---- .../normed/group/SemiNormedGroup.lean | 8 +- .../group/SemiNormedGroup/completion.lean | 2 +- .../normed/group/SemiNormedGroup/kernels.lean | 8 +- src/analysis/normed/group/add_circle.lean | 28 +- src/analysis/normed/group/add_torsor.lean | 18 +- src/analysis/normed/group/basic.lean | 536 +++++++++--------- src/analysis/normed/group/completion.lean | 2 +- src/analysis/normed/group/hom.lean | 192 +++---- src/analysis/normed/group/hom_completion.lean | 54 +- src/analysis/normed/group/infinite_sum.lean | 72 +-- src/analysis/normed/group/pointwise.lean | 4 +- src/analysis/normed/group/quotient.lean | 164 +++--- src/analysis/normed/order/basic.lean | 12 +- src/analysis/normed/order/lattice.lean | 30 +- src/analysis/normed_space/M_structure.lean | 40 +- src/analysis/normed_space/add_torsor.lean | 30 +- .../normed_space/affine_isometry.lean | 12 +- src/analysis/normed_space/algebra.lean | 6 +- src/analysis/normed_space/banach.lean | 84 +-- .../normed_space/banach_steinhaus.lean | 50 +- src/analysis/normed_space/basic.lean | 110 ++-- .../normed_space/bounded_linear_maps.lean | 94 +-- .../normed_space/compact_operator.lean | 14 +- src/analysis/normed_space/completion.lean | 2 +- .../normed_space/continuous_affine_map.lean | 40 +- src/analysis/normed_space/dual.lean | 46 +- src/analysis/normed_space/enorm.lean | 16 +- src/analysis/normed_space/exponential.lean | 14 +- src/analysis/normed_space/extend.lean | 26 +- src/analysis/normed_space/extr.lean | 40 +- .../normed_space/finite_dimension.lean | 114 ++-- .../normed_space/hahn_banach/extension.lean | 38 +- .../normed_space/indicator_function.lean | 10 +- src/analysis/normed_space/int.lean | 16 +- src/analysis/normed_space/is_R_or_C.lean | 28 +- .../normed_space/linear_isometry.lean | 26 +- src/analysis/normed_space/lp_equiv.lean | 6 +- src/analysis/normed_space/lp_space.lean | 194 +++---- src/analysis/normed_space/mazur_ulam.lean | 2 +- src/analysis/normed_space/multilinear.lean | 402 ++++++------- src/analysis/normed_space/operator_norm.lean | 430 +++++++------- src/analysis/normed_space/pi_Lp.lean | 60 +- src/analysis/normed_space/pointwise.lean | 30 +- src/analysis/normed_space/ray.lean | 30 +- src/analysis/normed_space/riesz_lemma.lean | 34 +- src/analysis/normed_space/spectrum.lean | 104 ++-- src/analysis/normed_space/star/basic.lean | 86 +-- src/analysis/normed_space/star/matrix.lean | 10 +- src/analysis/normed_space/star/spectrum.lean | 18 +- src/analysis/normed_space/units.lean | 62 +- src/analysis/normed_space/weak_dual.lean | 2 +- src/analysis/quaternion.lean | 6 +- src/analysis/schwartz_space.lean | 52 +- src/analysis/seminorm.lean | 30 +- src/analysis/special_functions/bernstein.lean | 22 +- src/analysis/special_functions/exp.lean | 32 +- src/analysis/special_functions/exp_deriv.lean | 2 +- src/analysis/special_functions/gamma.lean | 8 +- src/analysis/special_functions/gaussian.lean | 2 +- .../special_functions/japanese_bracket.lean | 34 +- src/analysis/special_functions/log/deriv.lean | 4 +- .../special_functions/non_integrable.lean | 38 +- src/analysis/special_functions/pow.lean | 2 +- src/analysis/specific_limits/normed.lean | 112 ++-- src/combinatorics/additive/behrend.lean | 2 +- src/data/complex/is_R_or_C.lean | 36 +- .../euclidean/angle/oriented/basic.lean | 32 +- src/geometry/euclidean/angle/sphere.lean | 8 +- .../euclidean/angle/unoriented/basic.lean | 58 +- .../angle/unoriented/right_angle.lean | 66 +-- src/geometry/euclidean/basic.lean | 10 +- src/geometry/euclidean/sphere.lean | 20 +- src/geometry/euclidean/triangle.lean | 42 +- src/geometry/manifold/complex.lean | 2 +- src/geometry/manifold/instances/sphere.lean | 72 +-- src/information_theory/hamming.lean | 4 +- .../affine_space/affine_subspace.lean | 22 +- .../constructions/borel_space.lean | 10 +- src/measure_theory/constructions/prod.lean | 30 +- .../covering/besicovitch_vector_space.lean | 140 ++--- .../covering/density_theorem.lean | 2 +- .../covering/differentiation.lean | 36 +- .../function/ae_eq_of_integral.lean | 30 +- .../conditional_expectation/basic.lean | 80 +-- .../conditional_expectation/real.lean | 18 +- .../function/continuous_map_dense.lean | 18 +- .../function/convergence_in_measure.lean | 4 +- src/measure_theory/function/jacobian.lean | 100 ++-- src/measure_theory/function/l1_space.lean | 144 ++--- src/measure_theory/function/l2_space.lean | 12 +- .../function/locally_integrable.lean | 2 +- src/measure_theory/function/lp_space.lean | 358 ++++++------ .../function/simple_func_dense_lp.lean | 46 +- .../function/strongly_measurable/basic.lean | 46 +- .../function/uniform_integrable.lean | 96 ++-- src/measure_theory/integral/bochner.lean | 128 ++--- .../integral/circle_integral.lean | 56 +- .../integral/circle_transform.lean | 2 +- .../integral/divergence_theorem.lean | 2 +- .../integral/integrable_on.lean | 6 +- .../integral/integral_eq_improper.lean | 40 +- .../integral/interval_integral.lean | 68 +-- src/measure_theory/integral/lebesgue.lean | 8 +- src/measure_theory/integral/set_integral.lean | 42 +- src/measure_theory/integral/set_to_l1.lean | 100 ++-- .../integral/torus_integral.lean | 14 +- .../measure/finite_measure.lean | 2 +- .../number_field/embeddings.lean | 8 +- src/number_theory/padics/hensel.lean | 226 ++++---- src/number_theory/padics/padic_integers.lean | 94 +-- src/number_theory/padics/padic_numbers.lean | 76 +-- src/number_theory/padics/ring_homs.lean | 24 +- src/probability/density.lean | 12 +- src/probability/ident_distrib.lean | 12 +- src/probability/integration.lean | 26 +- src/probability/martingale/convergence.lean | 8 +- src/probability/martingale/upcrossing.lean | 2 +- src/probability/variance.lean | 4 +- .../polynomial/cyclotomic/eval.lean | 12 +- src/set_theory/game/basic.lean | 6 +- src/set_theory/game/impartial.lean | 18 +- src/set_theory/game/nim.lean | 4 +- src/set_theory/game/pgame.lean | 60 +- .../algebra/module/finite_dimension.lean | 8 +- src/topology/algebra/polynomial.lean | 20 +- src/topology/algebra/uniform_convergence.lean | 2 +- src/topology/continuous_function/bounded.lean | 94 +-- src/topology/continuous_function/compact.lean | 28 +- src/topology/continuous_function/ideals.lean | 28 +- .../stone_weierstrass.lean | 26 +- .../continuous_function/weierstrass.lean | 2 +- .../continuous_function/zero_at_infty.lean | 2 +- src/topology/instances/ennreal.lean | 6 +- src/topology/tietze_extension.lean | 62 +- test/positivity.lean | 2 +- 211 files changed, 4978 insertions(+), 4978 deletions(-) diff --git a/archive/imo/imo1972_q5.lean b/archive/imo/imo1972_q5.lean index 5664210339a6c..21178464e58dd 100644 --- a/archive/imo/imo1972_q5.lean +++ b/archive/imo/imo1972_q5.lean @@ -16,49 +16,49 @@ Prove that `|g(x)| ≤ 1` for all `x`. -/ /-- -This proof begins by introducing the supremum of `f`, `k ≤ 1` as well as `k' = k / ∥g y∥`. We then +This proof begins by introducing the supremum of `f`, `k ≤ 1` as well as `k' = k / ‖g y‖`. We then suppose that the conclusion does not hold (`hneg`) and show that `k ≤ k'` (by -`2 * (∥f x∥ * ∥g y∥) ≤ 2 * k` obtained from the main hypothesis `hf1`) and that `k' < k` (obtained +`2 * (‖f x‖ * ‖g y‖) ≤ 2 * k` obtained from the main hypothesis `hf1`) and that `k' < k` (obtained from `hneg` directly), finally raising a contradiction with `k' < k'`. (Authored by Stanislas Polu inspired by Ruben Van de Velde). -/ example (f g : ℝ → ℝ) (hf1 : ∀ x, ∀ y, (f(x+y) + f(x-y)) = 2 * f(x) * g(y)) - (hf2 : ∀ y, ∥f(y)∥ ≤ 1) + (hf2 : ∀ y, ‖f(y)‖ ≤ 1) (hf3 : ∃ x, f(x) ≠ 0) (y : ℝ) : - ∥g(y)∥ ≤ 1 := + ‖g(y)‖ ≤ 1 := begin - set S := set.range (λ x, ∥f x∥), + set S := set.range (λ x, ‖f x‖), -- Introduce `k`, the supremum of `f`. let k : ℝ := Sup (S), - -- Show that `∥f x∥ ≤ k`. - have hk₁ : ∀ x, ∥f x∥ ≤ k, + -- Show that `‖f x‖ ≤ k`. + have hk₁ : ∀ x, ‖f x‖ ≤ k, { have h : bdd_above S, from ⟨1, set.forall_range_iff.mpr hf2⟩, intro x, exact le_cSup h (set.mem_range_self x), }, - -- Show that `2 * (∥f x∥ * ∥g y∥) ≤ 2 * k`. - have hk₂ : ∀ x, 2 * (∥f x∥ * ∥g y∥) ≤ 2 * k, + -- Show that `2 * (‖f x‖ * ‖g y‖) ≤ 2 * k`. + have hk₂ : ∀ x, 2 * (‖f x‖ * ‖g y‖) ≤ 2 * k, { intro x, - calc 2 * (∥f x∥ * ∥g y∥) - = ∥2 * f x * g y∥ : by simp [abs_mul, mul_assoc] - ... = ∥f (x + y) + f (x - y)∥ : by rw hf1 - ... ≤ ∥f (x + y)∥ + ∥f (x - y)∥ : norm_add_le _ _ + calc 2 * (‖f x‖ * ‖g y‖) + = ‖2 * f x * g y‖ : by simp [abs_mul, mul_assoc] + ... = ‖f (x + y) + f (x - y)‖ : by rw hf1 + ... ≤ ‖f (x + y)‖ + ‖f (x - y)‖ : norm_add_le _ _ ... ≤ k + k : add_le_add (hk₁ _) (hk₁ _) ... = 2 * k : (two_mul _).symm, }, -- Suppose the conclusion does not hold. by_contra' hneg, - set k' := k / ∥g y∥, + set k' := k / ‖g y‖, -- Demonstrate that `k' < k` using `hneg`. have H₁ : k' < k, { have h₁ : 0 < k, { obtain ⟨x, hx⟩ := hf3, calc 0 - < ∥f x∥ : norm_pos_iff.mpr hx + < ‖f x‖ : norm_pos_iff.mpr hx ... ≤ k : hk₁ x }, rw div_lt_iff, apply lt_mul_of_one_lt_right h₁ hneg, @@ -67,8 +67,8 @@ begin -- Demonstrate that `k ≤ k'` using `hk₂`. have H₂ : k ≤ k', { have h₁ : ∃ x : ℝ, x ∈ S, - { use ∥f 0∥, exact set.mem_range_self 0, }, - have h₂ : ∀ x, ∥f x∥ ≤ k', + { use ‖f 0‖, exact set.mem_range_self 0, }, + have h₂ : ∀ x, ‖f x‖ ≤ k', { intros x, rw le_div_iff, { apply (mul_le_mul_left zero_lt_two).mp (hk₂ x) }, @@ -95,28 +95,28 @@ This is a more concise version of the proof proposed by Ruben Van de Velde. -/ example (f g : ℝ → ℝ) (hf1 : ∀ x, ∀ y, (f (x+y) + f(x-y)) = 2 * f(x) * g(y)) - (hf2 : bdd_above (set.range (λ x, ∥f x∥))) + (hf2 : bdd_above (set.range (λ x, ‖f x‖))) (hf3 : ∃ x, f(x) ≠ 0) (y : ℝ) : - ∥g(y)∥ ≤ 1 := + ‖g(y)‖ ≤ 1 := begin obtain ⟨x, hx⟩ := hf3, - set k := ⨆ x, ∥f x∥, - have h : ∀ x, ∥f x∥ ≤ k := le_csupr hf2, + set k := ⨆ x, ‖f x‖, + have h : ∀ x, ‖f x‖ ≤ k := le_csupr hf2, by_contra' H, - have hgy : 0 < ∥g y∥, + have hgy : 0 < ‖g y‖, by linarith, have k_pos : 0 < k := lt_of_lt_of_le (norm_pos_iff.mpr hx) (h x), - have : k / ∥g y∥ < k := (div_lt_iff hgy).mpr (lt_mul_of_one_lt_right k_pos H), - have : k ≤ k / ∥g y∥, - { suffices : ∀ x, ∥f x∥ ≤ k / ∥g y∥, from csupr_le this, + have : k / ‖g y‖ < k := (div_lt_iff hgy).mpr (lt_mul_of_one_lt_right k_pos H), + have : k ≤ k / ‖g y‖, + { suffices : ∀ x, ‖f x‖ ≤ k / ‖g y‖, from csupr_le this, intro x, - suffices : 2 * (∥f x∥ * ∥g y∥) ≤ 2 * k, + suffices : 2 * (‖f x‖ * ‖g y‖) ≤ 2 * k, by { rwa [le_div_iff hgy, ←mul_le_mul_left (zero_lt_two : (0 : ℝ) < 2)] }, - calc 2 * (∥f x∥ * ∥g y∥) - = ∥2 * f x * g y∥ : by simp [abs_mul, mul_assoc] - ... = ∥f (x + y) + f (x - y)∥ : by rw hf1 - ... ≤ ∥f (x + y)∥ + ∥f (x - y)∥ : abs_add _ _ + calc 2 * (‖f x‖ * ‖g y‖) + = ‖2 * f x * g y‖ : by simp [abs_mul, mul_assoc] + ... = ‖f (x + y) + f (x - y)‖ : by rw hf1 + ... ≤ ‖f (x + y)‖ + ‖f (x - y)‖ : abs_add _ _ ... ≤ 2 * k : by linarith [h (x+y), h (x -y)] }, linarith, end diff --git a/counterexamples/phillips.lean b/counterexamples/phillips.lean index 1fc2dfafe2b21..95067c8ce8e41 100644 --- a/counterexamples/phillips.lean +++ b/counterexamples/phillips.lean @@ -375,7 +375,7 @@ section -/ lemma norm_indicator_le_one (s : set α) (x : α) : - ∥(indicator s (1 : α → ℝ)) x∥ ≤ 1 := + ‖(indicator s (1 : α → ℝ)) x‖ ≤ 1 := by { simp only [indicator, pi.one_apply], split_ifs; norm_num } /-- A functional in the dual space of bounded functions gives rise to a bounded additive measure, @@ -392,8 +392,8 @@ def _root_.continuous_linear_map.to_bounded_additive_measure by { ext x, simp [indicator_union_of_disjoint hst], }, rw [this, f.map_add], end, - exists_bound := ⟨∥f∥, λ s, begin - have I : ∥of_normed_add_comm_group_discrete (indicator s 1) 1 (norm_indicator_le_one s)∥ ≤ 1, + exists_bound := ⟨‖f‖, λ s, begin + have I : ‖of_normed_add_comm_group_discrete (indicator s 1) 1 (norm_indicator_le_one s)‖ ≤ 1, by apply norm_of_normed_add_comm_group_le _ zero_le_one, apply le_trans (f.le_op_norm _), simpa using mul_le_mul_of_nonneg_left I (norm_nonneg f), @@ -492,7 +492,7 @@ lemma countable_spf_mem (Hcont : #ℝ = aleph 1) (y : ℝ) : {x | y ∈ spf Hcon We construct a function `f` from `[0,1]` to a complete Banach space `B`, which is weakly measurable (i.e., for any continuous linear form `φ` on `B` the function `φ ∘ f` is measurable), bounded in -norm (i.e., for all `x`, one has `∥f x∥ ≤ 1`), and still `f` has no Pettis integral. +norm (i.e., for all `x`, one has `‖f x‖ ≤ 1`), and still `f` has no Pettis integral. This construction, due to Phillips, requires the continuum hypothesis. We will take for `B` the space of all bounded functions on `ℝ`, with the supremum norm (no measure here, we are really @@ -582,7 +582,7 @@ begin end /-- The function `f Hcont : ℝ → (discrete_copy ℝ →ᵇ ℝ)` is uniformly bounded by `1` in norm. -/ -lemma norm_bound (Hcont : #ℝ = aleph 1) (x : ℝ) : ∥f Hcont x∥ ≤ 1 := +lemma norm_bound (Hcont : #ℝ = aleph 1) (x : ℝ) : ‖f Hcont x‖ ≤ 1 := norm_of_normed_add_comm_group_le _ zero_le_one _ /-- The function `f Hcont : ℝ → (discrete_copy ℝ →ᵇ ℝ)` has no Pettis integral. -/ diff --git a/src/analysis/ODE/gronwall.lean b/src/analysis/ODE/gronwall.lean index a68f9966e0bf3..7e1fdf7ee5709 100644 --- a/src/analysis/ODE/gronwall.lean +++ b/src/analysis/ODE/gronwall.lean @@ -9,8 +9,8 @@ import analysis.special_functions.exp_deriv # Grönwall's inequality The main technical result of this file is the Grönwall-like inequality -`norm_le_gronwall_bound_of_norm_deriv_right_le`. It states that if `f : ℝ → E` satisfies `∥f a∥ ≤ δ` -and `∀ x ∈ [a, b), ∥f' x∥ ≤ K * ∥f x∥ + ε`, then for all `x ∈ [a, b]` we have `∥f x∥ ≤ δ * exp (K * +`norm_le_gronwall_bound_of_norm_deriv_right_le`. It states that if `f : ℝ → E` satisfies `‖f a‖ ≤ δ` +and `∀ x ∈ [a, b), ‖f' x‖ ≤ K * ‖f x‖ + ε`, then for all `x ∈ [a, b]` we have `‖f x‖ ≤ δ * exp (K * x) + (ε / K) * (exp (K * x) - 1)`. Then we use this inequality to prove some estimates on the possible rate of growth of the distance @@ -22,7 +22,7 @@ Sec. 4.5][HubbardWest-ode], where `norm_le_gronwall_bound_of_norm_deriv_right_le ## TODO -- Once we have FTC, prove an inequality for a function satisfying `∥f' x∥ ≤ K x * ∥f x∥ + ε`, +- Once we have FTC, prove an inequality for a function satisfying `‖f' x‖ ≤ K x * ‖f x‖ + ε`, or more generally `liminf_{y→x+0} (f y - f x)/(y - x) ≤ K x * f x + ε` with any sign of `K x` and `f x`. -/ @@ -101,7 +101,7 @@ the inequalities `f a ≤ δ` and `∀ x ∈ [a, b), liminf_{z→x+0} (f z - f x)/(z - x) ≤ K * (f x) + ε`, then `f x` is bounded by `gronwall_bound δ K ε (x - a)` on `[a, b]`. -See also `norm_le_gronwall_bound_of_norm_deriv_right_le` for a version bounding `∥f x∥`, +See also `norm_le_gronwall_bound_of_norm_deriv_right_le` for a version bounding `‖f x‖`, `f : ℝ → E`. -/ theorem le_gronwall_bound_of_liminf_deriv_right_le {f f' : ℝ → ℝ} {δ K ε : ℝ} {a b : ℝ} (hf : continuous_on f (Icc a b)) @@ -128,13 +128,13 @@ begin end /-- A Grönwall-like inequality: if `f : ℝ → E` is continuous on `[a, b]`, has right derivative -`f' x` at every point `x ∈ [a, b)`, and satisfies the inequalities `∥f a∥ ≤ δ`, -`∀ x ∈ [a, b), ∥f' x∥ ≤ K * ∥f x∥ + ε`, then `∥f x∥` is bounded by `gronwall_bound δ K ε (x - a)` +`f' x` at every point `x ∈ [a, b)`, and satisfies the inequalities `‖f a‖ ≤ δ`, +`∀ x ∈ [a, b), ‖f' x‖ ≤ K * ‖f x‖ + ε`, then `‖f x‖` is bounded by `gronwall_bound δ K ε (x - a)` on `[a, b]`. -/ theorem norm_le_gronwall_bound_of_norm_deriv_right_le {f f' : ℝ → E} {δ K ε : ℝ} {a b : ℝ} (hf : continuous_on f (Icc a b)) (hf' : ∀ x ∈ Ico a b, has_deriv_within_at f (f' x) (Ici x) x) - (ha : ∥f a∥ ≤ δ) (bound : ∀ x ∈ Ico a b, ∥f' x∥ ≤ K * ∥f x∥ + ε) : - ∀ x ∈ Icc a b, ∥f x∥ ≤ gronwall_bound δ K ε (x - a) := + (ha : ‖f a‖ ≤ δ) (bound : ∀ x ∈ Ico a b, ‖f' x‖ ≤ K * ‖f x‖ + ε) : + ∀ x ∈ Icc a b, ‖f x‖ ≤ gronwall_bound δ K ε (x - a) := le_gronwall_bound_of_liminf_deriv_right_le (continuous_norm.comp_continuous_on hf) (λ x hx r hr, (hf' x hx).liminf_right_slope_norm_le hr) ha bound diff --git a/src/analysis/ODE/picard_lindelof.lean b/src/analysis/ODE/picard_lindelof.lean index 5622696920f98..4e670e38341cb 100644 --- a/src/analysis/ODE/picard_lindelof.lean +++ b/src/analysis/ODE/picard_lindelof.lean @@ -51,7 +51,7 @@ structure is_picard_lindelof (hR : 0 ≤ R) (lipschitz : ∀ t ∈ Icc t_min t_max, lipschitz_on_with L (v t) (closed_ball x₀ R)) (cont : ∀ x ∈ closed_ball x₀ R, continuous_on (λ (t : ℝ), v t x) (Icc t_min t_max)) -(norm_le : ∀ (t ∈ Icc t_min t_max) (x ∈ closed_ball x₀ R), ∥v t x∥ ≤ C) +(norm_le : ∀ (t ∈ Icc t_min t_max) (x ∈ closed_ball x₀ R), ‖v t x‖ ≤ C) (C_mul_le_R : (C : ℝ) * linear_order.max (t_max - t₀) (t₀ - t_min) ≤ R) /-- This structure holds arguments of the Picard-Lipschitz (Cauchy-Lipschitz) theorem. It is part of @@ -100,7 +100,7 @@ have continuous_on (uncurry (flip v)) (closed_ball v.x₀ v.R ×ˢ Icc v.t_min v this.comp continuous_swap.continuous_on (preimage_swap_prod _ _).symm.subset lemma norm_le {t : ℝ} (ht : t ∈ Icc v.t_min v.t_max) {x : E} (hx : x ∈ closed_ball v.x₀ v.R) : - ∥v t x∥ ≤ v.C := + ‖v t x‖ ≤ v.C := v.is_pl.norm_le _ ht _ hx /-- The maximum of distances from `t₀` to the endpoints of `[t_min, t_max]`. -/ @@ -192,7 +192,7 @@ begin exact ⟨(v.proj x).2, f.mem_closed_ball _⟩ end -lemma norm_v_comp_le (t : ℝ) : ∥f.v_comp t∥ ≤ v.C := +lemma norm_v_comp_le (t : ℝ) : ‖f.v_comp t‖ ≤ v.C := v.norm_le (v.proj t).2 $ f.mem_closed_ball _ lemma dist_apply_le_dist (f₁ f₂ : fun_space v) (t : Icc v.t_min v.t_max) : @@ -262,7 +262,7 @@ begin simp only [dist_eq_norm, next_apply, add_sub_add_left_eq_sub, ← interval_integral.integral_sub (interval_integrable_v_comp _ _ _) (interval_integrable_v_comp _ _ _), norm_integral_eq_norm_integral_Ioc] at *, - calc ∥∫ τ in Ι (v.t₀ : ℝ) t, f₁.v_comp τ - f₂.v_comp τ∥ + calc ‖∫ τ in Ι (v.t₀ : ℝ) t, f₁.v_comp τ - f₂.v_comp τ‖ ≤ ∫ τ in Ι (v.t₀ : ℝ) t, v.L * ((v.L * |τ - v.t₀|) ^ n / n! * d) : begin refine norm_integral_le_of_norm_le (continuous.integrable_on_interval_oc _) _, @@ -338,7 +338,7 @@ end picard_lindelof lemma is_picard_lindelof.norm_le₀ {E : Type*} [normed_add_comm_group E] {v : ℝ → E → E} {t_min t₀ t_max : ℝ} {x₀ : E} {C R : ℝ} {L : ℝ≥0} - (hpl : is_picard_lindelof v t_min t₀ t_max x₀ L R C) : ∥v t₀ x₀∥ ≤ C := + (hpl : is_picard_lindelof v t_min t₀ t_max x₀ L R C) : ‖v t₀ x₀‖ ≤ C := hpl.norm_le t₀ hpl.ht₀ x₀ $ mem_closed_ball_self hpl.hR /-- Picard-Lindelöf (Cauchy-Lipschitz) theorem. -/ diff --git a/src/analysis/analytic/basic.lean b/src/analysis/analytic/basic.lean index 54c76bfbbb20a..d8f3bfea0cfeb 100644 --- a/src/analysis/analytic/basic.lean +++ b/src/analysis/analytic/basic.lean @@ -27,13 +27,13 @@ space is analytic, as well as the inverse on invertible operators. Let `p` be a formal multilinear series from `E` to `F`, i.e., `p n` is a multilinear map on `E^n` for `n : ℕ`. -* `p.radius`: the largest `r : ℝ≥0∞` such that `∥p n∥ * r^n` grows subexponentially. -* `p.le_radius_of_bound`, `p.le_radius_of_bound_nnreal`, `p.le_radius_of_is_O`: if `∥p n∥ * r ^ n` +* `p.radius`: the largest `r : ℝ≥0∞` such that `‖p n‖ * r^n` grows subexponentially. +* `p.le_radius_of_bound`, `p.le_radius_of_bound_nnreal`, `p.le_radius_of_is_O`: if `‖p n‖ * r ^ n` is bounded above, then `r ≤ p.radius`; * `p.is_o_of_lt_radius`, `p.norm_mul_pow_le_mul_pow_of_lt_radius`, `p.is_o_one_of_lt_radius`, `p.norm_mul_pow_le_of_lt_radius`, `p.nnnorm_mul_pow_le_of_lt_radius`: if `r < p.radius`, then - `∥p n∥ * r ^ n` tends to zero exponentially; -* `p.lt_radius_of_is_O`: if `r ≠ 0` and `∥p n∥ * r ^ n = O(a ^ n)` for some `-1 < a < 1`, then + `‖p n‖ * r ^ n` tends to zero exponentially; +* `p.lt_radius_of_is_O`: if `r ≠ 0` and `‖p n‖ * r ^ n = O(a ^ n)` for some `-1 < a < 1`, then `r < p.radius`; * `p.partial_sum n x`: the sum `∑_{i = 0}^{n-1} pᵢ xⁱ`. * `p.sum x`: the sum `∑'_{i = 0}^{∞} pᵢ xⁱ`. @@ -82,7 +82,7 @@ variables [topological_add_group E] [topological_add_group F] variables [has_continuous_const_smul 𝕜 E] [has_continuous_const_smul 𝕜 F] /-- Given a formal multilinear series `p` and a vector `x`, then `p.sum x` is the sum `Σ pₙ xⁿ`. A -priori, it only behaves well when `∥x∥ < p.radius`. -/ +priori, it only behaves well when `‖x‖ < p.radius`. -/ protected def sum (p : formal_multilinear_series 𝕜 E F) (x : E) : F := ∑' n : ℕ , p n (λ i, x) /-- Given a formal multilinear series `p` and a vector `x`, then `p.partial_sum n x` is the sum @@ -109,38 +109,38 @@ namespace formal_multilinear_series variables (p : formal_multilinear_series 𝕜 E F) {r : ℝ≥0} -/-- The radius of a formal multilinear series is the largest `r` such that the sum `Σ ∥pₙ∥ ∥y∥ⁿ` -converges for all `∥y∥ < r`. This implies that `Σ pₙ yⁿ` converges for all `∥y∥ < r`, but these +/-- The radius of a formal multilinear series is the largest `r` such that the sum `Σ ‖pₙ‖ ‖y‖ⁿ` +converges for all `‖y‖ < r`. This implies that `Σ pₙ yⁿ` converges for all `‖y‖ < r`, but these definitions are *not* equivalent in general. -/ def radius (p : formal_multilinear_series 𝕜 E F) : ℝ≥0∞ := -⨆ (r : ℝ≥0) (C : ℝ) (hr : ∀ n, ∥p n∥ * r ^ n ≤ C), (r : ℝ≥0∞) +⨆ (r : ℝ≥0) (C : ℝ) (hr : ∀ n, ‖p n‖ * r ^ n ≤ C), (r : ℝ≥0∞) -/-- If `∥pₙ∥ rⁿ` is bounded in `n`, then the radius of `p` is at least `r`. -/ -lemma le_radius_of_bound (C : ℝ) {r : ℝ≥0} (h : ∀ (n : ℕ), ∥p n∥ * r^n ≤ C) : +/-- If `‖pₙ‖ rⁿ` is bounded in `n`, then the radius of `p` is at least `r`. -/ +lemma le_radius_of_bound (C : ℝ) {r : ℝ≥0} (h : ∀ (n : ℕ), ‖p n‖ * r^n ≤ C) : (r : ℝ≥0∞) ≤ p.radius := le_supr_of_le r $ le_supr_of_le C $ (le_supr (λ _, (r : ℝ≥0∞)) h) -/-- If `∥pₙ∥ rⁿ` is bounded in `n`, then the radius of `p` is at least `r`. -/ -lemma le_radius_of_bound_nnreal (C : ℝ≥0) {r : ℝ≥0} (h : ∀ (n : ℕ), ∥p n∥₊ * r^n ≤ C) : +/-- If `‖pₙ‖ rⁿ` is bounded in `n`, then the radius of `p` is at least `r`. -/ +lemma le_radius_of_bound_nnreal (C : ℝ≥0) {r : ℝ≥0} (h : ∀ (n : ℕ), ‖p n‖₊ * r^n ≤ C) : (r : ℝ≥0∞) ≤ p.radius := p.le_radius_of_bound C $ λ n, by exact_mod_cast (h n) -/-- If `∥pₙ∥ rⁿ = O(1)`, as `n → ∞`, then the radius of `p` is at least `r`. -/ -lemma le_radius_of_is_O (h : (λ n, ∥p n∥ * r^n) =O[at_top] (λ n, (1 : ℝ))) : ↑r ≤ p.radius := +/-- If `‖pₙ‖ rⁿ = O(1)`, as `n → ∞`, then the radius of `p` is at least `r`. -/ +lemma le_radius_of_is_O (h : (λ n, ‖p n‖ * r^n) =O[at_top] (λ n, (1 : ℝ))) : ↑r ≤ p.radius := exists.elim (is_O_one_nat_at_top_iff.1 h) $ λ C hC, p.le_radius_of_bound C $ λ n, (le_abs_self _).trans (hC n) -lemma le_radius_of_eventually_le (C) (h : ∀ᶠ n in at_top, ∥p n∥ * r ^ n ≤ C) : ↑r ≤ p.radius := +lemma le_radius_of_eventually_le (C) (h : ∀ᶠ n in at_top, ‖p n‖ * r ^ n ≤ C) : ↑r ≤ p.radius := p.le_radius_of_is_O $ is_O.of_bound C $ h.mono $ λ n hn, by simpa -lemma le_radius_of_summable_nnnorm (h : summable (λ n, ∥p n∥₊ * r ^ n)) : ↑r ≤ p.radius := -p.le_radius_of_bound_nnreal (∑' n, ∥p n∥₊ * r ^ n) $ λ n, le_tsum' h _ +lemma le_radius_of_summable_nnnorm (h : summable (λ n, ‖p n‖₊ * r ^ n)) : ↑r ≤ p.radius := +p.le_radius_of_bound_nnreal (∑' n, ‖p n‖₊ * r ^ n) $ λ n, le_tsum' h _ -lemma le_radius_of_summable (h : summable (λ n, ∥p n∥ * r ^ n)) : ↑r ≤ p.radius := +lemma le_radius_of_summable (h : summable (λ n, ‖p n‖ * r ^ n)) : ↑r ≤ p.radius := p.le_radius_of_summable_nnnorm $ by { simp only [← coe_nnnorm] at h, exact_mod_cast h } lemma radius_eq_top_of_forall_nnreal_is_O - (h : ∀ r : ℝ≥0, (λ n, ∥p n∥ * r^n) =O[at_top] (λ n, (1 : ℝ))) : p.radius = ∞ := + (h : ∀ r : ℝ≥0, (λ n, ‖p n‖ * r^n) =O[at_top] (λ n, (1 : ℝ))) : p.radius = ∞ := ennreal.eq_top_of_forall_nnreal_le $ λ r, p.le_radius_of_is_O (h r) lemma radius_eq_top_of_eventually_eq_zero (h : ∀ᶠ n in at_top, p n = 0) : p.radius = ∞ := @@ -156,45 +156,45 @@ p.radius_eq_top_of_eventually_eq_zero $ mem_at_top_sets.2 (const_formal_multilinear_series 𝕜 E v).radius_eq_top_of_forall_image_add_eq_zero 1 (by simp [const_formal_multilinear_series]) -/-- For `r` strictly smaller than the radius of `p`, then `∥pₙ∥ rⁿ` tends to zero exponentially: -for some `0 < a < 1`, `∥p n∥ rⁿ = o(aⁿ)`. -/ +/-- For `r` strictly smaller than the radius of `p`, then `‖pₙ‖ rⁿ` tends to zero exponentially: +for some `0 < a < 1`, `‖p n‖ rⁿ = o(aⁿ)`. -/ lemma is_o_of_lt_radius (h : ↑r < p.radius) : - ∃ a ∈ Ioo (0 : ℝ) 1, (λ n, ∥p n∥ * r ^ n) =o[at_top] (pow a) := + ∃ a ∈ Ioo (0 : ℝ) 1, (λ n, ‖p n‖ * r ^ n) =o[at_top] (pow a) := begin - rw (tfae_exists_lt_is_o_pow (λ n, ∥p n∥ * r ^ n) 1).out 1 4, + rw (tfae_exists_lt_is_o_pow (λ n, ‖p n‖ * r ^ n) 1).out 1 4, simp only [radius, lt_supr_iff] at h, rcases h with ⟨t, C, hC, rt⟩, rw [ennreal.coe_lt_coe, ← nnreal.coe_lt_coe] at rt, have : 0 < (t : ℝ), from r.coe_nonneg.trans_lt rt, rw [← div_lt_one this] at rt, refine ⟨_, rt, C, or.inr zero_lt_one, λ n, _⟩, - calc |∥p n∥ * r ^ n| = (∥p n∥ * t ^ n) * (r / t) ^ n : + calc |‖p n‖ * r ^ n| = (‖p n‖ * t ^ n) * (r / t) ^ n : by field_simp [mul_right_comm, abs_mul, this.ne'] ... ≤ C * (r / t) ^ n : mul_le_mul_of_nonneg_right (hC n) (pow_nonneg (div_nonneg r.2 t.2) _) end -/-- For `r` strictly smaller than the radius of `p`, then `∥pₙ∥ rⁿ = o(1)`. -/ +/-- For `r` strictly smaller than the radius of `p`, then `‖pₙ‖ rⁿ = o(1)`. -/ lemma is_o_one_of_lt_radius (h : ↑r < p.radius) : - (λ n, ∥p n∥ * r ^ n) =o[at_top] (λ _, 1 : ℕ → ℝ) := + (λ n, ‖p n‖ * r ^ n) =o[at_top] (λ _, 1 : ℕ → ℝ) := let ⟨a, ha, hp⟩ := p.is_o_of_lt_radius h in hp.trans $ (is_o_pow_pow_of_lt_left ha.1.le ha.2).congr (λ n, rfl) one_pow -/-- For `r` strictly smaller than the radius of `p`, then `∥pₙ∥ rⁿ` tends to zero exponentially: -for some `0 < a < 1` and `C > 0`, `∥p n∥ * r ^ n ≤ C * a ^ n`. -/ +/-- For `r` strictly smaller than the radius of `p`, then `‖pₙ‖ rⁿ` tends to zero exponentially: +for some `0 < a < 1` and `C > 0`, `‖p n‖ * r ^ n ≤ C * a ^ n`. -/ lemma norm_mul_pow_le_mul_pow_of_lt_radius (h : ↑r < p.radius) : - ∃ (a ∈ Ioo (0 : ℝ) 1) (C > 0), ∀ n, ∥p n∥ * r^n ≤ C * a^n := + ∃ (a ∈ Ioo (0 : ℝ) 1) (C > 0), ∀ n, ‖p n‖ * r^n ≤ C * a^n := begin - rcases ((tfae_exists_lt_is_o_pow (λ n, ∥p n∥ * r ^ n) 1).out 1 5).mp (p.is_o_of_lt_radius h) + rcases ((tfae_exists_lt_is_o_pow (λ n, ‖p n‖ * r ^ n) 1).out 1 5).mp (p.is_o_of_lt_radius h) with ⟨a, ha, C, hC, H⟩, exact ⟨a, ha, C, hC, λ n, (le_abs_self _).trans (H n)⟩ end -/-- If `r ≠ 0` and `∥pₙ∥ rⁿ = O(aⁿ)` for some `-1 < a < 1`, then `r < p.radius`. -/ +/-- If `r ≠ 0` and `‖pₙ‖ rⁿ = O(aⁿ)` for some `-1 < a < 1`, then `r < p.radius`. -/ lemma lt_radius_of_is_O (h₀ : r ≠ 0) {a : ℝ} (ha : a ∈ Ioo (-1 : ℝ) 1) - (hp : (λ n, ∥p n∥ * r ^ n) =O[at_top] (pow a)) : + (hp : (λ n, ‖p n‖ * r ^ n) =O[at_top] (pow a)) : ↑r < p.radius := begin - rcases ((tfae_exists_lt_is_o_pow (λ n, ∥p n∥ * r ^ n) 1).out 2 5).mp ⟨a, ha, hp⟩ + rcases ((tfae_exists_lt_is_o_pow (λ n, ‖p n‖ * r ^ n) 1).out 2 5).mp ⟨a, ha, hp⟩ with ⟨a, ha, C, hC, hp⟩, rw [← pos_iff_ne_zero, ← nnreal.coe_pos] at h₀, lift a to ℝ≥0 using ha.1.le, @@ -207,39 +207,39 @@ begin exact (le_abs_self _).trans (hp n) end -/-- For `r` strictly smaller than the radius of `p`, then `∥pₙ∥ rⁿ` is bounded. -/ +/-- For `r` strictly smaller than the radius of `p`, then `‖pₙ‖ rⁿ` is bounded. -/ lemma norm_mul_pow_le_of_lt_radius (p : formal_multilinear_series 𝕜 E F) {r : ℝ≥0} - (h : (r : ℝ≥0∞) < p.radius) : ∃ C > 0, ∀ n, ∥p n∥ * r^n ≤ C := + (h : (r : ℝ≥0∞) < p.radius) : ∃ C > 0, ∀ n, ‖p n‖ * r^n ≤ C := let ⟨a, ha, C, hC, h⟩ := p.norm_mul_pow_le_mul_pow_of_lt_radius h in ⟨C, hC, λ n, (h n).trans $ mul_le_of_le_one_right hC.lt.le (pow_le_one _ ha.1.le ha.2.le)⟩ -/-- For `r` strictly smaller than the radius of `p`, then `∥pₙ∥ rⁿ` is bounded. -/ +/-- For `r` strictly smaller than the radius of `p`, then `‖pₙ‖ rⁿ` is bounded. -/ lemma norm_le_div_pow_of_pos_of_lt_radius (p : formal_multilinear_series 𝕜 E F) {r : ℝ≥0} - (h0 : 0 < r) (h : (r : ℝ≥0∞) < p.radius) : ∃ C > 0, ∀ n, ∥p n∥ ≤ C / r ^ n := + (h0 : 0 < r) (h : (r : ℝ≥0∞) < p.radius) : ∃ C > 0, ∀ n, ‖p n‖ ≤ C / r ^ n := let ⟨C, hC, hp⟩ := p.norm_mul_pow_le_of_lt_radius h in ⟨C, hC, λ n, iff.mpr (le_div_iff (pow_pos h0 _)) (hp n)⟩ -/-- For `r` strictly smaller than the radius of `p`, then `∥pₙ∥ rⁿ` is bounded. -/ +/-- For `r` strictly smaller than the radius of `p`, then `‖pₙ‖ rⁿ` is bounded. -/ lemma nnnorm_mul_pow_le_of_lt_radius (p : formal_multilinear_series 𝕜 E F) {r : ℝ≥0} - (h : (r : ℝ≥0∞) < p.radius) : ∃ C > 0, ∀ n, ∥p n∥₊ * r^n ≤ C := + (h : (r : ℝ≥0∞) < p.radius) : ∃ C > 0, ∀ n, ‖p n‖₊ * r^n ≤ C := let ⟨C, hC, hp⟩ := p.norm_mul_pow_le_of_lt_radius h in ⟨⟨C, hC.lt.le⟩, hC, by exact_mod_cast hp⟩ lemma le_radius_of_tendsto (p : formal_multilinear_series 𝕜 E F) {l : ℝ} - (h : tendsto (λ n, ∥p n∥ * r^n) at_top (𝓝 l)) : ↑r ≤ p.radius := + (h : tendsto (λ n, ‖p n‖ * r^n) at_top (𝓝 l)) : ↑r ≤ p.radius := p.le_radius_of_is_O (h.is_O_one _) lemma le_radius_of_summable_norm (p : formal_multilinear_series 𝕜 E F) - (hs : summable (λ n, ∥p n∥ * r^n)) : ↑r ≤ p.radius := + (hs : summable (λ n, ‖p n‖ * r^n)) : ↑r ≤ p.radius := p.le_radius_of_tendsto hs.tendsto_at_top_zero lemma not_summable_norm_of_radius_lt_nnnorm (p : formal_multilinear_series 𝕜 E F) {x : E} - (h : p.radius < ∥x∥₊) : ¬ summable (λ n, ∥p n∥ * ∥x∥^n) := + (h : p.radius < ‖x‖₊) : ¬ summable (λ n, ‖p n‖ * ‖x‖^n) := λ hs, not_le_of_lt h (p.le_radius_of_summable_norm hs) lemma summable_norm_mul_pow (p : formal_multilinear_series 𝕜 E F) {r : ℝ≥0} (h : ↑r < p.radius) : - summable (λ n : ℕ, ∥p n∥ * r ^ n) := + summable (λ n : ℕ, ‖p n‖ * r ^ n) := begin obtain ⟨a, ha : a ∈ Ioo (0 : ℝ) 1, C, hC : 0 < C, hp⟩ := p.norm_mul_pow_le_mul_pow_of_lt_radius h, exact summable_of_nonneg_of_le (λ n, mul_nonneg (norm_nonneg _) (pow_nonneg r.coe_nonneg _)) hp @@ -248,7 +248,7 @@ end lemma summable_norm_apply (p : formal_multilinear_series 𝕜 E F) {x : E} (hx : x ∈ emetric.ball (0 : E) p.radius) : - summable (λ n : ℕ, ∥p n (λ _, x)∥) := + summable (λ n : ℕ, ‖p n (λ _, x)‖) := begin rw mem_emetric_ball_zero_iff at hx, refine summable_of_nonneg_of_le (λ _, norm_nonneg _) (λ n, ((p n).le_op_norm _).trans_eq _) @@ -258,7 +258,7 @@ end lemma summable_nnnorm_mul_pow (p : formal_multilinear_series 𝕜 E F) {r : ℝ≥0} (h : ↑r < p.radius) : - summable (λ n : ℕ, ∥p n∥₊ * r ^ n) := + summable (λ n : ℕ, ‖p n‖₊ * r ^ n) := by { rw ← nnreal.summable_coe, push_cast, exact p.summable_norm_mul_pow h } protected lemma summable [complete_space F] @@ -267,11 +267,11 @@ protected lemma summable [complete_space F] summable_of_summable_norm (p.summable_norm_apply hx) lemma radius_eq_top_of_summable_norm (p : formal_multilinear_series 𝕜 E F) - (hs : ∀ r : ℝ≥0, summable (λ n, ∥p n∥ * r^n)) : p.radius = ∞ := + (hs : ∀ r : ℝ≥0, summable (λ n, ‖p n‖ * r^n)) : p.radius = ∞ := ennreal.eq_top_of_forall_nnreal_le (λ r, p.le_radius_of_summable_norm (hs r)) lemma radius_eq_top_iff_summable_norm (p : formal_multilinear_series 𝕜 E F) : - p.radius = ∞ ↔ ∀ r : ℝ≥0, summable (λ n, ∥p n∥ * r^n) := + p.radius = ∞ ↔ ∀ r : ℝ≥0, summable (λ n, ‖p n‖ * r^n) := begin split, { intros h r, @@ -285,9 +285,9 @@ begin { exact p.radius_eq_top_of_summable_norm } end -/-- If the radius of `p` is positive, then `∥pₙ∥` grows at most geometrically. -/ +/-- If the radius of `p` is positive, then `‖pₙ‖` grows at most geometrically. -/ lemma le_mul_pow_of_radius_pos (p : formal_multilinear_series 𝕜 E F) (h : 0 < p.radius) : - ∃ C r (hC : 0 < C) (hr : 0 < r), ∀ n, ∥p n∥ ≤ C * r ^ n := + ∃ C r (hC : 0 < C) (hr : 0 < r), ∀ n, ‖p n‖ ≤ C * r ^ n := begin rcases ennreal.lt_iff_exists_nnreal_btwn.1 h with ⟨r, r0, rlt⟩, have rpos : 0 < (r : ℝ), by simp [ennreal.coe_pos.1 r0], @@ -324,7 +324,7 @@ begin refine ennreal.le_of_forall_nnreal_lt (λ r hr, _), apply le_radius_of_is_O, apply (is_O.trans_is_o _ (p.is_o_one_of_lt_radius hr)).is_O, - refine is_O.mul (@is_O_with.is_O _ _ _ _ _ (∥f∥) _ _ _ _) (is_O_refl _ _), + refine is_O.mul (@is_O_with.is_O _ _ _ _ _ (‖f‖) _ _ _ _) (is_O_refl _ _), apply is_O_with.of_bound (eventually_of_forall (λ n, _)), simpa only [norm_norm] using f.norm_comp_continuous_multilinear_map_le (p n) end @@ -337,7 +337,7 @@ section variables {f g : E → F} {p pf pg : formal_multilinear_series 𝕜 E F} {x : E} {r r' : ℝ≥0∞} /-- Given a function `f : E → F` and a formal multilinear series `p`, we say that `f` has `p` as -a power series on the ball of radius `r > 0` around `x` if `f (x + y) = ∑' pₙ yⁿ` for all `∥y∥ < r`. +a power series on the ball of radius `r > 0` around `x` if `f (x + y) = ∑' pₙ yⁿ` for all `‖y‖ < r`. -/ structure has_fpower_series_on_ball (f : E → F) (p : formal_multilinear_series 𝕜 E F) (x : E) (r : ℝ≥0∞) : Prop := @@ -570,36 +570,36 @@ end /-- If a function admits a power series expansion, then it is exponentially close to the partial sums of this power series on strict subdisks of the disk of convergence. -This version provides an upper estimate that decreases both in `∥y∥` and `n`. See also +This version provides an upper estimate that decreases both in `‖y‖` and `n`. See also `has_fpower_series_on_ball.uniform_geometric_approx` for a weaker version. -/ lemma has_fpower_series_on_ball.uniform_geometric_approx' {r' : ℝ≥0} (hf : has_fpower_series_on_ball f p x r) (h : (r' : ℝ≥0∞) < r) : ∃ (a ∈ Ioo (0 : ℝ) 1) (C > 0), (∀ y ∈ metric.ball (0 : E) r', ∀ n, - ∥f (x + y) - p.partial_sum n y∥ ≤ C * (a * (∥y∥ / r')) ^ n) := + ‖f (x + y) - p.partial_sum n y‖ ≤ C * (a * (‖y‖ / r')) ^ n) := begin - obtain ⟨a, ha, C, hC, hp⟩ : ∃ (a ∈ Ioo (0 : ℝ) 1) (C > 0), ∀ n, ∥p n∥ * r' ^n ≤ C * a^n := + obtain ⟨a, ha, C, hC, hp⟩ : ∃ (a ∈ Ioo (0 : ℝ) 1) (C > 0), ∀ n, ‖p n‖ * r' ^n ≤ C * a^n := p.norm_mul_pow_le_mul_pow_of_lt_radius (h.trans_le hf.r_le), refine ⟨a, ha, C / (1 - a), div_pos hC (sub_pos.2 ha.2), λ y hy n, _⟩, - have yr' : ∥y∥ < r', by { rw ball_zero_eq at hy, exact hy }, + have yr' : ‖y‖ < r', by { rw ball_zero_eq at hy, exact hy }, have hr'0 : 0 < (r' : ℝ), from (norm_nonneg _).trans_lt yr', have : y ∈ emetric.ball (0 : E) r, { refine mem_emetric_ball_zero_iff.2 (lt_trans _ h), exact_mod_cast yr' }, rw [norm_sub_rev, ← mul_div_right_comm], - have ya : a * (∥y∥ / ↑r') ≤ a, + have ya : a * (‖y‖ / ↑r') ≤ a, from mul_le_of_le_one_right ha.1.le (div_le_one_of_le yr'.le r'.coe_nonneg), - suffices : ∥p.partial_sum n y - f (x + y)∥ ≤ C * (a * (∥y∥ / r')) ^ n / (1 - a * (∥y∥ / r')), + suffices : ‖p.partial_sum n y - f (x + y)‖ ≤ C * (a * (‖y‖ / r')) ^ n / (1 - a * (‖y‖ / r')), { refine this.trans _, apply_rules [div_le_div_of_le_left, sub_pos.2, div_nonneg, mul_nonneg, pow_nonneg, hC.lt.le, ha.1.le, norm_nonneg, nnreal.coe_nonneg, ha.2, (sub_le_sub_iff_left _).2]; apply_instance }, apply norm_sub_le_of_geometric_bound_of_has_sum (ya.trans_lt ha.2) _ (hf.has_sum this), assume n, - calc ∥(p n) (λ (i : fin n), y)∥ ≤ ∥p n∥ * (∏ i : fin n, ∥y∥) : + calc ‖(p n) (λ (i : fin n), y)‖ ≤ ‖p n‖ * (∏ i : fin n, ‖y‖) : continuous_multilinear_map.le_op_norm _ _ - ... = (∥p n∥ * r' ^ n) * (∥y∥ / r') ^ n : by field_simp [hr'0.ne', mul_right_comm] - ... ≤ (C * a ^ n) * (∥y∥ / r') ^ n : + ... = (‖p n‖ * r' ^ n) * (‖y‖ / r') ^ n : by field_simp [hr'0.ne', mul_right_comm] + ... ≤ (C * a ^ n) * (‖y‖ / r') ^ n : mul_le_mul_of_nonneg_right (hp n) (pow_nonneg (div_nonneg (norm_nonneg _) r'.coe_nonneg) _) - ... ≤ C * (a * (∥y∥ / r')) ^ n : by rw [mul_pow, mul_assoc] + ... ≤ C * (a * (‖y‖ / r')) ^ n : by rw [mul_pow, mul_assoc] end /-- If a function admits a power series expansion, then it is exponentially close to the partial @@ -607,13 +607,13 @@ sums of this power series on strict subdisks of the disk of convergence. -/ lemma has_fpower_series_on_ball.uniform_geometric_approx {r' : ℝ≥0} (hf : has_fpower_series_on_ball f p x r) (h : (r' : ℝ≥0∞) < r) : ∃ (a ∈ Ioo (0 : ℝ) 1) (C > 0), (∀ y ∈ metric.ball (0 : E) r', ∀ n, - ∥f (x + y) - p.partial_sum n y∥ ≤ C * a ^ n) := + ‖f (x + y) - p.partial_sum n y‖ ≤ C * a ^ n) := begin obtain ⟨a, ha, C, hC, hp⟩ : ∃ (a ∈ Ioo (0 : ℝ) 1) (C > 0), - (∀ y ∈ metric.ball (0 : E) r', ∀ n, ∥f (x + y) - p.partial_sum n y∥ ≤ C * (a * (∥y∥ / r')) ^ n), + (∀ y ∈ metric.ball (0 : E) r', ∀ n, ‖f (x + y) - p.partial_sum n y‖ ≤ C * (a * (‖y‖ / r')) ^ n), from hf.uniform_geometric_approx' h, refine ⟨a, ha, C, hC, λ y hy n, (hp y hy n).trans _⟩, - have yr' : ∥y∥ < r', by rwa ball_zero_eq at hy, + have yr' : ‖y‖ < r', by rwa ball_zero_eq at hy, refine mul_le_mul_of_nonneg_left (pow_le_pow_of_le_left _ _ _) hC.lt.le, exacts [mul_nonneg ha.1.le (div_nonneg (norm_nonneg y) r'.coe_nonneg), mul_le_of_le_one_right ha.1.le (div_le_one_of_le yr'.le r'.coe_nonneg)] @@ -621,12 +621,12 @@ end /-- Taylor formula for an analytic function, `is_O` version. -/ lemma has_fpower_series_at.is_O_sub_partial_sum_pow (hf : has_fpower_series_at f p x) (n : ℕ) : - (λ y : E, f (x + y) - p.partial_sum n y) =O[𝓝 0] (λ y, ∥y∥ ^ n) := + (λ y : E, f (x + y) - p.partial_sum n y) =O[𝓝 0] (λ y, ‖y‖ ^ n) := begin rcases hf with ⟨r, hf⟩, rcases ennreal.lt_iff_exists_nnreal_btwn.1 hf.r_pos with ⟨r', r'0, h⟩, obtain ⟨a, ha, C, hC, hp⟩ : ∃ (a ∈ Ioo (0 : ℝ) 1) (C > 0), - (∀ y ∈ metric.ball (0 : E) r', ∀ n, ∥f (x + y) - p.partial_sum n y∥ ≤ C * (a * (∥y∥ / r')) ^ n), + (∀ y ∈ metric.ball (0 : E) r', ∀ n, ‖f (x + y) - p.partial_sum n y‖ ≤ C * (a * (‖y‖ / r')) ^ n), from hf.uniform_geometric_approx' h, refine is_O_iff.2 ⟨C * (a / r') ^ n, _⟩, replace r'0 : 0 < (r' : ℝ), by exact_mod_cast r'0, @@ -636,24 +636,24 @@ end /-- If `f` has formal power series `∑ n, pₙ` on a ball of radius `r`, then for `y, z` in any smaller ball, the norm of the difference `f y - f z - p 1 (λ _, y - z)` is bounded above by -`C * (max ∥y - x∥ ∥z - x∥) * ∥y - z∥`. This lemma formulates this property using `is_O` and +`C * (max ‖y - x‖ ‖z - x‖) * ‖y - z‖`. This lemma formulates this property using `is_O` and `filter.principal` on `E × E`. -/ lemma has_fpower_series_on_ball.is_O_image_sub_image_sub_deriv_principal (hf : has_fpower_series_on_ball f p x r) (hr : r' < r) : (λ y : E × E, f y.1 - f y.2 - (p 1 (λ _, y.1 - y.2))) =O[𝓟 (emetric.ball (x, x) r')] - (λ y, ∥y - (x, x)∥ * ∥y.1 - y.2∥) := + (λ y, ‖y - (x, x)‖ * ‖y.1 - y.2‖) := begin lift r' to ℝ≥0 using ne_top_of_lt hr, rcases (zero_le r').eq_or_lt with rfl|hr'0, { simp only [is_O_bot, emetric.ball_zero, principal_empty, ennreal.coe_zero] }, obtain ⟨a, ha, C, hC : 0 < C, hp⟩ : - ∃ (a ∈ Ioo (0 : ℝ) 1) (C > 0), ∀ (n : ℕ), ∥p n∥ * ↑r' ^ n ≤ C * a ^ n, + ∃ (a ∈ Ioo (0 : ℝ) 1) (C > 0), ∀ (n : ℕ), ‖p n‖ * ↑r' ^ n ≤ C * a ^ n, from p.norm_mul_pow_le_mul_pow_of_lt_radius (hr.trans_le hf.r_le), simp only [← le_div_iff (pow_pos (nnreal.coe_pos.2 hr'0) _)] at hp, set L : E × E → ℝ := λ y, - (C * (a / r') ^ 2) * (∥y - (x, x)∥ * ∥y.1 - y.2∥) * (a / (1 - a) ^ 2 + 2 / (1 - a)), + (C * (a / r') ^ 2) * (‖y - (x, x)‖ * ‖y.1 - y.2‖) * (a / (1 - a) ^ 2 + 2 / (1 - a)), have hL : ∀ y ∈ emetric.ball (x, x) r', - ∥f y.1 - f y.2 - (p 1 (λ _, y.1 - y.2))∥ ≤ L y, + ‖f y.1 - f y.2 - (p 1 (λ _, y.1 - y.2))‖ ≤ L y, { intros y hy', have hy : y ∈ emetric.ball x r ×ˢ emetric.ball x r, { rw [emetric.ball_prod_same], exact emetric.ball_subset_ball hr.le hy' }, @@ -666,15 +666,15 @@ begin subsingleton.pi_single_eq, sub_sub_sub_cancel_right] }, rw [emetric.mem_ball, edist_eq_coe_nnnorm_sub, ennreal.coe_lt_coe] at hy', set B : ℕ → ℝ := λ n, - (C * (a / r') ^ 2) * (∥y - (x, x)∥ * ∥y.1 - y.2∥) * ((n + 2) * a ^ n), - have hAB : ∀ n, ∥A (n + 2)∥ ≤ B n := λ n, - calc ∥A (n + 2)∥ ≤ ∥p (n + 2)∥ * ↑(n + 2) * ∥y - (x, x)∥ ^ (n + 1) * ∥y.1 - y.2∥ : + (C * (a / r') ^ 2) * (‖y - (x, x)‖ * ‖y.1 - y.2‖) * ((n + 2) * a ^ n), + have hAB : ∀ n, ‖A (n + 2)‖ ≤ B n := λ n, + calc ‖A (n + 2)‖ ≤ ‖p (n + 2)‖ * ↑(n + 2) * ‖y - (x, x)‖ ^ (n + 1) * ‖y.1 - y.2‖ : by simpa only [fintype.card_fin, pi_norm_const (_ : E), prod.norm_def, pi.sub_def, prod.fst_sub, prod.snd_sub, sub_sub_sub_cancel_right] using (p $ n + 2).norm_image_sub_le (λ _, y.1 - x) (λ _, y.2 - x) - ... = ∥p (n + 2)∥ * ∥y - (x, x)∥ ^ n * (↑(n + 2) * ∥y - (x, x)∥ * ∥y.1 - y.2∥) : - by { rw [pow_succ ∥y - (x, x)∥], ring } - ... ≤ (C * a ^ (n + 2) / r' ^ (n + 2)) * r' ^ n * (↑(n + 2) * ∥y - (x, x)∥ * ∥y.1 - y.2∥) : + ... = ‖p (n + 2)‖ * ‖y - (x, x)‖ ^ n * (↑(n + 2) * ‖y - (x, x)‖ * ‖y.1 - y.2‖) : + by { rw [pow_succ ‖y - (x, x)‖], ring } + ... ≤ (C * a ^ (n + 2) / r' ^ (n + 2)) * r' ^ n * (↑(n + 2) * ‖y - (x, x)‖ * ‖y.1 - y.2‖) : by apply_rules [mul_le_mul_of_nonneg_right, mul_le_mul, hp, pow_le_pow_of_le_left, hy'.le, norm_nonneg, pow_nonneg, div_nonneg, mul_nonneg, nat.cast_nonneg, hC.le, r'.coe_nonneg, ha.1.le] @@ -683,11 +683,11 @@ begin have hBL : has_sum B (L y), { apply has_sum.mul_left, simp only [add_mul], - have : ∥a∥ < 1, by simp only [real.norm_eq_abs, abs_of_pos ha.1, ha.2], + have : ‖a‖ < 1, by simp only [real.norm_eq_abs, abs_of_pos ha.1, ha.2], convert (has_sum_coe_mul_geometric_of_norm_lt_1 this).add ((has_sum_geometric_of_norm_lt_1 this).mul_left 2) }, exact hA.norm_le_of_bounded hBL hAB }, - suffices : L =O[𝓟 (emetric.ball (x, x) r')] (λ y, ∥y - (x, x)∥ * ∥y.1 - y.2∥), + suffices : L =O[𝓟 (emetric.ball (x, x) r')] (λ y, ‖y - (x, x)‖ * ‖y.1 - y.2‖), { refine (is_O.of_bound 1 (eventually_principal.2 $ λ y hy, _)).trans this, rw one_mul, exact (hL y hy).trans (le_abs_self _) }, @@ -697,21 +697,21 @@ end /-- If `f` has formal power series `∑ n, pₙ` on a ball of radius `r`, then for `y, z` in any smaller ball, the norm of the difference `f y - f z - p 1 (λ _, y - z)` is bounded above by -`C * (max ∥y - x∥ ∥z - x∥) * ∥y - z∥`. -/ +`C * (max ‖y - x‖ ‖z - x‖) * ‖y - z‖`. -/ lemma has_fpower_series_on_ball.image_sub_sub_deriv_le (hf : has_fpower_series_on_ball f p x r) (hr : r' < r) : ∃ C, ∀ (y z ∈ emetric.ball x r'), - ∥f y - f z - (p 1 (λ _, y - z))∥ ≤ C * (max ∥y - x∥ ∥z - x∥) * ∥y - z∥ := + ‖f y - f z - (p 1 (λ _, y - z))‖ ≤ C * (max ‖y - x‖ ‖z - x‖) * ‖y - z‖ := by simpa only [is_O_principal, mul_assoc, norm_mul, norm_norm, prod.forall, emetric.mem_ball, prod.edist_eq, max_lt_iff, and_imp, @forall_swap (_ < _) E] using hf.is_O_image_sub_image_sub_deriv_principal hr /-- If `f` has formal power series `∑ n, pₙ` at `x`, then -`f y - f z - p 1 (λ _, y - z) = O(∥(y, z) - (x, x)∥ * ∥y - z∥)` as `(y, z) → (x, x)`. +`f y - f z - p 1 (λ _, y - z) = O(‖(y, z) - (x, x)‖ * ‖y - z‖)` as `(y, z) → (x, x)`. In particular, `f` is strictly differentiable at `x`. -/ lemma has_fpower_series_at.is_O_image_sub_norm_mul_norm_sub (hf : has_fpower_series_at f p x) : (λ y : E × E, f y.1 - f y.2 - (p 1 (λ _, y.1 - y.2))) =O[𝓝 (x, x)] - (λ y, ∥y - (x, x)∥ * ∥y.1 - y.2∥) := + (λ y, ‖y - (x, x)‖ * ‖y.1 - y.2‖) := begin rcases hf with ⟨r, hf⟩, rcases ennreal.lt_iff_exists_nnreal_btwn.1 hf.r_pos with ⟨r', r'0, h⟩, @@ -728,7 +728,7 @@ lemma has_fpower_series_on_ball.tendsto_uniformly_on {r' : ℝ≥0} (λ y, f (x + y)) at_top (metric.ball (0 : E) r') := begin obtain ⟨a, ha, C, hC, hp⟩ : ∃ (a ∈ Ioo (0 : ℝ) 1) (C > 0), - (∀ y ∈ metric.ball (0 : E) r', ∀ n, ∥f (x + y) - p.partial_sum n y∥ ≤ C * a ^ n), + (∀ y ∈ metric.ball (0 : E) r', ∀ n, ‖f (x + y) - p.partial_sum n y‖ ≤ C * a ^ n), from hf.uniform_geometric_approx h, refine metric.tendsto_uniformly_on_iff.2 (λ ε εpos, _), have L : tendsto (λ n, (C : ℝ) * a^n) at_top (𝓝 ((C : ℝ) * 0)) := @@ -838,7 +838,7 @@ section uniqueness open continuous_multilinear_map lemma asymptotics.is_O.continuous_multilinear_map_apply_eq_zero {n : ℕ} {p : E [×n]→L[𝕜] F} - (h : (λ y, p (λ i, y)) =O[𝓝 0] (λ y, ∥y∥ ^ (n + 1))) (y : E) : + (h : (λ y, p (λ i, y)) =O[𝓝 0] (λ y, ‖y‖ ^ (n + 1))) (y : E) : p (λ i, y) = 0 := begin obtain ⟨c, c_pos, hc⟩ := h.exists_pos, @@ -855,25 +855,25 @@ begin have h₀ := mul_pos c_pos (pow_pos hy (n.succ + 1)), obtain ⟨k, k_pos, k_norm⟩ := normed_field.exists_norm_lt 𝕜 (lt_min (mul_pos δ_pos (inv_pos.mpr hy)) (mul_pos ε_pos (inv_pos.mpr h₀))), - have h₁ : ∥k • y∥ < δ, + have h₁ : ‖k • y‖ < δ, { rw norm_smul, exact inv_mul_cancel_right₀ hy.ne.symm δ ▸ mul_lt_mul_of_pos_right (lt_of_lt_of_le k_norm (min_le_left _ _)) hy }, have h₂ := calc - ∥p (λ i, k • y)∥ ≤ c * ∥k • y∥ ^ (n.succ + 1) + ‖p (λ i, k • y)‖ ≤ c * ‖k • y‖ ^ (n.succ + 1) : by simpa only [norm_pow, norm_norm] using ht (k • y) (δε (mem_ball_zero_iff.mpr h₁)) - ... = ∥k∥ ^ n.succ * (∥k∥ * (c * ∥y∥ ^ (n.succ + 1))) + ... = ‖k‖ ^ n.succ * (‖k‖ * (c * ‖y‖ ^ (n.succ + 1))) : by { simp only [norm_smul, mul_pow], rw pow_succ, ring }, - have h₃ : ∥k∥ * (c * ∥y∥ ^ (n.succ + 1)) < ε, from inv_mul_cancel_right₀ h₀.ne.symm ε ▸ + have h₃ : ‖k‖ * (c * ‖y‖ ^ (n.succ + 1)) < ε, from inv_mul_cancel_right₀ h₀.ne.symm ε ▸ mul_lt_mul_of_pos_right (lt_of_lt_of_le k_norm (min_le_right _ _)) h₀, - calc ∥p (λ i, y)∥ = ∥(k⁻¹) ^ n.succ∥ * ∥p (λ i, k • y)∥ + calc ‖p (λ i, y)‖ = ‖(k⁻¹) ^ n.succ‖ * ‖p (λ i, k • y)‖ : by simpa only [inv_smul_smul₀ (norm_pos_iff.mp k_pos), norm_smul, finset.prod_const, finset.card_fin] using congr_arg norm (p.map_smul_univ (λ (i : fin n.succ), k⁻¹) (λ (i : fin n.succ), k • y)) - ... ≤ ∥(k⁻¹) ^ n.succ∥ * (∥k∥ ^ n.succ * (∥k∥ * (c * ∥y∥ ^ (n.succ + 1)))) + ... ≤ ‖(k⁻¹) ^ n.succ‖ * (‖k‖ ^ n.succ * (‖k‖ * (c * ‖y‖ ^ (n.succ + 1)))) : mul_le_mul_of_nonneg_left h₂ (norm_nonneg _) - ... = ∥(k⁻¹ * k) ^ n.succ∥ * (∥k∥ * (c * ∥y∥ ^ (n.succ + 1))) + ... = ‖(k⁻¹ * k) ^ n.succ‖ * (‖k‖ * (c * ‖y‖ ^ (n.succ + 1))) : by { rw ←mul_assoc, simp [norm_mul, mul_pow] } ... ≤ 0 + ε : by { rw inv_mul_cancel (norm_pos_iff.mp k_pos), simpa using h₃.le }, }, @@ -940,7 +940,7 @@ theorem has_fpower_series_on_ball.r_eq_top_of_exists {f : 𝕜 → E} {r : ℝ { r_le := ennreal.le_of_forall_pos_nnreal_lt $ λ r hr hr', let ⟨p', hp'⟩ := h' r hr in (h.exchange_radius hp').r_le, r_pos := ennreal.coe_lt_top, - has_sum := λ y hy, let ⟨r', hr'⟩ := exists_gt ∥y∥₊, ⟨p', hp'⟩ := h' r' hr'.ne_bot.bot_lt + has_sum := λ y hy, let ⟨r', hr'⟩ := exists_gt ‖y‖₊, ⟨p', hp'⟩ := h' r' hr'.ne_bot.bot_lt in (h.exchange_radius hp').has_sum $ mem_emetric_ball_zero_iff.mpr (ennreal.coe_lt_coe.2 hr') } end uniqueness @@ -994,17 +994,17 @@ continuous_multilinear_map.curry_fin_finset_apply_const _ _ _ _ _ @[simp] lemma norm_change_origin_series_term (k l : ℕ) (s : finset (fin (k + l))) (hs : s.card = l) : - ∥p.change_origin_series_term k l s hs∥ = ∥p (k + l)∥ := + ‖p.change_origin_series_term k l s hs‖ = ‖p (k + l)‖ := by simp only [change_origin_series_term, linear_isometry_equiv.norm_map] @[simp] lemma nnnorm_change_origin_series_term (k l : ℕ) (s : finset (fin (k + l))) (hs : s.card = l) : - ∥p.change_origin_series_term k l s hs∥₊ = ∥p (k + l)∥₊ := + ‖p.change_origin_series_term k l s hs‖₊ = ‖p (k + l)‖₊ := by simp only [change_origin_series_term, linear_isometry_equiv.nnnorm_map] lemma nnnorm_change_origin_series_term_apply_le (k l : ℕ) (s : finset (fin (k + l))) (hs : s.card = l) (x y : E) : - ∥p.change_origin_series_term k l s hs (λ _, x) (λ _, y)∥₊ ≤ ∥p (k + l)∥₊ * ∥x∥₊ ^ l * ∥y∥₊ ^ k := + ‖p.change_origin_series_term k l s hs (λ _, x) (λ _, y)‖₊ ≤ ‖p (k + l)‖₊ * ‖x‖₊ ^ l * ‖y‖₊ ^ k := begin rw [← p.nnnorm_change_origin_series_term k l s hs, ← fin.prod_const, ← fin.prod_const], apply continuous_multilinear_map.le_of_op_nnnorm_le, @@ -1021,13 +1021,13 @@ def change_origin_series (k : ℕ) : formal_multilinear_series 𝕜 E (E [×k] λ l, ∑ s : {s : finset (fin (k + l)) // finset.card s = l}, p.change_origin_series_term k l s s.2 lemma nnnorm_change_origin_series_le_tsum (k l : ℕ) : - ∥p.change_origin_series k l∥₊ ≤ - ∑' (x : {s : finset (fin (k + l)) // s.card = l}), ∥p (k + l)∥₊ := + ‖p.change_origin_series k l‖₊ ≤ + ∑' (x : {s : finset (fin (k + l)) // s.card = l}), ‖p (k + l)‖₊ := (nnnorm_sum_le _ _).trans_eq $ by simp only [tsum_fintype, nnnorm_change_origin_series_term] lemma nnnorm_change_origin_series_apply_le_tsum (k l : ℕ) (x : E) : - ∥p.change_origin_series k l (λ _, x)∥₊ ≤ - ∑' s : {s : finset (fin (k + l)) // s.card = l}, ∥p (k + l)∥₊ * ∥x∥₊ ^ l := + ‖p.change_origin_series k l (λ _, x)‖₊ ≤ + ∑' s : {s : finset (fin (k + l)) // s.card = l}, ‖p (k + l)‖₊ * ‖x‖₊ ^ l := begin rw [nnreal.tsum_mul_right, ← fin.prod_const], exact (p.change_origin_series k l).le_of_op_nnnorm_le _ @@ -1077,17 +1077,17 @@ with non-definitional equalities. -/ lemma change_origin_series_summable_aux₁ {r r' : ℝ≥0} (hr : (r + r' : ℝ≥0∞) < p.radius) : summable (λ s : Σ k l : ℕ, {s : finset (fin (k + l)) // s.card = l}, - ∥p (s.1 + s.2.1)∥₊ * r ^ s.2.1 * r' ^ s.1) := + ‖p (s.1 + s.2.1)‖₊ * r ^ s.2.1 * r' ^ s.1) := begin rw ← change_origin_index_equiv.symm.summable_iff, dsimp only [(∘), change_origin_index_equiv_symm_apply_fst, change_origin_index_equiv_symm_apply_snd_fst], have : ∀ n : ℕ, has_sum - (λ s : finset (fin n), ∥p (n - s.card + s.card)∥₊ * r ^ s.card * r' ^ (n - s.card)) - (∥p n∥₊ * (r + r') ^ n), + (λ s : finset (fin n), ‖p (n - s.card + s.card)‖₊ * r ^ s.card * r' ^ (n - s.card)) + (‖p n‖₊ * (r + r') ^ n), { intro n, -- TODO: why `simp only [tsub_add_cancel_of_le (card_finset_fin_le _)]` fails? - convert_to has_sum (λ s : finset (fin n), ∥p n∥₊ * (r ^ s.card * r' ^ (n - s.card))) _, + convert_to has_sum (λ s : finset (fin n), ‖p n‖₊ * (r ^ s.card * r' ^ (n - s.card))) _, { ext1 s, rw [tsub_add_cancel_of_le (card_finset_fin_le _), mul_assoc] }, rw ← fin.sum_pow_mul_eq_add_pow, exact (has_sum_fintype _).mul_left _ }, @@ -1097,7 +1097,7 @@ begin end lemma change_origin_series_summable_aux₂ (hr : (r : ℝ≥0∞) < p.radius) (k : ℕ) : - summable (λ s : Σ l : ℕ, {s : finset (fin (k + l)) // s.card = l}, ∥p (k + s.1)∥₊ * r ^ s.1) := + summable (λ s : Σ l : ℕ, {s : finset (fin (k + l)) // s.card = l}, ‖p (k + s.1)‖₊ * r ^ s.1) := begin rcases ennreal.lt_iff_exists_add_pos_lt.1 hr with ⟨r', h0, hr'⟩, simpa only [mul_inv_cancel_right₀ (pow_pos h0 _).ne'] @@ -1106,7 +1106,7 @@ begin end lemma change_origin_series_summable_aux₃ {r : ℝ≥0} (hr : ↑r < p.radius) (k : ℕ) : - summable (λ l : ℕ, ∥p.change_origin_series k l∥₊ * r ^ l) := + summable (λ l : ℕ, ‖p.change_origin_series k l‖₊ * r ^ l) := begin refine nnreal.summable_of_le (λ n, _) (nnreal.summable_sigma.1 $ p.change_origin_series_summable_aux₂ hr k).2, @@ -1119,9 +1119,9 @@ lemma le_change_origin_series_radius (k : ℕ) : ennreal.le_of_forall_nnreal_lt $ λ r hr, le_radius_of_summable_nnnorm _ (p.change_origin_series_summable_aux₃ hr k) -lemma nnnorm_change_origin_le (k : ℕ) (h : (∥x∥₊ : ℝ≥0∞) < p.radius) : - ∥p.change_origin x k∥₊ ≤ - ∑' s : Σ l : ℕ, {s : finset (fin (k + l)) // s.card = l}, ∥p (k + s.1)∥₊ * ∥x∥₊ ^ s.1 := +lemma nnnorm_change_origin_le (k : ℕ) (h : (‖x‖₊ : ℝ≥0∞) < p.radius) : + ‖p.change_origin x k‖₊ ≤ + ∑' s : Σ l : ℕ, {s : finset (fin (k + l)) // s.card = l}, ‖p (k + s.1)‖₊ * ‖x‖₊ ^ s.1 := begin refine tsum_of_nnnorm_bounded _ (λ l, p.nnnorm_change_origin_series_apply_le_tsum k l x), have := p.change_origin_series_summable_aux₂ h k, @@ -1129,17 +1129,17 @@ begin exact ((nnreal.summable_sigma.1 this).1 l).has_sum end -/-- The radius of convergence of `p.change_origin x` is at least `p.radius - ∥x∥`. In other words, +/-- The radius of convergence of `p.change_origin x` is at least `p.radius - ‖x‖`. In other words, `p.change_origin x` is well defined on the largest ball contained in the original ball of convergence.-/ -lemma change_origin_radius : p.radius - ∥x∥₊ ≤ (p.change_origin x).radius := +lemma change_origin_radius : p.radius - ‖x‖₊ ≤ (p.change_origin x).radius := begin refine ennreal.le_of_forall_pos_nnreal_lt (λ r h0 hr, _), rw [lt_tsub_iff_right, add_comm] at hr, - have hr' : (∥x∥₊ : ℝ≥0∞) < p.radius, from (le_add_right le_rfl).trans_lt hr, + have hr' : (‖x‖₊ : ℝ≥0∞) < p.radius, from (le_add_right le_rfl).trans_lt hr, apply le_radius_of_summable_nnnorm, - have : ∀ k : ℕ, ∥p.change_origin x k∥₊ * r ^ k ≤ - (∑' s : Σ l : ℕ, {s : finset (fin (k + l)) // s.card = l}, ∥p (k + s.1)∥₊ * ∥x∥₊ ^ s.1) * r ^ k, + have : ∀ k : ℕ, ‖p.change_origin x k‖₊ * r ^ k ≤ + (∑' s : Σ l : ℕ, {s : finset (fin (k + l)) // s.card = l}, ‖p (k + s.1)‖₊ * ‖x‖₊ ^ s.1) * r ^ k, from λ k, mul_le_mul_right' (p.nnnorm_change_origin_le k hr') (r ^ k), refine nnreal.summable_of_le this _, simpa only [← nnreal.tsum_mul_right] @@ -1158,7 +1158,7 @@ have _ := p.le_change_origin_series_radius k, ((p.change_origin_series k).has_fpower_series_on_ball (hr.trans_le this)).mono hr this /-- Summing the series `p.change_origin x` at a point `y` gives back `p (x + y)`-/ -theorem change_origin_eval (h : (∥x∥₊ + ∥y∥₊ : ℝ≥0∞) < p.radius) : +theorem change_origin_eval (h : (‖x‖₊ + ‖y‖₊ : ℝ≥0∞) < p.radius) : (p.change_origin x).sum y = (p.sum (x + y)) := begin have radius_pos : 0 < p.radius := lt_of_le_of_lt (zero_le _) h, @@ -1215,8 +1215,8 @@ variables [complete_space F] {f : E → F} {p : formal_multilinear_series 𝕜 E power series on any subball of this ball (even with a different center), given by `p.change_origin`. -/ theorem has_fpower_series_on_ball.change_origin - (hf : has_fpower_series_on_ball f p x r) (h : (∥y∥₊ : ℝ≥0∞) < r) : - has_fpower_series_on_ball f (p.change_origin y) (x + y) (r - ∥y∥₊) := + (hf : has_fpower_series_on_ball f p x r) (h : (‖y‖₊ : ℝ≥0∞) < r) : + has_fpower_series_on_ball f (p.change_origin y) (x + y) (r - ‖y‖₊) := { r_le := begin apply le_trans _ p.change_origin_radius, exact tsub_le_tsub hf.r_le le_rfl @@ -1238,7 +1238,7 @@ lemma has_fpower_series_on_ball.analytic_at_of_mem (hf : has_fpower_series_on_ball f p x r) (h : y ∈ emetric.ball x r) : analytic_at 𝕜 f y := begin - have : (∥y - x∥₊ : ℝ≥0∞) < r, by simpa [edist_eq_coe_nnnorm_sub] using h, + have : (‖y - x‖₊ : ℝ≥0∞) < r, by simpa [edist_eq_coe_nnnorm_sub] using h, have := hf.change_origin this, rw [add_sub_cancel'_right] at this, exact this.analytic_at @@ -1280,7 +1280,7 @@ begin refine ⟨p.radius ⊓ r.to_nnreal, by simp, _, _⟩, { simp only [r_pos.lt, lt_inf_iff, ennreal.coe_pos, real.to_nnreal_pos, and_true], obtain ⟨z, z_pos, le_z⟩ := normed_field.exists_norm_lt 𝕜 r_pos.lt, - have : (∥z∥₊ : ennreal) ≤ p.radius, + have : (‖z‖₊ : ennreal) ≤ p.radius, by { simp only [dist_zero_right] at h, apply formal_multilinear_series.le_radius_of_tendsto, convert tendsto_norm.comp (h le_z).summable.tendsto_at_top_zero, diff --git a/src/analysis/analytic/composition.lean b/src/analysis/analytic/composition.lean index 413bf60dcf096..eeeea5db90346 100644 --- a/src/analysis/analytic/composition.lean +++ b/src/analysis/analytic/composition.lean @@ -302,21 +302,21 @@ the norms of the relevant bits of `f` and `p`. -/ lemma comp_along_composition_bound {n : ℕ} (p : formal_multilinear_series 𝕜 E F) (c : composition n) (f : continuous_multilinear_map 𝕜 (λ (i : fin c.length), F) G) (v : fin n → E) : - ∥f.comp_along_composition p c v∥ ≤ - ∥f∥ * (∏ i, ∥p (c.blocks_fun i)∥) * (∏ i : fin n, ∥v i∥) := -calc ∥f.comp_along_composition p c v∥ = ∥f (p.apply_composition c v)∥ : rfl -... ≤ ∥f∥ * ∏ i, ∥p.apply_composition c v i∥ : continuous_multilinear_map.le_op_norm _ _ -... ≤ ∥f∥ * ∏ i, ∥p (c.blocks_fun i)∥ * - ∏ j : fin (c.blocks_fun i), ∥(v ∘ (c.embedding i)) j∥ : + ‖f.comp_along_composition p c v‖ ≤ + ‖f‖ * (∏ i, ‖p (c.blocks_fun i)‖) * (∏ i : fin n, ‖v i‖) := +calc ‖f.comp_along_composition p c v‖ = ‖f (p.apply_composition c v)‖ : rfl +... ≤ ‖f‖ * ∏ i, ‖p.apply_composition c v i‖ : continuous_multilinear_map.le_op_norm _ _ +... ≤ ‖f‖ * ∏ i, ‖p (c.blocks_fun i)‖ * + ∏ j : fin (c.blocks_fun i), ‖(v ∘ (c.embedding i)) j‖ : begin apply mul_le_mul_of_nonneg_left _ (norm_nonneg _), refine finset.prod_le_prod (λ i hi, norm_nonneg _) (λ i hi, _), apply continuous_multilinear_map.le_op_norm, end -... = ∥f∥ * (∏ i, ∥p (c.blocks_fun i)∥) * - ∏ i (j : fin (c.blocks_fun i)), ∥(v ∘ (c.embedding i)) j∥ : +... = ‖f‖ * (∏ i, ‖p (c.blocks_fun i)‖) * + ∏ i (j : fin (c.blocks_fun i)), ‖(v ∘ (c.embedding i)) j‖ : by rw [finset.prod_mul_distrib, mul_assoc] -... = ∥f∥ * (∏ i, ∥p (c.blocks_fun i)∥) * (∏ i : fin n, ∥v i∥) : +... = ‖f‖ * (∏ i, ‖p (c.blocks_fun i)‖) * (∏ i : fin n, ‖v i‖) : by { rw [← c.blocks_fin_equiv.prod_comp, ← finset.univ_sigma_univ, finset.prod_sigma], congr } @@ -325,7 +325,7 @@ the norms of the relevant bits of `q` and `p`. -/ lemma comp_along_composition_norm {n : ℕ} (q : formal_multilinear_series 𝕜 F G) (p : formal_multilinear_series 𝕜 E F) (c : composition n) : - ∥q.comp_along_composition p c∥ ≤ ∥q c.length∥ * ∏ i, ∥p (c.blocks_fun i)∥ := + ‖q.comp_along_composition p c‖ ≤ ‖q c.length‖ * ∏ i, ‖p (c.blocks_fun i)‖ := continuous_multilinear_map.op_norm_le_bound _ (mul_nonneg (norm_nonneg _) (finset.prod_nonneg (λ i hi, norm_nonneg _))) (comp_along_composition_bound _ _ _) @@ -333,7 +333,7 @@ continuous_multilinear_map.op_norm_le_bound _ lemma comp_along_composition_nnnorm {n : ℕ} (q : formal_multilinear_series 𝕜 F G) (p : formal_multilinear_series 𝕜 E F) (c : composition n) : - ∥q.comp_along_composition p c∥₊ ≤ ∥q c.length∥₊ * ∏ i, ∥p (c.blocks_fun i)∥₊ := + ‖q.comp_along_composition p c‖₊ ≤ ‖q c.length‖₊ * ∏ i, ‖p (c.blocks_fun i)‖₊ := by { rw ← nnreal.coe_le_coe, push_cast, exact q.comp_along_composition_norm p c } /-! @@ -439,17 +439,17 @@ theorem comp_summable_nnreal (q : formal_multilinear_series 𝕜 F G) (p : formal_multilinear_series 𝕜 E F) (hq : 0 < q.radius) (hp : 0 < p.radius) : ∃ r > (0 : ℝ≥0), - summable (λ i : Σ n, composition n, ∥q.comp_along_composition p i.2∥₊ * r ^ i.1) := + summable (λ i : Σ n, composition n, ‖q.comp_along_composition p i.2‖₊ * r ^ i.1) := begin - /- This follows from the fact that the growth rate of `∥qₙ∥` and `∥pₙ∥` is at most geometric, - giving a geometric bound on each `∥q.comp_along_composition p op∥`, together with the + /- This follows from the fact that the growth rate of `‖qₙ‖` and `‖pₙ‖` is at most geometric, + giving a geometric bound on each `‖q.comp_along_composition p op‖`, together with the fact that there are `2^(n-1)` compositions of `n`, giving at most a geometric loss. -/ rcases ennreal.lt_iff_exists_nnreal_btwn.1 (lt_min ennreal.zero_lt_one hq) with ⟨rq, rq_pos, hrq⟩, rcases ennreal.lt_iff_exists_nnreal_btwn.1 (lt_min ennreal.zero_lt_one hp) with ⟨rp, rp_pos, hrp⟩, simp only [lt_min_iff, ennreal.coe_lt_one_iff, ennreal.coe_pos] at hrp hrq rp_pos rq_pos, - obtain ⟨Cq, hCq0, hCq⟩ : ∃ Cq > 0, ∀ n, ∥q n∥₊ * rq^n ≤ Cq := + obtain ⟨Cq, hCq0, hCq⟩ : ∃ Cq > 0, ∀ n, ‖q n‖₊ * rq^n ≤ Cq := q.nnnorm_mul_pow_le_of_lt_radius hrq.2, - obtain ⟨Cp, hCp1, hCp⟩ : ∃ Cp ≥ 1, ∀ n, ∥p n∥₊ * rp^n ≤ Cp, + obtain ⟨Cp, hCp1, hCp⟩ : ∃ Cp ≥ 1, ∀ n, ‖p n‖₊ * rp^n ≤ Cp, { rcases p.nnnorm_mul_pow_le_of_lt_radius hrp.2 with ⟨Cp, -, hCp⟩, exact ⟨max Cp 1, le_max_right _ _, λ n, (hCp n).trans (le_max_left _ _)⟩ }, let r0 : ℝ≥0 := (4 * Cp)⁻¹, @@ -457,23 +457,23 @@ begin set r : ℝ≥0 := rp * rq * r0, have r_pos : 0 < r := mul_pos (mul_pos rp_pos rq_pos) r0_pos, have I : ∀ (i : Σ (n : ℕ), composition n), - ∥q.comp_along_composition p i.2∥₊ * r ^ i.1 ≤ Cq / 4 ^ i.1, + ‖q.comp_along_composition p i.2‖₊ * r ^ i.1 ≤ Cq / 4 ^ i.1, { rintros ⟨n, c⟩, have A, - calc ∥q c.length∥₊ * rq ^ n ≤ ∥q c.length∥₊* rq ^ c.length : + calc ‖q c.length‖₊ * rq ^ n ≤ ‖q c.length‖₊* rq ^ c.length : mul_le_mul' le_rfl (pow_le_pow_of_le_one rq.2 hrq.1.le c.length_le) ... ≤ Cq : hCq _, have B, - calc ((∏ i, ∥p (c.blocks_fun i)∥₊) * rp ^ n) - = ∏ i, ∥p (c.blocks_fun i)∥₊ * rp ^ c.blocks_fun i : + calc ((∏ i, ‖p (c.blocks_fun i)‖₊) * rp ^ n) + = ∏ i, ‖p (c.blocks_fun i)‖₊ * rp ^ c.blocks_fun i : by simp only [finset.prod_mul_distrib, finset.prod_pow_eq_pow_sum, c.sum_blocks_fun] ... ≤ ∏ i : fin c.length, Cp : finset.prod_le_prod' (λ i _, hCp _) ... = Cp ^ c.length : by simp ... ≤ Cp ^ n : pow_le_pow hCp1 c.length_le, - calc ∥q.comp_along_composition p c∥₊ * r ^ n - ≤ (∥q c.length∥₊ * ∏ i, ∥p (c.blocks_fun i)∥₊) * r ^ n : + calc ‖q.comp_along_composition p c‖₊ * r ^ n + ≤ (‖q c.length‖₊ * ∏ i, ‖p (c.blocks_fun i)‖₊) * r ^ n : mul_le_mul' (q.comp_along_composition_nnnorm p c) le_rfl - ... = (∥q c.length∥₊ * rq ^ n) * ((∏ i, ∥p (c.blocks_fun i)∥₊) * rp ^ n) * r0 ^ n : + ... = (‖q c.length‖₊ * rq ^ n) * ((∏ i, ‖p (c.blocks_fun i)‖₊) * rp ^ n) * r0 ^ n : by { simp only [r, mul_pow], ring } ... ≤ Cq * Cp ^ n * r0 ^ n : mul_le_mul' (mul_le_mul' A B) le_rfl ... = Cq / 4 ^ n : @@ -503,18 +503,18 @@ end summability over all compositions. -/ theorem le_comp_radius_of_summable (q : formal_multilinear_series 𝕜 F G) (p : formal_multilinear_series 𝕜 E F) (r : ℝ≥0) - (hr : summable (λ i : (Σ n, composition n), ∥q.comp_along_composition p i.2∥₊ * r ^ i.1)) : + (hr : summable (λ i : (Σ n, composition n), ‖q.comp_along_composition p i.2‖₊ * r ^ i.1)) : (r : ℝ≥0∞) ≤ (q.comp p).radius := begin refine le_radius_of_bound_nnreal _ - (∑' i : (Σ n, composition n), ∥comp_along_composition q p i.snd∥₊ * r ^ i.fst) (λ n, _), - calc ∥formal_multilinear_series.comp q p n∥₊ * r ^ n ≤ - ∑' (c : composition n), ∥comp_along_composition q p c∥₊ * r ^ n : + (∑' i : (Σ n, composition n), ‖comp_along_composition q p i.snd‖₊ * r ^ i.fst) (λ n, _), + calc ‖formal_multilinear_series.comp q p n‖₊ * r ^ n ≤ + ∑' (c : composition n), ‖comp_along_composition q p c‖₊ * r ^ n : begin rw [tsum_fintype, ← finset.sum_mul], exact mul_le_mul' (nnnorm_sum_le _ _) le_rfl end - ... ≤ ∑' (i : Σ (n : ℕ), composition n), ∥comp_along_composition q p i.snd∥₊ * r ^ i.fst : + ... ≤ ∑' (i : Σ (n : ℕ), composition n), ‖comp_along_composition q p i.snd‖₊ * r ^ i.fst : nnreal.tsum_comp_le_tsum_of_inj hr sigma_mk_injective end @@ -733,7 +733,7 @@ begin refine ⟨min rf' r, _⟩, refine ⟨le_trans (min_le_right rf' r) (formal_multilinear_series.le_comp_radius_of_summable q p r hr), min_pos, λ y hy, _⟩, - /- Let `y` satisfy `∥y∥ < min (r, rf', δ)`. We want to show that `g (f (x + y))` is the sum of + /- Let `y` satisfy `‖y‖ < min (r, rf', δ)`. We want to show that `g (f (x + y))` is the sum of `q.comp p` applied to `y`. -/ -- First, check that `y` is small enough so that estimates for `f` and `g` apply. have y_mem : y ∈ emetric.ball (0 : E) rf := @@ -793,10 +793,10 @@ begin { apply cauchy_seq_finset_of_norm_bounded _ (nnreal.summable_coe.2 hr) _, simp only [coe_nnnorm, nnreal.coe_mul, nnreal.coe_pow], rintros ⟨n, c⟩, - calc ∥(comp_along_composition q p c) (λ (j : fin n), y)∥ - ≤ ∥comp_along_composition q p c∥ * ∏ j : fin n, ∥y∥ : + calc ‖(comp_along_composition q p c) (λ (j : fin n), y)‖ + ≤ ‖comp_along_composition q p c‖ * ∏ j : fin n, ‖y‖ : by apply continuous_multilinear_map.le_op_norm - ... ≤ ∥comp_along_composition q p c∥ * (r : ℝ) ^ n : + ... ≤ ‖comp_along_composition q p c‖ * (r : ℝ) ^ n : begin apply mul_le_mul_of_nonneg_left _ (norm_nonneg _), rw [finset.prod_const, finset.card_fin], diff --git a/src/analysis/analytic/inverse.lean b/src/analysis/analytic/inverse.lean index b621409195275..91a9a365f5f3a 100644 --- a/src/analysis/analytic/inverse.lean +++ b/src/analysis/analytic/inverse.lean @@ -289,7 +289,7 @@ $$ Here, `q_{n-1}` can only appear in the term with `k = 2`, and it only appears twice, so there is hope this formula can lead to an at most geometric behavior. -Let `Qₙ = ∥qₙ∥`. Bounding `∥pₖ∥` with `C r^k` gives an inequality +Let `Qₙ = ‖qₙ‖`. Bounding `‖pₖ‖` with `C r^k` gives an inequality $$ Q_n ≤ C' \sum_{k=2}^n r^k \sum_{i_1 + \dotsc + i_k = n} Q_{i_1} \dotsm Q_{i_k}. $$ @@ -407,27 +407,27 @@ in the specific setup we are interesting in, by reducing to the general bound in `radius_right_inv_pos_of_radius_pos_aux1`. -/ lemma radius_right_inv_pos_of_radius_pos_aux2 {n : ℕ} (hn : 2 ≤ n + 1) (p : formal_multilinear_series 𝕜 E F) (i : E ≃L[𝕜] F) - {r a C : ℝ} (hr : 0 ≤ r) (ha : 0 ≤ a) (hC : 0 ≤ C) (hp : ∀ n, ∥p n∥ ≤ C * r ^ n) : - (∑ k in Ico 1 (n + 1), a ^ k * ∥p.right_inv i k∥) ≤ - ∥(i.symm : F →L[𝕜] E)∥ * a + ∥(i.symm : F →L[𝕜] E)∥ * C * ∑ k in Ico 2 (n + 1), - (r * ((∑ j in Ico 1 n, a ^ j * ∥p.right_inv i j∥))) ^ k := -let I := ∥(i.symm : F →L[𝕜] E)∥ in calc -∑ k in Ico 1 (n + 1), a ^ k * ∥p.right_inv i k∥ - = a * I + ∑ k in Ico 2 (n + 1), a ^ k * ∥p.right_inv i k∥ : + {r a C : ℝ} (hr : 0 ≤ r) (ha : 0 ≤ a) (hC : 0 ≤ C) (hp : ∀ n, ‖p n‖ ≤ C * r ^ n) : + (∑ k in Ico 1 (n + 1), a ^ k * ‖p.right_inv i k‖) ≤ + ‖(i.symm : F →L[𝕜] E)‖ * a + ‖(i.symm : F →L[𝕜] E)‖ * C * ∑ k in Ico 2 (n + 1), + (r * ((∑ j in Ico 1 n, a ^ j * ‖p.right_inv i j‖))) ^ k := +let I := ‖(i.symm : F →L[𝕜] E)‖ in calc +∑ k in Ico 1 (n + 1), a ^ k * ‖p.right_inv i k‖ + = a * I + ∑ k in Ico 2 (n + 1), a ^ k * ‖p.right_inv i k‖ : by simp only [linear_isometry_equiv.norm_map, pow_one, right_inv_coeff_one, nat.Ico_succ_singleton, sum_singleton, ← sum_Ico_consecutive _ one_le_two hn] ... = a * I + ∑ k in Ico 2 (n + 1), a ^ k * - ∥(i.symm : F →L[𝕜] E).comp_continuous_multilinear_map + ‖(i.symm : F →L[𝕜] E).comp_continuous_multilinear_map (∑ c in ({c | 1 < composition.length c}.to_finset : finset (composition k)), - p.comp_along_composition (p.right_inv i) c)∥ : + p.comp_along_composition (p.right_inv i) c)‖ : begin congr' 1, apply sum_congr rfl (λ j hj, _), rw [right_inv_coeff _ _ _ (mem_Ico.1 hj).1, norm_neg], end -... ≤ a * ∥(i.symm : F →L[𝕜] E)∥ + ∑ k in Ico 2 (n + 1), a ^ k * (I * +... ≤ a * ‖(i.symm : F →L[𝕜] E)‖ + ∑ k in Ico 2 (n + 1), a ^ k * (I * (∑ c in ({c | 1 < composition.length c}.to_finset : finset (composition k)), - C * r ^ c.length * ∏ j, ∥p.right_inv i (c.blocks_fun j)∥)) : + C * r ^ c.length * ∏ j, ‖p.right_inv i (c.blocks_fun j)‖)) : begin apply_rules [add_le_add, le_refl, sum_le_sum (λ j hj, _), mul_le_mul_of_nonneg_left, pow_nonneg, ha], @@ -441,17 +441,17 @@ begin end ... = I * a + I * C * ∑ k in Ico 2 (n + 1), a ^ k * (∑ c in ({c | 1 < composition.length c}.to_finset : finset (composition k)), - r ^ c.length * ∏ j, ∥p.right_inv i (c.blocks_fun j)∥) : + r ^ c.length * ∏ j, ‖p.right_inv i (c.blocks_fun j)‖) : begin - simp_rw [mul_assoc C, ← mul_sum, ← mul_assoc, mul_comm _ (∥↑i.symm∥), mul_assoc, ← mul_sum, + simp_rw [mul_assoc C, ← mul_sum, ← mul_assoc, mul_comm _ (‖↑i.symm‖), mul_assoc, ← mul_sum, ← mul_assoc, mul_comm _ C, mul_assoc, ← mul_sum], ring, end -... ≤ I * a + I * C * ∑ k in Ico 2 (n+1), (r * ((∑ j in Ico 1 n, a ^ j * ∥p.right_inv i j∥))) ^ k : +... ≤ I * a + I * C * ∑ k in Ico 2 (n+1), (r * ((∑ j in Ico 1 n, a ^ j * ‖p.right_inv i j‖))) ^ k : begin apply_rules [add_le_add, le_refl, mul_le_mul_of_nonneg_left, norm_nonneg, hC, mul_nonneg], simp_rw [mul_pow], - apply radius_right_inv_pos_of_radius_pos_aux1 n (λ k, ∥p.right_inv i k∥) + apply radius_right_inv_pos_of_radius_pos_aux1 n (λ k, ‖p.right_inv i k‖) (λ k, norm_nonneg _) hr ha, end @@ -460,9 +460,9 @@ also has a positive radius of convergence. -/ theorem radius_right_inv_pos_of_radius_pos (p : formal_multilinear_series 𝕜 E F) (i : E ≃L[𝕜] F) (hp : 0 < p.radius) : 0 < (p.right_inv i).radius := begin - obtain ⟨C, r, Cpos, rpos, ple⟩ : ∃ C r (hC : 0 < C) (hr : 0 < r), ∀ (n : ℕ), ∥p n∥ ≤ C * r ^ n := + obtain ⟨C, r, Cpos, rpos, ple⟩ : ∃ C r (hC : 0 < C) (hr : 0 < r), ∀ (n : ℕ), ‖p n‖ ≤ C * r ^ n := le_mul_pow_of_radius_pos p hp, - let I := ∥(i.symm : F →L[𝕜] E)∥, + let I := ‖(i.symm : F →L[𝕜] E)‖, -- choose `a` small enough to make sure that `∑_{k ≤ n} aᵏ Qₖ` will be controllable by -- induction obtain ⟨a, apos, ha1, ha2⟩ : ∃ a (apos : 0 < a), @@ -481,7 +481,7 @@ begin exact ⟨a, ha.1, ha.2.1.le, ha.2.2.le⟩ }, -- check by induction that the partial sums are suitably bounded, using the choice of `a` and the -- inductive control from Lemma `radius_right_inv_pos_of_radius_pos_aux2`. - let S := λ n, ∑ k in Ico 1 n, a ^ k * ∥p.right_inv i k∥, + let S := λ n, ∑ k in Ico 1 n, a ^ k * ‖p.right_inv i k‖, have IRec : ∀ n, 1 ≤ n → S n ≤ (I + 1) * a, { apply nat.le_induction, { simp only [S], @@ -519,14 +519,14 @@ begin by { apply lt_of_lt_of_le _ H, exact_mod_cast apos }, apply le_radius_of_bound _ ((I + 1) * a) (λ n, _), by_cases hn : n = 0, - { have : ∥p.right_inv i n∥ = ∥p.right_inv i 0∥, by congr; try { rw hn }, + { have : ‖p.right_inv i n‖ = ‖p.right_inv i 0‖, by congr; try { rw hn }, simp only [this, norm_zero, zero_mul, right_inv_coeff_zero], apply_rules [mul_nonneg, add_nonneg, norm_nonneg, zero_le_one, apos.le] }, { have one_le_n : 1 ≤ n := bot_lt_iff_ne_bot.2 hn, - calc ∥p.right_inv i n∥ * ↑a' ^ n = a ^ n * ∥p.right_inv i n∥ : mul_comm _ _ - ... ≤ ∑ k in Ico 1 (n + 1), a ^ k * ∥p.right_inv i k∥ : + calc ‖p.right_inv i n‖ * ↑a' ^ n = a ^ n * ‖p.right_inv i n‖ : mul_comm _ _ + ... ≤ ∑ k in Ico 1 (n + 1), a ^ k * ‖p.right_inv i k‖ : begin - have : ∀ k ∈ Ico 1 (n + 1), 0 ≤ a ^ k * ∥p.right_inv i k∥ := + have : ∀ k ∈ Ico 1 (n + 1), 0 ≤ a ^ k * ‖p.right_inv i k‖ := λ k hk, mul_nonneg (pow_nonneg apos.le _) (norm_nonneg _), exact single_le_sum this (by simp [one_le_n]), end diff --git a/src/analysis/analytic/radius_liminf.lean b/src/analysis/analytic/radius_liminf.lean index d3f1220e00408..f706addfdb95b 100644 --- a/src/analysis/analytic/radius_liminf.lean +++ b/src/analysis/analytic/radius_liminf.lean @@ -10,7 +10,7 @@ import analysis.special_functions.pow # Representation of `formal_multilinear_series.radius` as a `liminf` In this file we prove that the radius of convergence of a `formal_multilinear_series` is equal to -$\liminf_{n\to\infty} \frac{1}{\sqrt[n]{∥p n∥}}$. This lemma can't go to `basic.lean` because this +$\liminf_{n\to\infty} \frac{1}{\sqrt[n]{‖p n‖}}$. This lemma can't go to `basic.lean` because this would create a circular dependency once we redefine `exp` using `formal_multilinear_series`. -/ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] @@ -25,12 +25,12 @@ namespace formal_multilinear_series variables (p : formal_multilinear_series 𝕜 E F) /-- The radius of a formal multilinear series is equal to -$\liminf_{n\to\infty} \frac{1}{\sqrt[n]{∥p n∥}}$. The actual statement uses `ℝ≥0` and some +$\liminf_{n\to\infty} \frac{1}{\sqrt[n]{‖p n‖}}$. The actual statement uses `ℝ≥0` and some coercions. -/ -lemma radius_eq_liminf : p.radius = liminf (λ n, 1/((∥p n∥₊) ^ (1 / (n : ℝ)) : ℝ≥0)) at_top := +lemma radius_eq_liminf : p.radius = liminf (λ n, 1/((‖p n‖₊) ^ (1 / (n : ℝ)) : ℝ≥0)) at_top := begin have : ∀ (r : ℝ≥0) {n : ℕ}, 0 < n → - ((r : ℝ≥0∞) ≤ 1 / ↑(∥p n∥₊ ^ (1 / (n : ℝ))) ↔ ∥p n∥₊ * r ^ n ≤ 1), + ((r : ℝ≥0∞) ≤ 1 / ↑(‖p n‖₊ ^ (1 / (n : ℝ))) ↔ ‖p n‖₊ * r ^ n ≤ 1), { intros r n hn, have : 0 < (n : ℝ) := nat.cast_pos.2 hn, conv_lhs {rw [one_div, ennreal.le_inv_iff_mul_le, ← ennreal.coe_mul, @@ -38,7 +38,7 @@ begin nnreal.rpow_mul, ← nnreal.mul_rpow, ← nnreal.one_rpow (n⁻¹), nnreal.rpow_le_rpow_iff (inv_pos.2 this), mul_comm, nnreal.rpow_nat_cast] } }, apply le_antisymm; refine ennreal.le_of_forall_nnreal_lt (λ r hr, _), - { rcases ((tfae_exists_lt_is_o_pow (λ n, ∥p n∥ * r ^ n) 1).out 1 7).1 (p.is_o_of_lt_radius hr) + { rcases ((tfae_exists_lt_is_o_pow (λ n, ‖p n‖ * r ^ n) 1).out 1 7).1 (p.is_o_of_lt_radius hr) with ⟨a, ha, H⟩, refine le_Liminf_of_le (by apply_auto_param) (eventually_map.2 $ _), refine H.mp ((eventually_gt_at_top 0).mono $ λ n hn₀ hn, (this _ hn₀).2 diff --git a/src/analysis/analytic/uniqueness.lean b/src/analysis/analytic/uniqueness.lean index dce9486e217ba..65596baccb9b3 100644 --- a/src/analysis/analytic/uniqueness.lean +++ b/src/analysis/analytic/uniqueness.lean @@ -50,7 +50,7 @@ begin from emetric.mem_closure_iff.1 xu (r / 2) (ennreal.half_pos hp.r_pos.ne'), let q := p.change_origin (y - x), have has_series : has_fpower_series_on_ball f q y (r / 2), - { have A : (∥y - x∥₊ : ℝ≥0∞) < r / 2, by rwa [edist_comm, edist_eq_coe_nnnorm_sub] at hxy, + { have A : (‖y - x‖₊ : ℝ≥0∞) < r / 2, by rwa [edist_comm, edist_eq_coe_nnnorm_sub] at hxy, have := hp.change_origin (A.trans_le ennreal.half_le_self), simp only [add_sub_cancel'_right] at this, apply this.mono (ennreal.half_pos hp.r_pos.ne'), diff --git a/src/analysis/asymptotics/asymptotic_equivalent.lean b/src/analysis/asymptotics/asymptotic_equivalent.lean index ee0e721fa7c94..e0dc4078fc0e5 100644 --- a/src/analysis/asymptotics/asymptotic_equivalent.lean +++ b/src/analysis/asymptotics/asymptotic_equivalent.lean @@ -243,18 +243,18 @@ begin refine hφ.mp (huv.mp $ hCuv.mono $ λ x hCuvx huvx hφx, _), have key := - calc ∥φ x - 1∥ * ∥u x∥ - ≤ (c/2) / C * ∥u x∥ : mul_le_mul_of_nonneg_right hφx.le (norm_nonneg $ u x) - ... ≤ (c/2) / C * (C*∥v x∥) : mul_le_mul_of_nonneg_left hCuvx (div_pos (by linarith) hC).le - ... = c/2 * ∥v x∥ : by {field_simp [hC.ne.symm], ring}, - - calc ∥((λ (x : α), φ x • u x) - v) x∥ - = ∥(φ x - 1) • u x + (u x - v x)∥ : by simp [sub_smul, sub_add] - ... ≤ ∥(φ x - 1) • u x∥ + ∥u x - v x∥ : norm_add_le _ _ - ... = ∥φ x - 1∥ * ∥u x∥ + ∥u x - v x∥ : by rw norm_smul - ... ≤ c / 2 * ∥v x∥ + ∥u x - v x∥ : add_le_add_right key _ - ... ≤ c / 2 * ∥v x∥ + c / 2 * ∥v x∥ : add_le_add_left huvx _ - ... = c * ∥v x∥ : by ring, + calc ‖φ x - 1‖ * ‖u x‖ + ≤ (c/2) / C * ‖u x‖ : mul_le_mul_of_nonneg_right hφx.le (norm_nonneg $ u x) + ... ≤ (c/2) / C * (C*‖v x‖) : mul_le_mul_of_nonneg_left hCuvx (div_pos (by linarith) hC).le + ... = c/2 * ‖v x‖ : by {field_simp [hC.ne.symm], ring}, + + calc ‖((λ (x : α), φ x • u x) - v) x‖ + = ‖(φ x - 1) • u x + (u x - v x)‖ : by simp [sub_smul, sub_add] + ... ≤ ‖(φ x - 1) • u x‖ + ‖u x - v x‖ : norm_add_le _ _ + ... = ‖φ x - 1‖ * ‖u x‖ + ‖u x - v x‖ : by rw norm_smul + ... ≤ c / 2 * ‖v x‖ + ‖u x - v x‖ : add_le_add_right key _ + ... ≤ c / 2 * ‖v x‖ + c / 2 * ‖v x‖ : add_le_add_left huvx _ + ... = c * ‖v x‖ : by ring, end end smul diff --git a/src/analysis/asymptotics/asymptotics.lean b/src/analysis/asymptotics/asymptotics.lean index c37de7c3d5ee9..d3af26c25320c 100644 --- a/src/analysis/asymptotics/asymptotics.lean +++ b/src/analysis/asymptotics/asymptotics.lean @@ -27,7 +27,7 @@ instead. Often the ranges of `f` and `g` will be the real numbers, in which case the norm is the absolute value. In general, we have - `f =O[l] g ↔ (λ x, ∥f x∥) =O[l] (λ x, ∥g x∥)`, + `f =O[l] g ↔ (λ x, ‖f x‖) =O[l] (λ x, ‖g x‖)`, and similarly for `is_o`. But our setup allows us to use the notions e.g. with functions to the integers, rationals, complex numbers, or any normed vector space without mentioning the @@ -68,21 +68,21 @@ section defs /-! ### Definitions -/ /-- This version of the Landau notation `is_O_with C l f g` where `f` and `g` are two functions on -a type `α` and `l` is a filter on `α`, means that eventually for `l`, `∥f∥` is bounded by `C * ∥g∥`. -In other words, `∥f∥ / ∥g∥` is eventually bounded by `C`, modulo division by zero issues that are +a type `α` and `l` is a filter on `α`, means that eventually for `l`, `‖f‖` is bounded by `C * ‖g‖`. +In other words, `‖f‖ / ‖g‖` is eventually bounded by `C`, modulo division by zero issues that are avoided by this definition. Probably you want to use `is_O` instead of this relation. -/ @[irreducible] def is_O_with (c : ℝ) (l : filter α) (f : α → E) (g : α → F) : Prop := -∀ᶠ x in l, ∥ f x ∥ ≤ c * ∥ g x ∥ +∀ᶠ x in l, ‖ f x ‖ ≤ c * ‖ g x ‖ /-- Definition of `is_O_with`. We record it in a lemma as `is_O_with` is irreducible. -/ -lemma is_O_with_iff : is_O_with c l f g ↔ ∀ᶠ x in l, ∥f x∥ ≤ c * ∥g x∥ := by rw is_O_with +lemma is_O_with_iff : is_O_with c l f g ↔ ∀ᶠ x in l, ‖f x‖ ≤ c * ‖g x‖ := by rw is_O_with alias is_O_with_iff ↔ is_O_with.bound is_O_with.of_bound /-- The Landau notation `f =O[l] g` where `f` and `g` are two functions on a type `α` and `l` is -a filter on `α`, means that eventually for `l`, `∥f∥` is bounded by a constant multiple of `∥g∥`. -In other words, `∥f∥ / ∥g∥` is eventually bounded, modulo division by zero issues that are avoided +a filter on `α`, means that eventually for `l`, `‖f‖` is bounded by a constant multiple of `‖g‖`. +In other words, `‖f‖ / ‖g‖` is eventually bounded, modulo division by zero issues that are avoided by this definition. -/ @[irreducible] def is_O (l : filter α) (f : α → E) (g : α → F) : Prop := ∃ c : ℝ, is_O_with c l f g @@ -95,20 +95,20 @@ lemma is_O_iff_is_O_with : f =O[l] g ↔ ∃ c : ℝ, is_O_with c l f g := by rw /-- Definition of `is_O` in terms of filters. We record it in a lemma as we will set `is_O` to be irreducible at the end of this file. -/ -lemma is_O_iff : f =O[l] g ↔ ∃ c : ℝ, ∀ᶠ x in l, ∥f x∥ ≤ c * ∥g x∥ := +lemma is_O_iff : f =O[l] g ↔ ∃ c : ℝ, ∀ᶠ x in l, ‖f x‖ ≤ c * ‖g x‖ := by simp only [is_O, is_O_with] -lemma is_O.of_bound (c : ℝ) (h : ∀ᶠ x in l, ∥f x∥ ≤ c * ∥g x∥) : f =O[l] g := +lemma is_O.of_bound (c : ℝ) (h : ∀ᶠ x in l, ‖f x‖ ≤ c * ‖g x‖) : f =O[l] g := is_O_iff.2 ⟨c, h⟩ -lemma is_O.of_bound' (h : ∀ᶠ x in l, ∥f x∥ ≤ ∥g x∥) : f =O[l] g := +lemma is_O.of_bound' (h : ∀ᶠ x in l, ‖f x‖ ≤ ‖g x‖) : f =O[l] g := is_O.of_bound 1 $ by { simp_rw one_mul, exact h } -lemma is_O.bound : f =O[l] g → ∃ c : ℝ, ∀ᶠ x in l, ∥f x∥ ≤ c * ∥g x∥ := is_O_iff.1 +lemma is_O.bound : f =O[l] g → ∃ c : ℝ, ∀ᶠ x in l, ‖f x‖ ≤ c * ‖g x‖ := is_O_iff.1 /-- The Landau notation `f =o[l] g` where `f` and `g` are two functions on a type `α` and `l` is -a filter on `α`, means that eventually for `l`, `∥f∥` is bounded by an arbitrarily small constant -multiple of `∥g∥`. In other words, `∥f∥ / ∥g∥` tends to `0` along `l`, modulo division by zero +a filter on `α`, means that eventually for `l`, `‖f‖` is bounded by an arbitrarily small constant +multiple of `‖g‖`. In other words, `‖f‖ / ‖g‖` tends to `0` along `l`, modulo division by zero issues that are avoided by this definition. -/ @[irreducible] def is_o (l : filter α) (f : α → E) (g : α → F) : Prop := ∀ ⦃c : ℝ⦄, 0 < c → is_O_with c l f g @@ -123,12 +123,12 @@ alias is_o_iff_forall_is_O_with ↔ is_o.forall_is_O_with is_o.of_is_O_with /-- Definition of `is_o` in terms of filters. We record it in a lemma as we will set `is_o` to be irreducible at the end of this file. -/ -lemma is_o_iff : f =o[l] g ↔ ∀ ⦃c : ℝ⦄, 0 < c → ∀ᶠ x in l, ∥f x∥ ≤ c * ∥g x∥ := +lemma is_o_iff : f =o[l] g ↔ ∀ ⦃c : ℝ⦄, 0 < c → ∀ᶠ x in l, ‖f x‖ ≤ c * ‖g x‖ := by simp only [is_o, is_O_with] alias is_o_iff ↔ is_o.bound is_o.of_bound -lemma is_o.def (h : f =o[l] g) (hc : 0 < c) : ∀ᶠ x in l, ∥f x∥ ≤ c * ∥g x∥ := +lemma is_o.def (h : f =o[l] g) (hc : 0 < c) : ∀ᶠ x in l, ‖f x‖ ≤ c * ‖g x‖ := is_o_iff.1 h hc lemma is_o.def' (h : f =o[l] g) (hc : 0 < c) : is_O_with c l f g := @@ -148,7 +148,7 @@ lemma is_O.is_O_with : f =O[l] g → ∃ c : ℝ, is_O_with c l f g := is_O_iff_ theorem is_O_with.weaken (h : is_O_with c l f g') (hc : c ≤ c') : is_O_with c' l f g' := is_O_with.of_bound $ mem_of_superset h.bound $ λ x hx, -calc ∥f x∥ ≤ c * ∥g' x∥ : hx +calc ‖f x‖ ≤ c * ‖g' x‖ : hx ... ≤ _ : mul_le_mul_of_nonneg_right hc (norm_nonneg _) theorem is_O_with.exists_pos (h : is_O_with c l f g') : @@ -171,22 +171,22 @@ lemma is_O_iff_eventually_is_O_with : f =O[l] g' ↔ ∀ᶠ c in at_top, is_O_wi is_O_iff_is_O_with.trans ⟨λ ⟨c, hc⟩, mem_at_top_sets.2 ⟨c, λ c' hc', hc.weaken hc'⟩, λ h, h.exists⟩ -/-- `f = O(g)` if and only if `∀ᶠ x in l, ∥f x∥ ≤ c * ∥g x∥` for all sufficiently large `c`. -/ -lemma is_O_iff_eventually : f =O[l] g' ↔ ∀ᶠ c in at_top, ∀ᶠ x in l, ∥f x∥ ≤ c * ∥g' x∥ := +/-- `f = O(g)` if and only if `∀ᶠ x in l, ‖f x‖ ≤ c * ‖g x‖` for all sufficiently large `c`. -/ +lemma is_O_iff_eventually : f =O[l] g' ↔ ∀ᶠ c in at_top, ∀ᶠ x in l, ‖f x‖ ≤ c * ‖g' x‖ := is_O_iff_eventually_is_O_with.trans $ by simp only [is_O_with] lemma is_O.exists_mem_basis {ι} {p : ι → Prop} {s : ι → set α} (h : f =O[l] g') (hb : l.has_basis p s) : - ∃ (c : ℝ) (hc : 0 < c) (i : ι) (hi : p i), ∀ x ∈ s i, ∥f x∥ ≤ c * ∥g' x∥ := + ∃ (c : ℝ) (hc : 0 < c) (i : ι) (hi : p i), ∀ x ∈ s i, ‖f x‖ ≤ c * ‖g' x‖ := flip Exists₂.imp h.exists_pos $ λ c hc h, by simpa only [is_O_with_iff, hb.eventually_iff, exists_prop] using h -lemma is_O_with_inv (hc : 0 < c) : is_O_with c⁻¹ l f g ↔ ∀ᶠ x in l, c * ∥f x∥ ≤ ∥g x∥ := +lemma is_O_with_inv (hc : 0 < c) : is_O_with c⁻¹ l f g ↔ ∀ᶠ x in l, c * ‖f x‖ ≤ ‖g x‖ := by simp only [is_O_with, ← div_eq_inv_mul, le_div_iff' hc] -- We prove this lemma with strange assumptions to get two lemmas below automatically -lemma is_o_iff_nat_mul_le_aux (h₀ : (∀ x, 0 ≤ ∥f x∥) ∨ ∀ x, 0 ≤ ∥g x∥) : - f =o[l] g ↔ ∀ n : ℕ, ∀ᶠ x in l, ↑n * ∥f x∥ ≤ ∥g x∥ := +lemma is_o_iff_nat_mul_le_aux (h₀ : (∀ x, 0 ≤ ‖f x‖) ∨ ∀ x, 0 ≤ ‖g x‖) : + f =o[l] g ↔ ∀ n : ℕ, ∀ᶠ x in l, ↑n * ‖f x‖ ≤ ‖g x‖ := begin split, { rintro H (_|n), @@ -205,10 +205,10 @@ begin exact inv_pos.2 hn₀ } end -lemma is_o_iff_nat_mul_le : f =o[l] g' ↔ ∀ n : ℕ, ∀ᶠ x in l, ↑n * ∥f x∥ ≤ ∥g' x∥ := +lemma is_o_iff_nat_mul_le : f =o[l] g' ↔ ∀ n : ℕ, ∀ᶠ x in l, ↑n * ‖f x‖ ≤ ‖g' x‖ := is_o_iff_nat_mul_le_aux (or.inr $ λ x, norm_nonneg _) -lemma is_o_iff_nat_mul_le' : f' =o[l] g ↔ ∀ n : ℕ, ∀ᶠ x in l, ↑n * ∥f' x∥ ≤ ∥g x∥ := +lemma is_o_iff_nat_mul_le' : f' =o[l] g ↔ ∀ n : ℕ, ∀ᶠ x in l, ↑n * ‖f' x‖ ≤ ‖g x‖ := is_o_iff_nat_mul_le_aux (or.inl $ λ x, norm_nonneg _) /-! ### Subsingleton -/ @@ -343,9 +343,9 @@ theorem is_O_with.trans (hfg : is_O_with c l f g) (hgk : is_O_with c' l g k) (hc begin unfold is_O_with at *, filter_upwards [hfg, hgk] with x hx hx', - calc ∥f x∥ ≤ c * ∥g x∥ : hx - ... ≤ c * (c' * ∥k x∥) : mul_le_mul_of_nonneg_left hx' hc - ... = c * c' * ∥k x∥ : (mul_assoc _ _ _).symm + calc ‖f x‖ ≤ c * ‖g x‖ : hx + ... ≤ c * (c' * ‖k x‖) : mul_le_mul_of_nonneg_left hx' hc + ... = c * c' * ‖k x‖ : (mul_assoc _ _ _).symm end @[trans] theorem is_O.trans {f : α → E} {g : α → F'} {k : α → G} (hfg : f =O[l] g) @@ -386,27 +386,27 @@ let ⟨c, cpos, hc⟩ := hfg.exists_pos in hc.trans_is_o hgk cpos hfg.trans_is_O_with hgk.is_O_with one_pos lemma _root_.filter.eventually.trans_is_O {f : α → E} {g : α → F'} {k : α → G} - (hfg : ∀ᶠ x in l, ∥f x∥ ≤ ∥g x∥) (hgk : g =O[l] k) : f =O[l] k := + (hfg : ∀ᶠ x in l, ‖f x‖ ≤ ‖g x‖) (hgk : g =O[l] k) : f =O[l] k := (is_O.of_bound' hfg).trans hgk lemma _root_.filter.eventually.is_O {f : α → E} {g : α → ℝ} {l : filter α} - (hfg : ∀ᶠ x in l, ∥f x∥ ≤ g x) : f =O[l] g := + (hfg : ∀ᶠ x in l, ‖f x‖ ≤ g x) : f =O[l] g := is_O.of_bound' $ hfg.mono $ λ x hx, hx.trans $ real.le_norm_self _ section variable (l) -theorem is_O_with_of_le' (hfg : ∀ x, ∥f x∥ ≤ c * ∥g x∥) : is_O_with c l f g := +theorem is_O_with_of_le' (hfg : ∀ x, ‖f x‖ ≤ c * ‖g x‖) : is_O_with c l f g := is_O_with.of_bound $ univ_mem' hfg -theorem is_O_with_of_le (hfg : ∀ x, ∥f x∥ ≤ ∥g x∥) : is_O_with 1 l f g := +theorem is_O_with_of_le (hfg : ∀ x, ‖f x‖ ≤ ‖g x‖) : is_O_with 1 l f g := is_O_with_of_le' l $ λ x, by { rw one_mul, exact hfg x } -theorem is_O_of_le' (hfg : ∀ x, ∥f x∥ ≤ c * ∥g x∥) : f =O[l] g := +theorem is_O_of_le' (hfg : ∀ x, ‖f x‖ ≤ c * ‖g x‖) : f =O[l] g := (is_O_with_of_le' l hfg).is_O -theorem is_O_of_le (hfg : ∀ x, ∥f x∥ ≤ ∥g x∥) : f =O[l] g := +theorem is_O_of_le (hfg : ∀ x, ‖f x‖ ≤ ‖g x‖) : f =O[l] g := (is_O_with_of_le l hfg).is_O end @@ -416,17 +416,17 @@ is_O_with_of_le l $ λ _, le_rfl theorem is_O_refl (f : α → E) (l : filter α) : f =O[l] f := (is_O_with_refl f l).is_O -theorem is_O_with.trans_le (hfg : is_O_with c l f g) (hgk : ∀ x, ∥g x∥ ≤ ∥k x∥) (hc : 0 ≤ c) : +theorem is_O_with.trans_le (hfg : is_O_with c l f g) (hgk : ∀ x, ‖g x‖ ≤ ‖k x‖) (hc : 0 ≤ c) : is_O_with c l f k := (hfg.trans (is_O_with_of_le l hgk) hc).congr_const $ mul_one c -theorem is_O.trans_le (hfg : f =O[l] g') (hgk : ∀ x, ∥g' x∥ ≤ ∥k x∥) : f =O[l] k := +theorem is_O.trans_le (hfg : f =O[l] g') (hgk : ∀ x, ‖g' x‖ ≤ ‖k x‖) : f =O[l] k := hfg.trans (is_O_of_le l hgk) -theorem is_o.trans_le (hfg : f =o[l] g) (hgk : ∀ x, ∥g x∥ ≤ ∥k x∥) : f =o[l] k := +theorem is_o.trans_le (hfg : f =o[l] g) (hgk : ∀ x, ‖g x‖ ≤ ‖k x‖) : f =o[l] k := hfg.trans_is_O_with (is_O_with_of_le _ hgk) zero_lt_one -theorem is_o_irrefl' (h : ∃ᶠ x in l, ∥f' x∥ ≠ 0) : ¬f' =o[l] f' := +theorem is_o_irrefl' (h : ∃ᶠ x in l, ‖f' x‖ ≠ 0) : ¬f' =o[l] f' := begin intro ho, rcases ((ho.bound one_half_pos).and_frequently h).exists with ⟨x, hle, hne⟩, @@ -455,7 +455,7 @@ variables (c f g) end bot -@[simp] theorem is_O_with_pure {x} : is_O_with c (pure x) f g ↔ ∥f x∥ ≤ c * ∥g x∥ := is_O_with_iff +@[simp] theorem is_O_with_pure {x} : is_O_with c (pure x) f g ↔ ‖f x‖ ≤ c * ‖g x‖ := is_O_with_iff theorem is_O_with.sup (h : is_O_with c l f g) (h' : is_O_with c l' f g) : is_O_with c (l ⊔ l') f g := @@ -479,12 +479,12 @@ is_o.of_is_O_with $ λ c cpos, (h.forall_is_O_with cpos).sup (h'.forall_is_O_wit ⟨λ h, ⟨h.mono le_sup_left, h.mono le_sup_right⟩, λ h, h.1.sup h.2⟩ lemma is_O_with_insert [topological_space α] {x : α} {s : set α} {C : ℝ} {g : α → E} {g' : α → F} - (h : ∥g x∥ ≤ C * ∥g' x∥) : + (h : ‖g x‖ ≤ C * ‖g' x‖) : is_O_with C (𝓝[insert x s] x) g g' ↔ is_O_with C (𝓝[s] x) g g' := by simp_rw [is_O_with, nhds_within_insert, eventually_sup, eventually_pure, h, true_and] lemma is_O_with.insert [topological_space α] {x : α} {s : set α} {C : ℝ} {g : α → E} {g' : α → F} - (h1 : is_O_with C (𝓝[s] x) g g') (h2 : ∥g x∥ ≤ C * ∥g' x∥) : + (h1 : is_O_with C (𝓝[s] x) g g') (h2 : ‖g x‖ ≤ C * ‖g' x‖) : is_O_with C (𝓝[insert x s] x) g g' := (is_O_with_insert h2).mpr h1 @@ -508,7 +508,7 @@ section norm_abs variables {u v : α → ℝ} -@[simp] theorem is_O_with_norm_right : is_O_with c l f (λ x, ∥g' x∥) ↔ is_O_with c l f g' := +@[simp] theorem is_O_with_norm_right : is_O_with c l f (λ x, ‖g' x‖) ↔ is_O_with c l f g' := by simp only [is_O_with, norm_norm] @[simp] theorem is_O_with_abs_right : is_O_with c l f (λ x, |u x|) ↔ is_O_with c l f u := @@ -517,7 +517,7 @@ by simp only [is_O_with, norm_norm] alias is_O_with_norm_right ↔ is_O_with.of_norm_right is_O_with.norm_right alias is_O_with_abs_right ↔ is_O_with.of_abs_right is_O_with.abs_right -@[simp] theorem is_O_norm_right : f =O[l] (λ x, ∥g' x∥) ↔ f =O[l] g' := +@[simp] theorem is_O_norm_right : f =O[l] (λ x, ‖g' x‖) ↔ f =O[l] g' := by { unfold is_O, exact exists_congr (λ _, is_O_with_norm_right) } @[simp] theorem is_O_abs_right : f =O[l] (λ x, |u x|) ↔ f =O[l] u := @@ -526,7 +526,7 @@ by { unfold is_O, exact exists_congr (λ _, is_O_with_norm_right) } alias is_O_norm_right ↔ is_O.of_norm_right is_O.norm_right alias is_O_abs_right ↔ is_O.of_abs_right is_O.abs_right -@[simp] theorem is_o_norm_right : f =o[l] (λ x, ∥g' x∥) ↔ f =o[l] g' := +@[simp] theorem is_o_norm_right : f =o[l] (λ x, ‖g' x‖) ↔ f =o[l] g' := by { unfold is_o, exact forall₂_congr (λ _ _, is_O_with_norm_right) } @[simp] theorem is_o_abs_right : f =o[l] (λ x, |u x|) ↔ f =o[l] u := @@ -535,7 +535,7 @@ by { unfold is_o, exact forall₂_congr (λ _ _, is_O_with_norm_right) } alias is_o_norm_right ↔ is_o.of_norm_right is_o.norm_right alias is_o_abs_right ↔ is_o.of_abs_right is_o.abs_right -@[simp] theorem is_O_with_norm_left : is_O_with c l (λ x, ∥f' x∥) g ↔ is_O_with c l f' g := +@[simp] theorem is_O_with_norm_left : is_O_with c l (λ x, ‖f' x‖) g ↔ is_O_with c l f' g := by simp only [is_O_with, norm_norm] @[simp] theorem is_O_with_abs_left : is_O_with c l (λ x, |u x|) g ↔ is_O_with c l u g := @@ -544,7 +544,7 @@ by simp only [is_O_with, norm_norm] alias is_O_with_norm_left ↔ is_O_with.of_norm_left is_O_with.norm_left alias is_O_with_abs_left ↔ is_O_with.of_abs_left is_O_with.abs_left -@[simp] theorem is_O_norm_left : (λ x, ∥f' x∥) =O[l] g ↔ f' =O[l] g := +@[simp] theorem is_O_norm_left : (λ x, ‖f' x‖) =O[l] g ↔ f' =O[l] g := by { unfold is_O, exact exists_congr (λ _, is_O_with_norm_left) } @[simp] theorem is_O_abs_left : (λ x, |u x|) =O[l] g ↔ u =O[l] g := @@ -553,7 +553,7 @@ by { unfold is_O, exact exists_congr (λ _, is_O_with_norm_left) } alias is_O_norm_left ↔ is_O.of_norm_left is_O.norm_left alias is_O_abs_left ↔ is_O.of_abs_left is_O.abs_left -@[simp] theorem is_o_norm_left : (λ x, ∥f' x∥) =o[l] g ↔ f' =o[l] g := +@[simp] theorem is_o_norm_left : (λ x, ‖f' x‖) =o[l] g ↔ f' =o[l] g := by { unfold is_o, exact forall₂_congr (λ _ _, is_O_with_norm_left) } @[simp] theorem is_o_abs_left : (λ x, |u x|) =o[l] g ↔ u =o[l] g := @@ -562,7 +562,7 @@ by { unfold is_o, exact forall₂_congr (λ _ _, is_O_with_norm_left) } alias is_o_norm_left ↔ is_o.of_norm_left is_o.norm_left alias is_o_abs_left ↔ is_o.of_abs_left is_o.abs_left -theorem is_O_with_norm_norm : is_O_with c l (λ x, ∥f' x∥) (λ x, ∥g' x∥) ↔ is_O_with c l f' g' := +theorem is_O_with_norm_norm : is_O_with c l (λ x, ‖f' x‖) (λ x, ‖g' x‖) ↔ is_O_with c l f' g' := is_O_with_norm_left.trans is_O_with_norm_right theorem is_O_with_abs_abs : is_O_with c l (λ x, |u x|) (λ x, |v x|) ↔ is_O_with c l u v := @@ -571,7 +571,7 @@ is_O_with_abs_left.trans is_O_with_abs_right alias is_O_with_norm_norm ↔ is_O_with.of_norm_norm is_O_with.norm_norm alias is_O_with_abs_abs ↔ is_O_with.of_abs_abs is_O_with.abs_abs -theorem is_O_norm_norm : (λ x, ∥f' x∥) =O[l] (λ x, ∥g' x∥) ↔ f' =O[l] g' := +theorem is_O_norm_norm : (λ x, ‖f' x‖) =O[l] (λ x, ‖g' x‖) ↔ f' =O[l] g' := is_O_norm_left.trans is_O_norm_right theorem is_O_abs_abs : (λ x, |u x|) =O[l] (λ x, |v x|) ↔ u =O[l] v := @@ -580,7 +580,7 @@ is_O_abs_left.trans is_O_abs_right alias is_O_norm_norm ↔ is_O.of_norm_norm is_O.norm_norm alias is_O_abs_abs ↔ is_O.of_abs_abs is_O.abs_abs -theorem is_o_norm_norm : (λ x, ∥f' x∥) =o[l] (λ x, ∥g' x∥) ↔ f' =o[l] g' := +theorem is_o_norm_norm : (λ x, ‖f' x‖) =o[l] (λ x, ‖g' x‖) ↔ f' =o[l] g' := is_o_norm_left.trans is_o_norm_right theorem is_o_abs_abs : (λ x, |u x|) =o[l] (λ x, |v x|) ↔ u =o[l] v := @@ -723,8 +723,8 @@ variables {f₁ f₂ : α → E'} {g₁ g₂ : α → F'} theorem is_O_with.add (h₁ : is_O_with c₁ l f₁ g) (h₂ : is_O_with c₂ l f₂ g) : is_O_with (c₁ + c₂) l (λ x, f₁ x + f₂ x) g := by rw is_O_with at *; filter_upwards [h₁, h₂] with x hx₁ hx₂ using -calc ∥f₁ x + f₂ x∥ ≤ c₁ * ∥g x∥ + c₂ * ∥g x∥ : norm_add_le_of_le hx₁ hx₂ - ... = (c₁ + c₂) * ∥g x∥ : (add_mul _ _ _).symm +calc ‖f₁ x + f₂ x‖ ≤ c₁ * ‖g x‖ + c₂ * ‖g x‖ : norm_add_le_of_le hx₁ hx₂ + ... = (c₁ + c₂) * ‖g x‖ : (add_mul _ _ _).symm theorem is_O.add (h₁ : f₁ =O[l] g) (h₂ : f₂ =O[l] g) : (λ x, f₁ x + f₂ x) =O[l] g := let ⟨c₁, hc₁⟩ := h₁.is_O_with, ⟨c₂, hc₂⟩ := h₂.is_O_with in (hc₁.add hc₂).is_O @@ -734,7 +734,7 @@ is_o.of_is_O_with $ λ c cpos, ((h₁.forall_is_O_with $ half_pos cpos).add (h₂.forall_is_O_with $ half_pos cpos)).congr_const (add_halves c) theorem is_o.add_add (h₁ : f₁ =o[l] g₁) (h₂ : f₂ =o[l] g₂) : - (λ x, f₁ x + f₂ x) =o[l] (λ x, ∥g₁ x∥ + ∥g₂ x∥) := + (λ x, f₁ x + f₂ x) =o[l] (λ x, ‖g₁ x‖ + ‖g₂ x‖) := by refine (h₁.trans_le $ λ x, _).add (h₂.trans_le _); simp [abs_of_nonneg, add_nonneg] @@ -858,7 +858,7 @@ by simp only [is_O_with, exists_prop, true_and, norm_zero, mul_zero, norm_le_zer ⟨λ h, is_O_zero_right_iff.1 h.is_O, λ h, is_o.of_is_O_with $ λ c hc, is_O_with_zero_right_iff.2 h⟩ theorem is_O_with_const_const (c : E) {c' : F''} (hc' : c' ≠ 0) (l : filter α) : - is_O_with (∥c∥ / ∥c'∥) l (λ x : α, c) (λ x, c') := + is_O_with (‖c‖ / ‖c'‖) l (λ x : α, c) (λ x, c') := begin unfold is_O_with, apply univ_mem', @@ -885,33 +885,33 @@ calc f'' =O[pure x] g'' ↔ (λ y : α, f'' x) =O[pure x] (λ _, g'' x) : is_O_c end zero_const -@[simp] lemma is_O_with_top : is_O_with c ⊤ f g ↔ ∀ x, ∥f x∥ ≤ c * ∥g x∥ := by rw is_O_with; refl +@[simp] lemma is_O_with_top : is_O_with c ⊤ f g ↔ ∀ x, ‖f x‖ ≤ c * ‖g x‖ := by rw is_O_with; refl -@[simp] lemma is_O_top : f =O[⊤] g ↔ ∃ C, ∀ x, ∥f x∥ ≤ C * ∥g x∥ := by rw is_O_iff; refl +@[simp] lemma is_O_top : f =O[⊤] g ↔ ∃ C, ∀ x, ‖f x‖ ≤ C * ‖g x‖ := by rw is_O_iff; refl @[simp] lemma is_o_top : f'' =o[⊤] g'' ↔ ∀ x, f'' x = 0 := begin refine ⟨_, λ h, (is_o_zero g'' ⊤).congr (λ x, (h x).symm) (λ x, rfl)⟩, simp only [is_o_iff, eventually_top], refine λ h x, norm_le_zero_iff.1 _, - have : tendsto (λ c : ℝ, c * ∥g'' x∥) (𝓝[>] 0) (𝓝 0) := + have : tendsto (λ c : ℝ, c * ‖g'' x‖) (𝓝[>] 0) (𝓝 0) := ((continuous_id.mul continuous_const).tendsto' _ _ (zero_mul _)).mono_left inf_le_left, exact le_of_tendsto_of_tendsto tendsto_const_nhds this (eventually_nhds_within_iff.2 $ eventually_of_forall $ λ c hc, h hc x) end @[simp] lemma is_O_with_principal {s : set α} : - is_O_with c (𝓟 s) f g ↔ ∀ x ∈ s, ∥f x∥ ≤ c * ∥g x∥ := + is_O_with c (𝓟 s) f g ↔ ∀ x ∈ s, ‖f x‖ ≤ c * ‖g x‖ := by rw is_O_with; refl -lemma is_O_principal {s : set α} : f =O[𝓟 s] g ↔ ∃ c, ∀ x ∈ s, ∥f x∥ ≤ c * ∥g x∥ := +lemma is_O_principal {s : set α} : f =O[𝓟 s] g ↔ ∃ c, ∀ x ∈ s, ‖f x‖ ≤ c * ‖g x‖ := by rw is_O_iff; refl section variables (F) [has_one F] [norm_one_class F] -theorem is_O_with_const_one (c : E) (l : filter α) : is_O_with ∥c∥ l (λ x : α, c) (λ x, (1 : F)) := +theorem is_O_with_const_one (c : E) (l : filter α) : is_O_with ‖c‖ l (λ x : α, c) (λ x, (1 : F)) := by simp [is_O_with_iff] theorem is_O_const_one (c : E) (l : filter α) : (λ x : α, c) =O[l] (λ x, (1 : F)) := @@ -926,17 +926,17 @@ theorem is_o_const_iff_is_o_one {c : F''} (hc : c ≠ 0) : by simp only [is_o_iff, norm_one, mul_one, metric.nhds_basis_closed_ball.tendsto_right_iff, metric.mem_closed_ball, dist_zero_right] -@[simp] theorem is_O_one_iff : f =O[l] (λ x, 1 : α → F) ↔ is_bounded_under (≤) l (λ x, ∥f x∥) := +@[simp] theorem is_O_one_iff : f =O[l] (λ x, 1 : α → F) ↔ is_bounded_under (≤) l (λ x, ‖f x‖) := by { simp only [is_O_iff, norm_one, mul_one], refl } alias is_O_one_iff ↔ _ _root_.filter.is_bounded_under.is_O_one -@[simp] theorem is_o_one_left_iff : (λ x, 1 : α → F) =o[l] f ↔ tendsto (λ x, ∥f x∥) l at_top := -calc (λ x, 1 : α → F) =o[l] f ↔ ∀ n : ℕ, ∀ᶠ x in l, ↑n * ∥(1 : F)∥ ≤ ∥f x∥ : +@[simp] theorem is_o_one_left_iff : (λ x, 1 : α → F) =o[l] f ↔ tendsto (λ x, ‖f x‖) l at_top := +calc (λ x, 1 : α → F) =o[l] f ↔ ∀ n : ℕ, ∀ᶠ x in l, ↑n * ‖(1 : F)‖ ≤ ‖f x‖ : is_o_iff_nat_mul_le_aux $ or.inl $ λ x, by simp only [norm_one, zero_le_one] -... ↔ ∀ n : ℕ, true → ∀ᶠ x in l, ∥f x∥ ∈ Ici (n : ℝ) : +... ↔ ∀ n : ℕ, true → ∀ᶠ x in l, ‖f x‖ ∈ Ici (n : ℝ) : by simp only [norm_one, mul_one, true_implies_iff, mem_Ici] -... ↔ tendsto (λ x, ∥f x∥) l at_top : at_top_countable_basis_of_archimedean.1.tendsto_right_iff.symm +... ↔ tendsto (λ x, ‖f x‖) l at_top : at_top_countable_basis_of_archimedean.1.tendsto_right_iff.symm theorem _root_.filter.tendsto.is_O_one {c : E'} (h : tendsto f' l (𝓝 c)) : f' =O[l] (λ x, 1 : α → F) := @@ -966,14 +966,14 @@ h.norm.is_bounded_under_le.is_O_const hc lemma is_O.is_bounded_under_le {c : F} (h : f =O[l] (λ x, c)) : is_bounded_under (≤) l (norm ∘ f) := -let ⟨c', hc'⟩ := h.bound in ⟨c' * ∥c∥, eventually_map.2 hc'⟩ +let ⟨c', hc'⟩ := h.bound in ⟨c' * ‖c‖, eventually_map.2 hc'⟩ theorem is_O_const_of_ne {c : F''} (hc : c ≠ 0) : f =O[l] (λ x, c) ↔ is_bounded_under (≤) l (norm ∘ f) := ⟨λ h, h.is_bounded_under_le, λ h, h.is_O_const hc⟩ theorem is_O_const_iff {c : F''} : - f'' =O[l] (λ x, c) ↔ (c = 0 → f'' =ᶠ[l] 0) ∧ is_bounded_under (≤) l (λ x, ∥f'' x∥) := + f'' =O[l] (λ x, c) ↔ (c = 0 → f'' =ᶠ[l] 0) ∧ is_bounded_under (≤) l (λ x, ‖f'' x‖) := begin refine ⟨λ h, ⟨λ hc, is_O_zero_right_iff.1 (by rwa ← hc), h.is_bounded_under_le⟩, _⟩, rintro ⟨hcf, hf⟩, @@ -982,7 +982,7 @@ begin end theorem is_O_iff_is_bounded_under_le_div (h : ∀ᶠ x in l, g'' x ≠ 0) : - f =O[l] g'' ↔ is_bounded_under (≤) l (λ x, ∥f x∥ / ∥g'' x∥) := + f =O[l] g'' ↔ is_bounded_under (≤) l (λ x, ‖f x‖ / ‖g'' x‖) := begin simp only [is_O_iff, is_bounded_under, is_bounded, eventually_map], exact exists_congr (λ c, eventually_congr $ h.mono $ @@ -991,15 +991,15 @@ end /-- `(λ x, c) =O[l] f` if and only if `f` is bounded away from zero. -/ lemma is_O_const_left_iff_pos_le_norm {c : E''} (hc : c ≠ 0) : - (λ x, c) =O[l] f' ↔ ∃ b, 0 < b ∧ ∀ᶠ x in l, b ≤ ∥f' x∥ := + (λ x, c) =O[l] f' ↔ ∃ b, 0 < b ∧ ∀ᶠ x in l, b ≤ ‖f' x‖ := begin split, { intro h, rcases h.exists_pos with ⟨C, hC₀, hC⟩, - refine ⟨∥c∥ / C, div_pos (norm_pos_iff.2 hc) hC₀, _⟩, + refine ⟨‖c‖ / C, div_pos (norm_pos_iff.2 hc) hC₀, _⟩, exact hC.bound.mono (λ x, (div_le_iff' hC₀).2) }, { rintro ⟨b, hb₀, hb⟩, - refine is_O.of_bound (∥c∥ / b) (hb.mono $ λ x hx, _), + refine is_O.of_bound (‖c‖ / b) (hb.mono $ λ x hx, _), rw [div_mul_eq_mul_div, mul_div_assoc], exact le_mul_of_one_le_right (norm_nonneg _) ((one_le_div hb₀).2 hx) } end @@ -1021,14 +1021,14 @@ hfg.is_O.trans_tendsto hg /-! ### Multiplication by a constant -/ theorem is_O_with_const_mul_self (c : R) (f : α → R) (l : filter α) : - is_O_with ∥c∥ l (λ x, c * f x) f := + is_O_with ‖c‖ l (λ x, c * f x) f := is_O_with_of_le' _ $ λ x, norm_mul_le _ _ theorem is_O_const_mul_self (c : R) (f : α → R) (l : filter α) : (λ x, c * f x) =O[l] f := (is_O_with_const_mul_self c f l).is_O theorem is_O_with.const_mul_left {f : α → R} (h : is_O_with c l f g) (c' : R) : - is_O_with (∥c'∥ * c) l (λ x, c' * f x) g := + is_O_with (‖c'‖ * c) l (λ x, c' * f x) g := (is_O_with_const_mul_self c' f l).trans h (norm_nonneg c') theorem is_O.const_mul_left {f : α → R} (h : f =O[l] g) (c' : R) : @@ -1036,11 +1036,11 @@ theorem is_O.const_mul_left {f : α → R} (h : f =O[l] g) (c' : R) : let ⟨c, hc⟩ := h.is_O_with in (hc.const_mul_left c').is_O theorem is_O_with_self_const_mul' (u : Rˣ) (f : α → R) (l : filter α) : - is_O_with ∥(↑u⁻¹:R)∥ l f (λ x, ↑u * f x) := + is_O_with ‖(↑u⁻¹:R)‖ l f (λ x, ↑u * f x) := (is_O_with_const_mul_self ↑u⁻¹ _ l).congr_left $ λ x, u.inv_mul_cancel_left (f x) theorem is_O_with_self_const_mul (c : 𝕜) (hc : c ≠ 0) (f : α → 𝕜) (l : filter α) : - is_O_with ∥c∥⁻¹ l f (λ x, c * f x) := + is_O_with ‖c‖⁻¹ l f (λ x, c * f x) := (is_O_with_self_const_mul' (units.mk0 c hc) f l).congr_const $ norm_inv c @@ -1073,7 +1073,7 @@ is_o_const_mul_left_iff' $ is_unit.mk0 c hc theorem is_O_with.of_const_mul_right {g : α → R} {c : R} (hc' : 0 ≤ c') (h : is_O_with c' l f (λ x, c * g x)) : - is_O_with (c' * ∥c∥) l f g := + is_O_with (c' * ‖c‖) l f g := h.trans (is_O_with_const_mul_self c g l) hc' theorem is_O.of_const_mul_right {g : α → R} {c : R} (h : f =O[l] (λ x, c * g x)) : @@ -1082,12 +1082,12 @@ let ⟨c, cnonneg, hc⟩ := h.exists_nonneg in (hc.of_const_mul_right cnonneg).i theorem is_O_with.const_mul_right' {g : α → R} {u : Rˣ} {c' : ℝ} (hc' : 0 ≤ c') (h : is_O_with c' l f g) : - is_O_with (c' * ∥(↑u⁻¹:R)∥) l f (λ x, ↑u * g x) := + is_O_with (c' * ‖(↑u⁻¹:R)‖) l f (λ x, ↑u * g x) := h.trans (is_O_with_self_const_mul' _ _ _) hc' theorem is_O_with.const_mul_right {g : α → 𝕜} {c : 𝕜} (hc : c ≠ 0) {c' : ℝ} (hc' : 0 ≤ c') (h : is_O_with c' l f g) : - is_O_with (c' * ∥c∥⁻¹) l f (λ x, c * g x) := + is_O_with (c' * ‖c‖⁻¹) l f (λ x, c * g x) := h.trans (is_O_with_self_const_mul c hc g l) hc' theorem is_O.const_mul_right' {g : α → R} {c : R} (hc : is_unit c) (h : f =O[l] g) : @@ -1167,7 +1167,7 @@ theorem is_o.mul {f₁ f₂ : α → R} {g₁ g₂ : α → 𝕜} (h₁ : f₁ = h₁.mul_is_O h₂.is_O theorem is_O_with.pow' {f : α → R} {g : α → 𝕜} (h : is_O_with c l f g) : - ∀ n : ℕ, is_O_with (nat.cases_on n ∥(1 : R)∥ (λ n, c ^ (n + 1))) l (λ x, f x ^ n) (λ x, g x ^ n) + ∀ n : ℕ, is_O_with (nat.cases_on n ‖(1 : R)‖ (λ n, c ^ (n + 1))) l (λ x, f x ^ n) (λ x, g x ^ n) | 0 := by simpa using is_O_with_const_const (1 : R) (one_ne_zero' 𝕜) l | 1 := by simpa | (n + 2) := by simpa [pow_succ] using h.mul (is_O_with.pow' (n + 1)) @@ -1181,11 +1181,11 @@ theorem is_O_with.of_pow {n : ℕ} {f : α → 𝕜} {g : α → R} (h : is_O_wi (hn : n ≠ 0) (hc : c ≤ c' ^ n) (hc' : 0 ≤ c') : is_O_with c' l f g := is_O_with.of_bound $ (h.weaken hc).bound.mono $ λ x hx, le_of_pow_le_pow n (mul_nonneg hc' $ norm_nonneg _) hn.bot_lt $ - calc ∥f x∥ ^ n = ∥(f x) ^ n∥ : (norm_pow _ _).symm - ... ≤ c' ^ n * ∥(g x) ^ n∥ : hx - ... ≤ c' ^ n * ∥g x∥ ^ n : + calc ‖f x‖ ^ n = ‖(f x) ^ n‖ : (norm_pow _ _).symm + ... ≤ c' ^ n * ‖(g x) ^ n‖ : hx + ... ≤ c' ^ n * ‖g x‖ ^ n : mul_le_mul_of_nonneg_left (norm_pow_le' _ hn.bot_lt) (pow_nonneg hc' _) - ... = (c' * ∥g x∥) ^ n : (mul_pow _ _ _).symm + ... = (c' * ‖g x‖) ^ n : (mul_pow _ _ _).symm theorem is_O.pow {f : α → R} {g : α → 𝕜} (h : f =O[l] g) (n : ℕ) : (λ x, f x ^ n) =O[l] (λ x, g x ^ n) := @@ -1239,20 +1239,20 @@ section smul_const variables [normed_space 𝕜 E'] theorem is_O_with.const_smul_left (h : is_O_with c l f' g) (c' : 𝕜) : - is_O_with (∥c'∥ * c) l (λ x, c' • f' x) g := + is_O_with (‖c'‖ * c) l (λ x, c' • f' x) g := is_O_with.of_norm_left $ - by simpa only [← norm_smul, norm_norm] using h.norm_left.const_mul_left (∥c'∥) + by simpa only [← norm_smul, norm_norm] using h.norm_left.const_mul_left (‖c'‖) lemma is_O.const_smul_left (h : f' =O[l] g) (c : 𝕜) : (c • f') =O[l] g := let ⟨b, hb⟩ := h.is_O_with in (hb.const_smul_left _).is_O lemma is_o.const_smul_left (h : f' =o[l] g) (c : 𝕜) : (c • f') =o[l] g := -is_o.of_norm_left $ by simpa only [← norm_smul] using h.norm_left.const_mul_left (∥c∥) +is_o.of_norm_left $ by simpa only [← norm_smul] using h.norm_left.const_mul_left (‖c‖) theorem is_O_const_smul_left {c : 𝕜} (hc : c ≠ 0) : (λ x, c • f' x) =O[l] g ↔ f' =O[l] g := begin - have cne0 : ∥c∥ ≠ 0, from mt norm_eq_zero.mp hc, + have cne0 : ‖c‖ ≠ 0, from mt norm_eq_zero.mp hc, rw [←is_O_norm_left], simp only [norm_smul], rw [is_O_const_mul_left_iff cne0, is_O_norm_left], end @@ -1260,7 +1260,7 @@ end theorem is_o_const_smul_left {c : 𝕜} (hc : c ≠ 0) : (λ x, c • f' x) =o[l] g ↔ f' =o[l] g := begin - have cne0 : ∥c∥ ≠ 0, from mt norm_eq_zero.mp hc, + have cne0 : ‖c‖ ≠ 0, from mt norm_eq_zero.mp hc, rw [←is_o_norm_left], simp only [norm_smul], rw [is_o_const_mul_left_iff cne0, is_o_norm_left] end @@ -1268,7 +1268,7 @@ end theorem is_O_const_smul_right {c : 𝕜} (hc : c ≠ 0) : f =O[l] (λ x, c • f' x) ↔ f =O[l] f' := begin - have cne0 : ∥c∥ ≠ 0, from mt norm_eq_zero.mp hc, + have cne0 : ‖c‖ ≠ 0, from mt norm_eq_zero.mp hc, rw [←is_O_norm_right], simp only [norm_smul], rw [is_O_const_mul_right_iff cne0, is_O_norm_right] end @@ -1276,7 +1276,7 @@ end theorem is_o_const_smul_right {c : 𝕜} (hc : c ≠ 0) : f =o[l] (λ x, c • f' x) ↔ f =o[l] f' := begin - have cne0 : ∥c∥ ≠ 0, from mt norm_eq_zero.mp hc, + have cne0 : ‖c‖ ≠ 0, from mt norm_eq_zero.mp hc, rw [←is_o_norm_right], simp only [norm_smul], rw [is_o_const_mul_right_iff cne0, is_o_norm_right] end @@ -1377,7 +1377,7 @@ alias is_o_iff_tendsto' ↔ _ is_o_of_tendsto' alias is_o_iff_tendsto ↔ _ is_o_of_tendsto lemma is_o_const_left_of_ne {c : E''} (hc : c ≠ 0) : - (λ x, c) =o[l] g ↔ tendsto (λ x, ∥g x∥) l at_top := + (λ x, c) =o[l] g ↔ tendsto (λ x, ‖g x‖) l at_top := begin simp only [← is_o_one_left_iff ℝ], exact ⟨(is_O_const_const (1 : ℝ) hc l).trans_is_o, (is_O_const_one ℝ c l).trans_is_o⟩ @@ -1393,7 +1393,7 @@ end @[simp] theorem is_o_const_const_iff [ne_bot l] {d : E''} {c : F''} : (λ x, d) =o[l] (λ x, c) ↔ d = 0 := -have ¬tendsto (function.const α ∥c∥) l at_top, +have ¬tendsto (function.const α ‖c‖) l at_top, from not_tendsto_at_top_of_tendsto_nhds tendsto_const_nhds, by simp [function.const, this] @@ -1441,20 +1441,20 @@ section exists_mul_eq variables {u v : α → 𝕜} -/-- If `∥φ∥` is eventually bounded by `c`, and `u =ᶠ[l] φ * v`, then we have `is_O_with c u v l`. +/-- If `‖φ‖` is eventually bounded by `c`, and `u =ᶠ[l] φ * v`, then we have `is_O_with c u v l`. This does not require any assumptions on `c`, which is why we keep this version along with `is_O_with_iff_exists_eq_mul`. -/ -lemma is_O_with_of_eq_mul (φ : α → 𝕜) (hφ : ∀ᶠ x in l, ∥φ x∥ ≤ c) (h : u =ᶠ[l] φ * v) : +lemma is_O_with_of_eq_mul (φ : α → 𝕜) (hφ : ∀ᶠ x in l, ‖φ x‖ ≤ c) (h : u =ᶠ[l] φ * v) : is_O_with c l u v := begin unfold is_O_with, - refine h.symm.rw (λ x a, ∥a∥ ≤ c * ∥v x∥) (hφ.mono $ λ x hx, _), + refine h.symm.rw (λ x a, ‖a‖ ≤ c * ‖v x‖) (hφ.mono $ λ x hx, _), simp only [norm_mul, pi.mul_apply], exact mul_le_mul_of_nonneg_right hx (norm_nonneg _) end lemma is_O_with_iff_exists_eq_mul (hc : 0 ≤ c) : - is_O_with c l u v ↔ ∃ (φ : α → 𝕜) (hφ : ∀ᶠ x in l, ∥φ x∥ ≤ c), u =ᶠ[l] φ * v := + is_O_with c l u v ↔ ∃ (φ : α → 𝕜) (hφ : ∀ᶠ x in l, ‖φ x‖ ≤ c), u =ᶠ[l] φ * v := begin split, { intro h, @@ -1466,7 +1466,7 @@ begin end lemma is_O_with.exists_eq_mul (h : is_O_with c l u v) (hc : 0 ≤ c) : - ∃ (φ : α → 𝕜) (hφ : ∀ᶠ x in l, ∥φ x∥ ≤ c), u =ᶠ[l] φ * v := + ∃ (φ : α → 𝕜) (hφ : ∀ᶠ x in l, ‖φ x‖ ≤ c), u =ᶠ[l] φ * v := (is_O_with_iff_exists_eq_mul hc).mp h lemma is_O_iff_exists_eq_mul : @@ -1501,7 +1501,7 @@ end exists_mul_eq theorem div_is_bounded_under_of_is_O {α : Type*} {l : filter α} {f g : α → 𝕜} (h : f =O[l] g) : - is_bounded_under (≤) l (λ x, ∥f x / g x∥) := + is_bounded_under (≤) l (λ x, ‖f x / g x‖) := begin obtain ⟨c, h₀, hc⟩ := h.exists_nonneg, refine ⟨c, eventually_map.2 (hc.bound.mono (λ x hx, _))⟩, @@ -1511,7 +1511,7 @@ end theorem is_O_iff_div_is_bounded_under {α : Type*} {l : filter α} {f g : α → 𝕜} (hgf : ∀ᶠ x in l, g x = 0 → f x = 0) : - f =O[l] g ↔ is_bounded_under (≤) l (λ x, ∥f x / g x∥) := + f =O[l] g ↔ is_bounded_under (≤) l (λ x, ‖f x / g x‖) := begin refine ⟨div_is_bounded_under_of_is_O, λ h, _⟩, obtain ⟨c, hc⟩ := h, @@ -1547,7 +1547,7 @@ begin end theorem is_o_norm_pow_norm_pow {m n : ℕ} (h : m < n) : - (λ x : E', ∥x∥^n) =o[𝓝 0] (λ x, ∥x∥^m) := + (λ x : E', ‖x‖^n) =o[𝓝 0] (λ x, ‖x‖^m) := (is_o_pow_pow h).comp_tendsto tendsto_norm_zero theorem is_o_pow_id {n : ℕ} (h : 1 < n) : @@ -1555,21 +1555,21 @@ theorem is_o_pow_id {n : ℕ} (h : 1 < n) : by { convert is_o_pow_pow h, simp only [pow_one] } theorem is_o_norm_pow_id {n : ℕ} (h : 1 < n) : - (λ x : E', ∥x∥^n) =o[𝓝 0] (λ x, x) := + (λ x : E', ‖x‖^n) =o[𝓝 0] (λ x, x) := by simpa only [pow_one, is_o_norm_right] using @is_o_norm_pow_norm_pow E' _ _ _ h lemma is_O.eq_zero_of_norm_pow_within {f : E'' → F''} {s : set E''} {x₀ : E''} {n : ℕ} - (h : f =O[𝓝[s] x₀] λ x, ∥x - x₀∥ ^ n) (hx₀ : x₀ ∈ s) (hn : 0 < n) : f x₀ = 0 := + (h : f =O[𝓝[s] x₀] λ x, ‖x - x₀‖ ^ n) (hx₀ : x₀ ∈ s) (hn : 0 < n) : f x₀ = 0 := mem_of_mem_nhds_within hx₀ h.eq_zero_imp $ by simp_rw [sub_self, norm_zero, zero_pow hn] lemma is_O.eq_zero_of_norm_pow {f : E'' → F''} {x₀ : E''} {n : ℕ} - (h : f =O[𝓝 x₀] λ x, ∥x - x₀∥ ^ n) (hn : 0 < n) : f x₀ = 0 := + (h : f =O[𝓝 x₀] λ x, ‖x - x₀‖ ^ n) (hn : 0 < n) : f x₀ = 0 := by { rw [← nhds_within_univ] at h, exact h.eq_zero_of_norm_pow_within (mem_univ _) hn } lemma is_o_pow_sub_pow_sub (x₀ : E') {n m : ℕ} (h : n < m) : - (λ x, ∥x - x₀∥ ^ m) =o[𝓝 x₀] λ x, ∥x - x₀∥^n := + (λ x, ‖x - x₀‖ ^ m) =o[𝓝 x₀] λ x, ‖x - x₀‖^n := begin - have : tendsto (λ x, ∥x - x₀∥) (𝓝 x₀) (𝓝 0), + have : tendsto (λ x, ‖x - x₀‖) (𝓝 x₀) (𝓝 0), { apply tendsto_norm_zero.comp, rw ← sub_self x₀, exact tendsto_id.sub tendsto_const_nhds }, @@ -1577,7 +1577,7 @@ begin end lemma is_o_pow_sub_sub (x₀ : E') {m : ℕ} (h : 1 < m) : - (λ x, ∥x - x₀∥^m) =o[𝓝 x₀] λ x, x - x₀ := + (λ x, ‖x - x₀‖^m) =o[𝓝 x₀] λ x, x - x₀ := by simpa only [is_o_norm_right, pow_one] using is_o_pow_sub_pow_sub x₀ h theorem is_O_with.right_le_sub_of_lt_1 {f₁ f₂ : α → E'} (h : is_O_with c l f₁ f₂) (hc : c < 1) : @@ -1604,42 +1604,42 @@ theorem is_o.right_is_O_add {f₁ f₂ : α → E'} (h : f₁ =o[l] f₂) : ((h.def' one_half_pos).right_le_add_of_lt_1 one_half_lt_one).is_O /-- If `f x = O(g x)` along `cofinite`, then there exists a positive constant `C` such that -`∥f x∥ ≤ C * ∥g x∥` whenever `g x ≠ 0`. -/ +`‖f x‖ ≤ C * ‖g x‖` whenever `g x ≠ 0`. -/ theorem bound_of_is_O_cofinite (h : f =O[cofinite] g'') : - ∃ C > 0, ∀ ⦃x⦄, g'' x ≠ 0 → ∥f x∥ ≤ C * ∥g'' x∥ := + ∃ C > 0, ∀ ⦃x⦄, g'' x ≠ 0 → ‖f x‖ ≤ C * ‖g'' x‖ := begin rcases h.exists_pos with ⟨C, C₀, hC⟩, rw [is_O_with, eventually_cofinite] at hC, - rcases (hC.to_finset.image (λ x, ∥f x∥ / ∥g'' x∥)).exists_le with ⟨C', hC'⟩, - have : ∀ x, C * ∥g'' x∥ < ∥f x∥ → ∥f x∥ / ∥g'' x∥ ≤ C', by simpa using hC', + rcases (hC.to_finset.image (λ x, ‖f x‖ / ‖g'' x‖)).exists_le with ⟨C', hC'⟩, + have : ∀ x, C * ‖g'' x‖ < ‖f x‖ → ‖f x‖ / ‖g'' x‖ ≤ C', by simpa using hC', refine ⟨max C C', lt_max_iff.2 (or.inl C₀), λ x h₀, _⟩, rw [max_mul_of_nonneg _ _ (norm_nonneg _), le_max_iff, or_iff_not_imp_left, not_le], exact λ hx, (div_le_iff (norm_pos_iff.2 h₀)).1 (this _ hx) end theorem is_O_cofinite_iff (h : ∀ x, g'' x = 0 → f'' x = 0) : - f'' =O[cofinite] g'' ↔ ∃ C, ∀ x, ∥f'' x∥ ≤ C * ∥g'' x∥ := + f'' =O[cofinite] g'' ↔ ∃ C, ∀ x, ‖f'' x‖ ≤ C * ‖g'' x‖ := ⟨λ h', let ⟨C, C₀, hC⟩ := bound_of_is_O_cofinite h' in ⟨C, λ x, if hx : g'' x = 0 then by simp [h _ hx, hx] else hC hx⟩, λ h, (is_O_top.2 h).mono le_top⟩ theorem bound_of_is_O_nat_at_top {f : ℕ → E} {g'' : ℕ → E''} (h : f =O[at_top] g'') : - ∃ C > 0, ∀ ⦃x⦄, g'' x ≠ 0 → ∥f x∥ ≤ C * ∥g'' x∥ := + ∃ C > 0, ∀ ⦃x⦄, g'' x ≠ 0 → ‖f x‖ ≤ C * ‖g'' x‖ := bound_of_is_O_cofinite $ by rwa nat.cofinite_eq_at_top theorem is_O_nat_at_top_iff {f : ℕ → E''} {g : ℕ → F''} (h : ∀ x, g x = 0 → f x = 0) : - f =O[at_top] g ↔ ∃ C, ∀ x, ∥f x∥ ≤ C * ∥g x∥ := + f =O[at_top] g ↔ ∃ C, ∀ x, ‖f x‖ ≤ C * ‖g x‖ := by rw [← nat.cofinite_eq_at_top, is_O_cofinite_iff h] theorem is_O_one_nat_at_top_iff {f : ℕ → E''} : - f =O[at_top] (λ n, 1 : ℕ → ℝ) ↔ ∃ C, ∀ n, ∥f n∥ ≤ C := + f =O[at_top] (λ n, 1 : ℕ → ℝ) ↔ ∃ C, ∀ n, ‖f n‖ ≤ C := iff.trans (is_O_nat_at_top_iff (λ n h, (one_ne_zero h).elim)) $ by simp only [norm_one, mul_one] theorem is_O_with_pi {ι : Type*} [fintype ι] {E' : ι → Type*} [Π i, normed_add_comm_group (E' i)] {f : α → Π i, E' i} {C : ℝ} (hC : 0 ≤ C) : is_O_with C l f g' ↔ ∀ i, is_O_with C l (λ x, f x i) g' := -have ∀ x, 0 ≤ C * ∥g' x∥, from λ x, mul_nonneg hC (norm_nonneg _), +have ∀ x, 0 ≤ C * ‖g' x‖, from λ x, mul_nonneg hC (norm_nonneg _), by simp only [is_O_with_iff, pi_norm_le_iff_of_nonneg (this _), eventually_all] @[simp] theorem is_O_pi {ι : Type*} [fintype ι] {E' : ι → Type*} [Π i, normed_add_comm_group (E' i)] @@ -1665,7 +1665,7 @@ open asymptotics lemma summable_of_is_O {ι E} [normed_add_comm_group E] [complete_space E] {f : ι → E} {g : ι → ℝ} (hg : summable g) (h : f =O[cofinite] g) : summable f := let ⟨C, hC⟩ := h.is_O_with in -summable_of_norm_bounded_eventually (λ x, C * ∥g x∥) (hg.abs.mul_left _) hC.bound +summable_of_norm_bounded_eventually (λ x, C * ‖g x‖) (hg.abs.mul_left _) hC.bound lemma summable_of_is_O_nat {E} [normed_add_comm_group E] [complete_space E] {f : ℕ → E} {g : ℕ → ℝ} (hg : summable g) (h : f =O[at_top] g) : summable f := diff --git a/src/analysis/asymptotics/specific_asymptotics.lean b/src/analysis/asymptotics/specific_asymptotics.lean index 4fb0648f15a08..ef8af93edbb0a 100644 --- a/src/analysis/asymptotics/specific_asymptotics.lean +++ b/src/analysis/asymptotics/specific_asymptotics.lean @@ -88,7 +88,7 @@ begin end lemma asymptotics.is_O.trans_tendsto_norm_at_top {α : Type*} {u v : α → 𝕜} {l : filter α} - (huv : u =O[l] v) (hu : tendsto (λ x, ∥u x∥) l at_top) : tendsto (λ x, ∥v x∥) l at_top := + (huv : u =O[l] v) (hu : tendsto (λ x, ‖u x‖) l at_top) : tendsto (λ x, ‖v x‖) l at_top := begin rcases huv.exists_pos with ⟨c, hc, hcuv⟩, rw is_O_with at hcuv, @@ -109,24 +109,24 @@ lemma asymptotics.is_o.sum_range {α : Type*} [normed_add_comm_group α] (hg : 0 ≤ g) (h'g : tendsto (λ n, ∑ i in range n, g i) at_top at_top) : (λ n, ∑ i in range n, f i) =o[at_top] (λ n, ∑ i in range n, g i) := begin - have A : ∀ i, ∥g i∥ = g i := λ i, real.norm_of_nonneg (hg i), - have B : ∀ n, ∥∑ i in range n, g i∥ = ∑ i in range n, g i, + have A : ∀ i, ‖g i‖ = g i := λ i, real.norm_of_nonneg (hg i), + have B : ∀ n, ‖∑ i in range n, g i‖ = ∑ i in range n, g i, from λ n, by rwa [real.norm_eq_abs, abs_sum_of_nonneg'], apply is_o_iff.2 (λ ε εpos, _), - obtain ⟨N, hN⟩ : ∃ (N : ℕ), ∀ (b : ℕ), N ≤ b → ∥f b∥ ≤ ε / 2 * g b, + obtain ⟨N, hN⟩ : ∃ (N : ℕ), ∀ (b : ℕ), N ≤ b → ‖f b‖ ≤ ε / 2 * g b, by simpa only [A, eventually_at_top] using is_o_iff.mp h (half_pos εpos), have : (λ (n : ℕ), ∑ i in range N, f i) =o[at_top] (λ (n : ℕ), ∑ i in range n, g i), { apply is_o_const_left.2, exact or.inr (h'g.congr (λ n, (B n).symm)) }, filter_upwards [is_o_iff.1 this (half_pos εpos), Ici_mem_at_top N] with n hn Nn, - calc ∥∑ i in range n, f i∥ - = ∥∑ i in range N, f i + ∑ i in Ico N n, f i∥ : + calc ‖∑ i in range n, f i‖ + = ‖∑ i in range N, f i + ∑ i in Ico N n, f i‖ : by rw sum_range_add_sum_Ico _ Nn - ... ≤ ∥∑ i in range N, f i∥ + ∥∑ i in Ico N n, f i∥ : + ... ≤ ‖∑ i in range N, f i‖ + ‖∑ i in Ico N n, f i‖ : norm_add_le _ _ - ... ≤ ∥∑ i in range N, f i∥ + ∑ i in Ico N n, (ε / 2) * g i : + ... ≤ ‖∑ i in range N, f i‖ + ∑ i in Ico N n, (ε / 2) * g i : add_le_add le_rfl (norm_sum_le_of_le _ (λ i hi, hN _ (mem_Ico.1 hi).1)) - ... ≤ ∥∑ i in range N, f i∥ + ∑ i in range n, (ε / 2) * g i : + ... ≤ ‖∑ i in range N, f i‖ + ∑ i in range n, (ε / 2) * g i : begin refine add_le_add le_rfl _, apply sum_le_sum_of_subset_of_nonneg, @@ -135,12 +135,12 @@ begin { assume i hi hident, exact mul_nonneg (half_pos εpos).le (hg i) } end - ... ≤ (ε / 2) * ∥∑ i in range n, g i∥ + (ε / 2) * (∑ i in range n, g i) : + ... ≤ (ε / 2) * ‖∑ i in range n, g i‖ + (ε / 2) * (∑ i in range n, g i) : begin rw ← mul_sum, exact add_le_add hn (mul_le_mul_of_nonneg_left le_rfl (half_pos εpos).le), end - ... = ε * ∥(∑ i in range n, g i)∥ : by { simp [B], ring } + ... = ε * ‖(∑ i in range n, g i)‖ : by { simp [B], ring } end lemma asymptotics.is_o_sum_range_of_tendsto_zero {α : Type*} [normed_add_comm_group α] diff --git a/src/analysis/asymptotics/superpolynomial_decay.lean b/src/analysis/asymptotics/superpolynomial_decay.lean index 971ea9c30c514..4a9f43f38c4e6 100644 --- a/src/analysis/asymptotics/superpolynomial_decay.lean +++ b/src/analysis/asymptotics/superpolynomial_decay.lean @@ -280,12 +280,12 @@ variable [normed_linear_ordered_field β] variables (l k f) lemma superpolynomial_decay_iff_norm_tendsto_zero : - superpolynomial_decay l k f ↔ ∀ (n : ℕ), tendsto (λ (a : α), ∥(k a) ^ n * f a∥) l (𝓝 0) := + superpolynomial_decay l k f ↔ ∀ (n : ℕ), tendsto (λ (a : α), ‖(k a) ^ n * f a‖) l (𝓝 0) := ⟨λ h z, tendsto_zero_iff_norm_tendsto_zero.1 (h z), λ h z, tendsto_zero_iff_norm_tendsto_zero.2 (h z)⟩ lemma superpolynomial_decay_iff_superpolynomial_decay_norm : - superpolynomial_decay l k f ↔ superpolynomial_decay l (λ a, ∥k a∥) (λ a, ∥f a∥) := + superpolynomial_decay l k f ↔ superpolynomial_decay l (λ a, ‖k a‖) (λ a, ‖f a‖) := (superpolynomial_decay_iff_norm_tendsto_zero l k f).trans (by simp [superpolynomial_decay]) variables {l k} diff --git a/src/analysis/asymptotics/theta.lean b/src/analysis/asymptotics/theta.lean index 07a222068b723..7d4f9070084a4 100644 --- a/src/analysis/asymptotics/theta.lean +++ b/src/analysis/asymptotics/theta.lean @@ -74,17 +74,17 @@ h₁.1.trans_is_o h₂ (hf : f₁ =ᶠ[l] f₂) (h : f₂ =Θ[l] g) : f₁ =Θ[l] g := ⟨hf.trans_is_O h.1, h.2.trans_eventually_eq hf.symm⟩ -@[simp] lemma is_Theta_norm_left : (λ x, ∥f' x∥) =Θ[l] g ↔ f' =Θ[l] g := by simp [is_Theta] -@[simp] lemma is_Theta_norm_right : f =Θ[l] (λ x, ∥g' x∥) ↔ f =Θ[l] g' := by simp [is_Theta] +@[simp] lemma is_Theta_norm_left : (λ x, ‖f' x‖) =Θ[l] g ↔ f' =Θ[l] g := by simp [is_Theta] +@[simp] lemma is_Theta_norm_right : f =Θ[l] (λ x, ‖g' x‖) ↔ f =Θ[l] g' := by simp [is_Theta] alias is_Theta_norm_left ↔ is_Theta.of_norm_left is_Theta.norm_left alias is_Theta_norm_right ↔ is_Theta.of_norm_right is_Theta.norm_right -lemma is_Theta_of_norm_eventually_eq (h : (λ x, ∥f x∥) =ᶠ[l] (λ x, ∥g x∥)) : f =Θ[l] g := +lemma is_Theta_of_norm_eventually_eq (h : (λ x, ‖f x‖) =ᶠ[l] (λ x, ‖g x‖)) : f =Θ[l] g := ⟨is_O.of_bound 1 $ by simpa only [one_mul] using h.le, is_O.of_bound 1 $ by simpa only [one_mul] using h.symm.le⟩ -lemma is_Theta_of_norm_eventually_eq' {g : α → ℝ} (h : (λ x, ∥f' x∥) =ᶠ[l] g) : f' =Θ[l] g := +lemma is_Theta_of_norm_eventually_eq' {g : α → ℝ} (h : (λ x, ‖f' x‖) =ᶠ[l] g) : f' =Θ[l] g := is_Theta_of_norm_eventually_eq $ h.mono $ λ x hx, by simp only [← hx, norm_norm] lemma is_Theta.is_o_congr_left (h : f' =Θ[l] g') : f' =o[l] k ↔ g' =o[l] k := diff --git a/src/analysis/box_integral/basic.lean b/src/analysis/box_integral/basic.lean index 93a271902d31c..971bd2284073c 100644 --- a/src/analysis/box_integral/basic.lean +++ b/src/analysis/box_integral/basic.lean @@ -334,12 +334,12 @@ begin { rw [integral, dif_neg hgi] } end -/-- If `∥f x∥ ≤ g x` on `[l, u]` and `g` is integrable, then the norm of the integral of `f` is less +/-- If `‖f x‖ ≤ g x` on `[l, u]` and `g` is integrable, then the norm of the integral of `f` is less than or equal to the integral of `g`. -/ -lemma norm_integral_le_of_norm_le {g : ℝⁿ → ℝ} (hle : ∀ x ∈ I.Icc, ∥f x∥ ≤ g x) +lemma norm_integral_le_of_norm_le {g : ℝⁿ → ℝ} (hle : ∀ x ∈ I.Icc, ‖f x‖ ≤ g x) (μ : measure ℝⁿ) [is_locally_finite_measure μ] (hg : integrable I l g μ.to_box_additive.to_smul) : - ∥(integral I l f μ.to_box_additive.to_smul : E)∥ ≤ + ‖(integral I l f μ.to_box_additive.to_smul : E)‖ ≤ integral I l g μ.to_box_additive.to_smul := begin by_cases hfi : integrable.{u v v} I l f μ.to_box_additive.to_smul, @@ -352,9 +352,9 @@ begin exact integral_nonneg (λ x hx, (norm_nonneg _).trans (hle x hx)) μ } end -lemma norm_integral_le_of_le_const {c : ℝ} (hc : ∀ x ∈ I.Icc, ∥f x∥ ≤ c) +lemma norm_integral_le_of_le_const {c : ℝ} (hc : ∀ x ∈ I.Icc, ‖f x‖ ≤ c) (μ : measure ℝⁿ) [is_locally_finite_measure μ] : - ∥(integral I l f μ.to_box_additive.to_smul : E)∥ ≤ (μ I).to_real * c := + ‖(integral I l f μ.to_box_additive.to_smul : E)‖ ≤ (μ I).to_real * c := by simpa only [integral_const] using norm_integral_le_of_norm_le hc μ (integrable_const c) @@ -652,8 +652,8 @@ begin simp only [dist_eq_norm, integral_sum_sub_partitions _ _ h₁p h₂p, box_additive_map.to_smul_apply, ← smul_sub], have : ∀ J ∈ π₁.to_prepartition ⊓ π₂.to_prepartition, - ∥μ.to_box_additive J • (f ((π₁.inf_prepartition π₂.to_prepartition).tag J) - - f ((π₂.inf_prepartition π₁.to_prepartition).tag J))∥ ≤ μ.to_box_additive J * ε', + ‖μ.to_box_additive J • (f ((π₁.inf_prepartition π₂.to_prepartition).tag J) - + f ((π₂.inf_prepartition π₁.to_prepartition).tag J))‖ ≤ μ.to_box_additive J * ε', { intros J hJ, have : 0 ≤ μ.to_box_additive J, from ennreal.to_real_nonneg, rw [norm_smul, real.norm_eq_abs, abs_of_nonneg this, ← dist_eq_norm], diff --git a/src/analysis/box_integral/divergence_theorem.lean b/src/analysis/box_integral/divergence_theorem.lean index c23bc769cc82e..85b72e111e843 100644 --- a/src/analysis/box_integral/divergence_theorem.lean +++ b/src/analysis/box_integral/divergence_theorem.lean @@ -60,10 +60,10 @@ open measure_theory lemma norm_volume_sub_integral_face_upper_sub_lower_smul_le {f : ℝⁿ⁺¹ → E} {f' : ℝⁿ⁺¹ →L[ℝ] E} (hfc : continuous_on f I.Icc) {x : ℝⁿ⁺¹} (hxI : x ∈ I.Icc) {a : E} {ε : ℝ} (h0 : 0 < ε) - (hε : ∀ y ∈ I.Icc, ∥f y - a - f' (y - x)∥ ≤ ε * ∥y - x∥) {c : ℝ≥0} (hc : I.distortion ≤ c) : - ∥(∏ j, (I.upper j - I.lower j)) • f' (pi.single i 1) - + (hε : ∀ y ∈ I.Icc, ‖f y - a - f' (y - x)‖ ≤ ε * ‖y - x‖) {c : ℝ≥0} (hc : I.distortion ≤ c) : + ‖(∏ j, (I.upper j - I.lower j)) • f' (pi.single i 1) - (integral (I.face i) ⊥ (f ∘ i.insert_nth (I.upper i)) box_additive_map.volume - - integral (I.face i) ⊥ (f ∘ i.insert_nth (I.lower i)) box_additive_map.volume)∥ ≤ + integral (I.face i) ⊥ (f ∘ i.insert_nth (I.lower i)) box_additive_map.volume)‖ ≤ 2 * ε * c * ∏ j, (I.upper j - I.lower j) := begin /- **Plan of the proof**. The difference of the integrals of the affine function @@ -82,32 +82,32 @@ begin of the faces `x i = I.lower i` and `x i = I.upper i` is `(2 * ε * diam I.Icc)`-close to the value of `f'` on `pi.single i (I.upper i - I.lower i) = lᵢ • eᵢ`, where `lᵢ = I.upper i - I.lower i` is the length of `i`-th edge of `I` and `eᵢ = pi.single i 1` is the `i`-th unit vector. -/ - have : ∀ y ∈ (I.face i).Icc, ∥f' (pi.single i (I.upper i - I.lower i)) - - (f (i.insert_nth (I.upper i) y) - f (i.insert_nth (I.lower i) y))∥ ≤ 2 * ε * diam I.Icc, + have : ∀ y ∈ (I.face i).Icc, ‖f' (pi.single i (I.upper i - I.lower i)) - + (f (i.insert_nth (I.upper i) y) - f (i.insert_nth (I.lower i) y))‖ ≤ 2 * ε * diam I.Icc, { intros y hy, set g := λ y, f y - a - f' (y - x) with hg, - change ∀ y ∈ I.Icc, ∥g y∥ ≤ ε * ∥y - x∥ at hε, + change ∀ y ∈ I.Icc, ‖g y‖ ≤ ε * ‖y - x‖ at hε, clear_value g, obtain rfl : f = λ y, a + f' (y - x) + g y, by simp [hg], - convert_to ∥g (i.insert_nth (I.lower i) y) - g (i.insert_nth (I.upper i) y)∥ ≤ _, + convert_to ‖g (i.insert_nth (I.lower i) y) - g (i.insert_nth (I.upper i) y)‖ ≤ _, { congr' 1, have := fin.insert_nth_sub_same i (I.upper i) (I.lower i) y, simp only [← this, f'.map_sub], abel }, { have : ∀ z ∈ Icc (I.lower i) (I.upper i), i.insert_nth z y ∈ I.Icc, from λ z hz, I.maps_to_insert_nth_face_Icc hz hy, - replace hε : ∀ y ∈ I.Icc, ∥g y∥ ≤ ε * diam I.Icc, + replace hε : ∀ y ∈ I.Icc, ‖g y‖ ≤ ε * diam I.Icc, { intros y hy, refine (hε y hy).trans (mul_le_mul_of_nonneg_left _ h0.le), rw ← dist_eq_norm, exact dist_le_diam_of_mem I.is_compact_Icc.bounded hy hxI }, rw [two_mul, add_mul], exact norm_sub_le_of_le (hε _ (this _ Hl)) (hε _ (this _ Hu)) } }, - calc ∥(∏ j, (I.upper j - I.lower j)) • f' (pi.single i 1) - + calc ‖(∏ j, (I.upper j - I.lower j)) • f' (pi.single i 1) - (integral (I.face i) ⊥ (f ∘ i.insert_nth (I.upper i)) box_additive_map.volume - - integral (I.face i) ⊥ (f ∘ i.insert_nth (I.lower i)) box_additive_map.volume)∥ - = ∥integral.{0 u u} (I.face i) ⊥ + integral (I.face i) ⊥ (f ∘ i.insert_nth (I.lower i)) box_additive_map.volume)‖ + = ‖integral.{0 u u} (I.face i) ⊥ (λ (x : fin n → ℝ), f' (pi.single i (I.upper i - I.lower i)) - (f (i.insert_nth (I.upper i) x) - f (i.insert_nth (I.lower i) x))) - box_additive_map.volume∥ : + box_additive_map.volume‖ : begin rw [← integral_sub (Hi _ Hu) (Hi _ Hl), ← box.volume_face_mul i, mul_smul, ← box.volume_apply, ← box_additive_map.to_smul_apply, ← integral_const, ← box_additive_map.volume, @@ -173,8 +173,8 @@ begin because each of the integrals is close to `volume (J.face i) • f x`. TODO: there should be a shorter and more readable way to formalize this simple proof. -/ have : ∀ᶠ δ in 𝓝[>] (0 : ℝ), δ ∈ Ioc (0 : ℝ) (1 / 2) ∧ - (∀ y₁ y₂ ∈ closed_ball x δ ∩ I.Icc, ∥f y₁ - f y₂∥ ≤ ε / 2) ∧ - ((2 * δ) ^ (n + 1) * ∥f' x (pi.single i 1)∥ ≤ ε / 2), + (∀ y₁ y₂ ∈ closed_ball x δ ∩ I.Icc, ‖f y₁ - f y₂‖ ≤ ε / 2) ∧ + ((2 * δ) ^ (n + 1) * ‖f' x (pi.single i 1)‖ ≤ ε / 2), { refine eventually.and _ (eventually.and _ _), { exact Ioc_mem_nhds_within_Ioi ⟨le_rfl, one_half_pos⟩ }, { rcases ((nhds_within_has_basis nhds_basis_closed_ball _).tendsto_iff @@ -186,7 +186,7 @@ begin calc dist (f y₁) (f y₂) ≤ dist (f y₁) (f x) + dist (f y₂) (f x) : dist_triangle_right _ _ _ ... ≤ ε / 2 / 2 + ε / 2 / 2 : add_le_add (hδ₁ _ $ this hy₁) (hδ₁ _ $ this hy₂) ... = ε / 2 : add_halves _ }, - { have : continuous_within_at (λ δ, (2 * δ) ^ (n + 1) * ∥f' x (pi.single i 1)∥) + { have : continuous_within_at (λ δ, (2 * δ) ^ (n + 1) * ‖f' x (pi.single i 1)‖) (Ioi (0 : ℝ)) 0 := ((continuous_within_at_id.const_mul _).pow _).mul_const _, refine this.eventually (ge_mem_nhds _), simpa using half_pos ε0 } }, diff --git a/src/analysis/box_integral/integrability.lean b/src/analysis/box_integral/integrability.lean index 99f0439fc7155..d5ecdf76c0760 100644 --- a/src/analysis/box_integral/integrability.lean +++ b/src/analysis/box_integral/integrability.lean @@ -36,7 +36,7 @@ lemma has_integral_indicator_const (l : integration_params) (hl : l.bRiemann = f has_integral.{u v v} I l (s.indicator (λ _, y)) μ.to_box_additive.to_smul ((μ (s ∩ I)).to_real • y) := begin - refine has_integral_of_mul (∥y∥) (λ ε ε0, _), + refine has_integral_of_mul (‖y‖) (λ ε ε0, _), lift ε to ℝ≥0 using ε0.le, rw nnreal.coe_pos at ε0, /- First we choose a closed set `F ⊆ s ∩ I.Icc` and an open set `U ⊇ s` such that both `(s ∩ I.Icc) \ F` and `U \ s` have measuer less than `ε`. -/ @@ -98,14 +98,14 @@ lemma has_integral_zero_of_ae_eq_zero {l : integration_params} {I : box ι} {f : (hl : l.bRiemann = ff) : has_integral.{u v v} I l f μ.to_box_additive.to_smul 0 := begin - /- Each set `{x | n < ∥f x∥ ≤ n + 1}`, `n : ℕ`, has measure zero. We cover it by an open set of + /- Each set `{x | n < ‖f x‖ ≤ n + 1}`, `n : ℕ`, has measure zero. We cover it by an open set of measure less than `ε / 2 ^ n / (n + 1)`. Then the norm of the integral sum is less than `ε`. -/ refine has_integral_iff.2 (λ ε ε0, _), lift ε to ℝ≥0 using ε0.lt.le, rw [gt_iff_lt, nnreal.coe_pos] at ε0, rcases nnreal.exists_pos_sum_of_countable ε0.ne' ℕ with ⟨δ, δ0, c, hδc, hcε⟩, haveI := fact.mk (I.measure_coe_lt_top μ), change μ.restrict I {x | f x ≠ 0} = 0 at hf, - set N : (ι → ℝ) → ℕ := λ x, ⌈∥f x∥⌉₊, + set N : (ι → ℝ) → ℕ := λ x, ⌈‖f x‖⌉₊, have N0 : ∀ {x}, N x = 0 ↔ f x = 0, by { intro x, simp [N] }, have : ∀ n, ∃ U ⊇ N ⁻¹' {n}, is_open U ∧ μ.restrict I U < δ n / n, { refine λ n, (N ⁻¹' {n}).exists_is_open_lt_of_lt _ _, @@ -127,7 +127,7 @@ begin rintro n -, dsimp [integral_sum], have : ∀ J ∈ π.filter (λ J, N (π.tag J) = n), - ∥(μ ↑J).to_real • f (π.tag J)∥ ≤ (μ J).to_real * n, + ‖(μ ↑J).to_real • f (π.tag J)‖ ≤ (μ J).to_real * n, { intros J hJ, rw tagged_prepartition.mem_filter at hJ, rw [norm_smul, real.norm_eq_abs, abs_of_nonneg ennreal.to_real_nonneg], exact mul_le_mul_of_nonneg_left (hJ.2 ▸ nat.le_ceil _) ennreal.to_real_nonneg }, @@ -214,7 +214,7 @@ begin have hfi' := λ n, ((f n).has_box_integral μ I l hl).integrable, have hfgi : tendsto (λ n, (f n).integral (μ.restrict I)) at_top (𝓝 $ ∫ x in I, g x ∂μ), from tendsto_integral_approx_on_of_measurable_of_range_subset hg.measurable hgi _ subset.rfl, - have hfg_mono : ∀ x {m n}, m ≤ n → ∥f n x - g x∥ ≤ ∥f m x - g x∥, + have hfg_mono : ∀ x {m n}, m ≤ n → ‖f n x - g x‖ ≤ ‖f m x - g x‖, { intros x m n hmn, rw [← dist_eq_norm, ← dist_eq_norm, dist_nndist, dist_nndist, nnreal.coe_le_coe, ← ennreal.coe_le_coe, ← edist_nndist, ← edist_nndist], @@ -223,9 +223,9 @@ begin to `r`, the integral sum is `(μ I + 1 + 1) * ε`-close to the Bochner integral. -/ refine has_integral_of_mul ((μ I).to_real + 1 + 1) (λ ε ε0, _), lift ε to ℝ≥0 using ε0.le, rw nnreal.coe_pos at ε0, have ε0' := ennreal.coe_pos.2 ε0, - /- Choose `N` such that the integral of `∥f N x - g x∥` is less than or equal to `ε`. -/ - obtain ⟨N₀, hN₀⟩ : ∃ N : ℕ, ∫ x in I, ∥f N x - g x∥ ∂μ ≤ ε, - { have : tendsto (λ n, ∫⁻ x in I, ∥f n x - g x∥₊ ∂μ) at_top (𝓝 0), + /- Choose `N` such that the integral of `‖f N x - g x‖` is less than or equal to `ε`. -/ + obtain ⟨N₀, hN₀⟩ : ∃ N : ℕ, ∫ x in I, ‖f N x - g x‖ ∂μ ≤ ε, + { have : tendsto (λ n, ∫⁻ x in I, ‖f n x - g x‖₊ ∂μ) at_top (𝓝 0), from simple_func.tendsto_approx_on_range_L1_nnnorm hg.measurable hgi, refine (this.eventually (ge_mem_nhds ε0')).exists.imp (λ N hN, _), exact integral_coe_le_of_lintegral_coe_le hN }, @@ -286,13 +286,13 @@ begin hNxn J hJ], exact (hfi _).mono_set (prepartition.le_of_mem _ hJ) } }, { /- For the last jump, we use the fact that the distance between `f (Nx x) x` and `g x` is less - than or equal to the distance between `f N₀ x` and `g x` and the integral of `∥f N₀ x - g x∥` + than or equal to the distance between `f N₀ x` and `g x` and the integral of `‖f N₀ x - g x‖` is less than or equal to `ε`. -/ refine le_trans _ hN₀, have hfi : ∀ n (J ∈ π), integrable_on (f n) ↑J μ, from λ n J hJ, (hfi n).mono_set (π.le_of_mem' J hJ), have hgi : ∀ J ∈ π, integrable_on g ↑J μ, from λ J hJ, hgi.mono_set (π.le_of_mem' J hJ), - have hfgi : ∀ n (J ∈ π), integrable_on (λ x, ∥f n x - g x∥) J μ, + have hfgi : ∀ n (J ∈ π), integrable_on (λ x, ‖f n x - g x‖) J μ, from λ n J hJ, ((hfi n J hJ).sub (hgi J hJ)).norm, rw [← hπp.Union_eq, prepartition.Union_def', integral_finset_bUnion π.boxes (λ J hJ, J.measurable_set_coe) π.pairwise_disjoint hgi, diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index ef2787919a58a..0eda0104ed0fe 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -772,7 +772,7 @@ lemma iterated_fderiv_within_zero_eq_comp : iterated_fderiv_within 𝕜 0 f s = (continuous_multilinear_curry_fin0 𝕜 E F).symm ∘ f := rfl lemma norm_iterated_fderiv_within_zero : - ∥iterated_fderiv_within 𝕜 0 f s x∥ = ∥f x∥ := + ‖iterated_fderiv_within 𝕜 0 f s x‖ = ‖f x‖ := by rw [iterated_fderiv_within_zero_eq_comp, linear_isometry_equiv.norm_map] lemma iterated_fderiv_within_succ_apply_left {n : ℕ} (m : fin (n + 1) → E): @@ -788,8 +788,8 @@ lemma iterated_fderiv_within_succ_eq_comp_left {n : ℕ} : ∘ (fderiv_within 𝕜 (iterated_fderiv_within 𝕜 n f s) s) := rfl lemma norm_fderiv_within_iterated_fderiv_within {n : ℕ} : - ∥fderiv_within 𝕜 (iterated_fderiv_within 𝕜 n f s) s x∥ = - ∥iterated_fderiv_within 𝕜 (n + 1) f s x∥ := + ‖fderiv_within 𝕜 (iterated_fderiv_within 𝕜 n f s) s x‖ = + ‖iterated_fderiv_within 𝕜 (n + 1) f s x‖ := by rw [iterated_fderiv_within_succ_eq_comp_left, linear_isometry_equiv.norm_map] theorem iterated_fderiv_within_succ_apply_right {n : ℕ} @@ -832,8 +832,8 @@ lemma iterated_fderiv_within_succ_eq_comp_right {n : ℕ} (hs : unique_diff_on by { ext m, rw iterated_fderiv_within_succ_apply_right hs hx, refl } lemma norm_iterated_fderiv_within_fderiv_within {n : ℕ} (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) : - ∥iterated_fderiv_within 𝕜 n (fderiv_within 𝕜 f s) s x∥ = - ∥iterated_fderiv_within 𝕜 (n + 1) f s x∥ := + ‖iterated_fderiv_within 𝕜 n (fderiv_within 𝕜 f s) s x‖ = + ‖iterated_fderiv_within 𝕜 (n + 1) f s x‖ := by rw [iterated_fderiv_within_succ_eq_comp_right hs hx, linear_isometry_equiv.norm_map] @[simp] lemma iterated_fderiv_within_one_apply @@ -1459,7 +1459,7 @@ lemma iterated_fderiv_zero_eq_comp : iterated_fderiv 𝕜 0 f = (continuous_multilinear_curry_fin0 𝕜 E F).symm ∘ f := rfl lemma norm_iterated_fderiv_zero : - ∥iterated_fderiv 𝕜 0 f x∥ = ∥f x∥ := + ‖iterated_fderiv 𝕜 0 f x‖ = ‖f x‖ := by rw [iterated_fderiv_zero_eq_comp, linear_isometry_equiv.norm_map] lemma iterated_fderiv_with_zero_eq : @@ -1478,7 +1478,7 @@ lemma iterated_fderiv_succ_eq_comp_left {n : ℕ} : ∘ (fderiv 𝕜 (iterated_fderiv 𝕜 n f)) := rfl lemma norm_fderiv_iterated_fderiv {n : ℕ} : - ∥fderiv 𝕜 (iterated_fderiv 𝕜 n f) x∥ = ∥iterated_fderiv 𝕜 (n + 1) f x∥ := + ‖fderiv 𝕜 (iterated_fderiv 𝕜 n f) x‖ = ‖iterated_fderiv 𝕜 (n + 1) f x‖ := by rw [iterated_fderiv_succ_eq_comp_left, linear_isometry_equiv.norm_map] lemma iterated_fderiv_within_univ {n : ℕ} : @@ -1535,7 +1535,7 @@ lemma iterated_fderiv_succ_eq_comp_right {n : ℕ} : by { ext m, rw iterated_fderiv_succ_apply_right, refl } lemma norm_iterated_fderiv_fderiv {n : ℕ} : - ∥iterated_fderiv 𝕜 n (fderiv 𝕜 f) x∥ = ∥iterated_fderiv 𝕜 (n + 1) f x∥ := + ‖iterated_fderiv 𝕜 n (fderiv 𝕜 f) x‖ = ‖iterated_fderiv 𝕜 (n + 1) f x‖ := by rw [iterated_fderiv_succ_eq_comp_right, linear_isometry_equiv.norm_map] @[simp] lemma iterated_fderiv_one_apply (m : (fin 1) → E) : @@ -3226,12 +3226,12 @@ lemma cont_diff.has_strict_deriv_at hf.cont_diff_at.has_strict_deriv_at hn /-- If `f` has a formal Taylor series `p` up to order `1` on `{x} ∪ s`, where `s` is a convex set, -and `∥p x 1∥₊ < K`, then `f` is `K`-Lipschitz in a neighborhood of `x` within `s`. -/ +and `‖p x 1‖₊ < K`, then `f` is `K`-Lipschitz in a neighborhood of `x` within `s`. -/ lemma has_ftaylor_series_up_to_on.exists_lipschitz_on_with_of_nnnorm_lt {E F : Type*} [normed_add_comm_group E] [normed_space ℝ E] [normed_add_comm_group F] [normed_space ℝ F] {f : E → F} {p : E → formal_multilinear_series ℝ E F} {s : set E} {x : E} (hf : has_ftaylor_series_up_to_on 1 f p (insert x s)) (hs : convex ℝ s) (K : ℝ≥0) - (hK : ∥p x 1∥₊ < K) : + (hK : ‖p x 1‖₊ < K) : ∃ t ∈ 𝓝[s] x, lipschitz_on_with K f t := begin set f' := λ y, continuous_multilinear_curry_fin1 ℝ E F (p y 1), @@ -3240,7 +3240,7 @@ begin have hcont : continuous_within_at f' s x, from (continuous_multilinear_curry_fin1 ℝ E F).continuous_at.comp_continuous_within_at ((hf.cont _ le_rfl _ (mem_insert _ _)).mono (subset_insert x s)), - replace hK : ∥f' x∥₊ < K, by simpa only [linear_isometry_equiv.nnnorm_map], + replace hK : ‖f' x‖₊ < K, by simpa only [linear_isometry_equiv.nnnorm_map], exact hs.exists_nhds_within_lipschitz_on_with_of_has_fderiv_within_at_of_nnnorm_lt (eventually_nhds_within_iff.2 $ eventually_of_forall hder) hcont K hK end @@ -3271,10 +3271,10 @@ begin exact ⟨K, t, hst, hft⟩ end -/-- If `f` is `C^1` at `x` and `K > ∥fderiv 𝕂 f x∥`, then `f` is `K`-Lipschitz in a neighborhood of +/-- If `f` is `C^1` at `x` and `K > ‖fderiv 𝕂 f x‖`, then `f` is `K`-Lipschitz in a neighborhood of `x`. -/ lemma cont_diff_at.exists_lipschitz_on_with_of_nnnorm_lt {f : E' → F'} {x : E'} - (hf : cont_diff_at 𝕂 1 f x) (K : ℝ≥0) (hK : ∥fderiv 𝕂 f x∥₊ < K) : + (hf : cont_diff_at 𝕂 1 f x) (K : ℝ≥0) (hK : ‖fderiv 𝕂 f x‖₊ < K) : ∃ t ∈ 𝓝 x, lipschitz_on_with K f t := (hf.has_strict_fderiv_at le_rfl).exists_lipschitz_on_with_of_nnnorm_lt K hK diff --git a/src/analysis/calculus/deriv.lean b/src/analysis/calculus/deriv.lean index 6cce64afeeb56..0db145990c9b4 100644 --- a/src/analysis/calculus/deriv.lean +++ b/src/analysis/calculus/deriv.lean @@ -231,7 +231,7 @@ iff.rfl theorem has_deriv_at_filter_iff_tendsto : has_deriv_at_filter f f' x L ↔ - tendsto (λ x' : 𝕜, ∥x' - x∥⁻¹ * ∥f x' - f x - (x' - x) • f'∥) L (𝓝 0) := + tendsto (λ x' : 𝕜, ‖x' - x‖⁻¹ * ‖f x' - f x - (x' - x) • f'‖) L (𝓝 0) := has_fderiv_at_filter_iff_tendsto theorem has_deriv_within_at_iff_is_o : @@ -240,7 +240,7 @@ theorem has_deriv_within_at_iff_is_o : iff.rfl theorem has_deriv_within_at_iff_tendsto : has_deriv_within_at f f' s x ↔ - tendsto (λ x', ∥x' - x∥⁻¹ * ∥f x' - f x - (x' - x) • f'∥) (𝓝[s] x) (𝓝 0) := + tendsto (λ x', ‖x' - x‖⁻¹ * ‖f x' - f x - (x' - x) • f'‖) (𝓝[s] x) (𝓝 0) := has_fderiv_at_filter_iff_tendsto theorem has_deriv_at_iff_is_o : @@ -248,7 +248,7 @@ theorem has_deriv_at_iff_is_o : iff.rfl theorem has_deriv_at_iff_tendsto : has_deriv_at f f' x ↔ - tendsto (λ x', ∥x' - x∥⁻¹ * ∥f x' - f x - (x' - x) • f'∥) (𝓝 x) (𝓝 0) := + tendsto (λ x', ‖x' - x‖⁻¹ * ‖f x' - f x - (x' - x) • f'‖) (𝓝 x) (𝓝 0) := has_fderiv_at_filter_iff_tendsto theorem has_strict_deriv_at.has_deriv_at (h : has_strict_deriv_at f f' x) : @@ -1000,7 +1000,7 @@ has_fderiv_at_filter.is_O_sub h theorem has_deriv_at_filter.is_O_sub_rev (hf : has_deriv_at_filter f f' x L) (hf' : f' ≠ 0) : (λ x', x' - x) =O[L] (λ x', f x' - f x) := -suffices antilipschitz_with ∥f'∥₊⁻¹ (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f'), from hf.is_O_sub_rev this, +suffices antilipschitz_with ‖f'‖₊⁻¹ (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f'), from hf.is_O_sub_rev this, add_monoid_hom_class.antilipschitz_of_bound (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f') $ λ x, by simp [norm_smul, ← div_eq_inv_mul, mul_div_cancel _ (mt norm_eq_zero.1 hf')] @@ -1806,7 +1806,7 @@ htff'.of_local_left_inverse (f.symm.continuous_at ha) hf' (f.eventually_right_in lemma has_deriv_at.eventually_ne (h : has_deriv_at f f' x) (hf' : f' ≠ 0) : ∀ᶠ z in 𝓝[≠] x, f z ≠ f x := (has_deriv_at_iff_has_fderiv_at.1 h).eventually_ne - ⟨∥f'∥⁻¹, λ z, by field_simp [norm_smul, mt norm_eq_zero.1 hf']⟩ + ⟨‖f'‖⁻¹, λ z, by field_simp [norm_smul, mt norm_eq_zero.1 hf']⟩ lemma has_deriv_at.tendsto_punctured_nhds (h : has_deriv_at f f' x) (hf' : f' ≠ 0) : tendsto f (𝓝[≠] x) (𝓝[≠] (f x)) := @@ -2137,18 +2137,18 @@ open metric variables {E : Type u} [normed_add_comm_group E] [normed_space ℝ E] {f : ℝ → E} {f' : E} {s : set ℝ} {x r : ℝ} -/-- If `f` has derivative `f'` within `s` at `x`, then for any `r > ∥f'∥` the ratio -`∥f z - f x∥ / ∥z - x∥` is less than `r` in some neighborhood of `x` within `s`. +/-- If `f` has derivative `f'` within `s` at `x`, then for any `r > ‖f'‖` the ratio +`‖f z - f x‖ / ‖z - x‖` is less than `r` in some neighborhood of `x` within `s`. In other words, the limit superior of this ratio as `z` tends to `x` along `s` -is less than or equal to `∥f'∥`. -/ +is less than or equal to `‖f'‖`. -/ lemma has_deriv_within_at.limsup_norm_slope_le - (hf : has_deriv_within_at f f' s x) (hr : ∥f'∥ < r) : - ∀ᶠ z in 𝓝[s] x, ∥z - x∥⁻¹ * ∥f z - f x∥ < r := + (hf : has_deriv_within_at f f' s x) (hr : ‖f'‖ < r) : + ∀ᶠ z in 𝓝[s] x, ‖z - x‖⁻¹ * ‖f z - f x‖ < r := begin have hr₀ : 0 < r, from lt_of_le_of_lt (norm_nonneg f') hr, - have A : ∀ᶠ z in 𝓝[s \ {x}] x, ∥(z - x)⁻¹ • (f z - f x)∥ ∈ Iio r, + have A : ∀ᶠ z in 𝓝[s \ {x}] x, ‖(z - x)⁻¹ • (f z - f x)‖ ∈ Iio r, from (has_deriv_within_at_iff_tendsto_slope.1 hf).norm (is_open.mem_nhds is_open_Iio hr), - have B : ∀ᶠ z in 𝓝[{x}] x, ∥(z - x)⁻¹ • (f z - f x)∥ ∈ Iio r, + have B : ∀ᶠ z in 𝓝[{x}] x, ‖(z - x)⁻¹ • (f z - f x)‖ ∈ Iio r, from mem_of_superset self_mem_nhds_within (singleton_subset_iff.2 $ by simp [hr₀]), have C := mem_sup.2 ⟨A, B⟩, @@ -2158,16 +2158,16 @@ begin exact λ _, id end -/-- If `f` has derivative `f'` within `s` at `x`, then for any `r > ∥f'∥` the ratio -`(∥f z∥ - ∥f x∥) / ∥z - x∥` is less than `r` in some neighborhood of `x` within `s`. +/-- If `f` has derivative `f'` within `s` at `x`, then for any `r > ‖f'‖` the ratio +`(‖f z‖ - ‖f x‖) / ‖z - x‖` is less than `r` in some neighborhood of `x` within `s`. In other words, the limit superior of this ratio as `z` tends to `x` along `s` -is less than or equal to `∥f'∥`. +is less than or equal to `‖f'‖`. This lemma is a weaker version of `has_deriv_within_at.limsup_norm_slope_le` -where `∥f z∥ - ∥f x∥` is replaced by `∥f z - f x∥`. -/ +where `‖f z‖ - ‖f x‖` is replaced by `‖f z - f x‖`. -/ lemma has_deriv_within_at.limsup_slope_norm_le - (hf : has_deriv_within_at f f' s x) (hr : ∥f'∥ < r) : - ∀ᶠ z in 𝓝[s] x, ∥z - x∥⁻¹ * (∥f z∥ - ∥f x∥) < r := + (hf : has_deriv_within_at f f' s x) (hr : ‖f'‖ < r) : + ∀ᶠ z in 𝓝[s] x, ‖z - x‖⁻¹ * (‖f z‖ - ‖f x‖) < r := begin apply (hf.limsup_norm_slope_le hr).mono, assume z hz, @@ -2175,30 +2175,30 @@ begin exact inv_nonneg.2 (norm_nonneg _) end -/-- If `f` has derivative `f'` within `(x, +∞)` at `x`, then for any `r > ∥f'∥` the ratio -`∥f z - f x∥ / ∥z - x∥` is frequently less than `r` as `z → x+0`. +/-- If `f` has derivative `f'` within `(x, +∞)` at `x`, then for any `r > ‖f'‖` the ratio +`‖f z - f x‖ / ‖z - x‖` is frequently less than `r` as `z → x+0`. In other words, the limit inferior of this ratio as `z` tends to `x+0` -is less than or equal to `∥f'∥`. See also `has_deriv_within_at.limsup_norm_slope_le` +is less than or equal to `‖f'‖`. See also `has_deriv_within_at.limsup_norm_slope_le` for a stronger version using limit superior and any set `s`. -/ lemma has_deriv_within_at.liminf_right_norm_slope_le - (hf : has_deriv_within_at f f' (Ici x) x) (hr : ∥f'∥ < r) : - ∃ᶠ z in 𝓝[>] x, ∥z - x∥⁻¹ * ∥f z - f x∥ < r := + (hf : has_deriv_within_at f f' (Ici x) x) (hr : ‖f'‖ < r) : + ∃ᶠ z in 𝓝[>] x, ‖z - x‖⁻¹ * ‖f z - f x‖ < r := (hf.Ioi_of_Ici.limsup_norm_slope_le hr).frequently -/-- If `f` has derivative `f'` within `(x, +∞)` at `x`, then for any `r > ∥f'∥` the ratio -`(∥f z∥ - ∥f x∥) / (z - x)` is frequently less than `r` as `z → x+0`. +/-- If `f` has derivative `f'` within `(x, +∞)` at `x`, then for any `r > ‖f'‖` the ratio +`(‖f z‖ - ‖f x‖) / (z - x)` is frequently less than `r` as `z → x+0`. In other words, the limit inferior of this ratio as `z` tends to `x+0` -is less than or equal to `∥f'∥`. +is less than or equal to `‖f'‖`. See also * `has_deriv_within_at.limsup_norm_slope_le` for a stronger version using limit superior and any set `s`; * `has_deriv_within_at.liminf_right_norm_slope_le` for a stronger version using - `∥f z - f x∥` instead of `∥f z∥ - ∥f x∥`. -/ + `‖f z - f x‖` instead of `‖f z‖ - ‖f x‖`. -/ lemma has_deriv_within_at.liminf_right_slope_norm_le - (hf : has_deriv_within_at f f' (Ici x) x) (hr : ∥f'∥ < r) : - ∃ᶠ z in 𝓝[>] x, (z - x)⁻¹ * (∥f z∥ - ∥f x∥) < r := + (hf : has_deriv_within_at f f' (Ici x) x) (hr : ‖f'‖ < r) : + ∃ᶠ z in 𝓝[>] x, (z - x)⁻¹ * (‖f z‖ - ‖f x‖) < r := begin have := (hf.Ioi_of_Ici.limsup_slope_norm_le hr).frequently, refine this.mp (eventually.mono self_mem_nhds_within _), diff --git a/src/analysis/calculus/extend_deriv.lean b/src/analysis/calculus/extend_deriv.lean index 8915ac113fd2d..4e07fb7c59b69 100644 --- a/src/analysis/calculus/extend_deriv.lean +++ b/src/analysis/calculus/extend_deriv.lean @@ -43,29 +43,29 @@ begin { rw ← closure_closure at hx, exact has_fderiv_within_at_of_not_mem_closure hx }, push_neg at hx, rw [has_fderiv_within_at, has_fderiv_at_filter, asymptotics.is_o_iff], - /- One needs to show that `∥f y - f x - f' (y - x)∥ ≤ ε ∥y - x∥` for `y` close to `x` in `closure + /- One needs to show that `‖f y - f x - f' (y - x)‖ ≤ ε ‖y - x‖` for `y` close to `x` in `closure s`, where `ε` is an arbitrary positive constant. By continuity of the functions, it suffices to prove this for nearby points inside `s`. In a neighborhood of `x`, the derivative of `f` is arbitrarily close to `f'` by assumption. The mean value inequality completes the proof. -/ assume ε ε_pos, - obtain ⟨δ, δ_pos, hδ⟩ : ∃ δ > 0, ∀ y ∈ s, dist y x < δ → ∥fderiv ℝ f y - f'∥ < ε, + obtain ⟨δ, δ_pos, hδ⟩ : ∃ δ > 0, ∀ y ∈ s, dist y x < δ → ‖fderiv ℝ f y - f'‖ < ε, by simpa [dist_zero_right] using tendsto_nhds_within_nhds.1 h ε ε_pos, set B := ball x δ, - suffices : ∀ y ∈ B ∩ (closure s), ∥f y - f x - (f' y - f' x)∥ ≤ ε * ∥y - x∥, + suffices : ∀ y ∈ B ∩ (closure s), ‖f y - f x - (f' y - f' x)‖ ≤ ε * ‖y - x‖, from mem_nhds_within_iff.2 ⟨δ, δ_pos, λy hy, by simpa using this y hy⟩, - suffices : ∀ p : E × E, p ∈ closure ((B ∩ s) ×ˢ (B ∩ s)) → ∥f p.2 - f p.1 - (f' p.2 - f' p.1)∥ - ≤ ε * ∥p.2 - p.1∥, + suffices : ∀ p : E × E, p ∈ closure ((B ∩ s) ×ˢ (B ∩ s)) → ‖f p.2 - f p.1 - (f' p.2 - f' p.1)‖ + ≤ ε * ‖p.2 - p.1‖, { rw closure_prod_eq at this, intros y y_in, apply this ⟨x, y⟩, have : B ∩ closure s ⊆ closure (B ∩ s), from is_open_ball.inter_closure, exact ⟨this ⟨mem_ball_self δ_pos, hx⟩, this y_in⟩ }, - have key : ∀ p : E × E, p ∈ (B ∩ s) ×ˢ (B ∩ s) → ∥f p.2 - f p.1 - (f' p.2 - f' p.1)∥ - ≤ ε * ∥p.2 - p.1∥, + have key : ∀ p : E × E, p ∈ (B ∩ s) ×ˢ (B ∩ s) → ‖f p.2 - f p.1 - (f' p.2 - f' p.1)‖ + ≤ ε * ‖p.2 - p.1‖, { rintros ⟨u, v⟩ ⟨u_in, v_in⟩, have conv : convex ℝ (B ∩ s) := (convex_ball _ _).inter s_conv, have diff : differentiable_on ℝ f (B ∩ s) := f_diff.mono (inter_subset_right _ _), - have bound : ∀ z ∈ (B ∩ s), ∥fderiv_within ℝ f (B ∩ s) z - f'∥ ≤ ε, + have bound : ∀ z ∈ (B ∩ s), ‖fderiv_within ℝ f (B ∩ s) z - f'‖ ≤ ε, { intros z z_in, convert le_of_lt (hδ _ z_in.2 z_in.1), have op : is_open (B ∩ s) := is_open_ball.inter s_open, diff --git a/src/analysis/calculus/fderiv.lean b/src/analysis/calculus/fderiv.lean index 48279257ea2c8..02fb80dc1caf0 100644 --- a/src/analysis/calculus/fderiv.lean +++ b/src/analysis/calculus/fderiv.lean @@ -211,7 +211,7 @@ this fact, for functions having a derivative within a set. Its specific formulat tangent cone related discussions. -/ theorem has_fderiv_within_at.lim (h : has_fderiv_within_at f f' s x) {α : Type*} (l : filter α) {c : α → 𝕜} {d : α → E} {v : E} (dtop : ∀ᶠ n in l, x + d n ∈ s) - (clim : tendsto (λ n, ∥c n∥) l at_top) + (clim : tendsto (λ n, ‖c n‖) l at_top) (cdlim : tendsto (λ n, c n • d n) l (𝓝 v)) : tendsto (λn, c n • (f (x + d n) - f x)) l (𝓝 (f' v)) := begin @@ -266,8 +266,8 @@ section fderiv_properties theorem has_fderiv_at_filter_iff_tendsto : has_fderiv_at_filter f f' x L ↔ - tendsto (λ x', ∥x' - x∥⁻¹ * ∥f x' - f x - f' (x' - x)∥) L (𝓝 0) := -have h : ∀ x', ∥x' - x∥ = 0 → ∥f x' - f x - f' (x' - x)∥ = 0, from λ x' hx', + tendsto (λ x', ‖x' - x‖⁻¹ * ‖f x' - f x - f' (x' - x)‖) L (𝓝 0) := +have h : ∀ x', ‖x' - x‖ = 0 → ‖f x' - f x - f' (x' - x)‖ = 0, from λ x' hx', by { rw [sub_eq_zero.1 (norm_eq_zero.1 hx')], simp }, begin unfold has_fderiv_at_filter, @@ -276,11 +276,11 @@ begin end theorem has_fderiv_within_at_iff_tendsto : has_fderiv_within_at f f' s x ↔ - tendsto (λ x', ∥x' - x∥⁻¹ * ∥f x' - f x - f' (x' - x)∥) (𝓝[s] x) (𝓝 0) := + tendsto (λ x', ‖x' - x‖⁻¹ * ‖f x' - f x - f' (x' - x)‖) (𝓝[s] x) (𝓝 0) := has_fderiv_at_filter_iff_tendsto theorem has_fderiv_at_iff_tendsto : has_fderiv_at f f' x ↔ - tendsto (λ x', ∥x' - x∥⁻¹ * ∥f x' - f x - f' (x' - x)∥) (𝓝 x) (𝓝 0) := + tendsto (λ x', ‖x' - x‖⁻¹ * ‖f x' - f x - f' (x' - x)‖) (𝓝 x) (𝓝 0) := has_fderiv_at_filter_iff_tendsto theorem has_fderiv_at_iff_is_o_nhds_zero : has_fderiv_at f f' x ↔ @@ -292,24 +292,24 @@ end /-- Converse to the mean value inequality: if `f` is differentiable at `x₀` and `C`-lipschitz on a neighborhood of `x₀` then it its derivative at `x₀` has norm bounded by `C`. This version -only assumes that `∥f x - f x₀∥ ≤ C * ∥x - x₀∥` in a neighborhood of `x`. -/ +only assumes that `‖f x - f x₀‖ ≤ C * ‖x - x₀‖` in a neighborhood of `x`. -/ lemma has_fderiv_at.le_of_lip' {f : E → F} {f' : E →L[𝕜] F} {x₀ : E} (hf : has_fderiv_at f f' x₀) - {C : ℝ} (hC₀ : 0 ≤ C) (hlip : ∀ᶠ x in 𝓝 x₀, ∥f x - f x₀∥ ≤ C * ∥x - x₀∥) : ∥f'∥ ≤ C := + {C : ℝ} (hC₀ : 0 ≤ C) (hlip : ∀ᶠ x in 𝓝 x₀, ‖f x - f x₀‖ ≤ C * ‖x - x₀‖) : ‖f'‖ ≤ C := begin refine le_of_forall_pos_le_add (λ ε ε0, op_norm_le_of_nhds_zero _ _), exact add_nonneg hC₀ ε0.le, rw [← map_add_left_nhds_zero x₀, eventually_map] at hlip, filter_upwards [is_o_iff.1 (has_fderiv_at_iff_is_o_nhds_zero.1 hf) ε0, hlip] with y hy hyC, rw add_sub_cancel' at hyC, - calc ∥f' y∥ ≤ ∥f (x₀ + y) - f x₀∥ + ∥f (x₀ + y) - f x₀ - f' y∥ : norm_le_insert _ _ - ... ≤ C * ∥y∥ + ε * ∥y∥ : add_le_add hyC hy - ... = (C + ε) * ∥y∥ : (add_mul _ _ _).symm + calc ‖f' y‖ ≤ ‖f (x₀ + y) - f x₀‖ + ‖f (x₀ + y) - f x₀ - f' y‖ : norm_le_insert _ _ + ... ≤ C * ‖y‖ + ε * ‖y‖ : add_le_add hyC hy + ... = (C + ε) * ‖y‖ : (add_mul _ _ _).symm end /-- Converse to the mean value inequality: if `f` is differentiable at `x₀` and `C`-lipschitz on a neighborhood of `x₀` then it its derivative at `x₀` has norm bounded by `C`. -/ lemma has_fderiv_at.le_of_lip {f : E → F} {f' : E →L[𝕜] F} {x₀ : E} (hf : has_fderiv_at f f' x₀) - {s : set E} (hs : s ∈ 𝓝 x₀) {C : ℝ≥0} (hlip : lipschitz_on_with C f s) : ∥f'∥ ≤ C := + {s : set E} (hs : s ∈ 𝓝 x₀) {C : ℝ≥0} (hlip : lipschitz_on_with C f s) : ‖f'‖ ≤ C := begin refine hf.le_of_lip' C.coe_nonneg _, filter_upwards [hs] with x hx using hlip.norm_sub_le hx (mem_of_mem_nhds hs), @@ -384,10 +384,10 @@ protected lemma has_strict_fderiv_at.differentiable_at (hf : has_strict_fderiv_a differentiable_at 𝕜 f x := hf.has_fderiv_at.differentiable_at -/-- If `f` is strictly differentiable at `x` with derivative `f'` and `K > ∥f'∥₊`, then `f` is +/-- If `f` is strictly differentiable at `x` with derivative `f'` and `K > ‖f'‖₊`, then `f` is `K`-Lipschitz in a neighborhood of `x`. -/ lemma has_strict_fderiv_at.exists_lipschitz_on_with_of_nnnorm_lt (hf : has_strict_fderiv_at f f' x) - (K : ℝ≥0) (hK : ∥f'∥₊ < K) : ∃ s ∈ 𝓝 x, lipschitz_on_with K f s := + (K : ℝ≥0) (hK : ‖f'‖₊ < K) : ∃ s ∈ 𝓝 x, lipschitz_on_with K f s := begin have := hf.add_is_O_with (f'.is_O_with_comp _ _) hK, simp only [sub_add_cancel, is_O_with] at this, @@ -405,7 +405,7 @@ lemma has_strict_fderiv_at.exists_lipschitz_on_with (hf : has_strict_fderiv_at f /-- Directional derivative agrees with `has_fderiv`. -/ lemma has_fderiv_at.lim (hf : has_fderiv_at f f' x) (v : E) {α : Type*} {c : α → 𝕜} - {l : filter α} (hc : tendsto (λ n, ∥c n∥) l at_top) : + {l : filter α} (hc : tendsto (λ n, ‖c n‖) l at_top) : tendsto (λ n, (c n) • (f (x + (c n)⁻¹ • v) - f x)) l (𝓝 (f' v)) := begin refine (has_fderiv_within_at_univ.2 hf).lim _ univ_mem hc _, @@ -491,7 +491,7 @@ funext $ λ x, (h x).fderiv on a neighborhood of `x₀` then it its derivative at `x₀` has norm bounded by `C`. Version using `fderiv`. -/ lemma fderiv_at.le_of_lip {f : E → F} {x₀ : E} (hf : differentiable_at 𝕜 f x₀) - {s : set E} (hs : s ∈ 𝓝 x₀) {C : ℝ≥0} (hlip : lipschitz_on_with C f s) : ∥fderiv 𝕜 f x₀∥ ≤ C := + {s : set E} (hs : s ∈ 𝓝 x₀) {C : ℝ≥0} (hlip : lipschitz_on_with C f s) : ‖fderiv 𝕜 f x₀‖ ≤ C := hf.has_fderiv_at.le_of_lip hs hlip lemma has_fderiv_within_at.fderiv_within @@ -647,14 +647,14 @@ by by_cases hx : differentiable_within_at 𝕜 f t x; simp [fderiv_within_zero_of_not_differentiable_within_at, *] lemma asymptotics.is_O.has_fderiv_within_at {s : set E} {x₀ : E} {n : ℕ} - (h : f =O[𝓝[s] x₀] λ x, ∥x - x₀∥^n) (hx₀ : x₀ ∈ s) (hn : 1 < n) : + (h : f =O[𝓝[s] x₀] λ x, ‖x - x₀‖^n) (hx₀ : x₀ ∈ s) (hn : 1 < n) : has_fderiv_within_at f (0 : E →L[𝕜] F) s x₀ := by simp_rw [has_fderiv_within_at, has_fderiv_at_filter, h.eq_zero_of_norm_pow_within hx₀ $ zero_lt_one.trans hn, zero_apply, sub_zero, h.trans_is_o ((is_o_pow_sub_sub x₀ hn).mono nhds_within_le_nhds)] lemma asymptotics.is_O.has_fderiv_at {x₀ : E} {n : ℕ} - (h : f =O[𝓝 x₀] λ x, ∥x - x₀∥^n) (hn : 1 < n) : + (h : f =O[𝓝 x₀] λ x, ‖x - x₀‖^n) (hn : 1 < n) : has_fderiv_at f (0 : E →L[𝕜] F) x₀ := begin rw [← nhds_within_univ] at h, @@ -2158,10 +2158,10 @@ lemma is_bounded_bilinear_map.has_strict_fderiv_at (h : is_bounded_bilinear_map begin rw has_strict_fderiv_at, set T := (E × F) × (E × F), - have : (λ q : T, b (q.1 - q.2)) =o[𝓝 (p, p)] (λ q : T, ∥q.1 - q.2∥ * 1), + have : (λ q : T, b (q.1 - q.2)) =o[𝓝 (p, p)] (λ q : T, ‖q.1 - q.2‖ * 1), { refine (h.is_O'.comp_tendsto le_top).trans_is_o _, simp only [(∘)], - refine (is_O_refl (λ q : T, ∥q.1 - q.2∥) _).mul_is_o (is_o.norm_left $ (is_o_one_iff _).2 _), + refine (is_O_refl (λ q : T, ‖q.1 - q.2‖) _).mul_is_o (is_o.norm_left $ (is_o_one_iff _).2 _), rw [← sub_self p], exact continuous_at_fst.sub continuous_at_snd }, simp only [mul_one, is_o_norm_right] at this, @@ -2918,7 +2918,7 @@ lemma local_homeomorph.has_fderiv_at_symm (f : local_homeomorph E F) {f' : E ≃ htff'.of_local_left_inverse (f.symm.continuous_at ha) (f.eventually_right_inverse ha) lemma has_fderiv_within_at.eventually_ne (h : has_fderiv_within_at f f' s x) - (hf' : ∃ C, ∀ z, ∥z∥ ≤ C * ∥f' z∥) : + (hf' : ∃ C, ∀ z, ‖z‖ ≤ C * ‖f' z‖) : ∀ᶠ z in 𝓝[s \ {x}] x, f z ≠ f x := begin rw [nhds_within, diff_eq, ← inf_principal, ← inf_assoc, eventually_inf_principal], @@ -2928,7 +2928,7 @@ begin simpa [not_imp_not, sub_eq_zero] using (A.trans this.is_O_symm).eq_zero_imp end -lemma has_fderiv_at.eventually_ne (h : has_fderiv_at f f' x) (hf' : ∃ C, ∀ z, ∥z∥ ≤ C * ∥f' z∥) : +lemma has_fderiv_at.eventually_ne (h : has_fderiv_at f f' x) (hf' : ∃ C, ∀ z, ‖z‖ ≤ C * ‖f' z‖) : ∀ᶠ z in 𝓝[≠] x, f z ≠ f x := by simpa only [compl_eq_univ_diff] using (has_fderiv_within_at_univ.2 h).eventually_ne hf' @@ -2947,12 +2947,12 @@ variables {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] variables {f : E → F} {f' : E →L[ℝ] F} {x : E} theorem has_fderiv_at_filter_real_equiv {L : filter E} : - tendsto (λ x' : E, ∥x' - x∥⁻¹ * ∥f x' - f x - f' (x' - x)∥) L (𝓝 0) ↔ - tendsto (λ x' : E, ∥x' - x∥⁻¹ • (f x' - f x - f' (x' - x))) L (𝓝 0) := + tendsto (λ x' : E, ‖x' - x‖⁻¹ * ‖f x' - f x - f' (x' - x)‖) L (𝓝 0) ↔ + tendsto (λ x' : E, ‖x' - x‖⁻¹ • (f x' - f x - f' (x' - x))) L (𝓝 0) := begin symmetry, rw [tendsto_iff_norm_tendsto_zero], refine tendsto_congr (λ x', _), - have : ∥x' - x∥⁻¹ ≥ 0, from inv_nonneg.mpr (norm_nonneg _), + have : ‖x' - x‖⁻¹ ≥ 0, from inv_nonneg.mpr (norm_nonneg _), simp [norm_smul, abs_of_nonneg this] end diff --git a/src/analysis/calculus/fderiv_analytic.lean b/src/analysis/calculus/fderiv_analytic.lean index 42afc7b3f2353..58483589cadf0 100644 --- a/src/analysis/calculus/fderiv_analytic.lean +++ b/src/analysis/calculus/fderiv_analytic.lean @@ -30,7 +30,7 @@ lemma has_fpower_series_at.has_strict_fderiv_at (h : has_fpower_series_at f p x) has_strict_fderiv_at f (continuous_multilinear_curry_fin1 𝕜 E F (p 1)) x := begin refine h.is_O_image_sub_norm_mul_norm_sub.trans_is_o (is_o.of_norm_right _), - refine is_o_iff_exists_eq_mul.2 ⟨λ y, ∥y - (x, x)∥, _, eventually_eq.rfl⟩, + refine is_o_iff_exists_eq_mul.2 ⟨λ y, ‖y - (x, x)‖, _, eventually_eq.rfl⟩, refine (continuous_id.sub continuous_const).norm.tendsto' _ _ _, rw [_root_.id, sub_self, norm_zero] end @@ -64,12 +64,12 @@ lemma analytic_on.differentiable_on (h : analytic_on 𝕜 f s) : λ y hy, (h y hy).differentiable_within_at lemma has_fpower_series_on_ball.has_fderiv_at [complete_space F] - (h : has_fpower_series_on_ball f p x r) {y : E} (hy : (∥y∥₊ : ℝ≥0∞) < r) : + (h : has_fpower_series_on_ball f p x r) {y : E} (hy : (‖y‖₊ : ℝ≥0∞) < r) : has_fderiv_at f (continuous_multilinear_curry_fin1 𝕜 E F (p.change_origin y 1)) (x + y) := (h.change_origin hy).has_fpower_series_at.has_fderiv_at lemma has_fpower_series_on_ball.fderiv_eq [complete_space F] - (h : has_fpower_series_on_ball f p x r) {y : E} (hy : (∥y∥₊ : ℝ≥0∞) < r) : + (h : has_fpower_series_on_ball f p x r) {y : E} (hy : (‖y‖₊ : ℝ≥0∞) < r) : fderiv 𝕜 f (x + y) = continuous_multilinear_curry_fin1 𝕜 E F (p.change_origin y 1) := (h.has_fderiv_at hy).fderiv diff --git a/src/analysis/calculus/fderiv_measurable.lean b/src/analysis/calculus/fderiv_measurable.lean index ffdc0a0fc0a54..4a4a653356c2b 100644 --- a/src/analysis/calculus/fderiv_measurable.lean +++ b/src/analysis/calculus/fderiv_measurable.lean @@ -50,7 +50,7 @@ differentiability exactly says that the map is well approximated by `L`). This i For the other direction, the difficulty is that `L` in the union may depend on `ε, r, s`. The key point is that, in fact, it doesn't depend too much on them. First, if `x` belongs both to `A (L, r, ε)` and `A (L', r, ε)`, then `L` and `L'` have to be close on a shell, and thus -`∥L - L'∥` is bounded by `ε` (see `norm_sub_le_of_mem_A`). Assume now `x ∈ D`. If one has two maps +`‖L - L'‖` is bounded by `ε` (see `norm_sub_le_of_mem_A`). Assume now `x ∈ D`. If one has two maps `L` and `L'` such that `x` belongs to `A (L, r, ε)` and to `A (L', r', ε')`, one deduces that `L` is close to `L'` by arguing as follows. Consider another scale `s` smaller than `r` and `r'`. Take a linear map `L₁` that approximates `f` around `x` both at scales `r` and `s` w.r.t. `ε` (it exists as @@ -104,7 +104,7 @@ namespace fderiv_measurable_aux at scale `r` by the linear map `L`, up to an error `ε`. We tweak the definition to make sure that this is an open set.-/ def A (f : E → F) (L : E →L[𝕜] F) (r ε : ℝ) : set E := -{x | ∃ r' ∈ Ioc (r/2) r, ∀ y z ∈ ball x r', ∥f z - f y - L (z-y)∥ ≤ ε * r} +{x | ∃ r' ∈ Ioc (r/2) r, ∀ y z ∈ ball x r', ‖f z - f y - L (z-y)‖ ≤ ε * r} /-- The set `B f K r s ε` is the set of points `x` around which there exists a continuous linear map `L` belonging to `K` (a given set of continuous linear maps) that approximates well the @@ -143,7 +143,7 @@ end lemma le_of_mem_A {r ε : ℝ} {L : E →L[𝕜] F} {x : E} (hx : x ∈ A f L r ε) {y z : E} (hy : y ∈ closed_ball x (r/2)) (hz : z ∈ closed_ball x (r/2)) : - ∥f z - f y - L (z-y)∥ ≤ ε * r := + ‖f z - f y - L (z-y)‖ ≤ ε * r := begin rcases hx with ⟨r', r'mem, hr'⟩, exact hr' _ ((mem_closed_ball.1 hy).trans_lt r'mem.1) _ ((mem_closed_ball.1 hz).trans_lt r'mem.1) @@ -158,12 +158,12 @@ begin refine ⟨R, R_pos, λ r hr, _⟩, have : r ∈ Ioc (r/2) r := ⟨half_lt_self hr.1, le_rfl⟩, refine ⟨r, this, λ y hy z hz, _⟩, - calc ∥f z - f y - (fderiv 𝕜 f x) (z - y)∥ - = ∥(f z - f x - (fderiv 𝕜 f x) (z - x)) - (f y - f x - (fderiv 𝕜 f x) (y - x))∥ : + calc ‖f z - f y - (fderiv 𝕜 f x) (z - y)‖ + = ‖(f z - f x - (fderiv 𝕜 f x) (z - x)) - (f y - f x - (fderiv 𝕜 f x) (y - x))‖ : by { congr' 1, simp only [continuous_linear_map.map_sub], abel } - ... ≤ ∥(f z - f x - (fderiv 𝕜 f x) (z - x))∥ + ∥f y - f x - (fderiv 𝕜 f x) (y - x)∥ : + ... ≤ ‖(f z - f x - (fderiv 𝕜 f x) (z - x))‖ + ‖f y - f x - (fderiv 𝕜 f x) (y - x)‖ : norm_sub_le _ _ - ... ≤ ε / 2 * ∥z - x∥ + ε / 2 * ∥y - x∥ : + ... ≤ ε / 2 * ‖z - x‖ + ε / 2 * ‖y - x‖ : add_le_add (hR _ (lt_trans (mem_ball.1 hz) hr.2)) (hR _ (lt_trans (mem_ball.1 hy) hr.2)) ... ≤ ε / 2 * r + ε / 2 * r : add_le_add @@ -172,19 +172,19 @@ begin ... = ε * r : by ring end -lemma norm_sub_le_of_mem_A {c : 𝕜} (hc : 1 < ∥c∥) +lemma norm_sub_le_of_mem_A {c : 𝕜} (hc : 1 < ‖c‖) {r ε : ℝ} (hε : 0 < ε) (hr : 0 < r) {x : E} {L₁ L₂ : E →L[𝕜] F} - (h₁ : x ∈ A f L₁ r ε) (h₂ : x ∈ A f L₂ r ε) : ∥L₁ - L₂∥ ≤ 4 * ∥c∥ * ε := + (h₁ : x ∈ A f L₁ r ε) (h₂ : x ∈ A f L₂ r ε) : ‖L₁ - L₂‖ ≤ 4 * ‖c‖ * ε := begin - have : 0 ≤ 4 * ∥c∥ * ε := + have : 0 ≤ 4 * ‖c‖ * ε := mul_nonneg (mul_nonneg (by norm_num : (0 : ℝ) ≤ 4) (norm_nonneg _)) hε.le, refine op_norm_le_of_shell (half_pos hr) this hc _, assume y ley ylt, rw [div_div, div_le_iff' (mul_pos (by norm_num : (0 : ℝ) < 2) (zero_lt_one.trans hc))] at ley, - calc ∥(L₁ - L₂) y∥ - = ∥(f (x + y) - f x - L₂ ((x + y) - x)) - (f (x + y) - f x - L₁ ((x + y) - x))∥ : by simp - ... ≤ ∥(f (x + y) - f x - L₂ ((x + y) - x))∥ + ∥(f (x + y) - f x - L₁ ((x + y) - x))∥ : + calc ‖(L₁ - L₂) y‖ + = ‖(f (x + y) - f x - L₂ ((x + y) - x)) - (f (x + y) - f x - L₁ ((x + y) - x))‖ : by simp + ... ≤ ‖(f (x + y) - f x - L₂ ((x + y) - x))‖ + ‖(f (x + y) - f x - L₁ ((x + y) - x))‖ : norm_sub_le _ _ ... ≤ ε * r + ε * r : begin @@ -197,8 +197,8 @@ begin { simp only [dist_eq_norm, add_sub_cancel', mem_closed_ball, ylt.le] } }, end ... = 2 * ε * r : by ring - ... ≤ 2 * ε * (2 * ∥c∥ * ∥y∥) : mul_le_mul_of_nonneg_left ley (mul_nonneg (by norm_num) hε.le) - ... = 4 * ∥c∥ * ε * ∥y∥ : by ring + ... ≤ 2 * ε * (2 * ‖c‖ * ‖y‖) : mul_le_mul_of_nonneg_left ley (mul_nonneg (by norm_num) hε.le) + ... = 4 * ‖c‖ * ε * ‖y‖ : by ring end /-- Easy inclusion: a differentiability point with derivative in `K` belongs to `D f K`. -/ @@ -223,7 +223,7 @@ lemma D_subset_differentiable_set {K : set (E →L[𝕜] F)} (hK : is_complete K begin have P : ∀ {n : ℕ}, (0 : ℝ) < (1/2) ^ n := pow_pos (by norm_num), rcases normed_field.exists_one_lt_norm 𝕜 with ⟨c, hc⟩, - have cpos : 0 < ∥c∥ := lt_trans zero_lt_one hc, + have cpos : 0 < ‖c‖ := lt_trans zero_lt_one hc, assume x hx, have : ∀ (e : ℕ), ∃ (n : ℕ), ∀ p q, n ≤ p → n ≤ q → ∃ L ∈ K, x ∈ A f L ((1/2) ^ p) ((1/2) ^ e) ∩ A f L ((1/2) ^ q) ((1/2) ^ e), @@ -244,55 +244,55 @@ begin `2 ^ (- r)`. And `L e' p' r` is close to `L e' p' q'` as both approximate `f` at scale `2 ^ (- p')`. -/ have M : ∀ e p q e' p' q', n e ≤ p → n e ≤ q → n e' ≤ p' → n e' ≤ q' → e ≤ e' → - ∥L e p q - L e' p' q'∥ ≤ 12 * ∥c∥ * (1/2) ^ e, + ‖L e p q - L e' p' q'‖ ≤ 12 * ‖c‖ * (1/2) ^ e, { assume e p q e' p' q' hp hq hp' hq' he', let r := max (n e) (n e'), have I : ((1:ℝ)/2)^e' ≤ (1/2)^e := pow_le_pow_of_le_one (by norm_num) (by norm_num) he', - have J1 : ∥L e p q - L e p r∥ ≤ 4 * ∥c∥ * (1/2)^e, + have J1 : ‖L e p q - L e p r‖ ≤ 4 * ‖c‖ * (1/2)^e, { have I1 : x ∈ A f (L e p q) ((1 / 2) ^ p) ((1/2)^e) := (hn e p q hp hq).2.1, have I2 : x ∈ A f (L e p r) ((1 / 2) ^ p) ((1/2)^e) := (hn e p r hp (le_max_left _ _)).2.1, exact norm_sub_le_of_mem_A hc P P I1 I2 }, - have J2 : ∥L e p r - L e' p' r∥ ≤ 4 * ∥c∥ * (1/2)^e, + have J2 : ‖L e p r - L e' p' r‖ ≤ 4 * ‖c‖ * (1/2)^e, { have I1 : x ∈ A f (L e p r) ((1 / 2) ^ r) ((1/2)^e) := (hn e p r hp (le_max_left _ _)).2.2, have I2 : x ∈ A f (L e' p' r) ((1 / 2) ^ r) ((1/2)^e') := (hn e' p' r hp' (le_max_right _ _)).2.2, exact norm_sub_le_of_mem_A hc P P I1 (A_mono _ _ I I2) }, - have J3 : ∥L e' p' r - L e' p' q'∥ ≤ 4 * ∥c∥ * (1/2)^e, + have J3 : ‖L e' p' r - L e' p' q'‖ ≤ 4 * ‖c‖ * (1/2)^e, { have I1 : x ∈ A f (L e' p' r) ((1 / 2) ^ p') ((1/2)^e') := (hn e' p' r hp' (le_max_right _ _)).2.1, have I2 : x ∈ A f (L e' p' q') ((1 / 2) ^ p') ((1/2)^e') := (hn e' p' q' hp' hq').2.1, exact norm_sub_le_of_mem_A hc P P (A_mono _ _ I I1) (A_mono _ _ I I2) }, - calc ∥L e p q - L e' p' q'∥ - = ∥(L e p q - L e p r) + (L e p r - L e' p' r) + (L e' p' r - L e' p' q')∥ : + calc ‖L e p q - L e' p' q'‖ + = ‖(L e p q - L e p r) + (L e p r - L e' p' r) + (L e' p' r - L e' p' q')‖ : by { congr' 1, abel } - ... ≤ ∥L e p q - L e p r∥ + ∥L e p r - L e' p' r∥ + ∥L e' p' r - L e' p' q'∥ : + ... ≤ ‖L e p q - L e p r‖ + ‖L e p r - L e' p' r‖ + ‖L e' p' r - L e' p' q'‖ : le_trans (norm_add_le _ _) (add_le_add_right (norm_add_le _ _) _) - ... ≤ 4 * ∥c∥ * (1/2)^e + 4 * ∥c∥ * (1/2)^e + 4 * ∥c∥ * (1/2)^e : + ... ≤ 4 * ‖c‖ * (1/2)^e + 4 * ‖c‖ * (1/2)^e + 4 * ‖c‖ * (1/2)^e : by apply_rules [add_le_add] - ... = 12 * ∥c∥ * (1/2)^e : by ring }, + ... = 12 * ‖c‖ * (1/2)^e : by ring }, /- For definiteness, use `L0 e = L e (n e) (n e)`, to have a single sequence. We claim that this is a Cauchy sequence. -/ let L0 : ℕ → (E →L[𝕜] F) := λ e, L e (n e) (n e), have : cauchy_seq L0, { rw metric.cauchy_seq_iff', assume ε εpos, - obtain ⟨e, he⟩ : ∃ (e : ℕ), (1/2) ^ e < ε / (12 * ∥c∥) := + obtain ⟨e, he⟩ : ∃ (e : ℕ), (1/2) ^ e < ε / (12 * ‖c‖) := exists_pow_lt_of_lt_one (div_pos εpos (mul_pos (by norm_num) cpos)) (by norm_num), refine ⟨e, λ e' he', _⟩, rw [dist_comm, dist_eq_norm], - calc ∥L0 e - L0 e'∥ - ≤ 12 * ∥c∥ * (1/2)^e : M _ _ _ _ _ _ le_rfl le_rfl le_rfl le_rfl he' - ... < 12 * ∥c∥ * (ε / (12 * ∥c∥)) : + calc ‖L0 e - L0 e'‖ + ≤ 12 * ‖c‖ * (1/2)^e : M _ _ _ _ _ _ le_rfl le_rfl le_rfl le_rfl he' + ... < 12 * ‖c‖ * (ε / (12 * ‖c‖)) : mul_lt_mul' le_rfl he (le_of_lt P) (mul_pos (by norm_num) cpos) ... = ε : by { field_simp [(by norm_num : (12 : ℝ) ≠ 0), ne_of_gt cpos], ring } }, /- As it is Cauchy, the sequence `L0` converges, to a limit `f'` in `K`.-/ obtain ⟨f', f'K, hf'⟩ : ∃ f' ∈ K, tendsto L0 at_top (𝓝 f') := cauchy_seq_tendsto_of_is_complete hK (λ e, (hn e (n e) (n e) le_rfl le_rfl).1) this, - have Lf' : ∀ e p, n e ≤ p → ∥L e (n e) p - f'∥ ≤ 12 * ∥c∥ * (1/2)^e, + have Lf' : ∀ e p, n e ≤ p → ‖L e (n e) p - f'‖ ≤ 12 * ‖c‖ * (1/2)^e, { assume e p hp, apply le_of_tendsto (tendsto_const_nhds.sub hf').norm, rw eventually_at_top, @@ -305,21 +305,21 @@ begin this makes it possible to cover all scales, and thus to obtain a good linear approximation in the whole ball of radius `(1/2)^(n e)`. -/ assume ε εpos, - have pos : 0 < 4 + 12 * ∥c∥ := + have pos : 0 < 4 + 12 * ‖c‖ := add_pos_of_pos_of_nonneg (by norm_num) (mul_nonneg (by norm_num) (norm_nonneg _)), - obtain ⟨e, he⟩ : ∃ (e : ℕ), (1 / 2) ^ e < ε / (4 + 12 * ∥c∥) := + obtain ⟨e, he⟩ : ∃ (e : ℕ), (1 / 2) ^ e < ε / (4 + 12 * ‖c‖) := exists_pow_lt_of_lt_one (div_pos εpos pos) (by norm_num), rw eventually_nhds_iff_ball, refine ⟨(1/2) ^ (n e + 1), P, λ y hy, _⟩, -- We need to show that `f (x + y) - f x - f' y` is small. For this, we will work at scale - -- `k` where `k` is chosen with `∥y∥ ∼ 2 ^ (-k)`. + -- `k` where `k` is chosen with `‖y‖ ∼ 2 ^ (-k)`. by_cases y_pos : y = 0, {simp [y_pos] }, - have yzero : 0 < ∥y∥ := norm_pos_iff.mpr y_pos, - have y_lt : ∥y∥ < (1/2) ^ (n e + 1), by simpa using mem_ball_iff_norm.1 hy, - have yone : ∥y∥ ≤ 1 := + have yzero : 0 < ‖y‖ := norm_pos_iff.mpr y_pos, + have y_lt : ‖y‖ < (1/2) ^ (n e + 1), by simpa using mem_ball_iff_norm.1 hy, + have yone : ‖y‖ ≤ 1 := le_trans (y_lt.le) (pow_le_one _ (by norm_num) (by norm_num)), -- define the scale `k`. - obtain ⟨k, hk, h'k⟩ : ∃ (k : ℕ), (1/2) ^ (k + 1) < ∥y∥ ∧ ∥y∥ ≤ (1/2) ^ k := + obtain ⟨k, hk, h'k⟩ : ∃ (k : ℕ), (1/2) ^ (k + 1) < ‖y‖ ∧ ‖y‖ ≤ (1/2) ^ k := exists_nat_pow_near_of_lt_one yzero yone (by norm_num : (0 : ℝ) < 1/2) (by norm_num : (1 : ℝ)/2 < 1), -- the scale is large enough (as `y` is small enough) @@ -333,31 +333,31 @@ begin rw km at hk h'k, -- `f` is well approximated by `L e (n e) k` at the relevant scale -- (in fact, we use `m = k - 1` instead of `k` because of the precise definition of `A`). - have J1 : ∥f (x + y) - f x - L e (n e) m ((x + y) - x)∥ ≤ (1/2) ^ e * (1/2) ^ m, + have J1 : ‖f (x + y) - f x - L e (n e) m ((x + y) - x)‖ ≤ (1/2) ^ e * (1/2) ^ m, { apply le_of_mem_A (hn e (n e) m le_rfl m_ge).2.2, { simp only [mem_closed_ball, dist_self], exact div_nonneg (le_of_lt P) (zero_le_two) }, { simpa only [dist_eq_norm, add_sub_cancel', mem_closed_ball, pow_succ', mul_one_div] using h'k } }, - have J2 : ∥f (x + y) - f x - L e (n e) m y∥ ≤ 4 * (1/2) ^ e * ∥y∥ := calc - ∥f (x + y) - f x - L e (n e) m y∥ ≤ (1/2) ^ e * (1/2) ^ m : + have J2 : ‖f (x + y) - f x - L e (n e) m y‖ ≤ 4 * (1/2) ^ e * ‖y‖ := calc + ‖f (x + y) - f x - L e (n e) m y‖ ≤ (1/2) ^ e * (1/2) ^ m : by simpa only [add_sub_cancel'] using J1 ... = 4 * (1/2) ^ e * (1/2) ^ (m + 2) : by { field_simp, ring_exp } - ... ≤ 4 * (1/2) ^ e * ∥y∥ : + ... ≤ 4 * (1/2) ^ e * ‖y‖ : mul_le_mul_of_nonneg_left (le_of_lt hk) (mul_nonneg (by norm_num) (le_of_lt P)), -- use the previous estimates to see that `f (x + y) - f x - f' y` is small. - calc ∥f (x + y) - f x - f' y∥ - = ∥(f (x + y) - f x - L e (n e) m y) + (L e (n e) m - f') y∥ : + calc ‖f (x + y) - f x - f' y‖ + = ‖(f (x + y) - f x - L e (n e) m y) + (L e (n e) m - f') y‖ : congr_arg _ (by simp) - ... ≤ 4 * (1/2) ^ e * ∥y∥ + 12 * ∥c∥ * (1/2) ^ e * ∥y∥ : + ... ≤ 4 * (1/2) ^ e * ‖y‖ + 12 * ‖c‖ * (1/2) ^ e * ‖y‖ : norm_add_le_of_le J2 ((le_op_norm _ _).trans (mul_le_mul_of_nonneg_right (Lf' _ _ m_ge) (norm_nonneg _))) - ... = (4 + 12 * ∥c∥) * ∥y∥ * (1/2) ^ e : by ring - ... ≤ (4 + 12 * ∥c∥) * ∥y∥ * (ε / (4 + 12 * ∥c∥)) : + ... = (4 + 12 * ‖c‖) * ‖y‖ * (1/2) ^ e : by ring + ... ≤ (4 + 12 * ‖c‖) * ‖y‖ * (ε / (4 + 12 * ‖c‖)) : mul_le_mul_of_nonneg_left he.le (mul_nonneg (add_nonneg (by norm_num) (mul_nonneg (by norm_num) (norm_nonneg _))) (norm_nonneg _)) - ... = ε * ∥y∥ : by { field_simp [ne_of_gt pos], ring } }, + ... = ε * ‖y‖ : by { field_simp [ne_of_gt pos], ring } }, rw ← this.fderiv at f'K, exact ⟨this.differentiable_at, f'K⟩ end @@ -441,7 +441,7 @@ namespace right_deriv_measurable_aux at scale `r` by the linear map `h ↦ h • L`, up to an error `ε`. We tweak the definition to make sure that this is open on the right. -/ def A (f : ℝ → F) (L : F) (r ε : ℝ) : set ℝ := -{x | ∃ r' ∈ Ioc (r/2) r, ∀ y z ∈ Icc x (x + r'), ∥f z - f y - (z-y) • L∥ ≤ ε * r} +{x | ∃ r' ∈ Ioc (r/2) r, ∀ y z ∈ Icc x (x + r'), ‖f z - f y - (z-y) • L‖ ≤ ε * r} /-- The set `B f K r s ε` is the set of points `x` around which there exists a vector `L` belonging to `K` (a given set of vectors) such that `h • L` approximates well `f (x + h)` @@ -493,7 +493,7 @@ end lemma le_of_mem_A {r ε : ℝ} {L : F} {x : ℝ} (hx : x ∈ A f L r ε) {y z : ℝ} (hy : y ∈ Icc x (x + r/2)) (hz : z ∈ Icc x (x + r/2)) : - ∥f z - f y - (z-y) • L∥ ≤ ε * r := + ‖f z - f y - (z-y) • L‖ ≤ ε * r := begin rcases hx with ⟨r', r'mem, hr'⟩, have A : x + r / 2 ≤ x + r', by linarith [r'mem.1], @@ -510,14 +510,14 @@ begin refine ⟨m - x, by linarith [show x < m, from xm], λ r hr, _⟩, have : r ∈ Ioc (r/2) r := ⟨half_lt_self hr.1, le_rfl⟩, refine ⟨r, this, λ y hy z hz, _⟩, - calc ∥f z - f y - (z - y) • deriv_within f (Ici x) x∥ - = ∥(f z - f x - (z - x) • deriv_within f (Ici x) x) - - (f y - f x - (y - x) • deriv_within f (Ici x) x)∥ : + calc ‖f z - f y - (z - y) • deriv_within f (Ici x) x‖ + = ‖(f z - f x - (z - x) • deriv_within f (Ici x) x) + - (f y - f x - (y - x) • deriv_within f (Ici x) x)‖ : by { congr' 1, simp only [sub_smul], abel } - ... ≤ ∥f z - f x - (z - x) • deriv_within f (Ici x) x∥ - + ∥f y - f x - (y - x) • deriv_within f (Ici x) x∥ : + ... ≤ ‖f z - f x - (z - x) • deriv_within f (Ici x) x‖ + + ‖f y - f x - (y - x) • deriv_within f (Ici x) x‖ : norm_sub_le _ _ - ... ≤ ε / 2 * ∥z - x∥ + ε / 2 * ∥y - x∥ : + ... ≤ ε / 2 * ‖z - x‖ + ε / 2 * ‖y - x‖ : add_le_add (hm ⟨hz.1, hz.2.trans_lt (by linarith [hr.2])⟩) (hm ⟨hy.1, hy.2.trans_lt (by linarith [hr.2])⟩) ... ≤ ε / 2 * r + ε / 2 * r : @@ -535,15 +535,15 @@ end lemma norm_sub_le_of_mem_A {r x : ℝ} (hr : 0 < r) (ε : ℝ) {L₁ L₂ : F} - (h₁ : x ∈ A f L₁ r ε) (h₂ : x ∈ A f L₂ r ε) : ∥L₁ - L₂∥ ≤ 4 * ε := + (h₁ : x ∈ A f L₁ r ε) (h₂ : x ∈ A f L₂ r ε) : ‖L₁ - L₂‖ ≤ 4 * ε := begin - suffices H : ∥(r/2) • (L₁ - L₂)∥ ≤ (r / 2) * (4 * ε), + suffices H : ‖(r/2) • (L₁ - L₂)‖ ≤ (r / 2) * (4 * ε), by rwa [norm_smul, real.norm_of_nonneg (half_pos hr).le, mul_le_mul_left (half_pos hr)] at H, calc - ∥(r/2) • (L₁ - L₂)∥ - = ∥(f (x + r/2) - f x - (x + r/2 - x) • L₂) - (f (x + r/2) - f x - (x + r/2 - x) • L₁)∥ : + ‖(r/2) • (L₁ - L₂)‖ + = ‖(f (x + r/2) - f x - (x + r/2 - x) • L₂) - (f (x + r/2) - f x - (x + r/2 - x) • L₁)‖ : by simp [smul_sub] - ... ≤ ∥f (x + r/2) - f x - (x + r/2 - x) • L₂∥ + ∥f (x + r/2) - f x - (x + r/2 - x) • L₁∥ : + ... ≤ ‖f (x + r/2) - f x - (x + r/2 - x) • L₂‖ + ‖f (x + r/2) - f x - (x + r/2 - x) • L₁‖ : norm_sub_le _ _ ... ≤ ε * r + ε * r : begin @@ -598,32 +598,32 @@ begin `2 ^ (- r)`. And `L e' p' r` is close to `L e' p' q'` as both approximate `f` at scale `2 ^ (- p')`. -/ have M : ∀ e p q e' p' q', n e ≤ p → n e ≤ q → n e' ≤ p' → n e' ≤ q' → e ≤ e' → - ∥L e p q - L e' p' q'∥ ≤ 12 * (1/2) ^ e, + ‖L e p q - L e' p' q'‖ ≤ 12 * (1/2) ^ e, { assume e p q e' p' q' hp hq hp' hq' he', let r := max (n e) (n e'), have I : ((1:ℝ)/2)^e' ≤ (1/2)^e := pow_le_pow_of_le_one (by norm_num) (by norm_num) he', - have J1 : ∥L e p q - L e p r∥ ≤ 4 * (1/2)^e, + have J1 : ‖L e p q - L e p r‖ ≤ 4 * (1/2)^e, { have I1 : x ∈ A f (L e p q) ((1 / 2) ^ p) ((1/2)^e) := (hn e p q hp hq).2.1, have I2 : x ∈ A f (L e p r) ((1 / 2) ^ p) ((1/2)^e) := (hn e p r hp (le_max_left _ _)).2.1, exact norm_sub_le_of_mem_A P _ I1 I2 }, - have J2 : ∥L e p r - L e' p' r∥ ≤ 4 * (1/2)^e, + have J2 : ‖L e p r - L e' p' r‖ ≤ 4 * (1/2)^e, { have I1 : x ∈ A f (L e p r) ((1 / 2) ^ r) ((1/2)^e) := (hn e p r hp (le_max_left _ _)).2.2, have I2 : x ∈ A f (L e' p' r) ((1 / 2) ^ r) ((1/2)^e') := (hn e' p' r hp' (le_max_right _ _)).2.2, exact norm_sub_le_of_mem_A P _ I1 (A_mono _ _ I I2) }, - have J3 : ∥L e' p' r - L e' p' q'∥ ≤ 4 * (1/2)^e, + have J3 : ‖L e' p' r - L e' p' q'‖ ≤ 4 * (1/2)^e, { have I1 : x ∈ A f (L e' p' r) ((1 / 2) ^ p') ((1/2)^e') := (hn e' p' r hp' (le_max_right _ _)).2.1, have I2 : x ∈ A f (L e' p' q') ((1 / 2) ^ p') ((1/2)^e') := (hn e' p' q' hp' hq').2.1, exact norm_sub_le_of_mem_A P _ (A_mono _ _ I I1) (A_mono _ _ I I2) }, - calc ∥L e p q - L e' p' q'∥ - = ∥(L e p q - L e p r) + (L e p r - L e' p' r) + (L e' p' r - L e' p' q')∥ : + calc ‖L e p q - L e' p' q'‖ + = ‖(L e p q - L e p r) + (L e p r - L e' p' r) + (L e' p' r - L e' p' q')‖ : by { congr' 1, abel } - ... ≤ ∥L e p q - L e p r∥ + ∥L e p r - L e' p' r∥ + ∥L e' p' r - L e' p' q'∥ : + ... ≤ ‖L e p q - L e p r‖ + ‖L e p r - L e' p' r‖ + ‖L e' p' r - L e' p' q'‖ : le_trans (norm_add_le _ _) (add_le_add_right (norm_add_le _ _) _) ... ≤ 4 * (1/2)^e + 4 * (1/2)^e + 4 * (1/2)^e : by apply_rules [add_le_add] @@ -638,7 +638,7 @@ begin exists_pow_lt_of_lt_one (div_pos εpos (by norm_num)) (by norm_num), refine ⟨e, λ e' he', _⟩, rw [dist_comm, dist_eq_norm], - calc ∥L0 e - L0 e'∥ + calc ‖L0 e - L0 e'‖ ≤ 12 * (1/2)^e : M _ _ _ _ _ _ le_rfl le_rfl le_rfl le_rfl he' ... < 12 * (ε / 12) : mul_lt_mul' le_rfl he (le_of_lt P) (by norm_num) @@ -646,7 +646,7 @@ begin /- As it is Cauchy, the sequence `L0` converges, to a limit `f'` in `K`.-/ obtain ⟨f', f'K, hf'⟩ : ∃ f' ∈ K, tendsto L0 at_top (𝓝 f') := cauchy_seq_tendsto_of_is_complete hK (λ e, (hn e (n e) (n e) le_rfl le_rfl).1) this, - have Lf' : ∀ e p, n e ≤ p → ∥L e (n e) p - f'∥ ≤ 12 * (1/2)^e, + have Lf' : ∀ e p, n e ≤ p → ‖L e (n e) p - f'‖ ≤ 12 * (1/2)^e, { assume e p hp, apply le_of_tendsto (tendsto_const_nhds.sub hf').norm, rw eventually_at_top, @@ -666,7 +666,7 @@ begin zero_lt_one], filter_upwards [Icc_mem_nhds_within_Ici xmem] with y hy, -- We need to show that `f y - f x - f' (y - x)` is small. For this, we will work at scale - -- `k` where `k` is chosen with `∥y - x∥ ∼ 2 ^ (-k)`. + -- `k` where `k` is chosen with `‖y - x‖ ∼ 2 ^ (-k)`. rcases eq_or_lt_of_le hy.1 with rfl|xy, { simp only [sub_self, zero_smul, norm_zero, mul_zero]}, have yzero : 0 < y - x := sub_pos.2 xy, @@ -687,8 +687,8 @@ begin rw km at hk h'k, -- `f` is well approximated by `L e (n e) k` at the relevant scale -- (in fact, we use `m = k - 1` instead of `k` because of the precise definition of `A`). - have J : ∥f y - f x - (y - x) • L e (n e) m∥ ≤ 4 * (1/2) ^ e * ∥y - x∥ := calc - ∥f y - f x - (y - x) • L e (n e) m∥ ≤ (1/2) ^ e * (1/2) ^ m : + have J : ‖f y - f x - (y - x) • L e (n e) m‖ ≤ 4 * (1/2) ^ e * ‖y - x‖ := calc + ‖f y - f x - (y - x) • L e (n e) m‖ ≤ (1/2) ^ e * (1/2) ^ m : begin apply le_of_mem_A (hn e (n e) m le_rfl m_ge).2.2, { simp only [one_div, inv_pow, left_mem_Icc, le_add_iff_nonneg_right], @@ -699,16 +699,16 @@ begin ... = 4 * (1/2) ^ e * (1/2) ^ (m + 2) : by { field_simp, ring_exp } ... ≤ 4 * (1/2) ^ e * (y - x) : mul_le_mul_of_nonneg_left (le_of_lt hk) (mul_nonneg (by norm_num) (le_of_lt P)) - ... = 4 * (1/2) ^ e * ∥y - x∥ : by rw [real.norm_of_nonneg yzero.le], - calc ∥f y - f x - (y - x) • f'∥ - = ∥(f y - f x - (y - x) • L e (n e) m) + (y - x) • (L e (n e) m - f')∥ : + ... = 4 * (1/2) ^ e * ‖y - x‖ : by rw [real.norm_of_nonneg yzero.le], + calc ‖f y - f x - (y - x) • f'‖ + = ‖(f y - f x - (y - x) • L e (n e) m) + (y - x) • (L e (n e) m - f')‖ : by simp only [smul_sub, sub_add_sub_cancel] - ... ≤ 4 * (1/2) ^ e * ∥y - x∥ + ∥y - x∥ * (12 * (1/2) ^ e) : norm_add_le_of_le J + ... ≤ 4 * (1/2) ^ e * ‖y - x‖ + ‖y - x‖ * (12 * (1/2) ^ e) : norm_add_le_of_le J (by { rw [norm_smul], exact mul_le_mul_of_nonneg_left (Lf' _ _ m_ge) (norm_nonneg _) }) - ... = 16 * ∥y - x∥ * (1/2) ^ e : by ring - ... ≤ 16 * ∥y - x∥ * (ε / 16) : + ... = 16 * ‖y - x‖ * (1/2) ^ e : by ring + ... ≤ 16 * ‖y - x‖ * (ε / 16) : mul_le_mul_of_nonneg_left he.le (mul_nonneg (by norm_num) (norm_nonneg _)) - ... = ε * ∥y - x∥ : by ring }, + ... = ε * ‖y - x‖ : by ring }, rw ← this.deriv_within (unique_diff_on_Ici x x le_rfl) at f'K, exact ⟨this.differentiable_within_at, f'K⟩, end diff --git a/src/analysis/calculus/fderiv_symmetric.lean b/src/analysis/calculus/fderiv_symmetric.lean index 41b99929f84cc..34f00e42edab5 100644 --- a/src/analysis/calculus/fderiv_symmetric.lean +++ b/src/analysis/calculus/fderiv_symmetric.lean @@ -72,15 +72,15 @@ lemma convex.taylor_approx_two_segment (λ h : ℝ, f (x + h • v + h • w) - f (x + h • v) - h • f' x w - h^2 • f'' v w - (h^2/2) • f'' w w) =o[𝓝[>] 0] (λ h, h^2) := begin - -- it suffices to check that the expression is bounded by `ε * ((∥v∥ + ∥w∥) * ∥w∥) * h^2` for + -- it suffices to check that the expression is bounded by `ε * ((‖v‖ + ‖w‖) * ‖w‖) * h^2` for -- small enough `h`, for any positive `ε`. - apply is_o.trans_is_O (is_o_iff.2 (λ ε εpos, _)) (is_O_const_mul_self ((∥v∥ + ∥w∥) * ∥w∥) _ _), + apply is_o.trans_is_O (is_o_iff.2 (λ ε εpos, _)) (is_O_const_mul_self ((‖v‖ + ‖w‖) * ‖w‖) _ _), -- consider a ball of radius `δ` around `x` in which the Taylor approximation for `f''` is -- good up to `δ`. rw [has_fderiv_within_at, has_fderiv_at_filter, is_o_iff] at hx, rcases metric.mem_nhds_within_iff.1 (hx εpos) with ⟨δ, δpos, sδ⟩, - have E1 : ∀ᶠ h in 𝓝[>] (0:ℝ), h * (∥v∥ + ∥w∥) < δ, - { have : filter.tendsto (λ h, h * (∥v∥ + ∥w∥)) (𝓝[>] (0:ℝ)) (𝓝 (0 * (∥v∥ + ∥w∥))) := + have E1 : ∀ᶠ h in 𝓝[>] (0:ℝ), h * (‖v‖ + ‖w‖) < δ, + { have : filter.tendsto (λ h, h * (‖v‖ + ‖w‖)) (𝓝[>] (0:ℝ)) (𝓝 (0 * (‖v‖ + ‖w‖))) := (continuous_id.mul continuous_const).continuous_within_at, apply (tendsto_order.1 this).2 δ, simpa only [zero_mul] using δpos }, @@ -127,19 +127,19 @@ begin ring }, apply_rules [has_deriv_at.has_deriv_within_at, has_deriv_at.smul_const, has_deriv_at_id', has_deriv_at.pow, has_deriv_at.mul_const] } }, - -- check that `g'` is uniformly bounded, with a suitable bound `ε * ((∥v∥ + ∥w∥) * ∥w∥) * h^2`. - have g'_bound : ∀ t ∈ Ico (0 : ℝ) 1, ∥g' t∥ ≤ ε * ((∥v∥ + ∥w∥) * ∥w∥) * h^2, + -- check that `g'` is uniformly bounded, with a suitable bound `ε * ((‖v‖ + ‖w‖) * ‖w‖) * h^2`. + have g'_bound : ∀ t ∈ Ico (0 : ℝ) 1, ‖g' t‖ ≤ ε * ((‖v‖ + ‖w‖) * ‖w‖) * h^2, { assume t ht, - have I : ∥h • v + (t * h) • w∥ ≤ h * (∥v∥ + ∥w∥) := calc - ∥h • v + (t * h) • w∥ ≤ ∥h • v∥ + ∥(t * h) • w∥ : norm_add_le _ _ - ... = h * ∥v∥ + t * (h * ∥w∥) : + have I : ‖h • v + (t * h) • w‖ ≤ h * (‖v‖ + ‖w‖) := calc + ‖h • v + (t * h) • w‖ ≤ ‖h • v‖ + ‖(t * h) • w‖ : norm_add_le _ _ + ... = h * ‖v‖ + t * (h * ‖w‖) : by simp only [norm_smul, real.norm_eq_abs, hpos.le, abs_of_nonneg, abs_mul, ht.left, mul_assoc] - ... ≤ h * ∥v∥ + 1 * (h * ∥w∥) : + ... ≤ h * ‖v‖ + 1 * (h * ‖w‖) : add_le_add le_rfl (mul_le_mul_of_nonneg_right ht.2.le (mul_nonneg hpos.le (norm_nonneg _))) - ... = h * (∥v∥ + ∥w∥) : by ring, - calc ∥g' t∥ = ∥(f' (x + h • v + (t * h) • w) - f' x - f'' (h • v + (t * h) • w)) (h • w)∥ : + ... = h * (‖v‖ + ‖w‖) : by ring, + calc ‖g' t‖ = ‖(f' (x + h • v + (t * h) • w) - f' x - f'' (h • v + (t * h) • w)) (h • w)‖ : begin rw hg', have : h * (t * h) = t * (h * h), by ring, @@ -147,9 +147,9 @@ begin continuous_linear_map.add_apply, pi.smul_apply, smul_sub, smul_add, smul_smul, ← sub_sub, continuous_linear_map.coe_smul', pi.sub_apply, continuous_linear_map.map_smul, this] end - ... ≤ ∥f' (x + h • v + (t * h) • w) - f' x - f'' (h • v + (t * h) • w)∥ * ∥h • w∥ : + ... ≤ ‖f' (x + h • v + (t * h) • w) - f' x - f'' (h • v + (t * h) • w)‖ * ‖h • w‖ : continuous_linear_map.le_op_norm _ _ - ... ≤ (ε * ∥h • v + (t * h) • w∥) * (∥h • w∥) : + ... ≤ (ε * ‖h • v + (t * h) • w‖) * (‖h • w‖) : begin apply mul_le_mul_of_nonneg_right _ (norm_nonneg _), have H : x + h • v + (t * h) • w ∈ metric.ball x δ ∩ interior s, @@ -158,7 +158,7 @@ begin exact I.trans_lt hδ }, simpa only [mem_set_of_eq, add_assoc x, add_sub_cancel'] using sδ H, end - ... ≤ (ε * (∥h • v∥ + ∥h • w∥)) * (∥h • w∥) : + ... ≤ (ε * (‖h • v‖ + ‖h • w‖)) * (‖h • w‖) : begin apply mul_le_mul_of_nonneg_right _ (norm_nonneg _), apply mul_le_mul_of_nonneg_left _ (εpos.le), @@ -167,10 +167,10 @@ begin simp only [norm_smul, real.norm_eq_abs, abs_mul, abs_of_nonneg, ht.1, hpos.le, mul_assoc], exact mul_le_of_le_one_left (mul_nonneg hpos.le (norm_nonneg _)) ht.2.le, end - ... = ε * ((∥v∥ + ∥w∥) * ∥w∥) * h^2 : + ... = ε * ((‖v‖ + ‖w‖) * ‖w‖) * h^2 : by { simp only [norm_smul, real.norm_eq_abs, abs_mul, abs_of_nonneg, hpos.le], ring } }, -- conclude using the mean value inequality - have I : ∥g 1 - g 0∥ ≤ ε * ((∥v∥ + ∥w∥) * ∥w∥) * h^2, by simpa only [mul_one, sub_zero] using + have I : ‖g 1 - g 0‖ ≤ ε * ((‖v‖ + ‖w‖) * ‖w‖) * h^2, by simpa only [mul_one, sub_zero] using norm_image_sub_le_of_norm_deriv_le_segment' g_deriv g'_bound 1 (right_mem_Icc.2 zero_le_one), convert I using 1, { congr' 1, diff --git a/src/analysis/calculus/formal_multilinear_series.lean b/src/analysis/calculus/formal_multilinear_series.lean index 4cf449fdf27b6..4c7400f75d276 100644 --- a/src/analysis/calculus/formal_multilinear_series.lean +++ b/src/analysis/calculus/formal_multilinear_series.lean @@ -260,7 +260,7 @@ by rw [← mk_pi_field_coeff_eq p, continuous_multilinear_map.mk_pi_field_eq_zer @[simp] lemma apply_eq_pow_smul_coeff : p n (λ _, z) = z ^ n • p.coeff n := by simp -@[simp] lemma norm_apply_eq_norm_coef : ∥p n∥ = ∥coeff p n∥ := +@[simp] lemma norm_apply_eq_norm_coef : ‖p n‖ = ‖coeff p n‖ := by rw [← mk_pi_field_coeff_eq p, continuous_multilinear_map.norm_mk_pi_field] end coef diff --git a/src/analysis/calculus/inverse.lean b/src/analysis/calculus/inverse.lean index b9ea5c303d7f6..3ef859903176b 100644 --- a/src/analysis/calculus/inverse.lean +++ b/src/analysis/calculus/inverse.lean @@ -47,7 +47,7 @@ the inverse function, are formulated in `fderiv.lean`, `deriv.lean`, and `cont_d In the section about `approximates_linear_on` we introduce some `local notation` to make formulas shorter: -* by `N` we denote `∥f'⁻¹∥`; +* by `N` we denote `‖f'⁻¹‖`; * by `g` we denote the auxiliary contracting map `x ↦ x + f'.symm (y - f x)` used to prove that `{x | f x = y}` is nonempty. @@ -76,7 +76,7 @@ open continuous_linear_map (id) /-! ### Non-linear maps close to affine maps -In this section we study a map `f` such that `∥f x - f y - f' (x - y)∥ ≤ c * ∥x - y∥` on an open set +In this section we study a map `f` such that `‖f x - f y - f' (x - y)‖ ≤ c * ‖x - y‖` on an open set `s`, where `f' : E →L[𝕜] F` is a continuous linear map and `c` is suitably small. Maps of this type behave like `f a + f' (x - a)` near each `a ∈ s`. @@ -100,13 +100,13 @@ lemmas. This approach makes it possible -/ /-- We say that `f` approximates a continuous linear map `f'` on `s` with constant `c`, -if `∥f x - f y - f' (x - y)∥ ≤ c * ∥x - y∥` whenever `x, y ∈ s`. +if `‖f x - f y - f' (x - y)‖ ≤ c * ‖x - y‖` whenever `x, y ∈ s`. This predicate is defined to facilitate the splitting of the inverse function theorem into small lemmas. Some of these lemmas can be useful, e.g., to prove that the inverse function is defined on a specific set. -/ def approximates_linear_on (f : E → F) (f' : E →L[𝕜] F) (s : set E) (c : ℝ≥0) : Prop := -∀ (x ∈ s) (y ∈ s), ∥f x - f y - f' (x - y)∥ ≤ c * ∥x - y∥ +∀ (x ∈ s) (y ∈ s), ‖f x - f y - f' (x - y)‖ ≤ c * ‖x - y‖ @[simp] lemma approximates_linear_on_empty (f : E → F) (f' : E →L[𝕜] F) (c : ℝ≥0) : approximates_linear_on f f' ∅ c := @@ -153,7 +153,7 @@ begin end protected lemma lipschitz (hf : approximates_linear_on f f' s c) : - lipschitz_with (∥f'∥₊ + c) (s.restrict f) := + lipschitz_with (‖f'‖₊ + c) (s.restrict f) := by simpa only [restrict_apply, add_sub_cancel'_right] using (f'.lipschitz.restrict s).add hf.lipschitz_sub @@ -226,12 +226,12 @@ begin dist (f (g z)) y ≤ c * f'symm.nnnorm * dist (f z) y, { assume z hz hgz, set v := f'symm (y - f z) with hv, - calc dist (f (g z)) y = ∥f (z + v) - y∥ : by rw [dist_eq_norm] - ... = ∥f (z + v) - f z - f' v + f' v - (y - f z)∥ : by { congr' 1, abel } - ... = ∥f (z + v) - f z - f' ((z + v) - z)∥ : + calc dist (f (g z)) y = ‖f (z + v) - y‖ : by rw [dist_eq_norm] + ... = ‖f (z + v) - f z - f' v + f' v - (y - f z)‖ : by { congr' 1, abel } + ... = ‖f (z + v) - f z - f' ((z + v) - z)‖ : by simp only [continuous_linear_map.nonlinear_right_inverse.right_inv, add_sub_cancel', sub_add_cancel] - ... ≤ c * ∥(z + v) - z∥ : hf _ (hε hgz) _ (hε hz) + ... ≤ c * ‖(z + v) - z‖ : hf _ (hε hgz) _ (hε hz) ... ≤ c * (f'symm.nnnorm * dist (f z) y) : begin apply mul_le_mul_of_nonneg_left _ (nnreal.coe_nonneg c), simpa [hv, dist_eq_norm'] using f'symm.bound (y - f z), @@ -351,12 +351,12 @@ end locally_onto /-! From now on we assume that `f` approximates an invertible continuous linear map `f : E ≃L[𝕜] F`. -We also assume that either `E = {0}`, or `c < ∥f'⁻¹∥⁻¹`. We use `N` as an abbreviation for `∥f'⁻¹∥`. +We also assume that either `E = {0}`, or `c < ‖f'⁻¹‖⁻¹`. We use `N` as an abbreviation for `‖f'⁻¹‖`. -/ variables {f' : E ≃L[𝕜] F} {s : set E} {c : ℝ≥0} -local notation `N` := ∥(f'.symm : F →L[𝕜] E)∥₊ +local notation `N` := ‖(f'.symm : F →L[𝕜] E)‖₊ protected lemma antilipschitz (hf : approximates_linear_on f (f' : E →L[𝕜] F) s c) (hc : subsingleton E ∨ c < N⁻¹) : @@ -428,24 +428,24 @@ begin rcases (mem_image _ _ _).1 hx with ⟨x', x's, rfl⟩, rcases (mem_image _ _ _).1 hy with ⟨y', y's, rfl⟩, rw [← Af x', ← Af y', A.left_inv x's, A.left_inv y's], - calc ∥x' - y' - (f'.symm) (A x' - A y')∥ - ≤ N * ∥f' (x' - y' - (f'.symm) (A x' - A y'))∥ : + calc ‖x' - y' - (f'.symm) (A x' - A y')‖ + ≤ N * ‖f' (x' - y' - (f'.symm) (A x' - A y'))‖ : (f' : E →L[𝕜] F).bound_of_antilipschitz f'.antilipschitz _ - ... = N * ∥A y' - A x' - f' (y' - x')∥ : + ... = N * ‖A y' - A x' - f' (y' - x')‖ : begin congr' 2, simp only [continuous_linear_equiv.apply_symm_apply, continuous_linear_equiv.map_sub], abel, end - ... ≤ N * (c * ∥y' - x'∥) : + ... ≤ N * (c * ‖y' - x'‖) : mul_le_mul_of_nonneg_left (hf _ y's _ x's) (nnreal.coe_nonneg _) - ... ≤ N * (c * (((N⁻¹ - c)⁻¹ : ℝ≥0) * ∥A y' - A x'∥)) : + ... ≤ N * (c * (((N⁻¹ - c)⁻¹ : ℝ≥0) * ‖A y' - A x'‖)) : begin apply_rules [mul_le_mul_of_nonneg_left, nnreal.coe_nonneg], rw [← dist_eq_norm, ← dist_eq_norm], exact (hf.antilipschitz hc).le_mul_dist ⟨y', y's⟩ ⟨x', x's⟩, end - ... = (N * (N⁻¹ - c)⁻¹ * c : ℝ≥0) * ∥A x' - A y'∥ : + ... = (N * (N⁻¹ - c)⁻¹ * c : ℝ≥0) * ‖A x' - A y'‖ : by { simp only [norm_sub_rev, nonneg.coe_mul], ring } end @@ -485,7 +485,7 @@ lemma exists_homeomorph_extension {E : Type*} [normed_add_comm_group E] [normed_ {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] [finite_dimensional ℝ F] {s : set E} {f : E → F} {f' : E ≃L[ℝ] F} {c : ℝ≥0} (hf : approximates_linear_on f (f' : E →L[ℝ] F) s c) - (hc : subsingleton E ∨ lipschitz_extension_constant F * c < (∥(f'.symm : F →L[ℝ] E)∥₊)⁻¹) : + (hc : subsingleton E ∨ lipschitz_extension_constant F * c < (‖(f'.symm : F →L[ℝ] E)‖₊)⁻¹) : ∃ g : E ≃ₜ F, eq_on f g s := begin -- the difference `f - f'` is Lipschitz on `s`. It can be extended to a Lipschitz function `u` @@ -572,7 +572,7 @@ variables [cs : complete_space E] {f : E → F} {f' : E ≃L[𝕜] F} {a : E} lemma approximates_deriv_on_open_nhds (hf : has_strict_fderiv_at f (f' : E →L[𝕜] F) a) : ∃ (s : set E) (hs : a ∈ s ∧ is_open s), - approximates_linear_on f (f' : E →L[𝕜] F) s (∥(f'.symm : F →L[𝕜] E)∥₊⁻¹ / 2) := + approximates_linear_on f (f' : E →L[𝕜] F) s (‖(f'.symm : F →L[𝕜] E)‖₊⁻¹ / 2) := begin refine ((nhds_basis_opens a).exists_iff _).1 _, exact (λ s t, approximates_linear_on.mono_set), diff --git a/src/analysis/calculus/local_extr.lean b/src/analysis/calculus/local_extr.lean index b410bf11ec8b3..6694ec4d17df0 100644 --- a/src/analysis/calculus/local_extr.lean +++ b/src/analysis/calculus/local_extr.lean @@ -73,7 +73,7 @@ variables {E : Type u} [normed_add_comm_group E] [normed_space ℝ E] {f : E → {f' : E →L[ℝ] ℝ} /-- "Positive" tangent cone to `s` at `x`; the only difference from `tangent_cone_at` -is that we require `c n → ∞` instead of `∥c n∥ → ∞`. One can think about `pos_tangent_cone_at` +is that we require `c n → ∞` instead of `‖c n‖ → ∞`. One can think about `pos_tangent_cone_at` as `tangent_cone_at nnreal` but we have no theory of normed semifields yet. -/ def pos_tangent_cone_at (s : set E) (x : E) : set E := {y : E | ∃(c : ℕ → ℝ) (d : ℕ → E), (∀ᶠ n in at_top, x + d n ∈ s) ∧ @@ -119,7 +119,7 @@ lemma is_local_max_on.has_fderiv_within_at_nonpos {s : set E} (h : is_local_max_ f' y ≤ 0 := begin rcases hy with ⟨c, d, hd, hc, hcd⟩, - have hc' : tendsto (λ n, ∥c n∥) at_top at_top, + have hc' : tendsto (λ n, ‖c n‖) at_top at_top, from tendsto_at_top_mono (λ n, le_abs_self _) hc, refine le_of_tendsto (hf.lim at_top hd hc' hcd) _, replace hd : tendsto (λ n, a + d n) at_top (𝓝[s] (a + 0)), diff --git a/src/analysis/calculus/mean_value.lean b/src/analysis/calculus/mean_value.lean index 3a9e0bc673cae..f64fd61e91c71 100644 --- a/src/analysis/calculus/mean_value.lean +++ b/src/analysis/calculus/mean_value.lean @@ -20,20 +20,20 @@ In this file we prove the following facts: so they work both for real and complex derivatives. * `image_le_of*`, `image_norm_le_of_*` : several similar lemmas deducing `f x ≤ B x` or - `∥f x∥ ≤ B x` from upper estimates on `f'` or `∥f'∥`, respectively. These lemmas differ by + `‖f x‖ ≤ B x` from upper estimates on `f'` or `‖f'‖`, respectively. These lemmas differ by their assumptions: * `of_liminf_*` lemmas assume that limit inferior of some ratio is less than `B' x`; * `of_deriv_right_*`, `of_norm_deriv_right_*` lemmas assume that the right derivative or its norm is less than `B' x`; - * `of_*_lt_*` lemmas assume a strict inequality whenever `f x = B x` or `∥f x∥ = B x`; + * `of_*_lt_*` lemmas assume a strict inequality whenever `f x = B x` or `‖f x‖ = B x`; * `of_*_le_*` lemmas assume a non-strict inequality everywhere on `[a, b)`; * name of a lemma ends with `'` if (1) it assumes that `B` is continuous on `[a, b]` and has a right derivative at every point of `[a, b)`, and (2) the lemma has a counterpart assuming that `B` is differentiable everywhere on `ℝ` * `norm_image_sub_le_*_segment` : if derivative of `f` on `[a, b]` is bounded above - by a constant `C`, then `∥f x - f a∥ ≤ C * ∥x - a∥`; several versions deal with + by a constant `C`, then `‖f x - f a‖ ≤ C * ‖x - a‖`; several versions deal with right derivative and derivative within `[a, b]` (`has_deriv_within_at` or `deriv_within`). * `convex.is_const_of_fderiv_within_eq_zero` : if a function has derivative `0` on a convex set `s`, @@ -244,62 +244,62 @@ variables {f : ℝ → E} {a b : ℝ} /-- General fencing theorem for continuous functions with an estimate on the derivative. Let `f` and `B` be continuous functions on `[a, b]` such that -* `∥f a∥ ≤ B a`; +* `‖f a‖ ≤ B a`; * `B` has right derivative at every point of `[a, b)`; -* for each `x ∈ [a, b)` the right-side limit inferior of `(∥f z∥ - ∥f x∥) / (z - x)` +* for each `x ∈ [a, b)` the right-side limit inferior of `(‖f z‖ - ‖f x‖) / (z - x)` is bounded above by a function `f'`; -* we have `f' x < B' x` whenever `∥f x∥ = B x`. +* we have `f' x < B' x` whenever `‖f x‖ = B x`. -Then `∥f x∥ ≤ B x` everywhere on `[a, b]`. -/ +Then `‖f x‖ ≤ B x` everywhere on `[a, b]`. -/ lemma image_norm_le_of_liminf_right_slope_norm_lt_deriv_boundary {E : Type*} [normed_add_comm_group E] {f : ℝ → E} {f' : ℝ → ℝ} (hf : continuous_on f (Icc a b)) - -- `hf'` actually says `liminf (∥f z∥ - ∥f x∥) / (z - x) ≤ f' x` + -- `hf'` actually says `liminf (‖f z‖ - ‖f x‖) / (z - x) ≤ f' x` (hf' : ∀ x ∈ Ico a b, ∀ r, f' x < r → ∃ᶠ z in 𝓝[>] x, slope (norm ∘ f) x z < r) - {B B' : ℝ → ℝ} (ha : ∥f a∥ ≤ B a) (hB : continuous_on B (Icc a b)) + {B B' : ℝ → ℝ} (ha : ‖f a‖ ≤ B a) (hB : continuous_on B (Icc a b)) (hB' : ∀ x ∈ Ico a b, has_deriv_within_at B (B' x) (Ici x) x) - (bound : ∀ x ∈ Ico a b, ∥f x∥ = B x → f' x < B' x) : - ∀ ⦃x⦄, x ∈ Icc a b → ∥f x∥ ≤ B x := + (bound : ∀ x ∈ Ico a b, ‖f x‖ = B x → f' x < B' x) : + ∀ ⦃x⦄, x ∈ Icc a b → ‖f x‖ ≤ B x := image_le_of_liminf_slope_right_lt_deriv_boundary' (continuous_norm.comp_continuous_on hf) hf' ha hB hB' bound /-- General fencing theorem for continuous functions with an estimate on the norm of the derivative. Let `f` and `B` be continuous functions on `[a, b]` such that -* `∥f a∥ ≤ B a`; +* `‖f a‖ ≤ B a`; * `f` and `B` have right derivatives `f'` and `B'` respectively at every point of `[a, b)`; -* the norm of `f'` is strictly less than `B'` whenever `∥f x∥ = B x`. +* the norm of `f'` is strictly less than `B'` whenever `‖f x‖ = B x`. -Then `∥f x∥ ≤ B x` everywhere on `[a, b]`. We use one-sided derivatives in the assumptions +Then `‖f x‖ ≤ B x` everywhere on `[a, b]`. We use one-sided derivatives in the assumptions to make this theorem work for piecewise differentiable functions. -/ lemma image_norm_le_of_norm_deriv_right_lt_deriv_boundary' {f' : ℝ → E} (hf : continuous_on f (Icc a b)) (hf' : ∀ x ∈ Ico a b, has_deriv_within_at f (f' x) (Ici x) x) - {B B' : ℝ → ℝ} (ha : ∥f a∥ ≤ B a) (hB : continuous_on B (Icc a b)) + {B B' : ℝ → ℝ} (ha : ‖f a‖ ≤ B a) (hB : continuous_on B (Icc a b)) (hB' : ∀ x ∈ Ico a b, has_deriv_within_at B (B' x) (Ici x) x) - (bound : ∀ x ∈ Ico a b, ∥f x∥ = B x → ∥f' x∥ < B' x) : - ∀ ⦃x⦄, x ∈ Icc a b → ∥f x∥ ≤ B x := + (bound : ∀ x ∈ Ico a b, ‖f x‖ = B x → ‖f' x‖ < B' x) : + ∀ ⦃x⦄, x ∈ Icc a b → ‖f x‖ ≤ B x := image_norm_le_of_liminf_right_slope_norm_lt_deriv_boundary hf (λ x hx r hr, (hf' x hx).liminf_right_slope_norm_le hr) ha hB hB' bound /-- General fencing theorem for continuous functions with an estimate on the norm of the derivative. Let `f` and `B` be continuous functions on `[a, b]` such that -* `∥f a∥ ≤ B a`; +* `‖f a‖ ≤ B a`; * `f` has right derivative `f'` at every point of `[a, b)`; * `B` has derivative `B'` everywhere on `ℝ`; -* the norm of `f'` is strictly less than `B'` whenever `∥f x∥ = B x`. +* the norm of `f'` is strictly less than `B'` whenever `‖f x‖ = B x`. -Then `∥f x∥ ≤ B x` everywhere on `[a, b]`. We use one-sided derivatives in the assumptions +Then `‖f x‖ ≤ B x` everywhere on `[a, b]`. We use one-sided derivatives in the assumptions to make this theorem work for piecewise differentiable functions. -/ lemma image_norm_le_of_norm_deriv_right_lt_deriv_boundary {f' : ℝ → E} (hf : continuous_on f (Icc a b)) (hf' : ∀ x ∈ Ico a b, has_deriv_within_at f (f' x) (Ici x) x) - {B B' : ℝ → ℝ} (ha : ∥f a∥ ≤ B a) (hB : ∀ x, has_deriv_at B (B' x) x) - (bound : ∀ x ∈ Ico a b, ∥f x∥ = B x → ∥f' x∥ < B' x) : - ∀ ⦃x⦄, x ∈ Icc a b → ∥f x∥ ≤ B x := + {B B' : ℝ → ℝ} (ha : ‖f a‖ ≤ B a) (hB : ∀ x, has_deriv_at B (B' x) x) + (bound : ∀ x ∈ Ico a b, ‖f x‖ = B x → ‖f' x‖ < B' x) : + ∀ ⦃x⦄, x ∈ Icc a b → ‖f x‖ ≤ B x := image_norm_le_of_norm_deriv_right_lt_deriv_boundary' hf hf' ha (λ x hx, (hB x).continuous_at.continuous_within_at) (λ x hx, (hB x).has_deriv_within_at) bound @@ -307,51 +307,51 @@ image_norm_le_of_norm_deriv_right_lt_deriv_boundary' hf hf' ha /-- General fencing theorem for continuous functions with an estimate on the norm of the derivative. Let `f` and `B` be continuous functions on `[a, b]` such that -* `∥f a∥ ≤ B a`; +* `‖f a‖ ≤ B a`; * `f` and `B` have right derivatives `f'` and `B'` respectively at every point of `[a, b)`; -* we have `∥f' x∥ ≤ B x` everywhere on `[a, b)`. +* we have `‖f' x‖ ≤ B x` everywhere on `[a, b)`. -Then `∥f x∥ ≤ B x` everywhere on `[a, b]`. We use one-sided derivatives in the assumptions +Then `‖f x‖ ≤ B x` everywhere on `[a, b]`. We use one-sided derivatives in the assumptions to make this theorem work for piecewise differentiable functions. -/ lemma image_norm_le_of_norm_deriv_right_le_deriv_boundary' {f' : ℝ → E} (hf : continuous_on f (Icc a b)) (hf' : ∀ x ∈ Ico a b, has_deriv_within_at f (f' x) (Ici x) x) - {B B' : ℝ → ℝ} (ha : ∥f a∥ ≤ B a) (hB : continuous_on B (Icc a b)) + {B B' : ℝ → ℝ} (ha : ‖f a‖ ≤ B a) (hB : continuous_on B (Icc a b)) (hB' : ∀ x ∈ Ico a b, has_deriv_within_at B (B' x) (Ici x) x) - (bound : ∀ x ∈ Ico a b, ∥f' x∥ ≤ B' x) : - ∀ ⦃x⦄, x ∈ Icc a b → ∥f x∥ ≤ B x := + (bound : ∀ x ∈ Ico a b, ‖f' x‖ ≤ B' x) : + ∀ ⦃x⦄, x ∈ Icc a b → ‖f x‖ ≤ B x := image_le_of_liminf_slope_right_le_deriv_boundary (continuous_norm.comp_continuous_on hf) ha hB hB' $ (λ x hx r hr, (hf' x hx).liminf_right_slope_norm_le (lt_of_le_of_lt (bound x hx) hr)) /-- General fencing theorem for continuous functions with an estimate on the norm of the derivative. Let `f` and `B` be continuous functions on `[a, b]` such that -* `∥f a∥ ≤ B a`; +* `‖f a‖ ≤ B a`; * `f` has right derivative `f'` at every point of `[a, b)`; * `B` has derivative `B'` everywhere on `ℝ`; -* we have `∥f' x∥ ≤ B x` everywhere on `[a, b)`. +* we have `‖f' x‖ ≤ B x` everywhere on `[a, b)`. -Then `∥f x∥ ≤ B x` everywhere on `[a, b]`. We use one-sided derivatives in the assumptions +Then `‖f x‖ ≤ B x` everywhere on `[a, b]`. We use one-sided derivatives in the assumptions to make this theorem work for piecewise differentiable functions. -/ lemma image_norm_le_of_norm_deriv_right_le_deriv_boundary {f' : ℝ → E} (hf : continuous_on f (Icc a b)) (hf' : ∀ x ∈ Ico a b, has_deriv_within_at f (f' x) (Ici x) x) - {B B' : ℝ → ℝ} (ha : ∥f a∥ ≤ B a) (hB : ∀ x, has_deriv_at B (B' x) x) - (bound : ∀ x ∈ Ico a b, ∥f' x∥ ≤ B' x) : - ∀ ⦃x⦄, x ∈ Icc a b → ∥f x∥ ≤ B x := + {B B' : ℝ → ℝ} (ha : ‖f a‖ ≤ B a) (hB : ∀ x, has_deriv_at B (B' x) x) + (bound : ∀ x ∈ Ico a b, ‖f' x‖ ≤ B' x) : + ∀ ⦃x⦄, x ∈ Icc a b → ‖f x‖ ≤ B x := image_norm_le_of_norm_deriv_right_le_deriv_boundary' hf hf' ha (λ x hx, (hB x).continuous_at.continuous_within_at) (λ x hx, (hB x).has_deriv_within_at) bound /-- A function on `[a, b]` with the norm of the right derivative bounded by `C` -satisfies `∥f x - f a∥ ≤ C * (x - a)`. -/ +satisfies `‖f x - f a‖ ≤ C * (x - a)`. -/ theorem norm_image_sub_le_of_norm_deriv_right_le_segment {f' : ℝ → E} {C : ℝ} (hf : continuous_on f (Icc a b)) (hf' : ∀ x ∈ Ico a b, has_deriv_within_at f (f' x) (Ici x) x) - (bound : ∀x ∈ Ico a b, ∥f' x∥ ≤ C) : - ∀ x ∈ Icc a b, ∥f x - f a∥ ≤ C * (x - a) := + (bound : ∀x ∈ Ico a b, ‖f' x‖ ≤ C) : + ∀ x ∈ Icc a b, ‖f x - f a‖ ≤ C * (x - a) := begin let g := λ x, f x - f a, have hg : continuous_on g (Icc a b), from hf.sub continuous_on_const, @@ -367,12 +367,12 @@ begin end /-- A function on `[a, b]` with the norm of the derivative within `[a, b]` -bounded by `C` satisfies `∥f x - f a∥ ≤ C * (x - a)`, `has_deriv_within_at` +bounded by `C` satisfies `‖f x - f a‖ ≤ C * (x - a)`, `has_deriv_within_at` version. -/ theorem norm_image_sub_le_of_norm_deriv_le_segment' {f' : ℝ → E} {C : ℝ} (hf : ∀ x ∈ Icc a b, has_deriv_within_at f (f' x) (Icc a b) x) - (bound : ∀x ∈ Ico a b, ∥f' x∥ ≤ C) : - ∀ x ∈ Icc a b, ∥f x - f a∥ ≤ C * (x - a) := + (bound : ∀x ∈ Ico a b, ‖f' x‖ ≤ C) : + ∀ x ∈ Icc a b, ‖f x - f a‖ ≤ C * (x - a) := begin refine norm_image_sub_le_of_norm_deriv_right_le_segment (λ x hx, (hf x hx).continuous_within_at) (λ x hx, _) bound, @@ -380,32 +380,32 @@ begin end /-- A function on `[a, b]` with the norm of the derivative within `[a, b]` -bounded by `C` satisfies `∥f x - f a∥ ≤ C * (x - a)`, `deriv_within` +bounded by `C` satisfies `‖f x - f a‖ ≤ C * (x - a)`, `deriv_within` version. -/ theorem norm_image_sub_le_of_norm_deriv_le_segment {C : ℝ} (hf : differentiable_on ℝ f (Icc a b)) - (bound : ∀x ∈ Ico a b, ∥deriv_within f (Icc a b) x∥ ≤ C) : - ∀ x ∈ Icc a b, ∥f x - f a∥ ≤ C * (x - a) := + (bound : ∀x ∈ Ico a b, ‖deriv_within f (Icc a b) x‖ ≤ C) : + ∀ x ∈ Icc a b, ‖f x - f a‖ ≤ C * (x - a) := begin refine norm_image_sub_le_of_norm_deriv_le_segment' _ bound, exact λ x hx, (hf x hx).has_deriv_within_at end /-- A function on `[0, 1]` with the norm of the derivative within `[0, 1]` -bounded by `C` satisfies `∥f 1 - f 0∥ ≤ C`, `has_deriv_within_at` +bounded by `C` satisfies `‖f 1 - f 0‖ ≤ C`, `has_deriv_within_at` version. -/ theorem norm_image_sub_le_of_norm_deriv_le_segment_01' {f' : ℝ → E} {C : ℝ} (hf : ∀ x ∈ Icc (0:ℝ) 1, has_deriv_within_at f (f' x) (Icc (0:ℝ) 1) x) - (bound : ∀x ∈ Ico (0:ℝ) 1, ∥f' x∥ ≤ C) : - ∥f 1 - f 0∥ ≤ C := + (bound : ∀x ∈ Ico (0:ℝ) 1, ‖f' x‖ ≤ C) : + ‖f 1 - f 0‖ ≤ C := by simpa only [sub_zero, mul_one] using norm_image_sub_le_of_norm_deriv_le_segment' hf bound 1 (right_mem_Icc.2 zero_le_one) /-- A function on `[0, 1]` with the norm of the derivative within `[0, 1]` -bounded by `C` satisfies `∥f 1 - f 0∥ ≤ C`, `deriv_within` version. -/ +bounded by `C` satisfies `‖f 1 - f 0‖ ≤ C`, `deriv_within` version. -/ theorem norm_image_sub_le_of_norm_deriv_le_segment_01 {C : ℝ} (hf : differentiable_on ℝ f (Icc (0:ℝ) 1)) - (bound : ∀x ∈ Ico (0:ℝ) 1, ∥deriv_within f (Icc (0:ℝ) 1) x∥ ≤ C) : - ∥f 1 - f 0∥ ≤ C := + (bound : ∀x ∈ Ico (0:ℝ) 1, ‖deriv_within f (Icc (0:ℝ) 1) x‖ ≤ C) : + ‖f 1 - f 0‖ ≤ C := by simpa only [sub_zero, mul_one] using norm_image_sub_le_of_norm_deriv_le_segment hf bound 1 (right_mem_Icc.2 zero_le_one) @@ -420,7 +420,7 @@ theorem constant_of_deriv_within_zero (hdiff : differentiable_on ℝ f (Icc a b) (hderiv : ∀ x ∈ Ico a b, deriv_within f (Icc a b) x = 0) : ∀ x ∈ Icc a b, f x = f a := begin - have H : ∀ x ∈ Ico a b, ∥deriv_within f (Icc a b) x∥ ≤ 0 := + have H : ∀ x ∈ Ico a b, ‖deriv_within f (Icc a b) x‖ ≤ 0 := by simpa only [norm_le_zero_iff] using λ x hx, hderiv x hx, simpa only [zero_mul, norm_le_zero_iff, sub_eq_zero] using λ x hx, norm_image_sub_le_of_norm_deriv_le_segment hdiff H x hx, @@ -481,8 +481,8 @@ variables {f : E → G} {C : ℝ} {s : set E} {x y : E} {f' : E → E →L[𝕜 /-- The mean value theorem on a convex set: if the derivative of a function is bounded by `C`, then the function is `C`-Lipschitz. Version with `has_fderiv_within`. -/ theorem norm_image_sub_le_of_norm_has_fderiv_within_le - (hf : ∀ x ∈ s, has_fderiv_within_at f (f' x) s x) (bound : ∀x∈s, ∥f' x∥ ≤ C) - (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ∥f y - f x∥ ≤ C * ∥y - x∥ := + (hf : ∀ x ∈ s, has_fderiv_within_at f (f' x) s x) (bound : ∀x∈s, ‖f' x‖ ≤ C) + (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ‖f y - f x‖ ≤ C * ‖y - x‖ := begin letI : normed_space ℝ G := restrict_scalars.normed_space ℝ 𝕜 G, /- By composition with `t ↦ x + t • (y-x)`, we reduce to a statement for functions defined @@ -515,7 +515,7 @@ end `s`, then the function is `C`-Lipschitz on `s`. Version with `has_fderiv_within` and `lipschitz_on_with`. -/ theorem lipschitz_on_with_of_nnnorm_has_fderiv_within_le {C : ℝ≥0} - (hf : ∀ x ∈ s, has_fderiv_within_at f (f' x) s x) (bound : ∀x∈s, ∥f' x∥₊ ≤ C) + (hf : ∀ x ∈ s, has_fderiv_within_at f (f' x) s x) (bound : ∀x∈s, ‖f' x‖₊ ≤ C) (hs : convex ℝ s) : lipschitz_on_with C f s := begin rw lipschitz_on_with_iff_norm_sub_le, @@ -525,17 +525,17 @@ end /-- Let `s` be a convex set in a real normed vector space `E`, let `f : E → G` be a function differentiable within `s` in a neighborhood of `x : E` with derivative `f'`. Suppose that `f'` is -continuous within `s` at `x`. Then for any number `K : ℝ≥0` larger than `∥f' x∥₊`, `f` is +continuous within `s` at `x`. Then for any number `K : ℝ≥0` larger than `‖f' x‖₊`, `f` is `K`-Lipschitz on some neighborhood of `x` within `s`. See also `convex.exists_nhds_within_lipschitz_on_with_of_has_fderiv_within_at` for a version that claims existence of `K` instead of an explicit estimate. -/ lemma exists_nhds_within_lipschitz_on_with_of_has_fderiv_within_at_of_nnnorm_lt (hs : convex ℝ s) {f : E → G} (hder : ∀ᶠ y in 𝓝[s] x, has_fderiv_within_at f (f' y) s y) - (hcont : continuous_within_at f' s x) (K : ℝ≥0) (hK : ∥f' x∥₊ < K) : + (hcont : continuous_within_at f' s x) (K : ℝ≥0) (hK : ‖f' x‖₊ < K) : ∃ t ∈ 𝓝[s] x, lipschitz_on_with K f t := begin obtain ⟨ε, ε0, hε⟩ : - ∃ ε > 0, ball x ε ∩ s ⊆ {y | has_fderiv_within_at f (f' y) s y ∧ ∥f' y∥₊ < K}, + ∃ ε > 0, ball x ε ∩ s ⊆ {y | has_fderiv_within_at f (f' y) s y ∧ ‖f' y‖₊ < K}, from mem_nhds_within_iff.1 (hder.and $ hcont.nnnorm.eventually (gt_mem_nhds hK)), rw inter_comm at hε, refine ⟨s ∩ ball x ε, inter_mem_nhds_within _ (ball_mem_nhds _ ε0), _⟩, @@ -545,7 +545,7 @@ end /-- Let `s` be a convex set in a real normed vector space `E`, let `f : E → G` be a function differentiable within `s` in a neighborhood of `x : E` with derivative `f'`. Suppose that `f'` is -continuous within `s` at `x`. Then for any number `K : ℝ≥0` larger than `∥f' x∥₊`, `f` is Lipschitz +continuous within `s` at `x`. Then for any number `K : ℝ≥0` larger than `‖f' x‖₊`, `f` is Lipschitz on some neighborhood of `x` within `s`. See also `convex.exists_nhds_within_lipschitz_on_with_of_has_fderiv_within_at_of_nnnorm_lt` for a version with an explicit estimate on the Lipschitz constant. -/ @@ -559,8 +559,8 @@ lemma exists_nhds_within_lipschitz_on_with_of_has_fderiv_within_at /-- The mean value theorem on a convex set: if the derivative of a function within this set is bounded by `C`, then the function is `C`-Lipschitz. Version with `fderiv_within`. -/ theorem norm_image_sub_le_of_norm_fderiv_within_le - (hf : differentiable_on 𝕜 f s) (bound : ∀x∈s, ∥fderiv_within 𝕜 f s x∥ ≤ C) - (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ∥f y - f x∥ ≤ C * ∥y - x∥ := + (hf : differentiable_on 𝕜 f s) (bound : ∀x∈s, ‖fderiv_within 𝕜 f s x‖ ≤ C) + (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ‖f y - f x‖ ≤ C * ‖y - x‖ := hs.norm_image_sub_le_of_norm_has_fderiv_within_le (λ x hx, (hf x hx).has_fderiv_within_at) bound xs ys @@ -568,22 +568,22 @@ bound xs ys `s`, then the function is `C`-Lipschitz on `s`. Version with `fderiv_within` and `lipschitz_on_with`. -/ theorem lipschitz_on_with_of_nnnorm_fderiv_within_le {C : ℝ≥0} - (hf : differentiable_on 𝕜 f s) (bound : ∀ x ∈ s, ∥fderiv_within 𝕜 f s x∥₊ ≤ C) + (hf : differentiable_on 𝕜 f s) (bound : ∀ x ∈ s, ‖fderiv_within 𝕜 f s x‖₊ ≤ C) (hs : convex ℝ s) : lipschitz_on_with C f s:= hs.lipschitz_on_with_of_nnnorm_has_fderiv_within_le (λ x hx, (hf x hx).has_fderiv_within_at) bound /-- The mean value theorem on a convex set: if the derivative of a function is bounded by `C`, then the function is `C`-Lipschitz. Version with `fderiv`. -/ theorem norm_image_sub_le_of_norm_fderiv_le - (hf : ∀ x ∈ s, differentiable_at 𝕜 f x) (bound : ∀x∈s, ∥fderiv 𝕜 f x∥ ≤ C) - (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ∥f y - f x∥ ≤ C * ∥y - x∥ := + (hf : ∀ x ∈ s, differentiable_at 𝕜 f x) (bound : ∀x∈s, ‖fderiv 𝕜 f x‖ ≤ C) + (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ‖f y - f x‖ ≤ C * ‖y - x‖ := hs.norm_image_sub_le_of_norm_has_fderiv_within_le (λ x hx, (hf x hx).has_fderiv_at.has_fderiv_within_at) bound xs ys /-- The mean value theorem on a convex set: if the derivative of a function is bounded by `C` on `s`, then the function is `C`-Lipschitz on `s`. Version with `fderiv` and `lipschitz_on_with`. -/ theorem lipschitz_on_with_of_nnnorm_fderiv_le {C : ℝ≥0} - (hf : ∀ x ∈ s, differentiable_at 𝕜 f x) (bound : ∀x∈s, ∥fderiv 𝕜 f x∥₊ ≤ C) + (hf : ∀ x ∈ s, differentiable_at 𝕜 f x) (bound : ∀x∈s, ‖fderiv 𝕜 f x‖₊ ≤ C) (hs : convex ℝ s) : lipschitz_on_with C f s := hs.lipschitz_on_with_of_nnnorm_has_fderiv_within_le (λ x hx, (hf x hx).has_fderiv_at.has_fderiv_within_at) bound @@ -592,8 +592,8 @@ hs.lipschitz_on_with_of_nnnorm_has_fderiv_within_le the derivative and a fixed linear map, rather than a bound on the derivative itself. Version with `has_fderiv_within`. -/ theorem norm_image_sub_le_of_norm_has_fderiv_within_le' - (hf : ∀ x ∈ s, has_fderiv_within_at f (f' x) s x) (bound : ∀x∈s, ∥f' x - φ∥ ≤ C) - (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ∥f y - f x - φ (y - x)∥ ≤ C * ∥y - x∥ := + (hf : ∀ x ∈ s, has_fderiv_within_at f (f' x) s x) (bound : ∀x∈s, ‖f' x - φ‖ ≤ C) + (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ‖f y - f x - φ (y - x)‖ ≤ C * ‖y - x‖ := begin /- We subtract `φ` to define a new function `g` for which `g' = 0`, for which the previous theorem applies, `convex.norm_image_sub_le_of_norm_has_fderiv_within_le`. Then, we just need to glue @@ -601,23 +601,23 @@ begin let g := λy, f y - φ y, have hg : ∀ x ∈ s, has_fderiv_within_at g (f' x - φ) s x := λ x xs, (hf x xs).sub φ.has_fderiv_within_at, - calc ∥f y - f x - φ (y - x)∥ = ∥f y - f x - (φ y - φ x)∥ : by simp - ... = ∥(f y - φ y) - (f x - φ x)∥ : by abel - ... = ∥g y - g x∥ : by simp - ... ≤ C * ∥y - x∥ : convex.norm_image_sub_le_of_norm_has_fderiv_within_le hg bound hs xs ys, + calc ‖f y - f x - φ (y - x)‖ = ‖f y - f x - (φ y - φ x)‖ : by simp + ... = ‖(f y - φ y) - (f x - φ x)‖ : by abel + ... = ‖g y - g x‖ : by simp + ... ≤ C * ‖y - x‖ : convex.norm_image_sub_le_of_norm_has_fderiv_within_le hg bound hs xs ys, end /-- Variant of the mean value inequality on a convex set. Version with `fderiv_within`. -/ theorem norm_image_sub_le_of_norm_fderiv_within_le' - (hf : differentiable_on 𝕜 f s) (bound : ∀x∈s, ∥fderiv_within 𝕜 f s x - φ∥ ≤ C) - (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ∥f y - f x - φ (y - x)∥ ≤ C * ∥y - x∥ := + (hf : differentiable_on 𝕜 f s) (bound : ∀x∈s, ‖fderiv_within 𝕜 f s x - φ‖ ≤ C) + (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ‖f y - f x - φ (y - x)‖ ≤ C * ‖y - x‖ := hs.norm_image_sub_le_of_norm_has_fderiv_within_le' (λ x hx, (hf x hx).has_fderiv_within_at) bound xs ys /-- Variant of the mean value inequality on a convex set. Version with `fderiv`. -/ theorem norm_image_sub_le_of_norm_fderiv_le' - (hf : ∀ x ∈ s, differentiable_at 𝕜 f x) (bound : ∀x∈s, ∥fderiv 𝕜 f x - φ∥ ≤ C) - (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ∥f y - f x - φ (y - x)∥ ≤ C * ∥y - x∥ := + (hf : ∀ x ∈ s, differentiable_at 𝕜 f x) (bound : ∀x∈s, ‖fderiv 𝕜 f x - φ‖ ≤ C) + (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ‖f y - f x - φ (y - x)‖ ≤ C * ‖y - x‖ := hs.norm_image_sub_le_of_norm_has_fderiv_within_le' (λ x hx, (hf x hx).has_fderiv_at.has_fderiv_within_at) bound xs ys @@ -626,7 +626,7 @@ then it is a constant on this set. -/ theorem is_const_of_fderiv_within_eq_zero (hs : convex ℝ s) (hf : differentiable_on 𝕜 f s) (hf' : ∀ x ∈ s, fderiv_within 𝕜 f s x = 0) (hx : x ∈ s) (hy : y ∈ s) : f x = f y := -have bound : ∀ x ∈ s, ∥fderiv_within 𝕜 f s x∥ ≤ 0, +have bound : ∀ x ∈ s, ‖fderiv_within 𝕜 f s x‖ ≤ 0, 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 @@ -646,8 +646,8 @@ variables {f f' : 𝕜 → G} {s : set 𝕜} {x y : 𝕜} /-- The mean value theorem on a convex set in dimension 1: if the derivative of a function is bounded by `C`, then the function is `C`-Lipschitz. Version with `has_deriv_within`. -/ theorem norm_image_sub_le_of_norm_has_deriv_within_le {C : ℝ} - (hf : ∀ x ∈ s, has_deriv_within_at f (f' x) s x) (bound : ∀x∈s, ∥f' x∥ ≤ C) - (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ∥f y - f x∥ ≤ C * ∥y - x∥ := + (hf : ∀ x ∈ s, has_deriv_within_at f (f' x) s x) (bound : ∀x∈s, ‖f' x‖ ≤ C) + (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ‖f y - f x‖ ≤ C * ‖y - x‖ := convex.norm_image_sub_le_of_norm_has_fderiv_within_le (λ x hx, (hf x hx).has_fderiv_within_at) (λ x hx, le_trans (by simp) (bound x hx)) hs xs ys @@ -655,7 +655,7 @@ convex.norm_image_sub_le_of_norm_has_fderiv_within_le (λ x hx, (hf x hx).has_fd bounded by `C` on `s`, then the function is `C`-Lipschitz on `s`. Version with `has_deriv_within` and `lipschitz_on_with`. -/ theorem lipschitz_on_with_of_nnnorm_has_deriv_within_le {C : ℝ≥0} (hs : convex ℝ s) - (hf : ∀ x ∈ s, has_deriv_within_at f (f' x) s x) (bound : ∀x∈s, ∥f' x∥₊ ≤ C) : + (hf : ∀ x ∈ s, has_deriv_within_at f (f' x) s x) (bound : ∀x∈s, ‖f' x‖₊ ≤ C) : lipschitz_on_with C f s := convex.lipschitz_on_with_of_nnnorm_has_fderiv_within_le (λ x hx, (hf x hx).has_fderiv_within_at) (λ x hx, le_trans (by simp) (bound x hx)) hs @@ -663,8 +663,8 @@ convex.lipschitz_on_with_of_nnnorm_has_fderiv_within_le (λ x hx, (hf x hx).has_ /-- The mean value theorem on a convex set in dimension 1: if the derivative of a function within this set is bounded by `C`, then the function is `C`-Lipschitz. Version with `deriv_within` -/ theorem norm_image_sub_le_of_norm_deriv_within_le {C : ℝ} - (hf : differentiable_on 𝕜 f s) (bound : ∀x∈s, ∥deriv_within f s x∥ ≤ C) - (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ∥f y - f x∥ ≤ C * ∥y - x∥ := + (hf : differentiable_on 𝕜 f s) (bound : ∀x∈s, ‖deriv_within f s x‖ ≤ C) + (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ‖f y - f x‖ ≤ C * ‖y - x‖ := hs.norm_image_sub_le_of_norm_has_deriv_within_le (λ x hx, (hf x hx).has_deriv_within_at) bound xs ys @@ -672,15 +672,15 @@ bound xs ys bounded by `C` on `s`, then the function is `C`-Lipschitz on `s`. Version with `deriv_within` and `lipschitz_on_with`. -/ theorem lipschitz_on_with_of_nnnorm_deriv_within_le {C : ℝ≥0} (hs : convex ℝ s) - (hf : differentiable_on 𝕜 f s) (bound : ∀x∈s, ∥deriv_within f s x∥₊ ≤ C) : + (hf : differentiable_on 𝕜 f s) (bound : ∀x∈s, ‖deriv_within f s x‖₊ ≤ C) : lipschitz_on_with C f s := hs.lipschitz_on_with_of_nnnorm_has_deriv_within_le (λ x hx, (hf x hx).has_deriv_within_at) bound /-- The mean value theorem on a convex set in dimension 1: if the derivative of a function is bounded by `C`, then the function is `C`-Lipschitz. Version with `deriv`. -/ theorem norm_image_sub_le_of_norm_deriv_le {C : ℝ} - (hf : ∀ x ∈ s, differentiable_at 𝕜 f x) (bound : ∀x∈s, ∥deriv f x∥ ≤ C) - (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ∥f y - f x∥ ≤ C * ∥y - x∥ := + (hf : ∀ x ∈ s, differentiable_at 𝕜 f x) (bound : ∀x∈s, ‖deriv f x‖ ≤ C) + (hs : convex ℝ s) (xs : x ∈ s) (ys : y ∈ s) : ‖f y - f x‖ ≤ C * ‖y - x‖ := hs.norm_image_sub_le_of_norm_has_deriv_within_le (λ x hx, (hf x hx).has_deriv_at.has_deriv_within_at) bound xs ys @@ -688,7 +688,7 @@ hs.norm_image_sub_le_of_norm_has_deriv_within_le bounded by `C` on `s`, then the function is `C`-Lipschitz on `s`. Version with `deriv` and `lipschitz_on_with`. -/ theorem lipschitz_on_with_of_nnnorm_deriv_le {C : ℝ≥0} - (hf : ∀ x ∈ s, differentiable_at 𝕜 f x) (bound : ∀x∈s, ∥deriv f x∥₊ ≤ C) + (hf : ∀ x ∈ s, differentiable_at 𝕜 f x) (bound : ∀x∈s, ‖deriv f x‖₊ ≤ C) (hs : convex ℝ s) : lipschitz_on_with C f s := hs.lipschitz_on_with_of_nnnorm_has_deriv_within_le (λ x hx, (hf x hx).has_deriv_at.has_deriv_within_at) bound @@ -696,7 +696,7 @@ hs.lipschitz_on_with_of_nnnorm_has_deriv_within_le /-- The mean value theorem set in dimension 1: if the derivative of a function is bounded by `C`, then the function is `C`-Lipschitz. Version with `deriv` and `lipschitz_with`. -/ theorem _root_.lipschitz_with_of_nnnorm_deriv_le {C : ℝ≥0} (hf : differentiable 𝕜 f) - (bound : ∀ x, ∥deriv f x∥₊ ≤ C) : lipschitz_with C f := + (bound : ∀ x, ‖deriv f x‖₊ ≤ C) : lipschitz_with C f := lipschitz_on_univ.1 $ convex_univ.lipschitz_on_with_of_nnnorm_deriv_le (λ x hx, hf x) (λ x hx, bound x) @@ -1386,7 +1386,7 @@ begin rintros ⟨a, b⟩ h, rw [← ball_prod_same, prod_mk_mem_set_prod_eq] at h, -- exploit the choice of ε as the modulus of continuity of f' - have hf' : ∀ x' ∈ ball x ε, ∥f' x' - f' x∥ ≤ c, + have hf' : ∀ x' ∈ ball x ε, ‖f' x' - f' x‖ ≤ c, { intros x' H', rw ← dist_eq_norm, exact le_of_lt (hε H').2 }, -- apply mean value theorem letI : normed_space ℝ G := restrict_scalars.normed_space ℝ 𝕜 G, diff --git a/src/analysis/calculus/parametric_integral.lean b/src/analysis/calculus/parametric_integral.lean index dd7ed194fc022..0ba29b2f0fba4 100644 --- a/src/analysis/calculus/parametric_integral.lean +++ b/src/analysis/calculus/parametric_integral.lean @@ -62,7 +62,7 @@ variables {α : Type*} [measurable_space α] {μ : measure α} {𝕜 : Type*} [i {H : Type*} [normed_add_comm_group H] [normed_space 𝕜 H] /-- Differentiation under integral of `x ↦ ∫ F x a` at a given point `x₀`, assuming `F x₀` is -integrable, `∥F x a - F x₀ a∥ ≤ bound a * ∥x - x₀∥` for `x` in a ball around `x₀` for ae `a` with +integrable, `‖F x a - F x₀ a‖ ≤ bound a * ‖x - x₀‖` for `x` in a ball around `x₀` for ae `a` with integrable Lipschitz bound `bound` (with a ball radius independent of `a`), and `F x` is ae-measurable for `x` in the same ball. See `has_fderiv_at_integral_of_dominated_loc_of_lip` for a slightly less general but usually more useful version. -/ @@ -72,22 +72,22 @@ lemma has_fderiv_at_integral_of_dominated_loc_of_lip' {F : H → α → E} {F' : (hF_meas : ∀ x ∈ ball x₀ ε, ae_strongly_measurable (F x) μ) (hF_int : integrable (F x₀) μ) (hF'_meas : ae_strongly_measurable F' μ) - (h_lipsch : ∀ᵐ a ∂μ, ∀ x ∈ ball x₀ ε, ∥F x a - F x₀ a∥ ≤ bound a * ∥x - x₀∥) + (h_lipsch : ∀ᵐ a ∂μ, ∀ x ∈ ball x₀ ε, ‖F x a - F x₀ a‖ ≤ bound a * ‖x - x₀‖) (bound_integrable : integrable (bound : α → ℝ) μ) (h_diff : ∀ᵐ a ∂μ, has_fderiv_at (λ x, F x a) (F' a) x₀) : integrable F' μ ∧ has_fderiv_at (λ x, ∫ a, F x a ∂μ) (∫ a, F' a ∂μ) x₀ := begin have x₀_in : x₀ ∈ ball x₀ ε := mem_ball_self ε_pos, - have nneg : ∀ x, 0 ≤ ∥x - x₀∥⁻¹ := λ x, inv_nonneg.mpr (norm_nonneg _) , + have nneg : ∀ x, 0 ≤ ‖x - x₀‖⁻¹ := λ x, inv_nonneg.mpr (norm_nonneg _) , set b : α → ℝ := λ a, |bound a|, have b_int : integrable b μ := bound_integrable.norm, have b_nonneg : ∀ a, 0 ≤ b a := λ a, abs_nonneg _, - replace h_lipsch : ∀ᵐ a ∂μ, ∀ x ∈ ball x₀ ε, ∥F x a - F x₀ a∥ ≤ b a * ∥x - x₀∥, + replace h_lipsch : ∀ᵐ a ∂μ, ∀ x ∈ ball x₀ ε, ‖F x a - F x₀ a‖ ≤ b a * ‖x - x₀‖, from h_lipsch.mono (λ a ha x hx, (ha x hx).trans $ mul_le_mul_of_nonneg_right (le_abs_self _) (norm_nonneg _)), have hF_int' : ∀ x ∈ ball x₀ ε, integrable (F x) μ, { intros x x_in, - have : ∀ᵐ a ∂μ, ∥F x₀ a - F x a∥ ≤ ε * b a, + have : ∀ᵐ a ∂μ, ‖F x₀ a - F x a‖ ≤ ε * b a, { simp only [norm_sub_rev (F x₀ _)], refine h_lipsch.mono (λ a ha, (ha x x_in).trans _), rw mul_comm ε, @@ -96,7 +96,7 @@ begin exact integrable_of_norm_sub_le (hF_meas x x_in) hF_int (integrable.const_mul bound_integrable.norm ε) this }, have hF'_int : integrable F' μ, - { have : ∀ᵐ a ∂μ, ∥F' a∥ ≤ b a, + { have : ∀ᵐ a ∂μ, ‖F' a‖ ≤ b a, { apply (h_diff.and h_lipsch).mono, rintros a ⟨ha_diff, ha_lip⟩, refine ha_diff.le_of_lip' (b_nonneg a) (mem_of_superset (ball_mem_nhds _ ε_pos) $ ha_lip) }, @@ -104,8 +104,8 @@ begin refine ⟨hF'_int, _⟩, have h_ball: ball x₀ ε ∈ 𝓝 x₀ := ball_mem_nhds x₀ ε_pos, have : ∀ᶠ x in 𝓝 x₀, - ∥x - x₀∥⁻¹ * ∥∫ a, F x a ∂μ - ∫ a, F x₀ a ∂μ - (∫ a, F' a ∂μ) (x - x₀)∥ = - ∥∫ a, ∥x - x₀∥⁻¹ • (F x a - F x₀ a - F' a (x - x₀)) ∂μ∥, + ‖x - x₀‖⁻¹ * ‖∫ a, F x a ∂μ - ∫ a, F x₀ a ∂μ - (∫ a, F' a ∂μ) (x - x₀)‖ = + ‖∫ a, ‖x - x₀‖⁻¹ • (F x a - F x₀ a - F' a (x - x₀)) ∂μ‖, { apply mem_of_superset (ball_mem_nhds _ ε_pos), intros x x_in, rw [set.mem_set_of_eq, ← norm_smul_of_nonneg (nneg _), integral_smul, @@ -113,7 +113,7 @@ begin exacts [hF_int' x x_in, hF_int, (hF_int' x x_in).sub hF_int, hF'_int.apply_continuous_linear_map _] }, rw [has_fderiv_at_iff_tendsto, tendsto_congr' this, ← tendsto_zero_iff_norm_tendsto_zero, - ← show ∫ (a : α), ∥x₀ - x₀∥⁻¹ • (F x₀ a - F x₀ a - (F' a) (x₀ - x₀)) ∂μ = 0, by simp], + ← show ∫ (a : α), ‖x₀ - x₀‖⁻¹ • (F x₀ a - F x₀ a - (F' a) (x₀ - x₀)) ∂μ = 0, by simp], apply tendsto_integral_filter_of_dominated_convergence, { filter_upwards [h_ball] with _ x_in, apply ae_strongly_measurable.const_smul, @@ -122,28 +122,28 @@ begin intros x hx, apply (h_diff.and h_lipsch).mono, rintros a ⟨ha_deriv, ha_bound⟩, - show ∥∥x - x₀∥⁻¹ • (F x a - F x₀ a - F' a (x - x₀))∥ ≤ b a + ∥F' a∥, - replace ha_bound : ∥F x a - F x₀ a∥ ≤ b a * ∥x - x₀∥ := ha_bound x hx, - calc ∥∥x - x₀∥⁻¹ • (F x a - F x₀ a - F' a (x - x₀))∥ - = ∥∥x - x₀∥⁻¹ • (F x a - F x₀ a) - ∥x - x₀∥⁻¹ • F' a (x - x₀)∥ : by rw smul_sub - ... ≤ ∥∥x - x₀∥⁻¹ • (F x a - F x₀ a)∥ + ∥∥x - x₀∥⁻¹ • F' a (x - x₀)∥ : norm_sub_le _ _ - ... = ∥x - x₀∥⁻¹ * ∥F x a - F x₀ a∥ + ∥x - x₀∥⁻¹ * ∥F' a (x - x₀)∥ : + show ‖‖x - x₀‖⁻¹ • (F x a - F x₀ a - F' a (x - x₀))‖ ≤ b a + ‖F' a‖, + replace ha_bound : ‖F x a - F x₀ a‖ ≤ b a * ‖x - x₀‖ := ha_bound x hx, + calc ‖‖x - x₀‖⁻¹ • (F x a - F x₀ a - F' a (x - x₀))‖ + = ‖‖x - x₀‖⁻¹ • (F x a - F x₀ a) - ‖x - x₀‖⁻¹ • F' a (x - x₀)‖ : by rw smul_sub + ... ≤ ‖‖x - x₀‖⁻¹ • (F x a - F x₀ a)‖ + ‖‖x - x₀‖⁻¹ • F' a (x - x₀)‖ : norm_sub_le _ _ + ... = ‖x - x₀‖⁻¹ * ‖F x a - F x₀ a‖ + ‖x - x₀‖⁻¹ * ‖F' a (x - x₀)‖ : by { rw [norm_smul_of_nonneg, norm_smul_of_nonneg] ; exact nneg _} - ... ≤ ∥x - x₀∥⁻¹ * (b a * ∥x - x₀∥) + ∥x - x₀∥⁻¹ * (∥F' a∥ * ∥x - x₀∥) : add_le_add _ _ - ... ≤ b a + ∥F' a∥ : _, + ... ≤ ‖x - x₀‖⁻¹ * (b a * ‖x - x₀‖) + ‖x - x₀‖⁻¹ * (‖F' a‖ * ‖x - x₀‖) : add_le_add _ _ + ... ≤ b a + ‖F' a‖ : _, exact mul_le_mul_of_nonneg_left ha_bound (nneg _), apply mul_le_mul_of_nonneg_left ((F' a).le_op_norm _) (nneg _), - by_cases h : ∥x - x₀∥ = 0, + by_cases h : ‖x - x₀‖ = 0, { simpa [h] using add_nonneg (b_nonneg a) (norm_nonneg (F' a)) }, { field_simp [h] } }, { exact b_int.add hF'_int.norm }, { apply h_diff.mono, intros a ha, - suffices : tendsto (λ x, ∥x - x₀∥⁻¹ • (F x a - F x₀ a - F' a (x - x₀))) (𝓝 x₀) (𝓝 0), + suffices : tendsto (λ x, ‖x - x₀‖⁻¹ • (F x a - F x₀ a - F' a (x - x₀))) (𝓝 x₀) (𝓝 0), by simpa, rw tendsto_zero_iff_norm_tendsto_zero, - have : (λ x, ∥x - x₀∥⁻¹ * ∥F x a - F x₀ a - F' a (x - x₀)∥) = - λ x, ∥∥x - x₀∥⁻¹ • (F x a - F x₀ a - F' a (x - x₀))∥, + have : (λ x, ‖x - x₀‖⁻¹ * ‖F x a - F x₀ a - F' a (x - x₀)‖) = + λ x, ‖‖x - x₀‖⁻¹ • (F x a - F x₀ a - F' a (x - x₀))‖, { ext x, rw norm_smul_of_nonneg (nneg _) }, rwa [has_fderiv_at_iff_tendsto, this] at ha }, @@ -167,7 +167,7 @@ begin obtain ⟨δ, δ_pos, hδ⟩ : ∃ δ > 0, ∀ x ∈ ball x₀ δ, ae_strongly_measurable (F x) μ ∧ x ∈ ball x₀ ε, from eventually_nhds_iff_ball.mp (hF_meas.and (ball_mem_nhds x₀ ε_pos)), choose hδ_meas hδε using hδ, - replace h_lip : ∀ᵐ (a : α) ∂μ, ∀ x ∈ ball x₀ δ, ∥F x a - F x₀ a∥ ≤ |bound a| * ∥x - x₀∥, + replace h_lip : ∀ᵐ (a : α) ∂μ, ∀ x ∈ ball x₀ δ, ‖F x a - F x₀ a‖ ≤ |bound a| * ‖x - x₀‖, from h_lip.mono (λ a lip x hx, lip.norm_sub_le (hδε x hx) (mem_ball_self ε_pos)), replace bound_integrable := bound_integrable.norm, apply has_fderiv_at_integral_of_dominated_loc_of_lip' δ_pos; assumption @@ -183,7 +183,7 @@ lemma has_fderiv_at_integral_of_dominated_of_fderiv_le {F : H → α → E} {F' (hF_meas : ∀ᶠ x in 𝓝 x₀, ae_strongly_measurable (F x) μ) (hF_int : integrable (F x₀) μ) (hF'_meas : ae_strongly_measurable (F' x₀) μ) - (h_bound : ∀ᵐ a ∂μ, ∀ x ∈ ball x₀ ε, ∥F' x a∥ ≤ bound a) + (h_bound : ∀ᵐ a ∂μ, ∀ x ∈ ball x₀ ε, ‖F' x a‖ ≤ bound a) (bound_integrable : integrable (bound : α → ℝ) μ) (h_diff : ∀ᵐ a ∂μ, ∀ x ∈ ball x₀ ε, has_fderiv_at (λ x, F x a) (F' x a) x) : has_fderiv_at (λ x, ∫ a, F x a ∂μ) (∫ a, F' x₀ a ∂μ) x₀ := @@ -245,7 +245,7 @@ lemma has_deriv_at_integral_of_dominated_loc_of_deriv_le {F : 𝕜 → α → E} (hF_int : integrable (F x₀) μ) (hF'_meas : ae_strongly_measurable (F' x₀) μ) {bound : α → ℝ} - (h_bound : ∀ᵐ a ∂μ, ∀ x ∈ ball x₀ ε, ∥F' x a∥ ≤ bound a) + (h_bound : ∀ᵐ a ∂μ, ∀ x ∈ ball x₀ ε, ‖F' x a‖ ≤ bound a) (bound_integrable : integrable bound μ) (h_diff : ∀ᵐ a ∂μ, ∀ x ∈ ball x₀ ε, has_deriv_at (λ x, F x a) (F' x a) x) : (integrable (F' x₀) μ) ∧ has_deriv_at (λn, ∫ a, F n a ∂μ) (∫ a, F' x₀ a ∂μ) x₀ := diff --git a/src/analysis/calculus/parametric_interval_integral.lean b/src/analysis/calculus/parametric_interval_integral.lean index 1b8e62cef1690..03a19aa2a3649 100644 --- a/src/analysis/calculus/parametric_interval_integral.lean +++ b/src/analysis/calculus/parametric_interval_integral.lean @@ -55,7 +55,7 @@ lemma has_fderiv_at_integral_of_dominated_of_fderiv_le {F : H → ℝ → E} {F' (hF_meas : ∀ᶠ x in 𝓝 x₀, ae_strongly_measurable (F x) (μ.restrict (Ι a b))) (hF_int : interval_integrable (F x₀) μ a b) (hF'_meas : ae_strongly_measurable (F' x₀) (μ.restrict (Ι a b))) - (h_bound : ∀ᵐ t ∂μ, t ∈ Ι a b → ∀ x ∈ ball x₀ ε, ∥F' x t∥ ≤ bound t) + (h_bound : ∀ᵐ t ∂μ, t ∈ Ι a b → ∀ x ∈ ball x₀ ε, ‖F' x t‖ ≤ bound t) (bound_integrable : interval_integrable bound μ a b) (h_diff : ∀ᵐ t ∂μ, t ∈ Ι a b → ∀ x ∈ ball x₀ ε, has_fderiv_at (λ x, F x t) (F' x t) x) : has_fderiv_at (λ x, ∫ t in a..b, F x t ∂μ) (∫ t in a..b, F' x₀ t ∂μ) x₀ := @@ -98,7 +98,7 @@ lemma has_deriv_at_integral_of_dominated_loc_of_deriv_le {F : 𝕜 → ℝ → E (hF_meas : ∀ᶠ x in 𝓝 x₀, ae_strongly_measurable (F x) (μ.restrict (Ι a b))) (hF_int : interval_integrable (F x₀) μ a b) (hF'_meas : ae_strongly_measurable (F' x₀) (μ.restrict (Ι a b))) - (h_bound : ∀ᵐ t ∂μ, t ∈ Ι a b → ∀ x ∈ ball x₀ ε, ∥F' x t∥ ≤ bound t) + (h_bound : ∀ᵐ t ∂μ, t ∈ Ι a b → ∀ x ∈ ball x₀ ε, ‖F' x t‖ ≤ bound t) (bound_integrable : interval_integrable bound μ a b) (h_diff : ∀ᵐ t ∂μ, t ∈ Ι a b → ∀ x ∈ ball x₀ ε, has_deriv_at (λ x, F x t) (F' x t) x) : (interval_integrable (F' x₀) μ a b) ∧ diff --git a/src/analysis/calculus/tangent_cone.lean b/src/analysis/calculus/tangent_cone.lean index cfdc15a72bf7e..9e8ab6413dd15 100644 --- a/src/analysis/calculus/tangent_cone.lean +++ b/src/analysis/calculus/tangent_cone.lean @@ -42,7 +42,7 @@ variables {E : Type*} [add_comm_monoid E] [module 𝕜 E] [topological_space E] /-- The set of all tangent directions to the set `s` at the point `x`. -/ def tangent_cone_at (s : set E) (x : E) : set E := {y : E | ∃(c : ℕ → 𝕜) (d : ℕ → E), (∀ᶠ n in at_top, x + d n ∈ s) ∧ - (tendsto (λn, ∥c n∥) at_top at_top) ∧ (tendsto (λn, c n • d n) at_top (𝓝 y))} + (tendsto (λn, ‖c n‖) at_top at_top) ∧ (tendsto (λn, c n • d n) at_top (𝓝 y))} /-- A property ensuring that the tangent cone to `s` at `x` spans a dense subset of the whole space. The main role of this property is to ensure that the differential within `s` at `x` is unique, @@ -99,19 +99,19 @@ end /-- Auxiliary lemma ensuring that, under the assumptions defining the tangent cone, the sequence `d` tends to 0 at infinity. -/ lemma tangent_cone_at.lim_zero {α : Type*} (l : filter α) {c : α → 𝕜} {d : α → E} - (hc : tendsto (λn, ∥c n∥) l at_top) (hd : tendsto (λn, c n • d n) l (𝓝 y)) : + (hc : tendsto (λn, ‖c n‖) l at_top) (hd : tendsto (λn, c n • d n) l (𝓝 y)) : tendsto d l (𝓝 0) := begin - have A : tendsto (λn, ∥c n∥⁻¹) l (𝓝 0) := tendsto_inv_at_top_zero.comp hc, - have B : tendsto (λn, ∥c n • d n∥) l (𝓝 ∥y∥) := + have A : tendsto (λn, ‖c n‖⁻¹) l (𝓝 0) := tendsto_inv_at_top_zero.comp hc, + have B : tendsto (λn, ‖c n • d n‖) l (𝓝 ‖y‖) := (continuous_norm.tendsto _).comp hd, - have C : tendsto (λn, ∥c n∥⁻¹ * ∥c n • d n∥) l (𝓝 (0 * ∥y∥)) := A.mul B, + have C : tendsto (λn, ‖c n‖⁻¹ * ‖c n • d n‖) l (𝓝 (0 * ‖y‖)) := A.mul B, rw zero_mul at C, - have : ∀ᶠ n in l, ∥c n∥⁻¹ * ∥c n • d n∥ = ∥d n∥, + have : ∀ᶠ n in l, ‖c n‖⁻¹ * ‖c n • d n‖ = ‖d n‖, { apply (eventually_ne_of_tendsto_norm_at_top hc 0).mono (λn hn, _), rw [norm_smul, ← mul_assoc, inv_mul_cancel, one_mul], rwa [ne.def, norm_eq_zero] }, - have D : tendsto (λ n, ∥d n∥) l (𝓝 0) := + have D : tendsto (λ n, ‖d n‖) l (𝓝 0) := tendsto.congr' this C, rw tendsto_zero_iff_norm_tendsto_zero, exact D @@ -145,7 +145,7 @@ lemma subset_tangent_cone_prod_left {t : set F} {y : F} (ht : y ∈ closure t) : linear_map.inl 𝕜 E F '' (tangent_cone_at 𝕜 s x) ⊆ tangent_cone_at 𝕜 (s ×ˢ t) (x, y) := begin rintros _ ⟨v, ⟨c, d, hd, hc, hy⟩, rfl⟩, - have : ∀n, ∃d', y + d' ∈ t ∧ ∥c n • d'∥ < ((1:ℝ)/2)^n, + have : ∀n, ∃d', y + d' ∈ t ∧ ‖c n • d'‖ < ((1:ℝ)/2)^n, { assume n, rcases mem_closure_iff_nhds.1 ht _ (eventually_nhds_norm_smul_sub_lt (c n) y (pow_pos one_half_pos n)) with ⟨z, hz, hzt⟩, @@ -166,7 +166,7 @@ lemma subset_tangent_cone_prod_right {t : set F} {y : F} linear_map.inr 𝕜 E F '' (tangent_cone_at 𝕜 t y) ⊆ tangent_cone_at 𝕜 (s ×ˢ t) (x, y) := begin rintros _ ⟨w, ⟨c, d, hd, hc, hy⟩, rfl⟩, - have : ∀n, ∃d', x + d' ∈ s ∧ ∥c n • d'∥ < ((1:ℝ)/2)^n, + have : ∀n, ∃d', x + d' ∈ s ∧ ‖c n • d'‖ < ((1:ℝ)/2)^n, { assume n, rcases mem_closure_iff_nhds.1 hs _ (eventually_nhds_norm_smul_sub_lt (c n) x (pow_pos one_half_pos n)) with ⟨z, hz, hzs⟩, @@ -189,7 +189,7 @@ lemma maps_to_tangent_cone_pi {ι : Type*} [decidable_eq ι] {E : ι → Type*} (tangent_cone_at 𝕜 (set.pi univ s) x) := begin rintros w ⟨c, d, hd, hc, hy⟩, - have : ∀ n (j ≠ i), ∃ d', x j + d' ∈ s j ∧ ∥c n • d'∥ < (1 / 2 : ℝ) ^ n, + have : ∀ n (j ≠ i), ∃ d', x j + d' ∈ s j ∧ ‖c n • d'‖ < (1 / 2 : ℝ) ^ n, { assume n j hj, rcases mem_closure_iff_nhds.1 (hi j hj) _ (eventually_nhds_norm_smul_sub_lt (c n) (x j) (pow_pos one_half_pos n)) with ⟨z, hz, hzs⟩, @@ -219,8 +219,8 @@ begin { rw inv_pos, apply pow_pos, norm_num }, { apply inv_lt_one, apply one_lt_pow _ (nat.succ_ne_zero _), norm_num }, { simp only [d, sub_smul, smul_sub, one_smul], abel } }, - show filter.tendsto (λ (n : ℕ), ∥c n∥) filter.at_top filter.at_top, - { have : (λ (n : ℕ), ∥c n∥) = c, + show filter.tendsto (λ (n : ℕ), ‖c n‖) filter.at_top filter.at_top, + { have : (λ (n : ℕ), ‖c n‖) = c, by { ext n, exact abs_of_nonneg (pow_nonneg (by norm_num) _) }, rw this, exact (tendsto_pow_at_top_at_top_of_one_lt (by norm_num)).comp (tendsto_add_at_top_nat 1) }, diff --git a/src/analysis/calculus/taylor.lean b/src/analysis/calculus/taylor.lean index 56e059279c78c..31527b5521359 100644 --- a/src/analysis/calculus/taylor.lean +++ b/src/analysis/calculus/taylor.lean @@ -341,8 +341,8 @@ The difference of `f` and its `n`-th Taylor polynomial can be estimated by `C * (x - a)^(n+1) / n!` where `C` is a bound for the `n+1`-th iterated derivative of `f`. -/ lemma taylor_mean_remainder_bound {f : ℝ → E} {a b C x : ℝ} {n : ℕ} (hab : a ≤ b) (hf : cont_diff_on ℝ (n+1) f (Icc a b)) (hx : x ∈ Icc a b) - (hC : ∀ y ∈ Icc a b, ∥iterated_deriv_within (n + 1) f (Icc a b) y∥ ≤ C) : - ∥f x - taylor_within_eval f n (Icc a b) a x∥ ≤ C * (x - a)^(n+1) / n! := + (hC : ∀ y ∈ Icc a b, ‖iterated_deriv_within (n + 1) f (Icc a b) y‖ ≤ C) : + ‖f x - taylor_within_eval f n (Icc a b) a x‖ ≤ C * (x - a)^(n+1) / n! := begin rcases eq_or_lt_of_le hab with rfl|h, { rw [Icc_self, mem_singleton_iff] at hx, @@ -353,7 +353,7 @@ begin (unique_diff_on_Icc h), -- We can uniformly bound the derivative of the Taylor polynomial have h' : ∀ (y : ℝ) (hy : y ∈ Ico a x), - ∥((n! : ℝ)⁻¹ * (x - y) ^ n) • iterated_deriv_within (n + 1) f (Icc a b) y∥ + ‖((n! : ℝ)⁻¹ * (x - y) ^ n) • iterated_deriv_within (n + 1) f (Icc a b) y‖ ≤ (n! : ℝ)⁻¹ * |(x - a)|^n * C, { rintro y ⟨hay, hyx⟩, rw [norm_smul, real.norm_eq_abs], @@ -386,14 +386,14 @@ There exists a constant `C` such that for all `x ∈ Icc a b` the difference of Taylor polynomial can be estimated by `C * (x - a)^(n+1)`. -/ lemma exists_taylor_mean_remainder_bound {f : ℝ → E} {a b : ℝ} {n : ℕ} (hab : a ≤ b) (hf : cont_diff_on ℝ (n+1) f (Icc a b)) : - ∃ C, ∀ x ∈ Icc a b, ∥f x - taylor_within_eval f n (Icc a b) a x∥ ≤ C * (x - a)^(n+1) := + ∃ C, ∀ x ∈ Icc a b, ‖f x - taylor_within_eval f n (Icc a b) a x‖ ≤ C * (x - a)^(n+1) := begin rcases eq_or_lt_of_le hab with rfl|h, { refine ⟨0, λ x hx, _⟩, have : a = x, by simpa [← le_antisymm_iff] using hx, simp [← this] }, -- We estimate by the supremum of the norm of the iterated derivative - let g : ℝ → ℝ := λ y, ∥iterated_deriv_within (n + 1) f (Icc a b) y∥, + let g : ℝ → ℝ := λ y, ‖iterated_deriv_within (n + 1) f (Icc a b) y‖, use [has_Sup.Sup (g '' Icc a b) / n!], intros x hx, rw div_mul_eq_mul_div₀, diff --git a/src/analysis/calculus/uniform_limits_deriv.lean b/src/analysis/calculus/uniform_limits_deriv.lean index c4c10e998eaa9..d3ad4018fee8a 100644 --- a/src/analysis/calculus/uniform_limits_deriv.lean +++ b/src/analysis/calculus/uniform_limits_deriv.lean @@ -136,11 +136,11 @@ begin have hr : 0 < r, { simp [hR], }, have hr' : ∀ ⦃y : E⦄, y ∈ metric.ball x r → c y, { exact (λ y hy, hR' (lt_of_lt_of_le (metric.mem_ball.mp hy) (min_le_right _ _))), }, - have hxy : ∀ (y : E), y ∈ metric.ball x r → ∥y - x∥ < 1, + have hxy : ∀ (y : E), y ∈ metric.ball x r → ‖y - x‖ < 1, { intros y hy, rw [metric.mem_ball, dist_eq_norm] at hy, exact lt_of_lt_of_le hy (min_le_left _ _), }, - have hxyε : ∀ (y : E), y ∈ metric.ball x r → ε * ∥y - x∥ < ε, + have hxyε : ∀ (y : E), y ∈ metric.ball x r → ε * ‖y - x‖ < ε, { intros y hy, exact (mul_lt_iff_lt_one_right hε.lt).mpr (hxy y hy), }, @@ -211,7 +211,7 @@ begin (λ z hz, ((hf n.1 z hz).sub (hf n.2 z hz)).has_fderiv_within_at) (λ z hz, (hn z hz).le) (convex_ball x r) (metric.mem_ball_self hr) hy, refine lt_of_le_of_lt mvt _, - have : q * ∥y - x∥ < q * r, + have : q * ‖y - x‖ < q * r, { exact mul_lt_mul' rfl.le (by simpa only [dist_eq_norm] using metric.mem_ball.mp hy) (norm_nonneg _) hqpos, }, exact this.trans hq, }, @@ -260,15 +260,15 @@ begin end /-- If `f_n → g` pointwise and the derivatives `(f_n)' → h` _uniformly_ converge, then -in fact for a fixed `y`, the difference quotients `∥z - y∥⁻¹ • (f_n z - f_n y)` converge -_uniformly_ to `∥z - y∥⁻¹ • (g z - g y)` -/ +in fact for a fixed `y`, the difference quotients `‖z - y‖⁻¹ • (f_n z - f_n y)` converge +_uniformly_ to `‖z - y‖⁻¹ • (g z - g y)` -/ lemma difference_quotients_converge_uniformly (hf' : tendsto_uniformly_on_filter f' g' l (𝓝 x)) (hf : ∀ᶠ (n : ι × E) in (l ×ᶠ 𝓝 x), has_fderiv_at (f n.1) (f' n.1 n.2) n.2) (hfg : ∀ᶠ (y : E) in 𝓝 x, tendsto (λ n, f n y) l (𝓝 (g y))) : tendsto_uniformly_on_filter - (λ n : ι, λ y : E, (∥y - x∥⁻¹ : 𝕜) • (f n y - f n x)) - (λ y : E, (∥y - x∥⁻¹ : 𝕜) • (g y - g x)) + (λ n : ι, λ y : E, (‖y - x‖⁻¹ : 𝕜) • (f n y - f n x)) + (λ y : E, (‖y - x‖⁻¹ : 𝕜) • (g y - g x)) l (𝓝 x) := begin letI : normed_space ℝ E, from normed_space.restrict_scalars ℝ 𝕜 _, @@ -298,7 +298,7 @@ begin rw [← smul_sub, norm_smul, norm_inv, is_R_or_C.norm_coe_norm], refine lt_of_le_of_lt _ hqε, by_cases hyz' : x = y, { simp [hyz', hqpos.le], }, - have hyz : 0 < ∥y - x∥, + have hyz : 0 < ‖y - x‖, {rw norm_pos_iff, intros hy', exact hyz' (eq_of_sub_eq_zero hy').symm, }, rw [inv_mul_le_iff hyz, mul_comm, sub_sub_sub_comm], simp only [pi.zero_apply, dist_zero_left] at e, @@ -333,7 +333,7 @@ begin -- Introduce extra quantifier via curried filters suffices : tendsto - (λ (y : ι × E), ∥y.2 - x∥⁻¹ * ∥g y.2 - g x - (g' x) (y.2 - x)∥) (l.curry (𝓝 x)) (𝓝 0), + (λ (y : ι × E), ‖y.2 - x‖⁻¹ * ‖g y.2 - g x - (g' x) (y.2 - x)‖) (l.curry (𝓝 x)) (𝓝 0), { rw metric.tendsto_nhds at this ⊢, intros ε hε, specialize this ε hε, @@ -352,11 +352,11 @@ begin rw [←norm_norm, ←norm_inv,←@is_R_or_C.norm_of_real 𝕜 _ _, is_R_or_C.of_real_inv, ←norm_smul], }, rw ←tendsto_zero_iff_norm_tendsto_zero, - have : (λ a : ι × E, (∥a.2 - x∥⁻¹ : 𝕜) • (g a.2 - g x - (g' x) (a.2 - x))) = - (λ a : ι × E, (∥a.2 - x∥⁻¹ : 𝕜) • (g a.2 - g x - (f a.1 a.2 - f a.1 x))) + - (λ a : ι × E, (∥a.2 - x∥⁻¹ : 𝕜) • ((f a.1 a.2 - f a.1 x) - + have : (λ a : ι × E, (‖a.2 - x‖⁻¹ : 𝕜) • (g a.2 - g x - (g' x) (a.2 - x))) = + (λ a : ι × E, (‖a.2 - x‖⁻¹ : 𝕜) • (g a.2 - g x - (f a.1 a.2 - f a.1 x))) + + (λ a : ι × E, (‖a.2 - x‖⁻¹ : 𝕜) • ((f a.1 a.2 - f a.1 x) - ((f' a.1 x) a.2 - (f' a.1 x) x))) + - (λ a : ι × E, (∥a.2 - x∥⁻¹ : 𝕜) • ((f' a.1 x - g' x) (a.2 - x))), + (λ a : ι × E, (‖a.2 - x‖⁻¹ : 𝕜) • ((f' a.1 x - g' x) (a.2 - x))), { ext, simp only [pi.add_apply], rw [←smul_add, ←smul_add], congr, simp only [map_sub, sub_add_sub_cancel, continuous_linear_map.coe_sub', pi.sub_apply], }, simp_rw this, @@ -398,7 +398,7 @@ begin intros n, simp_rw [norm_smul, norm_inv, is_R_or_C.norm_coe_norm], by_cases hx : x = n.2, { simp [hx], }, - have hnx : 0 < ∥n.2 - x∥, + have hnx : 0 < ‖n.2 - x‖, { rw norm_pos_iff, intros hx', exact hx (eq_of_sub_eq_zero hx').symm, }, rw [inv_mul_le_iff hnx, mul_comm], simp only [function.comp_app, prod_map], diff --git a/src/analysis/complex/abs_max.lean b/src/analysis/complex/abs_max.lean index e69e16c11443f..61a9eae8571e3 100644 --- a/src/analysis/complex/abs_max.lean +++ b/src/analysis/complex/abs_max.lean @@ -22,29 +22,29 @@ conclusion (either equality of norms or of the values of the function). ### Theorems for any codomain Consider a function `f : E → F` that is complex differentiable on a set `s`, is continuous on its -closure, and `∥f x∥` has a maximum on `s` at `c`. We prove the following theorems. +closure, and `‖f x‖` has a maximum on `s` at `c`. We prove the following theorems. -- `complex.norm_eq_on_closed_ball_of_is_max_on`: if `s = metric.ball c r`, then `∥f x∥ = ∥f c∥` for +- `complex.norm_eq_on_closed_ball_of_is_max_on`: if `s = metric.ball c r`, then `‖f x‖ = ‖f c‖` for any `x` from the corresponding closed ball; - `complex.norm_eq_norm_of_is_max_on_of_ball_subset`: if `metric.ball c (dist w c) ⊆ s`, then - `∥f w∥ = ∥f c∥`; + `‖f w‖ = ‖f c‖`; - `complex.norm_eq_on_of_is_preconnected_of_is_max_on`: if `U` is an open (pre)connected set, `f` is - complex differentiable on `U`, and `∥f x∥` has a maximum on `U` at `c ∈ U`, then `∥f x∥ = ∥f c∥` + complex differentiable on `U`, and `‖f x‖` has a maximum on `U` at `c ∈ U`, then `‖f x‖ = ‖f c‖` for all `x ∈ U`; - `complex.norm_eq_on_closure_of_is_preconnected_of_is_max_on`: if `s` is open and (pre)connected - and `c ∈ s`, then `∥f x∥ = ∥f c∥` for all `x ∈ closure s`; + and `c ∈ s`, then `‖f x‖ = ‖f c‖` for all `x ∈ closure s`; - `complex.norm_eventually_eq_of_is_local_max`: if `f` is complex differentiable in a neighborhood - of `c` and `∥f x∥` has a local maximum at `c`, then `∥f x∥` is locally a constant in a + of `c` and `‖f x‖` has a local maximum at `c`, then `‖f x‖` is locally a constant in a neighborhood of `c`. ### Theorems for a strictly convex codomain If the codomain `F` is a strictly convex space, then in the lemmas from the previous section we can -prove `f w = f c` instead of `∥f w∥ = ∥f c∥`, see +prove `f w = f c` instead of `‖f w‖ = ‖f c‖`, see `complex.eq_on_of_is_preconnected_of_is_max_on_norm`, `complex.eq_on_closure_of_is_preconnected_of_is_max_on_norm`, `complex.eq_of_is_max_on_of_ball_subset`, `complex.eq_on_closed_ball_of_is_max_on_norm`, and @@ -58,11 +58,11 @@ this section `f g : E → F` are functions that are complex differentiable on a are continuous on its closure. We prove the following theorems. - `complex.exists_mem_frontier_is_max_on_norm`: If `E` is a finite dimensional space and `s` is a - nonempty bounded set, then there exists a point `z ∈ frontier s` such that `λ z, ∥f z∥` takes it + nonempty bounded set, then there exists a point `z ∈ frontier s` such that `λ z, ‖f z‖` takes it maximum value on `closure s` at `z`. -- `complex.norm_le_of_forall_mem_frontier_norm_le`: if `∥f z∥ ≤ C` for all `z ∈ frontier s`, then - `∥f z∥ ≤ C` for all `z ∈ s`; note that this theorem does not require `E` to be a finite +- `complex.norm_le_of_forall_mem_frontier_norm_le`: if `‖f z‖ ≤ C` for all `z ∈ frontier s`, then + `‖f z‖ ≤ C` for all `z ∈ s`; note that this theorem does not require `E` to be a finite dimensional space. - `complex.eq_on_closure_of_eq_on_frontier`: if `f x = g x` on the frontier of `s`, then `f x = g x` @@ -101,24 +101,24 @@ file. lemma norm_max_aux₁ [complete_space F] {f : ℂ → F} {z w : ℂ} (hd : diff_cont_on_cl ℂ f (ball z (dist w z))) (hz : is_max_on (norm ∘ f) (closed_ball z (dist w z)) z) : - ∥f w∥ = ∥f z∥ := + ‖f w‖ = ‖f z‖ := begin /- Consider a circle of radius `r = dist w z`. -/ set r : ℝ := dist w z, have hw : w ∈ closed_ball z r, from mem_closed_ball.2 le_rfl, - /- Assume the converse. Since `∥f w∥ ≤ ∥f z∥`, we have `∥f w∥ < ∥f z∥`. -/ + /- Assume the converse. Since `‖f w‖ ≤ ‖f z‖`, we have `‖f w‖ < ‖f z‖`. -/ refine (is_max_on_iff.1 hz _ hw).antisymm (not_lt.1 _), - rintro hw_lt : ∥f w∥ < ∥f z∥, + rintro hw_lt : ‖f w‖ < ‖f z‖, have hr : 0 < r, from dist_pos.2 (ne_of_apply_ne (norm ∘ f) hw_lt.ne), /- Due to Cauchy integral formula, it suffices to prove the following inequality. -/ - suffices : ∥∮ ζ in C(z, r), (ζ - z)⁻¹ • f ζ∥ < 2 * π * ∥f z∥, + suffices : ‖∮ ζ in C(z, r), (ζ - z)⁻¹ • f ζ‖ < 2 * π * ‖f z‖, { refine this.ne _, have A : ∮ ζ in C(z, r), (ζ - z)⁻¹ • f ζ = (2 * π * I : ℂ) • f z := hd.circle_integral_sub_inv_smul (mem_ball_self hr), simp [A, norm_smul, real.pi_pos.le] }, - suffices : ∥∮ ζ in C(z, r), (ζ - z)⁻¹ • f ζ∥ < 2 * π * r * (∥f z∥ / r), + suffices : ‖∮ ζ in C(z, r), (ζ - z)⁻¹ • f ζ‖ < 2 * π * r * (‖f z‖ / r), by rwa [mul_assoc, mul_div_cancel' _ hr.ne'] at this, - /- This inequality is true because `∥(ζ - z)⁻¹ • f ζ∥ ≤ ∥f z∥ / r` for all `ζ` on the circle and + /- This inequality is true because `‖(ζ - z)⁻¹ • f ζ‖ ≤ ‖f z‖ / r` for all `ζ` on the circle and this inequality is strict at `ζ = w`. -/ have hsub : sphere z r ⊆ closed_ball z r, from sphere_subset_closed_ball, refine circle_integral.norm_integral_lt_of_norm_le_const_of_lt hr _ _ ⟨w, rfl, _⟩, @@ -126,11 +126,11 @@ begin { refine ((continuous_on_id.sub continuous_on_const).inv₀ _).smul (hd.continuous_on_ball.mono hsub), exact λ ζ hζ, sub_ne_zero.2 (ne_of_mem_sphere hζ hr.ne') }, - show ∀ ζ ∈ sphere z r, ∥(ζ - z)⁻¹ • f ζ∥ ≤ ∥f z∥ / r, + show ∀ ζ ∈ sphere z r, ‖(ζ - z)⁻¹ • f ζ‖ ≤ ‖f z‖ / r, { rintros ζ (hζ : abs (ζ - z) = r), rw [le_div_iff hr, norm_smul, norm_inv, norm_eq_abs, hζ, mul_comm, mul_inv_cancel_left₀ hr.ne'], exact hz (hsub hζ) }, - show ∥(w - z)⁻¹ • f w∥ < ∥f z∥ / r, + show ‖(w - z)⁻¹ • f w‖ < ‖f z‖ / r, { rw [norm_smul, norm_inv, norm_eq_abs, ← div_eq_inv_mul], exact (div_lt_div_right hr).2 hw_lt } end @@ -141,10 +141,10 @@ Now we drop the assumption `complete_space F` by embedding `F` into its completi lemma norm_max_aux₂ {f : ℂ → F} {z w : ℂ} (hd : diff_cont_on_cl ℂ f (ball z (dist w z))) (hz : is_max_on (norm ∘ f) (closed_ball z (dist w z)) z) : - ∥f w∥ = ∥f z∥ := + ‖f w‖ = ‖f z‖ := begin set e : F →L[ℂ] F̂ := uniform_space.completion.to_complL, - have he : ∀ x, ∥e x∥ = ∥x∥, from uniform_space.completion.norm_coe, + have he : ∀ x, ‖e x‖ = ‖x‖, from uniform_space.completion.norm_coe, replace hz : is_max_on (norm ∘ (e ∘ f)) (closed_ball z (dist w z)) z, by simpa only [is_max_on, (∘), he] using hz, simpa only [he] using norm_max_aux₁ (e.differentiable.comp_diff_cont_on_cl hd) hz @@ -157,7 +157,7 @@ assumption `is_max_on (norm ∘ f) (ball z r) z`. lemma norm_max_aux₃ {f : ℂ → F} {z w : ℂ} {r : ℝ} (hr : dist w z = r) (hd : diff_cont_on_cl ℂ f (ball z r)) (hz : is_max_on (norm ∘ f) (ball z r) z) : - ∥f w∥ = ∥f z∥ := + ‖f w‖ = ‖f z‖ := begin subst r, rcases eq_or_ne w z with rfl|hne, { refl }, @@ -169,7 +169,7 @@ end ### Maximum modulus principle for any codomain If we do not assume that the codomain is a strictly convex space, then we can only claim that the -**norm** `∥f x∥` is locally constant. +**norm** `‖f x‖` is locally constant. -/ /-! @@ -177,18 +177,18 @@ Finally, we generalize the theorem from a disk in `ℂ` to a closed ball in any -/ /-- **Maximum modulus principle** on a closed ball: if `f : E → F` is continuous on a closed ball, -is complex differentiable on the corresponding open ball, and the norm `∥f w∥` takes its maximum -value on the open ball at its center, then the norm `∥f w∥` is constant on the closed ball. -/ +is complex differentiable on the corresponding open ball, and the norm `‖f w‖` takes its maximum +value on the open ball at its center, then the norm `‖f w‖` is constant on the closed ball. -/ lemma norm_eq_on_closed_ball_of_is_max_on {f : E → F} {z : E} {r : ℝ} (hd : diff_cont_on_cl ℂ f (ball z r)) (hz : is_max_on (norm ∘ f) (ball z r) z) : - eq_on (norm ∘ f) (const E ∥f z∥) (closed_ball z r) := + eq_on (norm ∘ f) (const E ‖f z‖) (closed_ball z r) := begin intros w hw, rw [mem_closed_ball, dist_comm] at hw, rcases eq_or_ne z w with rfl|hne, { refl }, set e : ℂ → E := line_map z w, have hde : differentiable ℂ e := (differentiable_id.smul_const (w - z)).add_const z, - suffices : ∥(f ∘ e) (1 : ℂ)∥ = ∥(f ∘ e) (0 : ℂ)∥, by simpa [e], + suffices : ‖(f ∘ e) (1 : ℂ)‖ = ‖(f ∘ e) (0 : ℂ)‖, by simpa [e], have hr : dist (1 : ℂ) 0 = 1, by simp, have hball : maps_to e (ball 0 1) (ball z r), { refine ((lipschitz_with_line_map z w).maps_to_ball @@ -200,18 +200,18 @@ end /-- **Maximum modulus principle**: if `f : E → F` is complex differentiable on a set `s`, the norm of `f` takes it maximum on `s` at `z`, and `w` is a point such that the closed ball with center `z` -and radius `dist w z` is included in `s`, then `∥f w∥ = ∥f z∥`. -/ +and radius `dist w z` is included in `s`, then `‖f w‖ = ‖f z‖`. -/ lemma norm_eq_norm_of_is_max_on_of_ball_subset {f : E → F} {s : set E} {z w : E} (hd : diff_cont_on_cl ℂ f s) (hz : is_max_on (norm ∘ f) s z) (hsub : ball z (dist w z) ⊆ s) : - ∥f w∥ = ∥f z∥ := + ‖f w‖ = ‖f z‖ := norm_eq_on_closed_ball_of_is_max_on (hd.mono hsub) (hz.on_subset hsub) (mem_closed_ball.2 le_rfl) /-- **Maximum modulus principle**: if `f : E → F` is complex differentiable in a neighborhood of `c` -and the norm `∥f z∥` has a local maximum at `c`, then `∥f z∥` is locally constant in a neighborhood +and the norm `‖f z‖` has a local maximum at `c`, then `‖f z‖` is locally constant in a neighborhood of `c`. -/ lemma norm_eventually_eq_of_is_local_max {f : E → F} {c : E} (hd : ∀ᶠ z in 𝓝 c, differentiable_at ℂ f z) (hc : is_local_max (norm ∘ f) c) : - ∀ᶠ y in 𝓝 c, ∥f y∥ = ∥f c∥ := + ∀ᶠ y in 𝓝 c, ‖f y‖ = ‖f c‖ := begin rcases nhds_basis_closed_ball.eventually_iff.1 (hd.and hc) with ⟨r, hr₀, hr⟩, exact nhds_basis_closed_ball.eventually_iff.2 ⟨r, hr₀, norm_eq_on_closed_ball_of_is_max_on @@ -232,36 +232,36 @@ end /-- **Maximum modulus principle** on a connected set. Let `U` be a (pre)connected open set in a complex normed space. Let `f : E → F` be a function that is complex differentiable on `U`. Suppose -that `∥f x∥` takes its maximum value on `U` at `c ∈ U`. Then `∥f x∥ = ∥f c∥` for all `x ∈ U`. -/ +that `‖f x‖` takes its maximum value on `U` at `c ∈ U`. Then `‖f x‖ = ‖f c‖` for all `x ∈ U`. -/ lemma norm_eq_on_of_is_preconnected_of_is_max_on {f : E → F} {U : set E} {c : E} (hc : is_preconnected U) (ho : is_open U) (hd : differentiable_on ℂ f U) (hcU : c ∈ U) (hm : is_max_on (norm ∘ f) U c) : - eq_on (norm ∘ f) (const E ∥f c∥) U := + eq_on (norm ∘ f) (const E ‖f c‖) U := begin set V := U ∩ {z | is_max_on (norm ∘ f) U z}, - have hV : ∀ x ∈ V, ∥f x∥ = ∥f c∥, from λ x hx, le_antisymm (hm hx.1) (hx.2 hcU), + have hV : ∀ x ∈ V, ‖f x‖ = ‖f c‖, from λ x hx, le_antisymm (hm hx.1) (hx.2 hcU), suffices : U ⊆ V, from λ x hx, hV x (this hx), have hVo : is_open V, { simpa only [ho.mem_nhds_iff, set_of_and, set_of_mem_eq] using is_open_set_of_mem_nhds_and_is_max_on_norm hd }, have hVne : (U ∩ V).nonempty := ⟨c, hcU, hcU, hm⟩, - set W := U ∩ {z | ∥f z∥ ≠ ∥f c∥}, + set W := U ∩ {z | ‖f z‖ ≠ ‖f c‖}, have hWo : is_open W, from hd.continuous_on.norm.preimage_open_of_open ho is_open_ne, have hdVW : disjoint V W, from disjoint_left.mpr (λ x hxV hxW, hxW.2 (hV x hxV)), have hUVW : U ⊆ V ∪ W, - from λ x hx, (eq_or_ne (∥f x∥) (∥f c∥)).imp (λ h, ⟨hx, λ y hy, (hm hy).out.trans_eq h.symm⟩) + from λ x hx, (eq_or_ne (‖f x‖) (‖f c‖)).imp (λ h, ⟨hx, λ y hy, (hm hy).out.trans_eq h.symm⟩) (and.intro hx), exact hc.subset_left_of_subset_union hVo hWo hdVW hUVW hVne, end /-- **Maximum modulus principle** on a connected set. Let `U` be a (pre)connected open set in a complex normed space. Let `f : E → F` be a function that is complex differentiable on `U` and is -continuous on its closure. Suppose that `∥f x∥` takes its maximum value on `U` at `c ∈ U`. Then -`∥f x∥ = ∥f c∥` for all `x ∈ closure U`. -/ +continuous on its closure. Suppose that `‖f x‖` takes its maximum value on `U` at `c ∈ U`. Then +`‖f x‖ = ‖f c‖` for all `x ∈ closure U`. -/ lemma norm_eq_on_closure_of_is_preconnected_of_is_max_on {f : E → F} {U : set E} {c : E} (hc : is_preconnected U) (ho : is_open U) (hd : diff_cont_on_cl ℂ f U) (hcU : c ∈ U) (hm : is_max_on (norm ∘ f) U c) : - eq_on (norm ∘ f) (const E ∥f c∥) (closure U) := + eq_on (norm ∘ f) (const E ‖f c‖) (closure U) := (norm_eq_on_of_is_preconnected_of_is_max_on hc ho hd.differentiable_on hcU hm).of_subset_closure hd.continuous_on.norm continuous_on_const subset_closure subset.rfl @@ -271,11 +271,11 @@ section strict_convex ### The case of a strictly convex codomain If the codomain `F` is a strictly convex space, then we can claim equalities like `f w = f z` -instead of `∥f w∥ = ∥f z∥`. +instead of `‖f w‖ = ‖f z‖`. Instead of repeating the proof starting with lemmas about integrals, we apply a corresponding lemma -above twice: for `f` and for `λ x, f x + f c`. Then we have `∥f w∥ = ∥f z∥` and -`∥f w + f z∥ = ∥f z + f z∥`, thus `∥f w + f z∥ = ∥f w∥ + ∥f z∥`. This is only possible if +above twice: for `f` and for `λ x, f x + f c`. Then we have `‖f w‖ = ‖f z‖` and +`‖f w + f z‖ = ‖f z + f z‖`, thus `‖f w + f z‖ = ‖f w‖ + ‖f z‖`. This is only possible if `f w = f z`, see `eq_of_norm_eq_of_norm_add_eq`. -/ @@ -283,7 +283,7 @@ variables [strict_convex_space ℝ F] /-- **Maximum modulus principle** on a connected set. Let `U` be a (pre)connected open set in a complex normed space. Let `f : E → F` be a function that is complex differentiable on `U`. Suppose -that `∥f x∥` takes its maximum value on `U` at `c ∈ U`. Then `f x = f c` for all `x ∈ U`. +that `‖f x‖` takes its maximum value on `U` at `c ∈ U`. Then `f x = f c` for all `x ∈ U`. TODO: change assumption from `is_max_on` to `is_local_max`. -/ lemma eq_on_of_is_preconnected_of_is_max_on_norm {f : E → F} {U : set E} {c : E} @@ -291,14 +291,14 @@ lemma eq_on_of_is_preconnected_of_is_max_on_norm {f : E → F} {U : set E} {c : (hm : is_max_on (norm ∘ f) U c) : eq_on f (const E (f c)) U := λ x hx, -have H₁ : ∥f x∥ = ∥f c∥, from norm_eq_on_of_is_preconnected_of_is_max_on hc ho hd hcU hm hx, -have H₂ : ∥f x + f c∥ = ∥f c + f c∥, +have H₁ : ‖f x‖ = ‖f c‖, from norm_eq_on_of_is_preconnected_of_is_max_on hc ho hd hcU hm hx, +have H₂ : ‖f x + f c‖ = ‖f c + f c‖, from norm_eq_on_of_is_preconnected_of_is_max_on hc ho (hd.add_const _) hcU hm.norm_add_self hx, eq_of_norm_eq_of_norm_add_eq H₁ $ by simp only [H₂, same_ray.rfl.norm_add, H₁] /-- **Maximum modulus principle** on a connected set. Let `U` be a (pre)connected open set in a complex normed space. Let `f : E → F` be a function that is complex differentiable on `U` and is -continuous on its closure. Suppose that `∥f x∥` takes its maximum value on `U` at `c ∈ U`. Then +continuous on its closure. Suppose that `‖f x‖` takes its maximum value on `U` at `c ∈ U`. Then `f x = f c` for all `x ∈ closure U`. -/ lemma eq_on_closure_of_is_preconnected_of_is_max_on_norm {f : E → F} {U : set E} {c : E} (hc : is_preconnected U) (ho : is_open U) (hd : diff_cont_on_cl ℂ f U) (hcU : c ∈ U) @@ -315,8 +315,8 @@ then `f w = f z`. -/ lemma eq_of_is_max_on_of_ball_subset {f : E → F} {s : set E} {z w : E} (hd : diff_cont_on_cl ℂ f s) (hz : is_max_on (norm ∘ f) s z) (hsub : ball z (dist w z) ⊆ s) : f w = f z := -have H₁ : ∥f w∥ = ∥f z∥, from norm_eq_norm_of_is_max_on_of_ball_subset hd hz hsub, -have H₂ : ∥f w + f z∥ = ∥f z + f z∥, +have H₁ : ‖f w‖ = ‖f z‖, from norm_eq_norm_of_is_max_on_of_ball_subset hd hz hsub, +have H₂ : ‖f w + f z‖ = ‖f z + f z‖, from norm_eq_norm_of_is_max_on_of_ball_subset (hd.add_const _) hz.norm_add_self hsub, eq_of_norm_eq_of_norm_add_eq H₁ $ by simp only [H₂, same_ray.rfl.norm_add, H₁] @@ -325,7 +325,7 @@ normed complex space to a strictly convex normed complex space has the following - it is continuous on a closed ball `metric.closed_ball z r`, - it is complex differentiable on the corresponding open ball; -- the norm `∥f w∥` takes its maximum value on the open ball at its center. +- the norm `‖f w‖` takes its maximum value on the open ball at its center. Then `f` is a constant on the closed ball. -/ lemma eq_on_closed_ball_of_is_max_on_norm {f : E → F} {z : E} {r : ℝ} @@ -334,7 +334,7 @@ lemma eq_on_closed_ball_of_is_max_on_norm {f : E → F} {z : E} {r : ℝ} λ x hx, eq_of_is_max_on_of_ball_subset hd hz $ ball_subset_ball hx /-- **Maximum modulus principle**: if `f : E → F` is complex differentiable in a neighborhood of `c` -and the norm `∥f z∥` has a local maximum at `c`, then `f` is locally constant in a neighborhood +and the norm `‖f z‖` has a local maximum at `c`, then `f` is locally constant in a neighborhood of `c`. -/ lemma eventually_eq_of_is_local_max_norm {f : E → F} {c : E} (hd : ∀ᶠ z in 𝓝 c, differentiable_at ℂ f z) (hc : is_local_max (norm ∘ f) c) : @@ -372,7 +372,7 @@ variable [nontrivial E] /-- **Maximum modulus principle**: if `f : E → F` is complex differentiable on a nonempty bounded set `U` and is continuous on its closure, then there exists a point `z ∈ frontier U` such that -`λ z, ∥f z∥` takes it maximum value on `closure U` at `z`. -/ +`λ z, ‖f z‖` takes it maximum value on `closure U` at `z`. -/ lemma exists_mem_frontier_is_max_on_norm [finite_dimensional ℂ E] {f : E → F} {U : set E} (hb : bounded U) (hne : U.nonempty) (hd : diff_cont_on_cl ℂ f U) : ∃ z ∈ frontier U, is_max_on (norm ∘ f) (closure U) z := @@ -391,11 +391,11 @@ begin end /-- **Maximum modulus principle**: if `f : E → F` is complex differentiable on a bounded set `U` and -`∥f z∥ ≤ C` for any `z ∈ frontier U`, then the same is true for any `z ∈ closure U`. -/ +`‖f z‖ ≤ C` for any `z ∈ frontier U`, then the same is true for any `z ∈ closure U`. -/ lemma norm_le_of_forall_mem_frontier_norm_le {f : E → F} {U : set E} (hU : bounded U) - (hd : diff_cont_on_cl ℂ f U) {C : ℝ} (hC : ∀ z ∈ frontier U, ∥f z∥ ≤ C) + (hd : diff_cont_on_cl ℂ f U) {C : ℝ} (hC : ∀ z ∈ frontier U, ‖f z‖ ≤ C) {z : E} (hz : z ∈ closure U) : - ∥f z∥ ≤ C := + ‖f z‖ ≤ C := begin rw [closure_eq_self_union_frontier, union_comm, mem_union] at hz, cases hz, { exact hC z hz }, @@ -410,8 +410,8 @@ begin from hd.comp hde.diff_cont_on_cl (maps_to_preimage _ _), have h₀ : (0 : ℂ) ∈ e ⁻¹' U, by simpa only [e, mem_preimage, line_map_apply_zero], rcases exists_mem_frontier_is_max_on_norm (hL.bounded_preimage hU) ⟨0, h₀⟩ hd with ⟨ζ, hζU, hζ⟩, - calc ∥f z∥ = ∥f (e 0)∥ : by simp only [e, line_map_apply_zero] - ... ≤ ∥f (e ζ)∥ : hζ (subset_closure h₀) + calc ‖f z‖ = ‖f (e 0)‖ : by simp only [e, line_map_apply_zero] + ... ≤ ‖f (e ζ)‖ : hζ (subset_closure h₀) ... ≤ C : hC _ (hde.continuous.frontier_preimage_subset _ hζU) end @@ -421,7 +421,7 @@ lemma eq_on_closure_of_eq_on_frontier {f g : E → F} {U : set E} (hU : bounded (hf : diff_cont_on_cl ℂ f U) (hg : diff_cont_on_cl ℂ g U) (hfg : eq_on f g (frontier U)) : eq_on f g (closure U) := begin - suffices H : ∀ z ∈ closure U, ∥(f - g) z∥ ≤ 0, by simpa [sub_eq_zero] using H, + suffices H : ∀ z ∈ closure U, ‖(f - g) z‖ ≤ 0, by simpa [sub_eq_zero] using H, refine λ z hz, norm_le_of_forall_mem_frontier_norm_le hU (hf.sub hg) (λ w hw, _) hz, simp [hfg hw] end diff --git a/src/analysis/complex/arg.lean b/src/analysis/complex/arg.lean index 34e540b1e5dd0..a1716af4733cb 100644 --- a/src/analysis/complex/arg.lean +++ b/src/analysis/complex/arg.lean @@ -57,7 +57,7 @@ same_ray_iff.mpr $ or.inr $ or.inr h lemma abs_add_eq (h : x.arg = y.arg) : (x + y).abs = x.abs + y.abs := (same_ray_of_arg_eq h).norm_add -lemma abs_sub_eq (h : x.arg = y.arg) : (x - y).abs = ∥x.abs - y.abs∥ := +lemma abs_sub_eq (h : x.arg = y.arg) : (x - y).abs = ‖x.abs - y.abs‖ := (same_ray_of_arg_eq h).norm_sub end complex diff --git a/src/analysis/complex/basic.lean b/src/analysis/complex/basic.lean index f65ab024a5a0e..2ae3893108310 100644 --- a/src/analysis/complex/basic.lean +++ b/src/analysis/complex/basic.lean @@ -36,7 +36,7 @@ open_locale complex_conjugate topological_space instance : has_norm ℂ := ⟨abs⟩ -@[simp] lemma norm_eq_abs (z : ℂ) : ∥z∥ = abs z := rfl +@[simp] lemma norm_eq_abs (z : ℂ) : ‖z‖ = abs z := rfl instance : normed_add_comm_group ℂ := add_group_norm.to_normed_add_comm_group @@ -53,8 +53,8 @@ instance : normed_field ℂ := instance : densely_normed_field ℂ := { lt_norm_lt := λ r₁ r₂ h₀ hr, let ⟨x, h⟩ := normed_field.exists_lt_norm_lt ℝ h₀ hr in - have this : ∥(∥x∥ : ℂ)∥ = ∥(∥x∥)∥, by simp only [norm_eq_abs, abs_of_real, real.norm_eq_abs], - ⟨∥x∥, by rwa [this, norm_norm]⟩ } + have this : ‖(‖x‖ : ℂ)‖ = ‖(‖x‖)‖, by simp only [norm_eq_abs, abs_of_real, real.norm_eq_abs], + ⟨‖x‖, by rwa [this, norm_norm]⟩ } instance {R : Type*} [normed_field R] [normed_algebra R ℝ] : normed_algebra R ℂ := { norm_smul_le := λ r x, begin @@ -113,17 +113,17 @@ by rw [nndist_comm, nndist_conj_self] @[simp] lemma comap_abs_nhds_zero : filter.comap abs (𝓝 0) = 𝓝 0 := comap_norm_nhds_zero -lemma norm_real (r : ℝ) : ∥(r : ℂ)∥ = ∥r∥ := abs_of_real _ +lemma norm_real (r : ℝ) : ‖(r : ℂ)‖ = ‖r‖ := abs_of_real _ -@[simp] lemma norm_rat (r : ℚ) : ∥(r : ℂ)∥ = |(r : ℝ)| := +@[simp] lemma norm_rat (r : ℚ) : ‖(r : ℂ)‖ = |(r : ℝ)| := by { rw ← of_real_rat_cast, exact norm_real _ } -@[simp] lemma norm_nat (n : ℕ) : ∥(n : ℂ)∥ = n := abs_of_nat _ +@[simp] lemma norm_nat (n : ℕ) : ‖(n : ℂ)‖ = n := abs_of_nat _ -@[simp] lemma norm_int {n : ℤ} : ∥(n : ℂ)∥ = |n| := +@[simp] lemma norm_int {n : ℤ} : ‖(n : ℂ)‖ = |n| := by simp [← rat.cast_coe_int] {single_pass := tt} -lemma norm_int_of_nonneg {n : ℤ} (hn : 0 ≤ n) : ∥(n : ℂ)∥ = n := +lemma norm_int_of_nonneg {n : ℤ} (hn : 0 ≤ n) : ‖(n : ℂ)‖ = n := by simp [hn] @[continuity] lemma continuous_abs : continuous abs := continuous_norm @@ -131,24 +131,24 @@ by simp [hn] @[continuity] lemma continuous_norm_sq : continuous norm_sq := by simpa [← norm_sq_eq_abs] using continuous_abs.pow 2 -@[simp, norm_cast] lemma nnnorm_real (r : ℝ) : ∥(r : ℂ)∥₊ = ∥r∥₊ := +@[simp, norm_cast] lemma nnnorm_real (r : ℝ) : ‖(r : ℂ)‖₊ = ‖r‖₊ := subtype.ext $ norm_real r -@[simp, norm_cast] lemma nnnorm_nat (n : ℕ) : ∥(n : ℂ)∥₊ = n := +@[simp, norm_cast] lemma nnnorm_nat (n : ℕ) : ‖(n : ℂ)‖₊ = n := subtype.ext $ by simp -@[simp, norm_cast] lemma nnnorm_int (n : ℤ) : ∥(n : ℂ)∥₊ = ∥n∥₊ := +@[simp, norm_cast] lemma nnnorm_int (n : ℤ) : ‖(n : ℂ)‖₊ = ‖n‖₊ := subtype.ext $ by simp only [coe_nnnorm, norm_int, int.norm_eq_abs] lemma nnnorm_eq_one_of_pow_eq_one {ζ : ℂ} {n : ℕ} (h : ζ ^ n = 1) (hn : n ≠ 0) : - ∥ζ∥₊ = 1 := + ‖ζ‖₊ = 1 := begin refine (@pow_left_inj nnreal _ _ _ _ zero_le' zero_le' hn.bot_lt).mp _, rw [←nnnorm_pow, h, nnnorm_one, one_pow], end lemma norm_eq_one_of_pow_eq_one {ζ : ℂ} {n : ℕ} (h : ζ ^ n = 1) (hn : n ≠ 0) : - ∥ζ∥ = 1 := + ‖ζ‖ = 1 := congr_arg coe (nnnorm_eq_one_of_pow_eq_one h hn) /-- The `abs` function on `ℂ` is proper. -/ @@ -172,12 +172,12 @@ def re_clm : ℂ →L[ℝ] ℝ := re_lm.mk_continuous 1 (λ x, by simp [abs_re_l @[simp] lemma re_clm_apply (z : ℂ) : (re_clm : ℂ → ℝ) z = z.re := rfl -@[simp] lemma re_clm_norm : ∥re_clm∥ = 1 := +@[simp] lemma re_clm_norm : ‖re_clm‖ = 1 := le_antisymm (linear_map.mk_continuous_norm_le _ zero_le_one _) $ -calc 1 = ∥re_clm 1∥ : by simp - ... ≤ ∥re_clm∥ : unit_le_op_norm _ _ (by simp) +calc 1 = ‖re_clm 1‖ : by simp + ... ≤ ‖re_clm‖ : unit_le_op_norm _ _ (by simp) -@[simp] lemma re_clm_nnnorm : ∥re_clm∥₊ = 1 := subtype.ext re_clm_norm +@[simp] lemma re_clm_nnnorm : ‖re_clm‖₊ = 1 := subtype.ext re_clm_norm /-- Continuous linear map version of the real part function, from `ℂ` to `ℝ`. -/ def im_clm : ℂ →L[ℝ] ℝ := im_lm.mk_continuous 1 (λ x, by simp [abs_im_le_abs]) @@ -188,12 +188,12 @@ def im_clm : ℂ →L[ℝ] ℝ := im_lm.mk_continuous 1 (λ x, by simp [abs_im_l @[simp] lemma im_clm_apply (z : ℂ) : (im_clm : ℂ → ℝ) z = z.im := rfl -@[simp] lemma im_clm_norm : ∥im_clm∥ = 1 := +@[simp] lemma im_clm_norm : ‖im_clm‖ = 1 := le_antisymm (linear_map.mk_continuous_norm_le _ zero_le_one _) $ -calc 1 = ∥im_clm I∥ : by simp - ... ≤ ∥im_clm∥ : unit_le_op_norm _ _ (by simp) +calc 1 = ‖im_clm I‖ : by simp + ... ≤ ‖im_clm‖ : unit_le_op_norm _ _ (by simp) -@[simp] lemma im_clm_nnnorm : ∥im_clm∥₊ = 1 := subtype.ext im_clm_norm +@[simp] lemma im_clm_nnnorm : ‖im_clm‖₊ = 1 := subtype.ext im_clm_norm lemma restrict_scalars_one_smul_right' (x : E) : continuous_linear_map.restrict_scalars ℝ ((1 : ℂ →L[ℂ] ℂ).smul_right x : ℂ →L[ℂ] E) = @@ -253,10 +253,10 @@ def conj_cle : ℂ ≃L[ℝ] ℂ := conj_lie @[simp] lemma conj_cle_apply (z : ℂ) : conj_cle z = conj z := rfl -@[simp] lemma conj_cle_norm : ∥(conj_cle : ℂ →L[ℝ] ℂ)∥ = 1 := +@[simp] lemma conj_cle_norm : ‖(conj_cle : ℂ →L[ℝ] ℂ)‖ = 1 := conj_lie.to_linear_isometry.norm_to_continuous_linear_map -@[simp] lemma conj_cle_nnorm : ∥(conj_cle : ℂ →L[ℝ] ℂ)∥₊ = 1 := subtype.ext conj_cle_norm +@[simp] lemma conj_cle_nnorm : ‖(conj_cle : ℂ →L[ℝ] ℂ)‖₊ = 1 := subtype.ext conj_cle_norm /-- Linear isometry version of the canonical embedding of `ℝ` in `ℂ`. -/ def of_real_li : ℝ →ₗᵢ[ℝ] ℂ := ⟨of_real_am.to_linear_map, norm_real⟩ @@ -281,9 +281,9 @@ def of_real_clm : ℝ →L[ℝ] ℂ := of_real_li.to_continuous_linear_map @[simp] lemma of_real_clm_apply (x : ℝ) : of_real_clm x = x := rfl -@[simp] lemma of_real_clm_norm : ∥of_real_clm∥ = 1 := of_real_li.norm_to_continuous_linear_map +@[simp] lemma of_real_clm_norm : ‖of_real_clm‖ = 1 := of_real_li.norm_to_continuous_linear_map -@[simp] lemma of_real_clm_nnnorm : ∥of_real_clm∥₊ = 1 := subtype.ext $ of_real_clm_norm +@[simp] lemma of_real_clm_nnnorm : ‖of_real_clm‖₊ = 1 := subtype.ext $ of_real_clm_norm noncomputable instance : is_R_or_C ℂ := { re := ⟨complex.re, complex.zero_re, complex.add_re⟩, @@ -315,7 +315,7 @@ lemma _root_.is_R_or_C.im_eq_complex_im : ⇑(is_R_or_C.im : ℂ →+ ℝ) = com section complex_order open_locale complex_order -lemma eq_coe_norm_of_nonneg {z : ℂ} (hz : 0 ≤ z) : z = ↑∥z∥ := +lemma eq_coe_norm_of_nonneg {z : ℂ} (hz : 0 ≤ z) : z = ↑‖z‖ := by rw [eq_re_of_real_le hz, is_R_or_C.norm_of_real, real.norm_of_nonneg (complex.le_def.2 hz).1] end complex_order diff --git a/src/analysis/complex/cauchy_integral.lean b/src/analysis/complex/cauchy_integral.lean index 2f7410cca4c6b..18fe7a8917912 100644 --- a/src/analysis/complex/cauchy_integral.lean +++ b/src/analysis/complex/cauchy_integral.lean @@ -271,9 +271,9 @@ integral_boundary_rect_eq_zero_of_continuous_on_of_differentiable_on f z w H.con H.mono $ inter_subset_inter (preimage_mono Ioo_subset_Icc_self) (preimage_mono Ioo_subset_Icc_self) -/-- If `f : ℂ → E` is continuous the closed annulus `r ≤ ∥z - c∥ ≤ R`, `0 < r ≤ R`, and is complex +/-- If `f : ℂ → E` is continuous the closed annulus `r ≤ ‖z - c‖ ≤ R`, `0 < r ≤ R`, and is complex differentiable at all but countably many points of its interior, then the integrals of -`f z / (z - c)` (formally, `(z - c)⁻¹ • f z`) over the circles `∥z - c∥ = r` and `∥z - c∥ = R` are +`f z / (z - c)` (formally, `(z - c)⁻¹ • f z`) over the circles `‖z - c‖ = r` and `‖z - c‖ = R` are equal to each other. -/ lemma circle_integral_sub_center_inv_smul_eq_of_differentiable_on_annulus_off_countable {c : ℂ} {r R : ℝ} (h0 : 0 < r) (hle : r ≤ R) {f : ℂ → E} {s : set ℂ} (hs : s.countable) @@ -310,8 +310,8 @@ begin end /-- **Cauchy-Goursat theorem** for an annulus. If `f : ℂ → E` is continuous on the closed annulus -`r ≤ ∥z - c∥ ≤ R`, `0 < r ≤ R`, and is complex differentiable at all but countably many points of -its interior, then the integrals of `f` over the circles `∥z - c∥ = r` and `∥z - c∥ = R` are equal +`r ≤ ‖z - c‖ ≤ R`, `0 < r ≤ R`, and is complex differentiable at all but countably many points of +its interior, then the integrals of `f` over the circles `‖z - c‖ = r` and `‖z - c‖ = R` are equal to each other. -/ lemma circle_integral_eq_of_differentiable_on_annulus_off_countable {c : ℂ} {r R : ℝ} (h0 : 0 < r) (hle : r ≤ R) {f : ℂ → E} {s : set ℂ} (hs : s.countable) @@ -329,7 +329,7 @@ calc ∮ z in C(c, R), f z = ∮ z in C(c, R), (z - c)⁻¹ • (z - c) • f z /-- **Cauchy integral formula** for the value at the center of a disc. If `f` is continuous on a punctured closed disc of radius `R`, is differentiable at all but countably many points of the interior of this disc, and has a limit `y` at the center of the disc, then the integral -$\oint_{∥z-c∥=R} \frac{f(z)}{z-c}\,dz$ is equal to $2πiy`. -/ +$\oint_{‖z-c‖=R} \frac{f(z)}{z-c}\,dz$ is equal to $2πiy`. -/ lemma circle_integral_sub_center_inv_smul_of_differentiable_on_off_countable_of_tendsto {c : ℂ} {R : ℝ} (h0 : 0 < R) {f : ℂ → E} {y : E} {s : set ℂ} (hs : s.countable) (hc : continuous_on f (closed_ball c R \ {c})) @@ -352,15 +352,15 @@ begin from λ z hz, ne_of_mem_of_not_mem hz (λ h, hr0.ne' $ dist_self c ▸ eq.symm h), /- The integral `∮ z in C(c, r), f z / (z - c)` does not depend on `0 < r ≤ R` and tends to `2πIy` as `r → 0`. -/ - calc ∥(∮ z in C(c, R), (z - c)⁻¹ • f z) - (2 * ↑π * I) • y∥ - = ∥(∮ z in C(c, r), (z - c)⁻¹ • f z) - ∮ z in C(c, r), (z - c)⁻¹ • y∥ : + calc ‖(∮ z in C(c, R), (z - c)⁻¹ • f z) - (2 * ↑π * I) • y‖ + = ‖(∮ z in C(c, r), (z - c)⁻¹ • f z) - ∮ z in C(c, r), (z - c)⁻¹ • y‖ : begin congr' 2, { exact circle_integral_sub_center_inv_smul_eq_of_differentiable_on_annulus_off_countable hr0 hrR hs (hc.mono hsub) (λ z hz, hd z ⟨hsub' hz.1, hz.2⟩) }, { simp [hr0.ne'] } end - ... = ∥∮ z in C(c, r), (z - c)⁻¹ • (f z - y)∥ : + ... = ‖∮ z in C(c, r), (z - c)⁻¹ • (f z - y)‖ : begin simp only [smul_sub], have hc' : continuous_on (λ z, (z - c)⁻¹) (sphere c r), @@ -394,7 +394,7 @@ circle_integral_sub_center_inv_smul_of_differentiable_on_off_countable_of_tendst (hc.continuous_at $ closed_ball_mem_nhds _ h0).continuous_within_at /-- **Cauchy-Goursat theorem** for a disk: if `f : ℂ → E` is continuous on a closed disk -`{z | ∥z - c∥ ≤ R}` and is complex differentiable at all but countably many points of its interior, +`{z | ‖z - c‖ ≤ R}` and is complex differentiable at all but countably many points of its interior, then the integral $\oint_{|z-c|=R}f(z)\,dz$ equals zero. -/ lemma circle_integral_eq_zero_of_differentiable_on_off_countable {R : ℝ} (h0 : 0 ≤ R) {f : ℂ → E} {c : ℂ} {s : set ℂ} (hs : s.countable) (hc : continuous_on f (closed_ball c R)) diff --git a/src/analysis/complex/conformal.lean b/src/analysis/complex/conformal.lean index bc434f9a0eb91..659ede8df10b2 100644 --- a/src/analysis/complex/conformal.lean +++ b/src/analysis/complex/conformal.lean @@ -45,9 +45,9 @@ variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [normed_spa lemma is_conformal_map_complex_linear {map : ℂ →L[ℂ] E} (nonzero : map ≠ 0) : is_conformal_map (map.restrict_scalars ℝ) := begin - have minor₁ : ∥map 1∥ ≠ 0, + have minor₁ : ‖map 1‖ ≠ 0, { simpa [ext_ring_iff] using nonzero }, - refine ⟨∥map 1∥, minor₁, ⟨∥map 1∥⁻¹ • map, _⟩, _⟩, + refine ⟨‖map 1‖, minor₁, ⟨‖map 1‖⁻¹ • map, _⟩, _⟩, { intros x, simp only [linear_map.smul_apply], have : x = x • 1 := by rw [smul_eq_mul, mul_one], diff --git a/src/analysis/complex/isometry.lean b/src/analysis/complex/isometry.lean index e153b1c16d6dd..7718c6c229a3b 100644 --- a/src/analysis/complex/isometry.lean +++ b/src/analysis/complex/isometry.lean @@ -90,7 +90,7 @@ end lemma linear_isometry.im_apply_eq_im {f : ℂ →ₗᵢ[ℝ] ℂ} (h : f 1 = 1) (z : ℂ) : z + conj z = f z + conj (f z) := begin - have : ∥f z - 1∥ = ∥z - 1∥ := by rw [← f.norm_map (z - 1), f.map_sub, h], + have : ‖f z - 1‖ = ‖z - 1‖ := by rw [← f.norm_map (z - 1), f.map_sub, h], apply_fun λ x, x ^ 2 at this, simp only [norm_eq_abs, ←norm_sq_eq_abs] at this, rw [←of_real_inj, ←mul_conj, ←mul_conj] at this, diff --git a/src/analysis/complex/liouville.lean b/src/analysis/complex/liouville.lean index 133d8fa3ea95b..95b25ddbf2e5a 100644 --- a/src/analysis/complex/liouville.lean +++ b/src/analysis/complex/liouville.lean @@ -48,13 +48,13 @@ begin end lemma norm_deriv_le_aux [complete_space F] {c : ℂ} {R C : ℝ} {f : ℂ → F} (hR : 0 < R) - (hf : diff_cont_on_cl ℂ f (ball c R)) (hC : ∀ z ∈ sphere c R, ∥f z∥ ≤ C) : - ∥deriv f c∥ ≤ C / R := + (hf : diff_cont_on_cl ℂ f (ball c R)) (hC : ∀ z ∈ sphere c R, ‖f z‖ ≤ C) : + ‖deriv f c‖ ≤ C / R := begin - have : ∀ z ∈ sphere c R, ∥(z - c) ^ (-2 : ℤ) • f z∥ ≤ C / (R * R), + have : ∀ z ∈ sphere c R, ‖(z - c) ^ (-2 : ℤ) • f z‖ ≤ C / (R * R), from λ z (hz : abs (z - c) = R), by simpa [-mul_inv_rev, norm_smul, hz, zpow_two, ←div_eq_inv_mul] using (div_le_div_right (mul_pos hR hR)).2 (hC z hz), - calc ∥deriv f c∥ = ∥(2 * π * I : ℂ)⁻¹ • ∮ z in C(c, R), (z - c) ^ (-2 : ℤ) • f z∥ : + calc ‖deriv f c‖ = ‖(2 * π * I : ℂ)⁻¹ • ∮ z in C(c, R), (z - c) ^ (-2 : ℤ) • f z‖ : congr_arg norm (deriv_eq_smul_circle_integral hR hf) ... ≤ R * (C / (R * R)) : circle_integral.norm_two_pi_I_inv_smul_integral_le_of_norm_le_const hR.le this @@ -65,14 +65,14 @@ end closure, and its values on the boundary circle of this disc are bounded from above by `C`, then the norm of its derivative at the center is at most `C / R`. -/ lemma norm_deriv_le_of_forall_mem_sphere_norm_le {c : ℂ} {R C : ℝ} {f : ℂ → F} (hR : 0 < R) - (hd : diff_cont_on_cl ℂ f (ball c R)) (hC : ∀ z ∈ sphere c R, ∥f z∥ ≤ C) : - ∥deriv f c∥ ≤ C / R := + (hd : diff_cont_on_cl ℂ f (ball c R)) (hC : ∀ z ∈ sphere c R, ‖f z‖ ≤ C) : + ‖deriv f c‖ ≤ C / R := begin set e : F →L[ℂ] F̂ := uniform_space.completion.to_complL, have : has_deriv_at (e ∘ f) (e (deriv f c)) c, from e.has_fderiv_at.comp_has_deriv_at c (hd.differentiable_at is_open_ball $ mem_ball_self hR).has_deriv_at, - calc ∥deriv f c∥ = ∥deriv (e ∘ f) c∥ : + calc ‖deriv f c‖ = ‖deriv (e ∘ f) c‖ : by { rw this.deriv, exact (uniform_space.completion.norm_coe _).symm } ... ≤ C / R : norm_deriv_le_aux hR (e.differentiable.comp_diff_cont_on_cl hd) @@ -85,12 +85,12 @@ lemma liouville_theorem_aux {f : ℂ → F} (hf : differentiable ℂ f) begin suffices : ∀ c, deriv f c = 0, from is_const_of_deriv_eq_zero hf this z w, clear z w, intro c, - obtain ⟨C, C₀, hC⟩ : ∃ C > (0 : ℝ), ∀ z, ∥f z∥ ≤ C, + obtain ⟨C, C₀, hC⟩ : ∃ C > (0 : ℝ), ∀ z, ‖f z‖ ≤ C, { rcases bounded_iff_forall_norm_le.1 hb with ⟨C, hC⟩, exact ⟨max C 1, lt_max_iff.2 (or.inr zero_lt_one), λ z, (hC (f z) (mem_range_self _)).trans (le_max_left _ _)⟩ }, refine norm_le_zero_iff.1 (le_of_forall_le_of_dense $ λ ε ε₀, _), - calc ∥deriv f c∥ ≤ C / (C / ε) : + calc ‖deriv f c‖ ≤ C / (C / ε) : norm_deriv_le_of_forall_mem_sphere_norm_le (div_pos C₀ ε₀) hf.diff_cont_on_cl (λ z _, hC z) ... = ε : div_div_cancel' C₀.lt.ne' end diff --git a/src/analysis/complex/open_mapping.lean b/src/analysis/complex/open_mapping.lean index f4648b5428e31..2c373492e781f 100644 --- a/src/analysis/complex/open_mapping.lean +++ b/src/analysis/complex/open_mapping.lean @@ -17,9 +17,9 @@ it is constant in a neighborhood of `z₀` or it maps any neighborhood of `z₀` its image `f z₀`. The results extend in higher dimension to `g : E → ℂ`. The proof of the local version on `ℂ` goes through two main steps: first, assuming that the function -is not constant around `z₀`, use the isolated zero principle to show that `∥f z∥` is bounded below +is not constant around `z₀`, use the isolated zero principle to show that `‖f z‖` is bounded below on a small `sphere z₀ r` around `z₀`, and then use the maximum principle applied to the auxiliary -function `(λ z, ∥f z - v∥)` to show that any `v` close enough to `f z₀` is in `f '' ball z₀ r`. That +function `(λ z, ‖f z - v‖)` to show that any `v` close enough to `f z₀` is in `f '' ball z₀ r`. That second step is implemented in `diff_cont_on_cl.ball_subset_image_closed_ball`. ## Main results @@ -38,23 +38,23 @@ variables {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] {U : set E} /-- If the modulus of a holomorphic function `f` is bounded below by `ε` on a circle, then its range contains a disk of radius `ε / 2`. -/ lemma diff_cont_on_cl.ball_subset_image_closed_ball (h : diff_cont_on_cl ℂ f (ball z₀ r)) - (hr : 0 < r) (hf : ∀ z ∈ sphere z₀ r, ε ≤ ∥f z - f z₀∥) (hz₀ : ∃ᶠ z in 𝓝 z₀, f z ≠ f z₀) : + (hr : 0 < r) (hf : ∀ z ∈ sphere z₀ r, ε ≤ ‖f z - f z₀‖) (hz₀ : ∃ᶠ z in 𝓝 z₀, f z ≠ f z₀) : ball (f z₀) (ε / 2) ⊆ f '' closed_ball z₀ r := begin /- This is a direct application of the maximum principle. Pick `v` close to `f z₀`, and look at - the function `λ z, ∥f z - v∥`: it is bounded below on the circle, and takes a small value at `z₀` + the function `λ z, ‖f z - v‖`: it is bounded below on the circle, and takes a small value at `z₀` so it is not constant on the disk, which implies that its infimum is equal to `0` and hence that `v` is in the range of `f`. -/ rintro v hv, have h1 : diff_cont_on_cl ℂ (λ z, f z - v) (ball z₀ r) := h.sub_const v, - have h2 : continuous_on (λ z, ∥f z - v∥) (closed_ball z₀ r), + have h2 : continuous_on (λ z, ‖f z - v‖) (closed_ball z₀ r), from continuous_norm.comp_continuous_on (closure_ball z₀ hr.ne.symm ▸ h1.continuous_on), have h3 : analytic_on ℂ f (ball z₀ r) := h.differentiable_on.analytic_on is_open_ball, - have h4 : ∀ z ∈ sphere z₀ r, ε / 2 ≤ ∥f z - v∥, - from λ z hz, by linarith [hf z hz, (show ∥v - f z₀∥ < ε / 2, from mem_ball.mp hv), + have h4 : ∀ z ∈ sphere z₀ r, ε / 2 ≤ ‖f z - v‖, + from λ z hz, by linarith [hf z hz, (show ‖v - f z₀‖ < ε / 2, from mem_ball.mp hv), norm_sub_sub_norm_sub_le_norm_sub (f z) v (f z₀)], - have h5 : ∥f z₀ - v∥ < ε / 2 := by simpa [← dist_eq_norm, dist_comm] using mem_ball.mp hv, - obtain ⟨z, hz1, hz2⟩ : ∃ z ∈ ball z₀ r, is_local_min (λ z, ∥f z - v∥) z, + have h5 : ‖f z₀ - v‖ < ε / 2 := by simpa [← dist_eq_norm, dist_comm] using mem_ball.mp hv, + obtain ⟨z, hz1, hz2⟩ : ∃ z ∈ ball z₀ r, is_local_min (λ z, ‖f z - v‖) z, from exists_local_min_mem_ball h2 (mem_closed_ball_self hr.le) (λ z hz, h5.trans_le (h4 z hz)), refine ⟨z, ball_subset_closed_ball hz1, sub_eq_zero.mp _⟩, have h6 := h1.differentiable_on.eventually_differentiable_at (is_open_ball.mem_nhds hz1), @@ -96,11 +96,11 @@ begin have h7 : ∀ z ∈ sphere z₀ r, f z ≠ f z₀, from λ z hz, h4 z (h5 (sphere_subset_closed_ball hz)) (ne_of_mem_sphere hz hr.ne.symm), have h8 : (sphere z₀ r).nonempty := normed_space.sphere_nonempty.mpr hr.le, - have h9 : continuous_on (λ x, ∥f x - f z₀∥) (sphere z₀ r), + have h9 : continuous_on (λ x, ‖f x - f z₀‖) (sphere z₀ r), from continuous_norm.comp_continuous_on ((h6.sub_const (f z₀)).continuous_on_ball.mono sphere_subset_closed_ball), obtain ⟨x, hx, hfx⟩ := (is_compact_sphere z₀ r).exists_forall_le h8 h9, - refine ⟨∥f x - f z₀∥ / 2, half_pos (norm_sub_pos_iff.mpr (h7 x hx)), _⟩, + refine ⟨‖f x - f z₀‖ / 2, half_pos (norm_sub_pos_iff.mpr (h7 x hx)), _⟩, exact (h6.ball_subset_image_closed_ball hr (λ z hz, hfx z hz) (not_eventually.mp h)).trans (image_subset f (closed_ball_subset_closed_ball inf_le_right)) end @@ -130,19 +130,19 @@ begin { left, -- If g is eventually constant along every direction, then it is eventually constant refine eventually_of_mem (ball_mem_nhds z₀ hr) (λ z hz, _), refine (eq_or_ne z z₀).cases_on (congr_arg g) (λ h', _), - replace h' : ∥z - z₀∥ ≠ 0 := by simpa only [ne.def, norm_eq_zero, sub_eq_zero], - let w : E := ∥z - z₀∥⁻¹ • (z - z₀), + replace h' : ‖z - z₀‖ ≠ 0 := by simpa only [ne.def, norm_eq_zero, sub_eq_zero], + let w : E := ‖z - z₀‖⁻¹ • (z - z₀), have h3 : ∀ t ∈ ball (0 : ℂ) r, gray w t = g z₀, { have e1 : is_preconnected (ball (0 : ℂ) r) := (convex_ball 0 r).is_preconnected, have e2 : w ∈ sphere (0 : E) 1 := by simp [w, norm_smul, h'], specialize h1 w e2, apply h1.eq_on_of_preconnected_of_eventually_eq analytic_on_const e1 (mem_ball_self hr), simpa [gray, ray] using h w e2 }, - have h4 : ∥z - z₀∥ < r := by simpa [dist_eq_norm] using mem_ball.mp hz, - replace h4 : ↑∥z - z₀∥ ∈ ball (0 : ℂ) r := by simpa only [mem_ball_zero_iff, norm_eq_abs, + have h4 : ‖z - z₀‖ < r := by simpa [dist_eq_norm] using mem_ball.mp hz, + replace h4 : ↑‖z - z₀‖ ∈ ball (0 : ℂ) r := by simpa only [mem_ball_zero_iff, norm_eq_abs, abs_of_real, abs_norm_eq_norm], simpa only [gray, ray, smul_smul, mul_inv_cancel h', one_smul, add_sub_cancel'_right, - function.comp_app, complex.coe_smul] using h3 ↑∥z - z₀∥ h4 }, + function.comp_app, complex.coe_smul] using h3 ↑‖z - z₀‖ h4 }, { right, -- Otherwise, it is open along at least one direction and that implies the result push_neg at h, obtain ⟨z, hz, hrz⟩ := h, diff --git a/src/analysis/complex/phragmen_lindelof.lean b/src/analysis/complex/phragmen_lindelof.lean index 984446957e8b2..b6b9903c6476a 100644 --- a/src/analysis/complex/phragmen_lindelof.lean +++ b/src/analysis/complex/phragmen_lindelof.lean @@ -63,7 +63,7 @@ lemma is_O_sub_exp_exp {a : ℝ} {f g : ℂ → E} {l : filter ℂ} {u : ℂ → ∃ (c < a) B, (f - g) =O[l] (λ z, expR (B * expR (c * |u z|))) := begin have : ∀ {c₁ c₂ B₁ B₂}, c₁ ≤ c₂ → 0 ≤ B₂ → B₁ ≤ B₂ → ∀ z, - ∥expR (B₁ * expR (c₁ * |u z|))∥ ≤ ∥expR (B₂ * expR (c₂ * |u z|))∥, + ‖expR (B₁ * expR (c₁ * |u z|))‖ ≤ ‖expR (B₂ * expR (c₂ * |u z|))‖, { intros c₁ c₂ B₁ B₂ hc hB₀ hB z, rw [real.norm_eq_abs, real.norm_eq_abs, real.abs_exp, real.abs_exp, real.exp_le_exp], exact mul_le_mul hB (real.exp_le_exp.2 $ mul_le_mul_of_nonneg_right hc $ abs_nonneg _) @@ -108,27 +108,27 @@ variables [normed_space ℂ E] {a b C : ℝ} {f g : ℂ → E} {z : ℂ} Let `f : ℂ → E` be a function such that * `f` is differentiable on `U` and is continuous on its closure; -* `∥f z∥` is bounded from above by `A * exp(B * exp(c * |re z|))` on `U` for some `c < π / (b - a)`; -* `∥f z∥` is bounded from above by a constant `C` on the boundary of `U`. +* `‖f z‖` is bounded from above by `A * exp(B * exp(c * |re z|))` on `U` for some `c < π / (b - a)`; +* `‖f z‖` is bounded from above by a constant `C` on the boundary of `U`. -Then `∥f z∥` is bounded by the same constant on the closed strip +Then `‖f z‖` is bounded by the same constant on the closed strip `{z : ℂ | a ≤ im z ≤ b}`. Moreover, it suffices to verify the second assumption only for sufficiently large values of `|re z|`. -/ lemma horizontal_strip (hfd : diff_cont_on_cl ℂ f (im ⁻¹' Ioo a b)) (hB : ∃ (c < π / (b - a)) B, f =O[comap (has_abs.abs ∘ re) at_top ⊓ 𝓟 (im ⁻¹' Ioo a b)] (λ z, expR (B * expR (c * |z.re|)))) - (hle_a : ∀ z : ℂ, im z = a → ∥f z∥ ≤ C) (hle_b : ∀ z, im z = b → ∥f z∥ ≤ C) + (hle_a : ∀ z : ℂ, im z = a → ‖f z‖ ≤ C) (hle_b : ∀ z, im z = b → ‖f z‖ ≤ C) (hza : a ≤ im z) (hzb : im z ≤ b) : - ∥f z∥ ≤ C := + ‖f z‖ ≤ C := begin -- If `im z = a` or `im z = b`, then we apply `hle_a` or `hle_b`, otherwise `im z ∈ Ioo a b`. rw le_iff_eq_or_lt at hza hzb, cases hza with hza hza, { exact hle_a _ hza.symm }, cases hzb with hzb hzb, { exact hle_b _ hzb }, -- WLOG, `0 < C`. - suffices : ∀ C' : ℝ, 0 < C' → (∀ w : ℂ, im w = a → ∥f w∥ ≤ C') → - (∀ w : ℂ, im w = b → ∥f w∥ ≤ C') → ∥f z∥ ≤ C', + suffices : ∀ C' : ℝ, 0 < C' → (∀ w : ℂ, im w = a → ‖f w‖ ≤ C') → + (∀ w : ℂ, im w = b → ‖f w‖ ≤ C') → ‖f z‖ ≤ C', { refine le_of_forall_le_of_dense (λ C' hC', this C' _ (λ w hw, _) (λ w hw, _)), { refine ((norm_nonneg (f (a * I))).trans (hle_a _ _)).trans_lt hC', rw [mul_I_im, of_real_re] }, @@ -150,14 +150,14 @@ begin have hb' : d * b < π / 2, from (lt_div_iff hb).1 hd, set aff : ℂ → ℂ := λ w, d * (w - a * I), set g : ℝ → ℂ → ℂ := λ ε w, exp (ε * (exp (aff w) + exp (-aff w))), - /- Since `g ε z → 1` as `ε → 0⁻`, it suffices to prove that `∥g ε z • f z∥ ≤ C` + /- Since `g ε z → 1` as `ε → 0⁻`, it suffices to prove that `‖g ε z • f z‖ ≤ C` for all negative `ε`. -/ - suffices : ∀ᶠ ε : ℝ in 𝓝[<] 0, ∥g ε z • f z∥ ≤ C, + suffices : ∀ᶠ ε : ℝ in 𝓝[<] 0, ‖g ε z • f z‖ ≤ C, { refine le_of_tendsto (tendsto.mono_left _ nhds_within_le_nhds) this, apply ((continuous_of_real.mul continuous_const).cexp.smul continuous_const).norm.tendsto', simp, apply_instance }, filter_upwards [self_mem_nhds_within] with ε ε₀, change ε < 0 at ε₀, - -- An upper estimate on `∥g ε w∥` that will be used in two branches of the proof. + -- An upper estimate on `‖g ε w‖` that will be used in two branches of the proof. obtain ⟨δ, δ₀, hδ⟩ : ∃ δ : ℝ, δ < 0 ∧ ∀ ⦃w⦄, im w ∈ Icc (a - b) (a + b) → abs (g ε w) ≤ expR (δ * expR (d * |re w|)), { refine ⟨ε * real.cos (d * b), mul_neg_of_neg_of_pos ε₀ (real.cos_pos_of_mem_Ioo $ abs_lt.1 $ @@ -177,14 +177,14 @@ begin /- Our apriori estimate on `f` implies that `g ε w • f w → 0` as `|w.re| → ∞` along the strip. In particular, its norm is less than or equal to `C` for sufficiently large `|w.re|`. -/ obtain ⟨R, hzR, hR⟩ : ∃ R : ℝ, |z.re| < R ∧ ∀ w, |re w| = R → im w ∈ Ioo (a - b) (a + b) → - ∥g ε w • f w∥ ≤ C, + ‖g ε w • f w‖ ≤ C, { refine ((eventually_gt_at_top _).and _).exists, rcases hO.exists_pos with ⟨A, hA₀, hA⟩, simp only [is_O_with_iff, eventually_inf_principal, eventually_comap, mem_Ioo, ← abs_lt, mem_preimage, (∘), real.norm_eq_abs, abs_of_pos (real.exp_pos _)] at hA, suffices : tendsto (λ R, expR (δ * expR (d * R) + B * expR (c * R) + real.log A)) at_top (𝓝 0), { filter_upwards [this.eventually (ge_mem_nhds hC₀), hA] with R hR Hle w hre him, - calc ∥g ε w • f w∥ ≤ expR (δ * expR (d * R) + B * expR (c * R) + real.log A) : _ + calc ‖g ε w • f w‖ ≤ expR (δ * expR (d * R) + B * expR (c * R) + real.log A) : _ ... ≤ C : hR, rw [norm_smul, real.exp_add, ← hre, real.exp_add, real.exp_log hA₀, mul_assoc, mul_comm _ A], exact mul_le_mul (hδ $ Ioo_subset_Icc_self him) (Hle _ hre him) (norm_nonneg _) @@ -203,7 +203,7 @@ begin have hR₀ : 0 < R, from (_root_.abs_nonneg _).trans_lt hzR, /- Finally, we apply the bounded version of the maximum modulus principle to the rectangle `(-R, R) × (a - b, a + b)`. The function is bounded by `C` on the horizontal sides by assumption - (and because `∥g ε w∥ ≤ 1`) and on the vertical sides by the choice of `R`. -/ + (and because `‖g ε w‖ ≤ 1`) and on the vertical sides by the choice of `R`. -/ have hgd : differentiable ℂ (g ε), from ((((differentiable_id.sub_const _).const_mul _).cexp.add ((differentiable_id.sub_const _).const_mul _).neg.cexp).const_mul _).cexp, @@ -229,7 +229,7 @@ end Let `f : ℂ → E` be a function such that * `f` is differentiable on `U` and is continuous on its closure; -* `∥f z∥` is bounded from above by `A * exp(B * exp(c * |re z|))` on `U` for some `c < π / (b - a)`; +* `‖f z‖` is bounded from above by `A * exp(B * exp(c * |re z|))` on `U` for some `c < π / (b - a)`; * `f z = 0` on the boundary of `U`. Then `f` is equal to zero on the closed strip `{z : ℂ | a ≤ im z ≤ b}`. @@ -246,7 +246,7 @@ lemma eq_zero_on_horizontal_strip (hd : diff_cont_on_cl ℂ f (im ⁻¹' Ioo a b Let `f g : ℂ → E` be functions such that * `f` and `g` are differentiable on `U` and are continuous on its closure; -* `∥f z∥` and `∥g z∥` are bounded from above by `A * exp(B * exp(c * |re z|))` on `U` for some +* `‖f z‖` and `‖g z‖` are bounded from above by `A * exp(B * exp(c * |re z|))` on `U` for some `c < π / (b - a)`; * `f z = g z` on the boundary of `U`. @@ -271,21 +271,21 @@ lemma eq_on_horizontal_strip {g : ℂ → E} (hdf : diff_cont_on_cl ℂ f (im Let `f : ℂ → E` be a function such that * `f` is differentiable on `U` and is continuous on its closure; -* `∥f z∥` is bounded from above by `A * exp(B * exp(c * |im z|))` on `U` for some `c < π / (b - a)`; -* `∥f z∥` is bounded from above by a constant `C` on the boundary of `U`. +* `‖f z‖` is bounded from above by `A * exp(B * exp(c * |im z|))` on `U` for some `c < π / (b - a)`; +* `‖f z‖` is bounded from above by a constant `C` on the boundary of `U`. -Then `∥f z∥` is bounded by the same constant on the closed strip +Then `‖f z‖` is bounded by the same constant on the closed strip `{z : ℂ | a ≤ re z ≤ b}`. Moreover, it suffices to verify the second assumption only for sufficiently large values of `|im z|`. -/ lemma vertical_strip (hfd : diff_cont_on_cl ℂ f (re ⁻¹' Ioo a b)) (hB : ∃ (c < π / (b - a)) B, f =O[comap (has_abs.abs ∘ im) at_top ⊓ 𝓟 (re ⁻¹' Ioo a b)] (λ z, expR (B * expR (c * |z.im|)))) - (hle_a : ∀ z : ℂ, re z = a → ∥f z∥ ≤ C) (hle_b : ∀ z, re z = b → ∥f z∥ ≤ C) + (hle_a : ∀ z : ℂ, re z = a → ‖f z‖ ≤ C) (hle_b : ∀ z, re z = b → ‖f z‖ ≤ C) (hza : a ≤ re z) (hzb : re z ≤ b) : - ∥f z∥ ≤ C := + ‖f z‖ ≤ C := begin - suffices : ∥(λ z, f (z * (-I))) (z * I)∥ ≤ C, by simpa [mul_assoc] using this, + suffices : ‖(λ z, f (z * (-I))) (z * I)‖ ≤ C, by simpa [mul_assoc] using this, have H : maps_to (λ z, z * (-I)) (im ⁻¹' Ioo a b) (re ⁻¹' Ioo a b), { intros z hz, simpa using hz }, refine horizontal_strip (hfd.comp (differentiable_id.mul_const _).diff_cont_on_cl H) @@ -303,7 +303,7 @@ end Let `f : ℂ → E` be a function such that * `f` is differentiable on `U` and is continuous on its closure; -* `∥f z∥` is bounded from above by `A * exp(B * exp(c * |im z|))` on `U` for some `c < π / (b - a)`; +* `‖f z‖` is bounded from above by `A * exp(B * exp(c * |im z|))` on `U` for some `c < π / (b - a)`; * `f z = 0` on the boundary of `U`. Then `f` is equal to zero on the closed strip `{z : ℂ | a ≤ re z ≤ b}`. @@ -320,7 +320,7 @@ lemma eq_zero_on_vertical_strip (hd : diff_cont_on_cl ℂ f (re ⁻¹' Ioo a b)) Let `f g : ℂ → E` be functions such that * `f` and `g` are differentiable on `U` and are continuous on its closure; -* `∥f z∥` and `∥g z∥` are bounded from above by `A * exp(B * exp(c * |im z|))` on `U` for some +* `‖f z‖` and `‖g z‖` are bounded from above by `A * exp(B * exp(c * |im z|))` on `U` for some `c < π / (b - a)`; * `f z = g z` on the boundary of `U`. @@ -344,17 +344,17 @@ lemma eq_on_vertical_strip {g : ℂ → E} (hdf : diff_cont_on_cl ℂ f (re ⁻ /-- **Phragmen-Lindelöf principle** in the first quadrant. Let `f : ℂ → E` be a function such that * `f` is differentiable in the open first quadrant and is continuous on its closure; -* `∥f z∥` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open first quadrant +* `‖f z‖` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open first quadrant for some `c < 2`; -* `∥f z∥` is bounded from above by a constant `C` on the boundary of the first quadrant. +* `‖f z‖` is bounded from above by a constant `C` on the boundary of the first quadrant. -Then `∥f z∥` is bounded from above by the same constant on the closed first quadrant. -/ +Then `‖f z‖` is bounded from above by the same constant on the closed first quadrant. -/ lemma quadrant_I (hd : diff_cont_on_cl ℂ f (Ioi 0 ×ℂ Ioi 0)) (hB : ∃ (c < (2 : ℝ)) B, f =O[comap complex.abs at_top ⊓ 𝓟 (Ioi 0 ×ℂ Ioi 0)] (λ z, expR (B * (abs z) ^ c))) - (hre : ∀ x : ℝ, 0 ≤ x → ∥f x∥ ≤ C) (him : ∀ x : ℝ, 0 ≤ x → ∥f (x * I)∥ ≤ C) + (hre : ∀ x : ℝ, 0 ≤ x → ‖f x‖ ≤ C) (him : ∀ x : ℝ, 0 ≤ x → ‖f (x * I)‖ ≤ C) (hz_re : 0 ≤ z.re) (hz_im : 0 ≤ z.im) : - ∥f z∥ ≤ C := + ‖f z‖ ≤ C := begin -- The case `z = 0` is trivial. rcases eq_or_ne z 0 with rfl|hzne, { exact hre 0 le_rfl }, @@ -365,7 +365,7 @@ begin exact ⟨arg_nonneg_iff.2 hz_im, arg_le_pi_div_two_iff.2 (or.inl hz_re)⟩ }, clear hz_re hz_im hzne, -- We are going to apply `phragmen_lindelof.horizontal_strip` to `f ∘ complex.exp` and `ζ`. - change ∥(f ∘ exp) ζ∥ ≤ C, + change ‖(f ∘ exp) ζ‖ ≤ C, have H : maps_to exp (im ⁻¹' Ioo 0 (π / 2)) (Ioi 0 ×ℂ Ioi 0), { intros z hz, rw [mem_re_prod_im, exp_re, exp_im, mem_Ioi, mem_Ioi], @@ -414,7 +414,7 @@ end /-- **Phragmen-Lindelöf principle** in the first quadrant. Let `f : ℂ → E` be a function such that * `f` is differentiable in the open first quadrant and is continuous on its closure; -* `∥f z∥` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open first quadrant +* `‖f z‖` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open first quadrant for some `A`, `B`, and `c < 2`; * `f` is equal to zero on the boundary of the first quadrant. @@ -430,7 +430,7 @@ lemma eq_zero_on_quadrant_I (hd : diff_cont_on_cl ℂ f (Ioi 0 ×ℂ Ioi 0)) /-- **Phragmen-Lindelöf principle** in the first quadrant. Let `f g : ℂ → E` be functions such that * `f` and `g` are differentiable in the open first quadrant and are continuous on its closure; -* `∥f z∥` and `∥g z∥` are bounded from above by `A * exp(B * (abs z) ^ c)` on the open first +* `‖f z‖` and `‖g z‖` are bounded from above by `A * exp(B * (abs z) ^ c)` on the open first quadrant for some `A`, `B`, and `c < 2`; * `f` is equal to `g` on the boundary of the first quadrant. @@ -449,21 +449,21 @@ lemma eq_on_quadrant_I (hdf : diff_cont_on_cl ℂ f (Ioi 0 ×ℂ Ioi 0)) /-- **Phragmen-Lindelöf principle** in the second quadrant. Let `f : ℂ → E` be a function such that * `f` is differentiable in the open second quadrant and is continuous on its closure; -* `∥f z∥` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open second quadrant +* `‖f z‖` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open second quadrant for some `c < 2`; -* `∥f z∥` is bounded from above by a constant `C` on the boundary of the second quadrant. +* `‖f z‖` is bounded from above by a constant `C` on the boundary of the second quadrant. -Then `∥f z∥` is bounded from above by the same constant on the closed second quadrant. -/ +Then `‖f z‖` is bounded from above by the same constant on the closed second quadrant. -/ lemma quadrant_II (hd : diff_cont_on_cl ℂ f (Iio 0 ×ℂ Ioi 0)) (hB : ∃ (c < (2 : ℝ)) B, f =O[comap complex.abs at_top ⊓ 𝓟 (Iio 0 ×ℂ Ioi 0)] (λ z, expR (B * (abs z) ^ c))) - (hre : ∀ x : ℝ, x ≤ 0 → ∥f x∥ ≤ C) (him : ∀ x : ℝ, 0 ≤ x → ∥f (x * I)∥ ≤ C) + (hre : ∀ x : ℝ, x ≤ 0 → ‖f x‖ ≤ C) (him : ∀ x : ℝ, 0 ≤ x → ‖f (x * I)‖ ≤ C) (hz_re : z.re ≤ 0) (hz_im : 0 ≤ z.im) : - ∥f z∥ ≤ C := + ‖f z‖ ≤ C := begin obtain ⟨z, rfl⟩ : ∃ z', z' * I = z, from ⟨z / I, div_mul_cancel _ I_ne_zero⟩, simp only [mul_I_re, mul_I_im, neg_nonpos] at hz_re hz_im, - change ∥(f ∘ (* I)) z∥ ≤ C, + change ‖(f ∘ (* I)) z‖ ≤ C, have H : maps_to (* I) (Ioi 0 ×ℂ Ioi 0) (Iio 0 ×ℂ Ioi 0), { intros w hw, simpa only [mem_re_prod_im, mul_I_re, mul_I_im, neg_lt_zero, mem_Iio] using hw.symm }, @@ -478,7 +478,7 @@ end /-- **Phragmen-Lindelöf principle** in the second quadrant. Let `f : ℂ → E` be a function such that * `f` is differentiable in the open second quadrant and is continuous on its closure; -* `∥f z∥` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open second quadrant +* `‖f z‖` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open second quadrant for some `A`, `B`, and `c < 2`; * `f` is equal to zero on the boundary of the second quadrant. @@ -494,7 +494,7 @@ lemma eq_zero_on_quadrant_II (hd : diff_cont_on_cl ℂ f (Iio 0 ×ℂ Ioi 0)) /-- **Phragmen-Lindelöf principle** in the second quadrant. Let `f g : ℂ → E` be functions such that * `f` and `g` are differentiable in the open second quadrant and are continuous on its closure; -* `∥f z∥` and `∥g z∥` are bounded from above by `A * exp(B * (abs z) ^ c)` on the open second +* `‖f z‖` and `‖g z‖` are bounded from above by `A * exp(B * (abs z) ^ c)` on the open second quadrant for some `A`, `B`, and `c < 2`; * `f` is equal to `g` on the boundary of the second quadrant. @@ -513,21 +513,21 @@ lemma eq_on_quadrant_II (hdf : diff_cont_on_cl ℂ f (Iio 0 ×ℂ Ioi 0)) /-- **Phragmen-Lindelöf principle** in the third quadrant. Let `f : ℂ → E` be a function such that * `f` is differentiable in the open third quadrant and is continuous on its closure; -* `∥f z∥` is bounded from above by `A * exp (B * (abs z) ^ c)` on the open third quadrant +* `‖f z‖` is bounded from above by `A * exp (B * (abs z) ^ c)` on the open third quadrant for some `c < 2`; -* `∥f z∥` is bounded from above by a constant `C` on the boundary of the third quadrant. +* `‖f z‖` is bounded from above by a constant `C` on the boundary of the third quadrant. -Then `∥f z∥` is bounded from above by the same constant on the closed third quadrant. -/ +Then `‖f z‖` is bounded from above by the same constant on the closed third quadrant. -/ lemma quadrant_III (hd : diff_cont_on_cl ℂ f (Iio 0 ×ℂ Iio 0)) (hB : ∃ (c < (2 : ℝ)) B, f =O[comap complex.abs at_top ⊓ 𝓟 (Iio 0 ×ℂ Iio 0)] (λ z, expR (B * (abs z) ^ c))) - (hre : ∀ x : ℝ, x ≤ 0 → ∥f x∥ ≤ C) (him : ∀ x : ℝ, x ≤ 0 → ∥f (x * I)∥ ≤ C) + (hre : ∀ x : ℝ, x ≤ 0 → ‖f x‖ ≤ C) (him : ∀ x : ℝ, x ≤ 0 → ‖f (x * I)‖ ≤ C) (hz_re : z.re ≤ 0) (hz_im : z.im ≤ 0) : - ∥f z∥ ≤ C := + ‖f z‖ ≤ C := begin obtain ⟨z, rfl⟩ : ∃ z', -z' = z, from ⟨-z, neg_neg z⟩, simp only [neg_re, neg_im, neg_nonpos] at hz_re hz_im, - change ∥(f ∘ has_neg.neg) z∥ ≤ C, + change ‖(f ∘ has_neg.neg) z‖ ≤ C, have H : maps_to has_neg.neg (Ioi 0 ×ℂ Ioi 0) (Iio 0 ×ℂ Iio 0), { intros w hw, simpa only [mem_re_prod_im, neg_re, neg_im, neg_lt_zero, mem_Iio] using hw }, @@ -545,7 +545,7 @@ end /-- **Phragmen-Lindelöf principle** in the third quadrant. Let `f : ℂ → E` be a function such that * `f` is differentiable in the open third quadrant and is continuous on its closure; -* `∥f z∥` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open third quadrant +* `‖f z‖` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open third quadrant for some `A`, `B`, and `c < 2`; * `f` is equal to zero on the boundary of the third quadrant. @@ -561,7 +561,7 @@ lemma eq_zero_on_quadrant_III (hd : diff_cont_on_cl ℂ f (Iio 0 ×ℂ Iio 0)) /-- **Phragmen-Lindelöf principle** in the third quadrant. Let `f g : ℂ → E` be functions such that * `f` and `g` are differentiable in the open third quadrant and are continuous on its closure; -* `∥f z∥` and `∥g z∥` are bounded from above by `A * exp(B * (abs z) ^ c)` on the open third +* `‖f z‖` and `‖g z‖` are bounded from above by `A * exp(B * (abs z) ^ c)` on the open third quadrant for some `A`, `B`, and `c < 2`; * `f` is equal to `g` on the boundary of the third quadrant. @@ -580,21 +580,21 @@ lemma eq_on_quadrant_III (hdf : diff_cont_on_cl ℂ f (Iio 0 ×ℂ Iio 0)) /-- **Phragmen-Lindelöf principle** in the fourth quadrant. Let `f : ℂ → E` be a function such that * `f` is differentiable in the open fourth quadrant and is continuous on its closure; -* `∥f z∥` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open fourth quadrant +* `‖f z‖` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open fourth quadrant for some `c < 2`; -* `∥f z∥` is bounded from above by a constant `C` on the boundary of the fourth quadrant. +* `‖f z‖` is bounded from above by a constant `C` on the boundary of the fourth quadrant. -Then `∥f z∥` is bounded from above by the same constant on the closed fourth quadrant. -/ +Then `‖f z‖` is bounded from above by the same constant on the closed fourth quadrant. -/ lemma quadrant_IV (hd : diff_cont_on_cl ℂ f (Ioi 0 ×ℂ Iio 0)) (hB : ∃ (c < (2 : ℝ)) B, f =O[comap complex.abs at_top ⊓ 𝓟 (Ioi 0 ×ℂ Iio 0)] (λ z, expR (B * (abs z) ^ c))) - (hre : ∀ x : ℝ, 0 ≤ x → ∥f x∥ ≤ C) (him : ∀ x : ℝ, x ≤ 0 → ∥f (x * I)∥ ≤ C) + (hre : ∀ x : ℝ, 0 ≤ x → ‖f x‖ ≤ C) (him : ∀ x : ℝ, x ≤ 0 → ‖f (x * I)‖ ≤ C) (hz_re : 0 ≤ z.re) (hz_im : z.im ≤ 0) : - ∥f z∥ ≤ C := + ‖f z‖ ≤ C := begin obtain ⟨z, rfl⟩ : ∃ z', -z' = z, from ⟨-z, neg_neg z⟩, simp only [neg_re, neg_im, neg_nonpos, neg_nonneg] at hz_re hz_im, - change ∥(f ∘ has_neg.neg) z∥ ≤ C, + change ‖(f ∘ has_neg.neg) z‖ ≤ C, have H : maps_to has_neg.neg (Iio 0 ×ℂ Ioi 0) (Ioi 0 ×ℂ Iio 0), { intros w hw, simpa only [mem_re_prod_im, neg_re, neg_im, neg_lt_zero, neg_pos, mem_Ioi, mem_Iio] using hw }, @@ -612,7 +612,7 @@ end /-- **Phragmen-Lindelöf principle** in the fourth quadrant. Let `f : ℂ → E` be a function such that * `f` is differentiable in the open fourth quadrant and is continuous on its closure; -* `∥f z∥` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open fourth quadrant +* `‖f z‖` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open fourth quadrant for some `A`, `B`, and `c < 2`; * `f` is equal to zero on the boundary of the fourth quadrant. @@ -628,7 +628,7 @@ lemma eq_zero_on_quadrant_IV (hd : diff_cont_on_cl ℂ f (Ioi 0 ×ℂ Iio 0)) /-- **Phragmen-Lindelöf principle** in the fourth quadrant. Let `f g : ℂ → E` be functions such that * `f` and `g` are differentiable in the open fourth quadrant and are continuous on its closure; -* `∥f z∥` and `∥g z∥` are bounded from above by `A * exp(B * (abs z) ^ c)` on the open fourth +* `‖f z‖` and `‖g z‖` are bounded from above by `A * exp(B * (abs z) ^ c)` on the open fourth quadrant for some `A`, `B`, and `c < 2`; * `f` is equal to `g` on the boundary of the fourth quadrant. @@ -651,24 +651,24 @@ lemma eq_on_quadrant_IV (hdf : diff_cont_on_cl ℂ f (Ioi 0 ×ℂ Iio 0)) /-- **Phragmen-Lindelöf principle** in the right half-plane. Let `f : ℂ → E` be a function such that * `f` is differentiable in the open right half-plane and is continuous on its closure; -* `∥f z∥` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open right half-plane +* `‖f z‖` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open right half-plane for some `c < 2`; -* `∥f z∥` is bounded from above by a constant `C` on the imaginary axis; +* `‖f z‖` is bounded from above by a constant `C` on the imaginary axis; * `f x → 0` as `x : ℝ` tends to infinity. -Then `∥f z∥` is bounded from above by the same constant on the closed right half-plane. +Then `‖f z‖` is bounded from above by the same constant on the closed right half-plane. See also `phragmen_lindelof.right_half_plane_of_bounded_on_real` for a stronger version. -/ lemma right_half_plane_of_tendsto_zero_on_real (hd : diff_cont_on_cl ℂ f {z | 0 < z.re}) (hexp : ∃ (c < (2 : ℝ)) B, f =O[comap complex.abs at_top ⊓ 𝓟 {z | 0 < z.re}] (λ z, expR (B * (abs z) ^ c))) - (hre : tendsto (λ x : ℝ, f x) at_top (𝓝 0)) (him : ∀ x : ℝ, ∥f (x * I)∥ ≤ C) (hz : 0 ≤ z.re) : - ∥f z∥ ≤ C := + (hre : tendsto (λ x : ℝ, f x) at_top (𝓝 0)) (him : ∀ x : ℝ, ‖f (x * I)‖ ≤ C) (hz : 0 ≤ z.re) : + ‖f z‖ ≤ C := begin /- We are going to apply the Phragmen-Lindelöf principle in the first and fourth quadrants. - The lemmas immediately imply that for any upper estimate `C'` on `∥f x∥`, `x : ℝ`, `0 ≤ x`, + The lemmas immediately imply that for any upper estimate `C'` on `‖f x‖`, `x : ℝ`, `0 ≤ x`, the number `max C C'` is an upper estimate on `f` in the whole right half-plane. -/ revert z, - have hle : ∀ C', (∀ x : ℝ, 0 ≤ x → ∥f x∥ ≤ C') → ∀ z : ℂ, 0 ≤ z.re → ∥f z∥ ≤ max C C', + have hle : ∀ C', (∀ x : ℝ, 0 ≤ x → ‖f x‖ ≤ C') → ∀ z : ℂ, 0 ≤ z.re → ‖f z‖ ≤ max C C', { intros C' hC' z hz, cases le_total z.im 0, { refine quadrant_IV (hd.mono $ λ _, and.left) (Exists₃.imp (λ c hc B hO, _) hexp) @@ -679,9 +679,9 @@ begin (λ x hx, (hC' x hx).trans $ le_max_right _ _) (λ x hx, (him x).trans (le_max_left _ _)) hz h, exact hO.mono (inf_le_inf_left _ $ principal_mono.2 $ λ _, and.left) } }, - -- Since `f` is continuous on `Ici 0` and `∥f x∥` tends to zero as `x → ∞`, - -- the norm `∥f x∥` takes its maximum value at some `x₀ : ℝ`. - obtain ⟨x₀, hx₀, hmax⟩ : ∃ x : ℝ, 0 ≤ x ∧ ∀ y : ℝ, 0 ≤ y → ∥f y∥ ≤ ∥f x∥, + -- Since `f` is continuous on `Ici 0` and `‖f x‖` tends to zero as `x → ∞`, + -- the norm `‖f x‖` takes its maximum value at some `x₀ : ℝ`. + obtain ⟨x₀, hx₀, hmax⟩ : ∃ x : ℝ, 0 ≤ x ∧ ∀ y : ℝ, 0 ≤ y → ‖f y‖ ≤ ‖f x‖, { have hfc : continuous_on (λ x : ℝ, f x) (Ici 0), { refine hd.continuous_on.comp continuous_of_real.continuous_on (λ x hx, _), rwa closure_set_of_lt_re }, @@ -689,22 +689,22 @@ begin { refine ⟨0, le_rfl, λ y hy, _⟩, rw [h₀ y hy, h₀ 0 le_rfl] }, push_neg at h₀, rcases h₀ with ⟨x₀, hx₀, hne⟩, - have hlt : ∥(0 : E)∥ < ∥f x₀∥, by rwa [norm_zero, norm_pos_iff], - suffices : ∀ᶠ x : ℝ in cocompact ℝ ⊓ 𝓟 (Ici 0), ∥f x∥ ≤ ∥f x₀∥, + have hlt : ‖(0 : E)‖ < ‖f x₀‖, by rwa [norm_zero, norm_pos_iff], + suffices : ∀ᶠ x : ℝ in cocompact ℝ ⊓ 𝓟 (Ici 0), ‖f x‖ ≤ ‖f x₀‖, by simpa only [exists_prop] using hfc.norm.exists_forall_ge' is_closed_Ici hx₀ this, rw [real.cocompact_eq, inf_sup_right, (disjoint_at_bot_principal_Ici (0 : ℝ)).eq_bot, bot_sup_eq], exact (hre.norm.eventually $ ge_mem_nhds hlt).filter_mono inf_le_left }, - cases le_or_lt (∥f x₀∥) C, - { -- If `∥f x₀∥ ≤ C`, then `hle` implies the required estimate + cases le_or_lt (‖f x₀‖) C, + { -- If `‖f x₀‖ ≤ C`, then `hle` implies the required estimate simpa only [max_eq_left h] using hle _ hmax }, - { -- Otherwise, `∥f z∥ ≤ ∥f x₀∥` for all `z` in the right half-plane due to `hle`. + { -- Otherwise, `‖f z‖ ≤ ‖f x₀‖` for all `z` in the right half-plane due to `hle`. replace hmax : is_max_on (norm ∘ f) {z | 0 < z.re} x₀, { rintros z (hz : 0 < z.re), simpa [max_eq_right h.le] using hle _ hmax _ hz.le }, -- Due to the maximum modulus principle applied to the closed ball of radius `x₀.re`, - -- `∥f 0∥ = ∥f x₀∥`. - have : ∥f 0∥ = ∥f x₀∥, + -- `‖f 0‖ = ‖f x₀‖`. + have : ‖f 0‖ = ‖f x₀‖, { apply norm_eq_norm_of_is_max_on_of_ball_subset hd hmax, -- move to a lemma? intros z hz, @@ -715,7 +715,7 @@ begin ... ≤ |x₀ - z.re| : le_abs_self _ ... = |(z - x₀).re| : by rw [sub_re, of_real_re, _root_.abs_sub_comm] ... ≤ abs (z - x₀) : abs_re_le_abs _ }, - -- Thus we have `C < ∥f x₀∥ = ∥f 0∥ ≤ C`. Contradiction completes the proof. + -- Thus we have `C < ‖f x₀‖ = ‖f 0‖ ≤ C`. Contradiction completes the proof. refine (h.not_le $ this ▸ _).elim, simpa using him 0 } end @@ -723,32 +723,32 @@ end /-- **Phragmen-Lindelöf principle** in the right half-plane. Let `f : ℂ → E` be a function such that * `f` is differentiable in the open right half-plane and is continuous on its closure; -* `∥f z∥` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open right half-plane +* `‖f z‖` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open right half-plane for some `c < 2`; -* `∥f z∥` is bounded from above by a constant `C` on the imaginary axis; -* `∥f x∥` is bounded from above by a constant for large real values of `x`. +* `‖f z‖` is bounded from above by a constant `C` on the imaginary axis; +* `‖f x‖` is bounded from above by a constant for large real values of `x`. -Then `∥f z∥` is bounded from above by `C` on the closed right half-plane. +Then `‖f z‖` is bounded from above by `C` on the closed right half-plane. See also `phragmen_lindelof.right_half_plane_of_tendsto_zero_on_real` for a weaker version. -/ lemma right_half_plane_of_bounded_on_real (hd : diff_cont_on_cl ℂ f {z | 0 < z.re}) (hexp : ∃ (c < (2 : ℝ)) B, f =O[comap complex.abs at_top ⊓ 𝓟 {z | 0 < z.re}] (λ z, expR (B * (abs z) ^ c))) - (hre : is_bounded_under (≤) at_top (λ x : ℝ, ∥f x∥)) - (him : ∀ x : ℝ, ∥f (x * I)∥ ≤ C) (hz : 0 ≤ z.re) : - ∥f z∥ ≤ C := + (hre : is_bounded_under (≤) at_top (λ x : ℝ, ‖f x‖)) + (him : ∀ x : ℝ, ‖f (x * I)‖ ≤ C) (hz : 0 ≤ z.re) : + ‖f z‖ ≤ C := begin -- For each `ε < 0`, the function `λ z, exp (ε * z) • f z` satisfies assumptions of - -- `right_half_plane_of_tendsto_zero_on_real`, hence `∥exp (ε * z) • f z∥ ≤ C` for all `ε < 0`. + -- `right_half_plane_of_tendsto_zero_on_real`, hence `‖exp (ε * z) • f z‖ ≤ C` for all `ε < 0`. -- Taking the limit as `ε → 0`, we obtain the required inequality. - suffices : ∀ᶠ ε : ℝ in 𝓝[<] 0, ∥exp (ε * z) • f z∥ ≤ C, + suffices : ∀ᶠ ε : ℝ in 𝓝[<] 0, ‖exp (ε * z) • f z‖ ≤ C, { refine le_of_tendsto (tendsto.mono_left _ nhds_within_le_nhds) this, apply ((continuous_of_real.mul continuous_const).cexp.smul continuous_const).norm.tendsto', simp, apply_instance }, filter_upwards [self_mem_nhds_within] with ε ε₀, change ε < 0 at ε₀, - set g : ℂ → E := λ z, exp (ε * z) • f z, change ∥g z∥ ≤ C, + set g : ℂ → E := λ z, exp (ε * z) • f z, change ‖g z‖ ≤ C, replace hd : diff_cont_on_cl ℂ g {z : ℂ | 0 < z.re}, from (differentiable_id.const_mul _).cexp.diff_cont_on_cl.smul hd, - have hgn : ∀ z, ∥g z∥ = expR (ε * z.re) * ∥f z∥, + have hgn : ∀ z, ‖g z‖ = expR (ε * z.re) * ‖f z‖, { intro z, rw [norm_smul, norm_eq_abs, abs_exp, of_real_mul_re] }, refine right_half_plane_of_tendsto_zero_on_real hd _ _ (λ y, _) hz, { refine Exists₃.imp (λ c hc B hO, (is_O.of_bound 1 _).trans hO) hexp, @@ -767,19 +767,19 @@ end /-- **Phragmen-Lindelöf principle** in the right half-plane. Let `f : ℂ → E` be a function such that * `f` is differentiable in the open right half-plane and is continuous on its closure; -* `∥f z∥` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open right half-plane +* `‖f z‖` is bounded from above by `A * exp(B * (abs z) ^ c)` on the open right half-plane for some `c < 2`; -* `∥f z∥` is bounded from above by a constant on the imaginary axis; +* `‖f z‖` is bounded from above by a constant on the imaginary axis; * `f x`, `x : ℝ`, tends to zero superexponentially fast as `x → ∞`: - for any natural `n`, `exp (n * x) * ∥f x∥` tends to zero as `x → ∞`. + for any natural `n`, `exp (n * x) * ‖f x‖` tends to zero as `x → ∞`. Then `f` is equal to zero on the closed right half-plane. -/ lemma eq_zero_on_right_half_plane_of_superexponential_decay (hd : diff_cont_on_cl ℂ f {z | 0 < z.re}) (hexp : ∃ (c < (2 : ℝ)) B, f =O[comap complex.abs at_top ⊓ 𝓟 {z | 0 < z.re}] (λ z, expR (B * (abs z) ^ c))) - (hre : superpolynomial_decay at_top expR (λ x, ∥f x∥)) - (him : ∃ C, ∀ x : ℝ, ∥f (x * I)∥ ≤ C) : + (hre : superpolynomial_decay at_top expR (λ x, ‖f x‖)) + (him : ∃ C, ∀ x : ℝ, ‖f (x * I)‖ ≤ C) : eq_on f 0 {z : ℂ | 0 ≤ z.re} := begin rcases him with ⟨C, hC⟩, @@ -789,11 +789,11 @@ begin continuous_on_const subset_closure subset.rfl }, -- Consider $g_n(z)=e^{nz}f(z)$. set g : ℕ → ℂ → E := λ n z, (exp z) ^ n • f z, - have hg : ∀ n z, ∥g n z∥ = (expR z.re) ^ n * ∥f z∥, + have hg : ∀ n z, ‖g n z‖ = (expR z.re) ^ n * ‖f z‖, { intros n z, simp only [norm_smul, norm_eq_abs, complex.abs_pow, abs_exp] }, intros z hz, -- Since `e^{nz} → ∞` as `n → ∞`, it suffices to show that each `g_n` is bounded from above by `C` - suffices H : ∀ n : ℕ, ∥g n z∥ ≤ C, + suffices H : ∀ n : ℕ, ‖g n z‖ ≤ C, { contrapose! H, simp only [hg], exact (((tendsto_pow_at_top_at_top_of_one_lt (real.one_lt_exp_iff.2 hz)).at_top_mul @@ -825,11 +825,11 @@ end that * `f` and `g` are differentiable in the open right half-plane and are continuous on its closure; -* `∥f z∥` and `∥g z∥` are bounded from above by `A * exp(B * (abs z) ^ c)` on the open right +* `‖f z‖` and `‖g z‖` are bounded from above by `A * exp(B * (abs z) ^ c)` on the open right half-plane for some `c < 2`; -* `∥f z∥` and `∥g z∥` are bounded from above by constants on the imaginary axis; +* `‖f z‖` and `‖g z‖` are bounded from above by constants on the imaginary axis; * `f x - g x`, `x : ℝ`, tends to zero superexponentially fast as `x → ∞`: - for any natural `n`, `exp (n * x) * ∥f x - g x∥` tends to zero as `x → ∞`. + for any natural `n`, `exp (n * x) * ‖f x - g x‖` tends to zero as `x → ∞`. Then `f` is equal to `g` on the closed right half-plane. -/ lemma eq_on_right_half_plane_of_superexponential_decay {g : ℂ → E} @@ -838,8 +838,8 @@ lemma eq_on_right_half_plane_of_superexponential_decay {g : ℂ → E} (λ z, expR (B * (abs z) ^ c))) (hgexp : ∃ (c < (2 : ℝ)) B, g =O[comap complex.abs at_top ⊓ 𝓟 {z | 0 < z.re}] (λ z, expR (B * (abs z) ^ c))) - (hre : superpolynomial_decay at_top expR (λ x, ∥f x - g x∥)) - (hfim : ∃ C, ∀ x : ℝ, ∥f (x * I)∥ ≤ C) (hgim : ∃ C, ∀ x : ℝ, ∥g (x * I)∥ ≤ C) : + (hre : superpolynomial_decay at_top expR (λ x, ‖f x - g x‖)) + (hfim : ∃ C, ∀ x : ℝ, ‖f (x * I)‖ ≤ C) (hgim : ∃ C, ∀ x : ℝ, ‖g (x * I)‖ ≤ C) : eq_on f g {z : ℂ | 0 ≤ z.re} := begin suffices : eq_on (f - g) 0 {z : ℂ | 0 ≤ z.re}, diff --git a/src/analysis/complex/polynomial.lean b/src/analysis/complex/polynomial.lean index bc7c47964700c..d3a4710d9c922 100644 --- a/src/analysis/complex/polynomial.lean +++ b/src/analysis/complex/polynomial.lean @@ -29,7 +29,7 @@ begin exact degree_C_le }, { obtain ⟨z₀, h₀⟩ := f.exists_forall_norm_le, simp only [bounded_iff_forall_norm_le, set.forall_range_iff, norm_inv], - exact ⟨∥eval z₀ f∥⁻¹, λ z, inv_le_inv_of_le (norm_pos_iff.2 $ hf z₀) (h₀ z)⟩ }, + exact ⟨‖eval z₀ f‖⁻¹, λ z, inv_le_inv_of_le (norm_pos_iff.2 $ hf z₀) (h₀ z)⟩ }, end instance is_alg_closed : is_alg_closed ℂ := diff --git a/src/analysis/complex/removable_singularity.lean b/src/analysis/complex/removable_singularity.lean index 816caed16fc9a..ce0694306c384 100644 --- a/src/analysis/complex/removable_singularity.lean +++ b/src/analysis/complex/removable_singularity.lean @@ -101,7 +101,7 @@ lemma differentiable_on_update_lim_of_bdd_above {f : ℂ → E} {s : set ℂ} {c (hb : bdd_above (norm ∘ f '' (s \ {c}))) : differentiable_on ℂ (update f c (lim (𝓝[≠] c) f)) s := differentiable_on_update_lim_of_is_o hc hd $ is_bounded_under.is_o_sub_self_inv $ - let ⟨C, hC⟩ := hb in ⟨C + ∥f c∥, eventually_map.2 $ mem_nhds_within_iff_exists_mem_nhds_inter.2 + let ⟨C, hC⟩ := hb in ⟨C + ‖f c‖, eventually_map.2 $ mem_nhds_within_iff_exists_mem_nhds_inter.2 ⟨s, hc, λ z hz, norm_sub_le_of_le (hC $ mem_image_of_mem _ hz) le_rfl⟩⟩ /-- **Removable singularity** theorem: if a function `f : ℂ → E` is complex differentiable on a @@ -122,7 +122,7 @@ end bounded on a punctured neighborhood of `c`, then `f` has a limit at `c`. -/ lemma tendsto_lim_of_differentiable_on_punctured_nhds_of_bounded_under {f : ℂ → E} {c : ℂ} (hd : ∀ᶠ z in 𝓝[≠] c, differentiable_at ℂ f z) - (hb : is_bounded_under (≤) (𝓝[≠] c) (λ z, ∥f z - f c∥)) : + (hb : is_bounded_under (≤) (𝓝[≠] c) (λ z, ‖f z - f c‖)) : tendsto f (𝓝[≠] c) (𝓝 $ lim (𝓝[≠] c) f) := tendsto_lim_of_differentiable_on_punctured_nhds_of_is_o hd hb.is_o_sub_self_inv diff --git a/src/analysis/complex/roots_of_unity.lean b/src/analysis/complex/roots_of_unity.lean index ccb8fdaf78d87..ac3d0e1e87b18 100644 --- a/src/analysis/complex/roots_of_unity.lean +++ b/src/analysis/complex/roots_of_unity.lean @@ -96,10 +96,10 @@ end end complex lemma is_primitive_root.norm'_eq_one {ζ : ℂ} {n : ℕ} (h : is_primitive_root ζ n) (hn : n ≠ 0) : - ∥ζ∥ = 1 := complex.norm_eq_one_of_pow_eq_one h.pow_eq_one hn + ‖ζ‖ = 1 := complex.norm_eq_one_of_pow_eq_one h.pow_eq_one hn lemma is_primitive_root.nnnorm_eq_one {ζ : ℂ} {n : ℕ} (h : is_primitive_root ζ n) (hn : n ≠ 0) : - ∥ζ∥₊ = 1 := subtype.ext $ h.norm'_eq_one hn + ‖ζ‖₊ = 1 := subtype.ext $ h.norm'_eq_one hn lemma is_primitive_root.arg_ext {n m : ℕ} {ζ μ : ℂ} (hζ : is_primitive_root ζ n) (hμ : is_primitive_root μ m) (hn : n ≠ 0) (hm : m ≠ 0) (h : ζ.arg = μ.arg) : ζ = μ := diff --git a/src/analysis/complex/schwarz.lean b/src/analysis/complex/schwarz.lean index 469986d422a44..39c31bbd89302 100644 --- a/src/analysis/complex/schwarz.lean +++ b/src/analysis/complex/schwarz.lean @@ -59,10 +59,10 @@ variables {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] {R R₁ R /-- An auxiliary lemma for `complex.norm_dslope_le_div_of_maps_to_ball`. -/ lemma schwarz_aux {f : ℂ → ℂ} (hd : differentiable_on ℂ f (ball c R₁)) (h_maps : maps_to f (ball c R₁) (ball (f c) R₂)) (hz : z ∈ ball c R₁) : - ∥dslope f c z∥ ≤ R₂ / R₁ := + ‖dslope f c z‖ ≤ R₂ / R₁ := begin have hR₁ : 0 < R₁, from nonempty_ball.1 ⟨z, hz⟩, - suffices : ∀ᶠ r in 𝓝[<] R₁, ∥dslope f c z∥ ≤ R₂ / r, + suffices : ∀ᶠ r in 𝓝[<] R₁, ‖dslope f c z‖ ≤ R₂ / r, { refine ge_of_tendsto _ this, exact (tendsto_const_nhds.div tendsto_id hR₁.ne').mono_left nhds_within_le_nhds }, rw mem_ball at hz, @@ -87,16 +87,16 @@ end /-- Two cases of the **Schwarz Lemma** (derivative and distance), merged together. -/ lemma norm_dslope_le_div_of_maps_to_ball (hd : differentiable_on ℂ f (ball c R₁)) (h_maps : maps_to f (ball c R₁) (ball (f c) R₂)) (hz : z ∈ ball c R₁) : - ∥dslope f c z∥ ≤ R₂ / R₁ := + ‖dslope f c z‖ ≤ R₂ / R₁ := begin have hR₁ : 0 < R₁, from nonempty_ball.1 ⟨z, hz⟩, have hR₂ : 0 < R₂, from nonempty_ball.1 ⟨f z, h_maps hz⟩, cases eq_or_ne (dslope f c z) 0 with hc hc, { rw [hc, norm_zero], exact div_nonneg hR₂.le hR₁.le }, rcases exists_dual_vector ℂ _ hc with ⟨g, hg, hgf⟩, - have hg' : ∥g∥₊ = 1, from nnreal.eq hg, - have hg₀ : ∥g∥₊ ≠ 0, by simpa only [hg'] using one_ne_zero, - calc ∥dslope f c z∥ = ∥dslope (g ∘ f) c z∥ : + have hg' : ‖g‖₊ = 1, from nnreal.eq hg, + have hg₀ : ‖g‖₊ ≠ 0, by simpa only [hg'] using one_ne_zero, + calc ‖dslope f c z‖ = ‖dslope (g ∘ f) c z‖ : begin rw [g.dslope_comp, hgf, is_R_or_C.norm_of_real, norm_norm], exact λ _, hd.differentiable_at (ball_mem_nhds _ hR₁) @@ -110,18 +110,18 @@ begin end /-- Equality case in the **Schwarz Lemma**: in the setup of `norm_dslope_le_div_of_maps_to_ball`, if -`∥dslope f c z₀∥ = R₂ / R₁` holds at a point in the ball then the map `f` is affine. -/ +`‖dslope f c z₀‖ = R₂ / R₁` holds at a point in the ball then the map `f` is affine. -/ lemma affine_of_maps_to_ball_of_exists_norm_dslope_eq_div [complete_space E] [strict_convex_space ℝ E] (hd : differentiable_on ℂ f (ball c R₁)) (h_maps : set.maps_to f (ball c R₁) (ball (f c) R₂)) (h_z₀ : z₀ ∈ ball c R₁) - (h_eq : ∥dslope f c z₀∥ = R₂ / R₁) : + (h_eq : ‖dslope f c z₀‖ = R₂ / R₁) : set.eq_on f (λ z, f c + (z - c) • dslope f c z₀) (ball c R₁) := begin set g := dslope f c, rintro z hz, by_cases z = c, { simp [h] }, have h_R₁ : 0 < R₁ := nonempty_ball.mp ⟨_, h_z₀⟩, - have g_le_div : ∀ z ∈ ball c R₁, ∥g z∥ ≤ R₂ / R₁, + have g_le_div : ∀ z ∈ ball c R₁, ‖g z‖ ≤ R₂ / R₁, from λ z hz, norm_dslope_le_div_of_maps_to_ball hd h_maps hz, have g_max : is_max_on (norm ∘ g) (ball c R₁) z₀, from is_max_on_iff.mpr (λ z hz, by simpa [h_eq] using g_le_div z hz), @@ -135,8 +135,8 @@ end lemma affine_of_maps_to_ball_of_exists_norm_dslope_eq_div' [complete_space E] [strict_convex_space ℝ E] (hd : differentiable_on ℂ f (ball c R₁)) (h_maps : set.maps_to f (ball c R₁) (ball (f c) R₂)) - (h_z₀ : ∃ z₀ ∈ ball c R₁, ∥dslope f c z₀∥ = R₂ / R₁) : - ∃ C : E, ∥C∥ = R₂ / R₁ ∧ set.eq_on f (λ z, f c + (z - c) • C) (ball c R₁) := + (h_z₀ : ∃ z₀ ∈ ball c R₁, ‖dslope f c z₀‖ = R₂ / R₁) : + ∃ C : E, ‖C‖ = R₂ / R₁ ∧ set.eq_on f (λ z, f c + (z - c) • C) (ball c R₁) := let ⟨z₀, h_z₀, h_eq⟩ := h_z₀ in ⟨dslope f c z₀, h_eq, affine_of_maps_to_ball_of_exists_norm_dslope_eq_div hd h_maps h_z₀ h_eq⟩ @@ -145,7 +145,7 @@ let ⟨z₀, h_z₀, h_eq⟩ := h_z₀ in `f` at `c` is at most the ratio `R₂ / R₁`. -/ lemma norm_deriv_le_div_of_maps_to_ball (hd : differentiable_on ℂ f (ball c R₁)) (h_maps : maps_to f (ball c R₁) (ball (f c) R₂)) (h₀ : 0 < R₁) : - ∥deriv f c∥ ≤ R₂ / R₁ := + ‖deriv f c‖ ≤ R₂ / R₁ := by simpa only [dslope_same] using norm_dslope_le_div_of_maps_to_ball hd h_maps (mem_ball_self h₀) /-- The **Schwarz Lemma**: if `f : ℂ → E` sends an open disk with center `c` and radius `R₁` to an diff --git a/src/analysis/convex/gauge.lean b/src/analysis/convex/gauge.lean index ba0ed614635c8..2f53e33cfa4c7 100644 --- a/src/analysis/convex/gauge.lean +++ b/src/analysis/convex/gauge.lean @@ -288,7 +288,7 @@ end linear_ordered_field section is_R_or_C variables [is_R_or_C 𝕜] [module 𝕜 E] [is_scalar_tower ℝ 𝕜 E] -lemma gauge_norm_smul (hs : balanced 𝕜 s) (r : 𝕜) (x : E) : gauge s (∥r∥ • x) = gauge s (r • x) := +lemma gauge_norm_smul (hs : balanced 𝕜 s) (r : 𝕜) (x : E) : gauge s (‖r‖ • x) = gauge s (r • x) := begin rw @is_R_or_C.real_smul_eq_coe_smul 𝕜, obtain rfl | hr := eq_or_ne r 0, @@ -300,7 +300,7 @@ begin end /-- If `s` is balanced, then the Minkowski functional is ℂ-homogeneous. -/ -lemma gauge_smul (hs : balanced 𝕜 s) (r : 𝕜) (x : E) : gauge s (r • x) = ∥r∥ * gauge s x := +lemma gauge_smul (hs : balanced 𝕜 s) (r : 𝕜) (x : E) : gauge s (r • x) = ‖r‖ * gauge s x := by { rw [←smul_eq_mul, ←gauge_smul_of_nonneg (norm_nonneg r), gauge_norm_smul hs], apply_instance } end is_R_or_C @@ -436,18 +436,18 @@ end add_comm_group section norm variables [seminormed_add_comm_group E] [normed_space ℝ E] {s : set E} {r : ℝ} {x : E} -lemma gauge_unit_ball (x : E) : gauge (metric.ball (0 : E) 1) x = ∥x∥ := +lemma gauge_unit_ball (x : E) : gauge (metric.ball (0 : E) 1) x = ‖x‖ := begin obtain rfl | hx := eq_or_ne x 0, { rw [norm_zero, gauge_zero] }, refine (le_of_forall_pos_le_add $ λ ε hε, _).antisymm _, - { have : 0 < ∥x∥ + ε := by positivity, + { have : 0 < ‖x‖ + ε := by positivity, refine gauge_le_of_mem this.le _, rw [smul_ball this.ne', smul_zero, real.norm_of_nonneg this.le, mul_one, mem_ball_zero_iff], exact lt_add_of_pos_right _ hε }, refine le_gauge_of_not_mem balanced_ball_zero.star_convex (absorbent_ball_zero zero_lt_one).absorbs (λ h, _), - obtain hx' | hx' := eq_or_ne (∥x∥) 0, + obtain hx' | hx' := eq_or_ne (‖x‖) 0, { rw hx' at h, exact hx (zero_smul_set_subset _ h) }, { rw [mem_smul_set_iff_inv_smul_mem₀ hx', mem_ball_zero_iff, norm_smul, norm_inv, norm_norm, @@ -455,7 +455,7 @@ begin exact lt_irrefl _ h } end -lemma gauge_ball (hr : 0 < r) (x : E) : gauge (metric.ball (0 : E) r) x = ∥x∥ / r := +lemma gauge_ball (hr : 0 < r) (x : E) : gauge (metric.ball (0 : E) r) x = ‖x‖ / r := begin rw [←smul_unit_ball_of_pos hr, gauge_smul_left, pi.smul_apply, gauge_unit_ball, smul_eq_mul, abs_of_nonneg hr.le, div_eq_inv_mul], @@ -463,7 +463,7 @@ begin exact λ _, id, end -lemma mul_gauge_le_norm (hs : metric.ball (0 : E) r ⊆ s) : r * gauge s x ≤ ∥x∥ := +lemma mul_gauge_le_norm (hs : metric.ball (0 : E) r ⊆ s) : r * gauge s x ≤ ‖x‖ := begin obtain hr | hr := le_or_lt r 0, { exact (mul_nonpos_of_nonpos_of_nonneg hr $ gauge_nonneg _).trans (norm_nonneg _) }, diff --git a/src/analysis/convex/integral.lean b/src/analysis/convex/integral.lean index 9c95f27474cd2..2bb94caed20b2 100644 --- a/src/analysis/convex/integral.lean +++ b/src/analysis/convex/integral.lean @@ -315,12 +315,12 @@ lemma strict_concave_on.ae_eq_const_or_lt_map_average [is_finite_measure μ] by simpa only [pi.neg_apply, average_neg, neg_lt_neg_iff] using hg.neg.ae_eq_const_or_map_average_lt hgc.neg hsc hfs hfi hgi.neg -/-- If `E` is a strictly convex normed space and `f : α → E` is a function such that `∥f x∥ ≤ C` +/-- If `E` is a strictly convex normed space and `f : α → E` is a function such that `‖f x‖ ≤ C` a.e., then either this function is a.e. equal to its average value, or the norm of its average value is strictly less than `C`. -/ lemma ae_eq_const_or_norm_average_lt_of_norm_le_const [strict_convex_space ℝ E] - (h_le : ∀ᵐ x ∂μ, ∥f x∥ ≤ C) : - (f =ᵐ[μ] const α ⨍ x, f x ∂μ) ∨ ∥⨍ x, f x ∂μ∥ < C := + (h_le : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : + (f =ᵐ[μ] const α ⨍ x, f x ∂μ) ∨ ‖⨍ x, f x ∂μ‖ < C := begin cases le_or_lt C 0 with hC0 hC0, { have : f =ᵐ[μ] 0, from h_le.mono (λ x hx, norm_le_zero_iff.1 (hx.trans hC0)), @@ -336,12 +336,12 @@ begin is_closed_ball h_le hfi end -/-- If `E` is a strictly convex normed space and `f : α → E` is a function such that `∥f x∥ ≤ C` +/-- If `E` is a strictly convex normed space and `f : α → E` is a function such that `‖f x‖ ≤ C` a.e., then either this function is a.e. equal to its average value, or the norm of its integral is strictly less than `(μ univ).to_real * C`. -/ lemma ae_eq_const_or_norm_integral_lt_of_norm_le_const [strict_convex_space ℝ E] - [is_finite_measure μ] (h_le : ∀ᵐ x ∂μ, ∥f x∥ ≤ C) : - (f =ᵐ[μ] const α ⨍ x, f x ∂μ) ∨ ∥∫ x, f x ∂μ∥ < (μ univ).to_real * C := + [is_finite_measure μ] (h_le : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : + (f =ᵐ[μ] const α ⨍ x, f x ∂μ) ∨ ‖∫ x, f x ∂μ‖ < (μ univ).to_real * C := begin cases eq_or_ne μ 0 with h₀ h₀, { left, simp [h₀] }, have hμ : 0 < (μ univ).to_real, @@ -351,12 +351,12 @@ begin ← div_eq_inv_mul, div_lt_iff' hμ] at H end -/-- If `E` is a strictly convex normed space and `f : α → E` is a function such that `∥f x∥ ≤ C` +/-- If `E` is a strictly convex normed space and `f : α → E` is a function such that `‖f x‖ ≤ C` a.e. on a set `t` of finite measure, then either this function is a.e. equal to its average value on `t`, or the norm of its integral over `t` is strictly less than `(μ t).to_real * C`. -/ lemma ae_eq_const_or_norm_set_integral_lt_of_norm_le_const [strict_convex_space ℝ E] - (ht : μ t ≠ ∞) (h_le : ∀ᵐ x ∂μ.restrict t, ∥f x∥ ≤ C) : - (f =ᵐ[μ.restrict t] const α ⨍ x in t, f x ∂μ) ∨ ∥∫ x in t, f x ∂μ∥ < (μ t).to_real * C := + (ht : μ t ≠ ∞) (h_le : ∀ᵐ x ∂μ.restrict t, ‖f x‖ ≤ C) : + (f =ᵐ[μ.restrict t] const α ⨍ x in t, f x ∂μ) ∨ ‖∫ x in t, f x ∂μ‖ < (μ t).to_real * C := begin haveI := fact.mk ht.lt_top, rw [← restrict_apply_univ], diff --git a/src/analysis/convex/strict_convex_space.lean b/src/analysis/convex/strict_convex_space.lean index bc9fa523c305d..8abf23959c21e 100644 --- a/src/analysis/convex/strict_convex_space.lean +++ b/src/analysis/convex/strict_convex_space.lean @@ -42,7 +42,7 @@ We also provide several lemmas that can be used as alternative constructors for - `strict_convex_space.of_strict_convex_closed_unit_ball`: if `closed_ball (0 : E) 1` is strictly convex, then `E` is a strictly convex space; -- `strict_convex_space.of_norm_add`: if `∥x + y∥ = ∥x∥ + ∥y∥` implies `same_ray ℝ x y` for all +- `strict_convex_space.of_norm_add`: if `‖x + y‖ = ‖x‖ + ‖y‖` implies `same_ray ℝ x y` for all nonzero `x y : E`, then `E` is a strictly convex space. ## Implementation notes @@ -88,11 +88,11 @@ lemma strict_convex_space.of_strict_convex_closed_unit_ball strict_convex_space 𝕜 E := ⟨λ r hr, by simpa only [smul_closed_unit_ball_of_nonneg hr.le] using h.smul r⟩ -/-- Strict convexity is equivalent to `∥a • x + b • y∥ < 1` for all `x` and `y` of norm at most `1` +/-- Strict convexity is equivalent to `‖a • x + b • y‖ < 1` for all `x` and `y` of norm at most `1` and all strictly positive `a` and `b` such that `a + b = 1`. This lemma shows that it suffices to check this for points of norm one and some `a`, `b` such that `a + b = 1`. -/ lemma strict_convex_space.of_norm_combo_lt_one - (h : ∀ x y : E, ∥x∥ = 1 → ∥y∥ = 1 → x ≠ y → ∃ a b : ℝ, a + b = 1 ∧ ∥a • x + b • y∥ < 1) : + (h : ∀ x y : E, ‖x‖ = 1 → ‖y‖ = 1 → x ≠ y → ∃ a b : ℝ, a + b = 1 ∧ ‖a • x + b • y‖ < 1) : strict_convex_space ℝ E := begin refine strict_convex_space.of_strict_convex_closed_unit_ball ℝ @@ -106,8 +106,8 @@ begin end lemma strict_convex_space.of_norm_combo_ne_one - (h : ∀ x y : E, ∥x∥ = 1 → ∥y∥ = 1 → x ≠ y → - ∃ a b : ℝ, 0 ≤ a ∧ 0 ≤ b ∧ a + b = 1 ∧ ∥a • x + b • y∥ ≠ 1) : + (h : ∀ x y : E, ‖x‖ = 1 → ‖y‖ = 1 → x ≠ y → + ∃ a b : ℝ, 0 ≤ a ∧ 0 ≤ b ∧ a + b = 1 ∧ ‖a • x + b • y‖ ≠ 1) : strict_convex_space ℝ E := begin refine strict_convex_space.of_strict_convex_closed_unit_ball ℝ @@ -120,7 +120,7 @@ begin end lemma strict_convex_space.of_norm_add_ne_two - (h : ∀ ⦃x y : E⦄, ∥x∥ = 1 → ∥y∥ = 1 → x ≠ y → ∥x + y∥ ≠ 2) : + (h : ∀ ⦃x y : E⦄, ‖x‖ = 1 → ‖y‖ = 1 → x ≠ y → ‖x + y‖ ≠ 2) : strict_convex_space ℝ E := begin refine strict_convex_space.of_norm_combo_ne_one @@ -131,15 +131,15 @@ begin end lemma strict_convex_space.of_pairwise_sphere_norm_ne_two - (h : (sphere (0 : E) 1).pairwise $ λ x y, ∥x + y∥ ≠ 2) : + (h : (sphere (0 : E) 1).pairwise $ λ x y, ‖x + y‖ ≠ 2) : strict_convex_space ℝ E := strict_convex_space.of_norm_add_ne_two $ λ x y hx hy, h (mem_sphere_zero_iff_norm.2 hx) (mem_sphere_zero_iff_norm.2 hy) -/-- If `∥x + y∥ = ∥x∥ + ∥y∥` implies that `x y : E` are in the same ray, then `E` is a strictly +/-- If `‖x + y‖ = ‖x‖ + ‖y‖` implies that `x y : E` are in the same ray, then `E` is a strictly convex space. See also a more -/ lemma strict_convex_space.of_norm_add - (h : ∀ x y : E, ∥x∥ = 1 → ∥y∥ = 1 → ∥x + y∥ = 2 → same_ray ℝ x y) : + (h : ∀ x y : E, ‖x‖ = 1 → ‖y‖ = 1 → ‖x + y‖ = 2 → same_ray ℝ x y) : strict_convex_space ℝ E := begin refine strict_convex_space.of_pairwise_sphere_norm_ne_two (λ x hx y hy, mt $ λ h₂, _), @@ -169,21 +169,21 @@ lemma open_segment_subset_ball_of_ne (hx : x ∈ closed_ball z r) (hy : y ∈ cl /-- If `x` and `y` are two distinct vectors of norm at most `r`, then a convex combination of `x` and `y` with positive coefficients has norm strictly less than `r`. -/ -lemma norm_combo_lt_of_ne (hx : ∥x∥ ≤ r) (hy : ∥y∥ ≤ r) (hne : x ≠ y) (ha : 0 < a) (hb : 0 < b) - (hab : a + b = 1) : ∥a • x + b • y∥ < r := +lemma norm_combo_lt_of_ne (hx : ‖x‖ ≤ r) (hy : ‖y‖ ≤ r) (hne : x ≠ y) (ha : 0 < a) (hb : 0 < b) + (hab : a + b = 1) : ‖a • x + b • y‖ < r := begin simp only [← mem_ball_zero_iff, ← mem_closed_ball_zero_iff] at hx hy ⊢, exact combo_mem_ball_of_ne hx hy hne ha hb hab end -/-- In a strictly convex space, if `x` and `y` are not in the same ray, then `∥x + y∥ < ∥x∥ + -∥y∥`. -/ -lemma norm_add_lt_of_not_same_ray (h : ¬same_ray ℝ x y) : ∥x + y∥ < ∥x∥ + ∥y∥ := +/-- In a strictly convex space, if `x` and `y` are not in the same ray, then `‖x + y‖ < ‖x‖ + +‖y‖`. -/ +lemma norm_add_lt_of_not_same_ray (h : ¬same_ray ℝ x y) : ‖x + y‖ < ‖x‖ + ‖y‖ := begin simp only [same_ray_iff_inv_norm_smul_eq, not_or_distrib, ← ne.def] at h, rcases h with ⟨hx, hy, hne⟩, rw ← norm_pos_iff at hx hy, - have hxy : 0 < ∥x∥ + ∥y∥ := add_pos hx hy, + have hxy : 0 < ‖x‖ + ‖y‖ := add_pos hx hy, have := combo_mem_ball_of_ne (inv_norm_smul_mem_closed_unit_ball x) (inv_norm_smul_mem_closed_unit_ball y) hne (div_pos hx hxy) (div_pos hy hxy) (by rw [← add_div, div_self hxy.ne']), @@ -192,13 +192,13 @@ begin real.norm_of_nonneg (inv_pos.2 hxy).le, ← div_eq_inv_mul, div_lt_one hxy] at this end -lemma lt_norm_sub_of_not_same_ray (h : ¬same_ray ℝ x y) : ∥x∥ - ∥y∥ < ∥x - y∥ := +lemma lt_norm_sub_of_not_same_ray (h : ¬same_ray ℝ x y) : ‖x‖ - ‖y‖ < ‖x - y‖ := begin nth_rewrite 0 ←sub_add_cancel x y at ⊢ h, exact sub_lt_iff_lt_add.2 (norm_add_lt_of_not_same_ray $ λ H', h $ H'.add_left same_ray.rfl), end -lemma abs_lt_norm_sub_of_not_same_ray (h : ¬same_ray ℝ x y) : |∥x∥ - ∥y∥| < ∥x - y∥ := +lemma abs_lt_norm_sub_of_not_same_ray (h : ¬same_ray ℝ x y) : |‖x‖ - ‖y‖| < ‖x - y‖ := begin refine abs_sub_lt_iff.2 ⟨lt_norm_sub_of_not_same_ray h, _⟩, rw norm_sub_rev, @@ -207,23 +207,23 @@ end /-- In a strictly convex space, two vectors `x`, `y` are in the same ray if and only if the triangle inequality for `x` and `y` becomes an equality. -/ -lemma same_ray_iff_norm_add : same_ray ℝ x y ↔ ∥x + y∥ = ∥x∥ + ∥y∥ := +lemma same_ray_iff_norm_add : same_ray ℝ x y ↔ ‖x + y‖ = ‖x‖ + ‖y‖ := ⟨same_ray.norm_add, λ h, not_not.1 $ λ h', (norm_add_lt_of_not_same_ray h').ne h⟩ /-- If `x` and `y` are two vectors in a strictly convex space have the same norm and the norm of their sum is equal to the sum of their norms, then they are equal. -/ -lemma eq_of_norm_eq_of_norm_add_eq (h₁ : ∥x∥ = ∥y∥) (h₂ : ∥x + y∥ = ∥x∥ + ∥y∥) : x = y := +lemma eq_of_norm_eq_of_norm_add_eq (h₁ : ‖x‖ = ‖y‖) (h₂ : ‖x + y‖ = ‖x‖ + ‖y‖) : x = y := (same_ray_iff_norm_add.mpr h₂).eq_of_norm_eq h₁ /-- In a strictly convex space, two vectors `x`, `y` are not in the same ray if and only if the triangle inequality for `x` and `y` is strict. -/ -lemma not_same_ray_iff_norm_add_lt : ¬ same_ray ℝ x y ↔ ∥x + y∥ < ∥x∥ + ∥y∥ := +lemma not_same_ray_iff_norm_add_lt : ¬ same_ray ℝ x y ↔ ‖x + y‖ < ‖x‖ + ‖y‖ := same_ray_iff_norm_add.not.trans (norm_add_le _ _).lt_iff_ne.symm -lemma same_ray_iff_norm_sub : same_ray ℝ x y ↔ ∥x - y∥ = |∥x∥ - ∥y∥| := +lemma same_ray_iff_norm_sub : same_ray ℝ x y ↔ ‖x - y‖ = |‖x‖ - ‖y‖| := ⟨same_ray.norm_sub, λ h, not_not.1 $ λ h', (abs_lt_norm_sub_of_not_same_ray h').ne' h⟩ -lemma not_same_ray_iff_abs_lt_norm_sub : ¬ same_ray ℝ x y ↔ |∥x∥ - ∥y∥| < ∥x - y∥ := +lemma not_same_ray_iff_abs_lt_norm_sub : ¬ same_ray ℝ x y ↔ |‖x‖ - ‖y‖| < ‖x - y‖ := same_ray_iff_norm_sub.not.trans $ ne_comm.trans (abs_norm_sub_norm_le _ _).lt_iff_ne.symm /-- In a strictly convex space, the triangle inequality turns into an equality if and only if the @@ -232,7 +232,7 @@ lemma dist_add_dist_eq_iff : dist x y + dist y z = dist x z ↔ y ∈ [x -[ℝ] by simp only [mem_segment_iff_same_ray, same_ray_iff_norm_add, dist_eq_norm', sub_add_sub_cancel', eq_comm] -lemma norm_midpoint_lt_iff (h : ∥x∥ = ∥y∥) : ∥(1/2 : ℝ) • (x + y)∥ < ∥x∥ ↔ x ≠ y := +lemma norm_midpoint_lt_iff (h : ‖x‖ = ‖y‖) : ‖(1/2 : ℝ) • (x + y)‖ < ‖x‖ ↔ x ≠ y := by rw [norm_smul, real.norm_of_nonneg (one_div_nonneg.2 zero_le_two), ←inv_eq_one_div, ←div_eq_inv_mul, div_lt_iff (@zero_lt_two ℝ _ _), mul_two, ←not_same_ray_iff_of_norm_eq h, not_same_ray_iff_norm_add_lt, h] diff --git a/src/analysis/convex/topology.lean b/src/analysis/convex/topology.lean index 90f12a2b14dcc..c5a430c021e32 100644 --- a/src/analysis/convex/topology.lean +++ b/src/analysis/convex/topology.lean @@ -339,8 +339,8 @@ variables [seminormed_add_comm_group E] [normed_space ℝ E] {s t : set E} and `convex_on_univ_norm`. -/ lemma convex_on_norm (hs : convex ℝ s) : convex_on ℝ s norm := ⟨hs, λ x hx y hy a b ha hb hab, - calc ∥a • x + b • y∥ ≤ ∥a • x∥ + ∥b • y∥ : norm_add_le _ _ - ... = a * ∥x∥ + b * ∥y∥ + calc ‖a • x + b • y‖ ≤ ‖a • x‖ + ‖b • y‖ : norm_add_le _ _ + ... = a * ‖x‖ + b * ‖y‖ : by rw [norm_smul, norm_smul, real.norm_of_nonneg ha, real.norm_of_nonneg hb]⟩ /-- The norm on a real normed space is convex on the whole space. See also `seminorm.convex_on` diff --git a/src/analysis/convex/uniform.lean b/src/analysis/convex/uniform.lean index baf324e368ea9..5a7fc0b019ddf 100644 --- a/src/analysis/convex/uniform.lean +++ b/src/analysis/convex/uniform.lean @@ -9,10 +9,10 @@ import analysis.convex.strict_convex_space # Uniformly convex spaces This file defines uniformly convex spaces, which are real normed vector spaces in which for all -strictly positive `ε`, there exists some strictly positive `δ` such that `ε ≤ ∥x - y∥` implies -`∥x + y∥ ≤ 2 - δ` for all `x` and `y` of norm at most than `1`. This means that the triangle +strictly positive `ε`, there exists some strictly positive `δ` such that `ε ≤ ‖x - y‖` implies +`‖x + y‖ ≤ 2 - δ` for all `x` and `y` of norm at most than `1`. This means that the triangle inequality is strict with a uniform bound, as opposed to strictly convex spaces where the triangle -inequality is strict but not necessarily uniformly (`∥x + y∥ < ∥x∥ + ∥y∥` for all `x` and `y` not in +inequality is strict but not necessarily uniformly (`‖x + y‖ < ‖x‖ + ‖y‖` for all `x` and `y` not in the same ray). ## Main declarations @@ -33,13 +33,13 @@ open set metric open_locale convex pointwise /-- A *uniformly convex space* is a real normed space where the triangle inequality is strict with a -uniform bound. Namely, over the `x` and `y` of norm `1`, `∥x + y∥` is uniformly bounded above -by a constant `< 2` when `∥x - y∥` is uniformly bounded below by a positive constant. +uniform bound. Namely, over the `x` and `y` of norm `1`, `‖x + y‖` is uniformly bounded above +by a constant `< 2` when `‖x - y‖` is uniformly bounded below by a positive constant. See also `uniform_convex_space.of_uniform_convex_closed_unit_ball`. -/ class uniform_convex_space (E : Type*) [seminormed_add_comm_group E] : Prop := (uniform_convex : ∀ ⦃ε : ℝ⦄, 0 < ε → ∃ δ, 0 < δ ∧ - ∀ ⦃x : E⦄, ∥x∥ = 1 → ∀ ⦃y⦄, ∥y∥ = 1 → ε ≤ ∥x - y∥ → ∥x + y∥ ≤ 2 - δ) + ∀ ⦃x : E⦄, ‖x‖ = 1 → ∀ ⦃y⦄, ‖y‖ = 1 → ε ≤ ‖x - y‖ → ‖x + y‖ ≤ 2 - δ) variables {E : Type*} @@ -47,36 +47,36 @@ section seminormed_add_comm_group variables (E) [seminormed_add_comm_group E] [uniform_convex_space E] {ε : ℝ} lemma exists_forall_sphere_dist_add_le_two_sub (hε : 0 < ε) : - ∃ δ, 0 < δ ∧ ∀ ⦃x : E⦄, ∥x∥ = 1 → ∀ ⦃y⦄, ∥y∥ = 1 → ε ≤ ∥x - y∥ → ∥x + y∥ ≤ 2 - δ := + ∃ δ, 0 < δ ∧ ∀ ⦃x : E⦄, ‖x‖ = 1 → ∀ ⦃y⦄, ‖y‖ = 1 → ε ≤ ‖x - y‖ → ‖x + y‖ ≤ 2 - δ := uniform_convex_space.uniform_convex hε variables [normed_space ℝ E] lemma exists_forall_closed_ball_dist_add_le_two_sub (hε : 0 < ε) : - ∃ δ, 0 < δ ∧ ∀ ⦃x : E⦄, ∥x∥ ≤ 1 → ∀ ⦃y⦄, ∥y∥ ≤ 1 → ε ≤ ∥x - y∥ → ∥x + y∥ ≤ 2 - δ := + ∃ δ, 0 < δ ∧ ∀ ⦃x : E⦄, ‖x‖ ≤ 1 → ∀ ⦃y⦄, ‖y‖ ≤ 1 → ε ≤ ‖x - y‖ → ‖x + y‖ ≤ 2 - δ := begin have hε' : 0 < ε / 3 := div_pos hε zero_lt_three, obtain ⟨δ, hδ, h⟩ := exists_forall_sphere_dist_add_le_two_sub E hε', set δ' := min (1/2) (min (ε/3) $ δ/3), refine ⟨δ', lt_min one_half_pos $ lt_min hε' (div_pos hδ zero_lt_three), λ x hx y hy hxy, _⟩, - obtain hx' | hx' := le_or_lt (∥x∥) (1 - δ'), + obtain hx' | hx' := le_or_lt (‖x‖) (1 - δ'), { exact (norm_add_le_of_le hx' hy).trans (sub_add_eq_add_sub _ _ _).le }, - obtain hy' | hy' := le_or_lt (∥y∥) (1 - δ'), + obtain hy' | hy' := le_or_lt (‖y‖) (1 - δ'), { exact (norm_add_le_of_le hx hy').trans (add_sub_assoc _ _ _).ge }, have hδ' : 0 < 1 - δ' := sub_pos_of_lt (min_lt_of_left_lt one_half_lt_one), - have h₁ : ∀ z : E, 1 - δ' < ∥z∥ → ∥∥z∥⁻¹ • z∥ = 1, + have h₁ : ∀ z : E, 1 - δ' < ‖z‖ → ‖‖z‖⁻¹ • z‖ = 1, { rintro z hz, rw [norm_smul_of_nonneg (inv_nonneg.2 $ norm_nonneg _), inv_mul_cancel (hδ'.trans hz).ne'] }, - have h₂ : ∀ z : E, ∥z∥ ≤ 1 → 1 - δ' ≤ ∥z∥ → ∥∥z∥⁻¹ • z - z∥ ≤ δ', + have h₂ : ∀ z : E, ‖z‖ ≤ 1 → 1 - δ' ≤ ‖z‖ → ‖‖z‖⁻¹ • z - z‖ ≤ δ', { rintro z hz hδz, nth_rewrite 2 ←one_smul ℝ z, rwa [←sub_smul, norm_smul_of_nonneg (sub_nonneg_of_le $ one_le_inv (hδ'.trans_le hδz) hz), sub_mul, inv_mul_cancel (hδ'.trans_le hδz).ne', one_mul, sub_le_comm] }, - set x' := ∥x∥⁻¹ • x, - set y' := ∥y∥⁻¹ • y, - have hxy' : ε/3 ≤ ∥x' - y'∥ := + set x' := ‖x‖⁻¹ • x, + set y' := ‖y‖⁻¹ • y, + have hxy' : ε/3 ≤ ‖x' - y'‖ := calc ε/3 = ε - (ε/3 + ε/3) : by ring - ... ≤ ∥x - y∥ - (∥x' - x∥ + ∥y' - y∥) : sub_le_sub hxy (add_le_add + ... ≤ ‖x - y‖ - (‖x' - x‖ + ‖y' - y‖) : sub_le_sub hxy (add_le_add ((h₂ _ hx hx'.le).trans $ min_le_of_right_le $ min_le_left _ _) $ (h₂ _ hy hy'.le).trans $ min_le_of_right_le $ min_le_left _ _) ... ≤ _ : begin @@ -84,7 +84,7 @@ begin rw [sub_le_iff_le_add, norm_sub_rev _ x, ←add_assoc, this], exact norm_add₃_le _ _ _, end, - calc ∥x + y∥ ≤ ∥x' + y'∥ + ∥x' - x∥ + ∥y' - y∥ : begin + calc ‖x + y‖ ≤ ‖x' + y'‖ + ‖x' - x‖ + ‖y' - y‖ : begin have : ∀ x' y', x + y = x' + y' + (x - x') + (y - y') := λ _ _, by abel, rw [norm_sub_rev, norm_sub_rev y', this], exact norm_add₃_le _ _ _, @@ -101,7 +101,7 @@ begin end lemma exists_forall_closed_ball_dist_add_le_two_mul_sub (hε : 0 < ε) (r : ℝ) : - ∃ δ, 0 < δ ∧ ∀ ⦃x : E⦄, ∥x∥ ≤ r → ∀ ⦃y⦄, ∥y∥ ≤ r → ε ≤ ∥x - y∥ → ∥x + y∥ ≤ 2 * r - δ := + ∃ δ, 0 < δ ∧ ∀ ⦃x : E⦄, ‖x‖ ≤ r → ∀ ⦃y⦄, ‖y‖ ≤ r → ε ≤ ‖x - y‖ → ‖x + y‖ ≤ 2 * r - δ := begin obtain hr | hr := le_or_lt r 0, { exact ⟨1, one_pos, λ x hx y hy h, (hε.not_le $ h.trans $ (norm_sub_le _ _).trans $ diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index 1a48310740f6e..a1e9a93302f95 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -103,7 +103,7 @@ variables [add_group G] [topological_space G] lemma has_compact_support.convolution_integrand_bound_right (hcg : has_compact_support g) (hg : continuous g) {x t : G} {s : set G} (hx : x ∈ s) : - ∥L (f t) (g (x - t))∥ ≤ (- tsupport g + s).indicator (λ t, ∥L∥ * ∥f t∥ * (⨆ i, ∥g i∥)) t := + ‖L (f t) (g (x - t))‖ ≤ (- tsupport g + s).indicator (λ t, ‖L‖ * ‖f t‖ * (⨆ i, ‖g i‖)) t := begin refine le_indicator (λ t ht, _) (λ t ht, _) t, { refine (L.le_op_norm₂ _ _).trans _, @@ -122,7 +122,7 @@ L.continuous₂.comp₂ continuous_const $ hg.comp $ continuous_id.sub continuou lemma has_compact_support.convolution_integrand_bound_left (hcf : has_compact_support f) (hf : continuous f) {x t : G} {s : set G} (hx : x ∈ s) : - ∥L (f (x - t)) (g t)∥ ≤ (- tsupport f + s).indicator (λ t, ∥L∥ * (⨆ i, ∥f i∥) * ∥g t∥) t := + ‖L (f (x - t)) (g t)‖ ≤ (- tsupport f + s).indicator (λ t, ‖L‖ * (⨆ i, ‖f i‖) * ‖g t‖) t := by { convert hcf.convolution_integrand_bound_right L.flip hf hx, simp_rw [L.op_norm_flip, mul_right_comm] } @@ -182,7 +182,7 @@ integrable on the support of the integrand, and that both functions are strongly Note: we could weaken the measurability condition to hold only for `μ.restrict s`. -/ lemma bdd_above.convolution_exists_at' {x₀ : G} - {s : set G} (hbg : bdd_above ((λ i, ∥g i∥) '' ((λ t, - t + x₀) ⁻¹' s))) + {s : set G} (hbg : bdd_above ((λ i, ‖g i‖) '' ((λ t, - t + x₀) ⁻¹' s))) (hs : measurable_set s) (h2s : support (λ t, L (f t) (g (x₀ - t))) ⊆ s) (hf : integrable_on f s μ) (hmf : ae_strongly_measurable f μ) @@ -191,7 +191,7 @@ lemma bdd_above.convolution_exists_at' {x₀ : G} begin set s' := (λ t, - t + x₀) ⁻¹' s, have : ∀ᵐ (t : G) ∂μ, - ∥L (f t) (g (x₀ - t))∥ ≤ s.indicator (λ t, ∥L∥ * ∥f t∥ * ⨆ i : s', ∥g i∥) t, + ‖L (f t) (g (x₀ - t))‖ ≤ s.indicator (λ t, ‖L‖ * ‖f t‖ * ⨆ i : s', ‖g i‖) t, { refine eventually_of_forall _, refine le_indicator (λ t ht, _) (λ t ht, _), { refine (L.le_op_norm₂ _ _).trans _, @@ -206,14 +206,14 @@ begin { exact hmf.convolution_integrand_snd' L hmg } end -/-- If `∥f∥ *[μ] ∥g∥` exists, then `f *[L, μ] g` exists. -/ +/-- If `‖f‖ *[μ] ‖g‖` exists, then `f *[L, μ] g` exists. -/ lemma convolution_exists_at.of_norm' {x₀ : G} - (h : convolution_exists_at (λ x, ∥f x∥) (λ x, ∥g x∥) x₀ (mul ℝ ℝ) μ) + (h : convolution_exists_at (λ x, ‖f x‖) (λ x, ‖g x‖) x₀ (mul ℝ ℝ) μ) (hmf : ae_strongly_measurable f μ) (hmg : ae_strongly_measurable g $ map (λ t, x₀ - t) μ) : convolution_exists_at f g x₀ L μ := begin - refine (h.const_mul ∥L∥).mono' (hmf.convolution_integrand_snd' L hmg) + refine (h.const_mul ‖L‖).mono' (hmf.convolution_integrand_snd' L hmg) (eventually_of_forall $ λ x, _), rw [mul_apply', ← mul_assoc], apply L.le_op_norm₂, @@ -234,9 +234,9 @@ lemma measure_theory.ae_strongly_measurable.convolution_integrand_swap_snd (hf.mono' (quasi_measure_preserving_sub_left_of_right_invariant μ x).absolutely_continuous) .convolution_integrand_swap_snd' L hg -/-- If `∥f∥ *[μ] ∥g∥` exists, then `f *[L, μ] g` exists. -/ +/-- If `‖f‖ *[μ] ‖g‖` exists, then `f *[L, μ] g` exists. -/ lemma convolution_exists_at.of_norm {x₀ : G} - (h : convolution_exists_at (λ x, ∥f x∥) (λ x, ∥g x∥) x₀ (mul ℝ ℝ) μ) + (h : convolution_exists_at (λ x, ‖f x‖) (λ x, ‖g x‖) x₀ (mul ℝ ℝ) μ) (hmf : ae_strongly_measurable f μ) (hmg : ae_strongly_measurable g μ) : convolution_exists_at f g x₀ L μ := @@ -260,13 +260,13 @@ lemma measure_theory.integrable.convolution_integrand (hf : integrable f ν) (hg begin have h_meas : ae_strongly_measurable (λ (p : G × G), L (f p.2) (g (p.1 - p.2))) (μ.prod ν) := hf.ae_strongly_measurable.convolution_integrand L hg.ae_strongly_measurable, - have h2_meas : ae_strongly_measurable (λ (y : G), ∫ (x : G), ∥L (f y) (g (x - y))∥ ∂μ) ν := + have h2_meas : ae_strongly_measurable (λ (y : G), ∫ (x : G), ‖L (f y) (g (x - y))‖ ∂μ) ν := h_meas.prod_swap.norm.integral_prod_right', simp_rw [integrable_prod_iff' h_meas], refine ⟨eventually_of_forall (λ t, (L (f t)).integrable_comp (hg.comp_sub_right t)), _⟩, refine integrable.mono' _ h2_meas (eventually_of_forall $ - λ t, (_ : _ ≤ ∥L∥ * ∥f t∥ * ∫ x, ∥g (x - t)∥ ∂μ)), - { simp_rw [integral_sub_right_eq_self (λ t, ∥ g t ∥)], + λ t, (_ : _ ≤ ‖L‖ * ‖f t‖ * ∫ x, ‖g (x - t)‖ ∂μ)), + { simp_rw [integral_sub_right_eq_self (λ t, ‖ g t ‖)], exact (hf.norm.const_mul _).mul_const _ }, { simp_rw [← integral_mul_left], rw [real.norm_of_nonneg], @@ -332,7 +332,7 @@ integrable on the support of the integrand, and that both functions are strongly This is a variant of `bdd_above.convolution_exists_at'` in an abelian group with a left-invariant measure. This allows us to state the boundedness and measurability of `g` in a more natural way. -/ lemma bdd_above.convolution_exists_at [sigma_finite μ] {x₀ : G} - {s : set G} (hbg : bdd_above ((λ i, ∥g i∥) '' ((λ t, x₀ - t) ⁻¹' s))) + {s : set G} (hbg : bdd_above ((λ i, ‖g i‖) '' ((λ t, x₀ - t) ⁻¹' s))) (hs : measurable_set s) (h2s : support (λ t, L (f t) (g (x₀ - t))) ⊆ s) (hf : integrable_on f s μ) (hmf : ae_strongly_measurable f μ) @@ -502,7 +502,7 @@ begin let K' := - tsupport g + K, have hK' : is_compact K' := hcg.neg.add hK, have : ∀ᶠ x in 𝓝 x₀, ∀ᵐ (t : G) ∂μ, - ∥L (f t) (g (x - t))∥ ≤ K'.indicator (λ t, ∥L∥ * ∥f t∥ * (⨆ i, ∥g i∥)) t := + ‖L (f t) (g (x - t))‖ ≤ K'.indicator (λ t, ‖L‖ * ‖f t‖ * (⨆ i, ‖g i‖)) t := eventually_of_mem h2K (λ x hx, eventually_of_forall $ λ t, hcg.convolution_integrand_bound_right L hg hx), refine continuous_at_of_dominated _ this _ _, @@ -517,12 +517,12 @@ end /-- The convolution is continuous if one function is integrable and the other is bounded and continuous. -/ lemma bdd_above.continuous_convolution_right_of_integrable - (hbg : bdd_above (range (λ x, ∥g x∥))) (hf : integrable f μ) (hg : continuous g) : + (hbg : bdd_above (range (λ x, ‖g x‖))) (hf : integrable f μ) (hg : continuous g) : continuous (f ⋆[L, μ] g) := begin refine continuous_iff_continuous_at.mpr (λ x₀, _), have : ∀ᶠ x in 𝓝 x₀, ∀ᵐ (t : G) ∂μ, - ∥L (f t) (g (x - t))∥ ≤ ∥L∥ * ∥f t∥ * (⨆ i, ∥g i∥), + ‖L (f t) (g (x - t))‖ ≤ ‖L‖ * ‖f t‖ * (⨆ i, ‖g i‖), { refine eventually_of_forall (λ x, eventually_of_forall $ λ t, _), refine (L.le_op_norm₂ _ _).trans _, exact mul_le_mul_of_nonneg_left (le_csupr hbg $ x - t) @@ -608,7 +608,7 @@ lemma has_compact_support.continuous_convolution_left [locally_compact_space G] by { rw [← convolution_flip], exact hcf.continuous_convolution_right L.flip hg hf } lemma bdd_above.continuous_convolution_left_of_integrable - (hbf : bdd_above (range (λ x, ∥f x∥))) (hf : continuous f) (hg : integrable g μ) : + (hbf : bdd_above (range (λ x, ‖f x‖))) (hf : continuous f) (hg : integrable g μ) : continuous (f ⋆[L, μ] g) := by { rw [← convolution_flip], exact hbf.continuous_convolution_right_of_integrable L.flip hg hf } @@ -660,18 +660,18 @@ lemma dist_convolution_le' {x₀ : G} {R ε : ℝ} {z₀ : E'} (hf : support f ⊆ ball (0 : G) R) (hmg : ae_strongly_measurable g μ) (hg : ∀ x ∈ ball x₀ R, dist (g x) z₀ ≤ ε) : - dist ((f ⋆[L, μ] g : G → F) x₀) (∫ t, L (f t) z₀ ∂μ) ≤ ∥L∥ * ∫ x, ∥f x∥ ∂μ * ε := + dist ((f ⋆[L, μ] g : G → F) x₀) (∫ t, L (f t) z₀ ∂μ) ≤ ‖L‖ * ∫ x, ‖f x‖ ∂μ * ε := begin have hfg : convolution_exists_at f g x₀ L μ, { refine bdd_above.convolution_exists_at L _ metric.is_open_ball.measurable_set (subset_trans _ hf) hif.integrable_on hif.ae_strongly_measurable hmg, swap, { refine λ t, mt (λ ht : f t = 0, _), simp_rw [ht, L.map_zero₂] }, rw [bdd_above_def], - refine ⟨∥z₀∥ + ε, _⟩, + refine ⟨‖z₀‖ + ε, _⟩, rintro _ ⟨x, hx, rfl⟩, refine norm_le_norm_add_const_of_dist_le (hg x _), rwa [mem_ball_iff_norm, norm_sub_rev, ← mem_ball_zero_iff] }, - have h2 : ∀ t, dist (L (f t) (g (x₀ - t))) (L (f t) z₀) ≤ ∥L (f t)∥ * ε, + have h2 : ∀ t, dist (L (f t) (g (x₀ - t))) (L (f t) z₀) ≤ ‖L (f t)‖ * ε, { intro t, by_cases ht : t ∈ support f, { have h2t := hf ht, rw [mem_ball_zero_iff] at h2t, @@ -688,7 +688,7 @@ begin (eventually_of_forall h2)).trans _, rw [integral_mul_right], refine mul_le_mul_of_nonneg_right _ hε, - have h3 : ∀ t, ∥L (f t)∥ ≤ ∥L∥ * ∥f t∥, + have h3 : ∀ t, ‖L (f t)‖ ≤ ‖L‖ * ‖f t‖, { intros t, exact L.le_op_norm (f t) }, refine (integral_mono (L.integrable_comp hif).norm (hif.norm.const_mul _) h3).trans_eq _, @@ -894,16 +894,16 @@ calc ((f ⋆[L, ν] g) ⋆[L₂, μ] k) x₀ /-- Convolution is associative. This requires that * all maps are a.e. strongly measurable w.r.t one of the measures * `f ⋆[L, ν] g` exists almost everywhere -* `∥g∥ ⋆[μ] ∥k∥` exists almost everywhere -* `∥f∥ ⋆[ν] (∥g∥ ⋆[μ] ∥k∥)` exists at `x₀` -/ +* `‖g‖ ⋆[μ] ‖k‖` exists almost everywhere +* `‖f‖ ⋆[ν] (‖g‖ ⋆[μ] ‖k‖)` exists at `x₀` -/ lemma convolution_assoc (hL : ∀ (x : E) (y : E') (z : E''), L₂ (L x y) z = L₃ x (L₄ y z)) {x₀ : G} (hf : ae_strongly_measurable f ν) (hg : ae_strongly_measurable g μ) (hk : ae_strongly_measurable k μ) (hfg : ∀ᵐ y ∂μ, convolution_exists_at f g y L ν) - (hgk : ∀ᵐ x ∂ν, convolution_exists_at (λ x, ∥g x∥) (λ x, ∥k x∥) x (mul ℝ ℝ) μ) - (hfgk : convolution_exists_at (λ x, ∥f x∥) ((λ x, ∥g x∥) ⋆[mul ℝ ℝ, μ] (λ x, ∥k x∥)) + (hgk : ∀ᵐ x ∂ν, convolution_exists_at (λ x, ‖g x‖) (λ x, ‖k x‖) x (mul ℝ ℝ) μ) + (hfgk : convolution_exists_at (λ x, ‖f x‖) ((λ x, ‖g x‖) ⋆[mul ℝ ℝ, μ] (λ x, ‖k x‖)) x₀ (mul ℝ ℝ) ν) : ((f ⋆[L, ν] g) ⋆[L₂, μ] k) x₀ = (f ⋆[L₃, ν] (g ⋆[L₄, μ] k)) x₀ := begin @@ -919,7 +919,7 @@ begin ((measurable_const.sub measurable_snd).sub measurable_fst) (eventually_of_forall $ λ y, _), dsimp only, exact quasi_measure_preserving_sub_left_of_right_invariant μ _ }, - have h2_meas : ae_strongly_measurable (λ y, ∫ x, ∥L₃ (f y) (L₄ (g x) (k (x₀ - y - x)))∥ ∂μ) ν := + have h2_meas : ae_strongly_measurable (λ y, ∫ x, ‖L₃ (f y) (L₄ (g x) (k (x₀ - y - x)))‖ ∂μ) ν := h_meas.prod_swap.norm.integral_prod_right', have h3 : map (λ z : G × G, (z.1 - z.2, z.2)) (μ.prod ν) = μ.prod ν := (measure_preserving_sub_prod μ ν).map_eq, @@ -931,9 +931,9 @@ begin simp_rw [integrable_prod_iff' h_meas], refine ⟨((quasi_measure_preserving_sub_left_of_right_invariant ν x₀).ae hgk).mono (λ t ht, (L₃ (f t)).integrable_comp $ ht.of_norm L₄ hg hk), _⟩, - refine (hfgk.const_mul (∥L₃∥ * ∥L₄∥)).mono' h2_meas + refine (hfgk.const_mul (‖L₃‖ * ‖L₄‖)).mono' h2_meas (((quasi_measure_preserving_sub_left_of_right_invariant ν x₀).ae hgk).mono $ λ t ht, _), - { simp_rw [convolution_def, mul_apply', mul_mul_mul_comm ∥L₃∥ ∥L₄∥, ← integral_mul_left], + { simp_rw [convolution_def, mul_apply', mul_mul_mul_comm ‖L₃‖ ‖L₄‖, ← integral_mul_left], rw [real.norm_of_nonneg], { refine integral_mono_of_nonneg (eventually_of_forall $ λ t, norm_nonneg _) ((ht.const_mul _).const_mul _) (eventually_of_forall $ λ s, _), @@ -992,7 +992,7 @@ begin have hK' : is_compact K' := (hcg.fderiv 𝕜).neg.add (is_compact_closed_ball x₀ 1), refine has_fderiv_at_integral_of_dominated_of_fderiv_le zero_lt_one h1 _ (h2 x₀) _ _ _, - { exact K'.indicator (λ t, ∥L'∥ * ∥f t∥ * (⨆ x, ∥fderiv 𝕜 g x∥)) }, + { exact K'.indicator (λ t, ‖L'‖ * ‖f t‖ * (⨆ x, ‖fderiv 𝕜 g x‖)) }, { exact hcg.convolution_exists_right L hf hg.continuous x₀ }, { refine eventually_of_forall (λ t x hx, _), exact (hcg.fderiv 𝕜).convolution_integrand_bound_right L' diff --git a/src/analysis/fourier.lean b/src/analysis/fourier.lean index 8c568d14f027d..8a1f3b608c08c 100644 --- a/src/analysis/fourier.lean +++ b/src/analysis/fourier.lean @@ -242,12 +242,12 @@ by simpa using hilbert_basis.has_sum_repr fourier_series f /-- **Parseval's identity**: the sum of the squared norms of the Fourier coefficients equals the `L2` norm of the function. -/ lemma tsum_sq_fourier_series_repr (f : Lp ℂ 2 haar_circle) : - ∑' i : ℤ, ∥fourier_series.repr f i∥ ^ 2 = ∫ t : circle, ∥f t∥ ^ 2 ∂ haar_circle := + ∑' i : ℤ, ‖fourier_series.repr f i‖ ^ 2 = ∫ t : circle, ‖f t‖ ^ 2 ∂ haar_circle := begin - have H₁ : ∥fourier_series.repr f∥ ^ 2 = ∑' i, ∥fourier_series.repr f i∥ ^ 2, + have H₁ : ‖fourier_series.repr f‖ ^ 2 = ∑' i, ‖fourier_series.repr f i‖ ^ 2, { exact_mod_cast lp.norm_rpow_eq_tsum _ (fourier_series.repr f), norm_num }, - have H₂ : ∥fourier_series.repr f∥ ^ 2 = ∥f∥ ^2 := by simp, + have H₂ : ‖fourier_series.repr f‖ ^ 2 = ‖f‖ ^2 := by simp, have H₃ := congr_arg is_R_or_C.re (@L2.inner_def circle ℂ ℂ _ _ _ _ f f), rw ← integral_re at H₃, { simp only [← norm_sq_eq_inner] at H₃, diff --git a/src/analysis/inner_product_space/adjoint.lean b/src/analysis/inner_product_space/adjoint.lean index e03af80401fc1..044b0de10850d 100644 --- a/src/analysis/inner_product_space/adjoint.lean +++ b/src/analysis/inner_product_space/adjoint.lean @@ -81,7 +81,7 @@ begin rw [adjoint_aux_inner_right, adjoint_aux_inner_left], end -@[simp] lemma adjoint_aux_norm (A : E →L[𝕜] F) : ∥adjoint_aux A∥ = ∥A∥ := +@[simp] lemma adjoint_aux_norm (A : E →L[𝕜] F) : ‖adjoint_aux A‖ = ‖A‖ := begin refine le_antisymm _ _, { refine continuous_linear_map.op_norm_le_bound _ (norm_nonneg _) (λ x, _), @@ -123,20 +123,20 @@ begin simp only [adjoint_inner_right, continuous_linear_map.coe_comp', function.comp_app], end -lemma apply_norm_sq_eq_inner_adjoint_left (A : E →L[𝕜] E) (x : E) : ∥A x∥^2 = re ⟪(A† * A) x, x⟫ := +lemma apply_norm_sq_eq_inner_adjoint_left (A : E →L[𝕜] E) (x : E) : ‖A x‖^2 = re ⟪(A† * A) x, x⟫ := have h : ⟪(A† * A) x, x⟫ = ⟪A x, A x⟫ := by { rw [←adjoint_inner_left], refl }, by rw [h, ←inner_self_eq_norm_sq _] lemma apply_norm_eq_sqrt_inner_adjoint_left (A : E →L[𝕜] E) (x : E) : - ∥A x∥ = real.sqrt (re ⟪(A† * A) x, x⟫) := + ‖A x‖ = real.sqrt (re ⟪(A† * A) x, x⟫) := by rw [←apply_norm_sq_eq_inner_adjoint_left, real.sqrt_sq (norm_nonneg _)] -lemma apply_norm_sq_eq_inner_adjoint_right (A : E →L[𝕜] E) (x : E) : ∥A x∥^2 = re ⟪x, (A† * A) x⟫ := +lemma apply_norm_sq_eq_inner_adjoint_right (A : E →L[𝕜] E) (x : E) : ‖A x‖^2 = re ⟪x, (A† * A) x⟫ := have h : ⟪x, (A† * A) x⟫ = ⟪A x, A x⟫ := by { rw [←adjoint_inner_right], refl }, by rw [h, ←inner_self_eq_norm_sq _] lemma apply_norm_eq_sqrt_inner_adjoint_right (A : E →L[𝕜] E) (x : E) : - ∥A x∥ = real.sqrt (re ⟪x, (A† * A) x⟫) := + ‖A x‖ = real.sqrt (re ⟪x, (A† * A) x⟫) := by rw [←apply_norm_sq_eq_inner_adjoint_right, real.sqrt_sq (norm_nonneg _)] /-- The adjoint is unique: a map `A` is the adjoint of `B` iff it satisfies `⟪A x, y⟫ = ⟪x, B y⟫` @@ -190,17 +190,17 @@ instance : cstar_ring (E →L[𝕜] E) := intros A, rw [star_eq_adjoint], refine le_antisymm _ _, - { calc ∥A† * A∥ ≤ ∥A†∥ * ∥A∥ : op_norm_comp_le _ _ - ... = ∥A∥ * ∥A∥ : by rw [linear_isometry_equiv.norm_map] }, + { calc ‖A† * A‖ ≤ ‖A†‖ * ‖A‖ : op_norm_comp_le _ _ + ... = ‖A‖ * ‖A‖ : by rw [linear_isometry_equiv.norm_map] }, { rw [←sq, ←real.sqrt_le_sqrt_iff (norm_nonneg _), real.sqrt_sq (norm_nonneg _)], refine op_norm_le_bound _ (real.sqrt_nonneg _) (λ x, _), have := calc - re ⟪(A† * A) x, x⟫ ≤ ∥(A† * A) x∥ * ∥x∥ : re_inner_le_norm _ _ - ... ≤ ∥A† * A∥ * ∥x∥ * ∥x∥ : mul_le_mul_of_nonneg_right + re ⟪(A† * A) x, x⟫ ≤ ‖(A† * A) x‖ * ‖x‖ : re_inner_le_norm _ _ + ... ≤ ‖A† * A‖ * ‖x‖ * ‖x‖ : mul_le_mul_of_nonneg_right (le_op_norm _ _) (norm_nonneg _), - calc ∥A x∥ = real.sqrt (re ⟪(A† * A) x, x⟫) : by rw [apply_norm_eq_sqrt_inner_adjoint_left] - ... ≤ real.sqrt (∥A† * A∥ * ∥x∥ * ∥x∥) : real.sqrt_le_sqrt this - ... = real.sqrt (∥A† * A∥) * ∥x∥ + calc ‖A x‖ = real.sqrt (re ⟪(A† * A) x, x⟫) : by rw [apply_norm_eq_sqrt_inner_adjoint_left] + ... ≤ real.sqrt (‖A† * A‖ * ‖x‖ * ‖x‖) : real.sqrt_le_sqrt this + ... = real.sqrt (‖A† * A‖) * ‖x‖ : by rw [mul_assoc, real.sqrt_mul (norm_nonneg _), real.sqrt_mul_self (norm_nonneg _)] } end⟩ diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index c97e33017d0b3..329833a78ef01 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -93,14 +93,14 @@ end notations /-- An inner product space is a vector space with an additional operation called inner product. The norm could be derived from the inner product, instead we require the existence of a norm and -the fact that `∥x∥^2 = re ⟪x, x⟫` to be able to put instances on `𝕂` or product +the fact that `‖x‖^2 = re ⟪x, x⟫` to be able to put instances on `𝕂` or product spaces. To construct a norm from an inner product, see `inner_product_space.of_core`. -/ class inner_product_space (𝕜 : Type*) (E : Type*) [is_R_or_C 𝕜] extends normed_add_comm_group E, normed_space 𝕜 E, has_inner 𝕜 E := -(norm_sq_eq_inner : ∀ (x : E), ∥x∥^2 = re (inner x x)) +(norm_sq_eq_inner : ∀ (x : E), ‖x‖^2 = re (inner x x)) (conj_sym : ∀ x y, conj (inner y x) = inner x y) (add_left : ∀ x y z, inner (x + y) z = inner x z + inner y z) (smul_left : ∀ x y r, inner (r • x) y = (conj r) * inner x y) @@ -298,19 +298,19 @@ def to_has_norm : has_norm F := local attribute [instance] to_has_norm -lemma norm_eq_sqrt_inner (x : F) : ∥x∥ = sqrt (re ⟪x, x⟫) := rfl +lemma norm_eq_sqrt_inner (x : F) : ‖x‖ = sqrt (re ⟪x, x⟫) := rfl -lemma inner_self_eq_norm_mul_norm (x : F) : re ⟪x, x⟫ = ∥x∥ * ∥x∥ := +lemma inner_self_eq_norm_mul_norm (x : F) : re ⟪x, x⟫ = ‖x‖ * ‖x‖ := by rw [norm_eq_sqrt_inner, ←sqrt_mul inner_self_nonneg (re ⟪x, x⟫), sqrt_mul_self inner_self_nonneg] -lemma sqrt_norm_sq_eq_norm {x : F} : sqrt (norm_sqF x) = ∥x∥ := rfl +lemma sqrt_norm_sq_eq_norm {x : F} : sqrt (norm_sqF x) = ‖x‖ := rfl /-- Cauchy–Schwarz inequality with norm -/ -lemma abs_inner_le_norm (x y : F) : abs ⟪x, y⟫ ≤ ∥x∥ * ∥y∥ := +lemma abs_inner_le_norm (x y : F) : abs ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := nonneg_le_nonneg_of_sq_le_sq (mul_nonneg (sqrt_nonneg _) (sqrt_nonneg _)) begin - have H : ∥x∥ * ∥y∥ * (∥x∥ * ∥y∥) = re ⟪y, y⟫ * re ⟪x, x⟫, + have H : ‖x‖ * ‖y‖ * (‖x‖ * ‖y‖) = re ⟪y, y⟫ * re ⟪x, x⟫, { simp only [inner_self_eq_norm_mul_norm], ring, }, rw H, conv @@ -327,11 +327,11 @@ add_group_norm.to_normed_add_comm_group map_zero' := by simp only [sqrt_zero, inner_zero_right, map_zero], neg' := λ x, by simp only [inner_neg_left, neg_neg, inner_neg_right], add_le' := λ x y, begin - have h₁ : abs ⟪x, y⟫ ≤ ∥x∥ * ∥y∥ := abs_inner_le_norm _ _, + have h₁ : abs ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := abs_inner_le_norm _ _, have h₂ : re ⟪x, y⟫ ≤ abs ⟪x, y⟫ := re_le_abs _, - have h₃ : re ⟪x, y⟫ ≤ ∥x∥ * ∥y∥ := by linarith, - have h₄ : re ⟪y, x⟫ ≤ ∥x∥ * ∥y∥ := by rwa [←inner_conj_sym, conj_re], - have : ∥x + y∥ * ∥x + y∥ ≤ (∥x∥ + ∥y∥) * (∥x∥ + ∥y∥), + have h₃ : re ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := by linarith, + have h₄ : re ⟪y, x⟫ ≤ ‖x‖ * ‖y‖ := by rwa [←inner_conj_sym, conj_re], + have : ‖x + y‖ * ‖x + y‖ ≤ (‖x‖ + ‖y‖) * (‖x‖ + ‖y‖), { simp only [←inner_self_eq_norm_mul_norm, inner_add_add_self, mul_add, mul_comm, map_add], linarith }, exact nonneg_le_nonneg_of_sq_le_sq (add_nonneg (sqrt_nonneg _) (sqrt_nonneg _)) this, @@ -367,7 +367,7 @@ begin letI : normed_space 𝕜 F := @inner_product_space.of_core.to_normed_space 𝕜 F _ _ _ c, exact { norm_sq_eq_inner := λ x, begin - have h₁ : ∥x∥^2 = (sqrt (re (c.inner x x))) ^ 2 := rfl, + have h₁ : ‖x‖^2 = (sqrt (re (c.inner x x))) ^ 2 := rfl, have h₂ : 0 ≤ re (c.inner x x) := inner_product_space.of_core.inner_self_nonneg, simp [h₁, sq_sqrt, h₂], end, @@ -516,9 +516,9 @@ by { have h := @inner_self_nonpos ℝ F _ _ x, simpa using h } @[simp] lemma inner_self_re_to_K {x : E} : (re ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := by rw is_R_or_C.ext_iff; exact ⟨by simp, by simp [inner_self_nonneg_im]⟩ -lemma inner_self_eq_norm_sq_to_K (x : E) : ⟪x, x⟫ = (∥x∥ ^ 2 : 𝕜) := +lemma inner_self_eq_norm_sq_to_K (x : E) : ⟪x, x⟫ = (‖x‖ ^ 2 : 𝕜) := begin - suffices : (is_R_or_C.re ⟪x, x⟫ : 𝕜) = ∥x∥ ^ 2, + suffices : (is_R_or_C.re ⟪x, x⟫ : 𝕜) = ‖x‖ ^ 2, { simpa [inner_self_re_to_K] using this }, exact_mod_cast (norm_sq_eq_inner x).symm end @@ -688,7 +688,7 @@ include 𝕜 /-- An orthonormal set of vectors in an `inner_product_space` -/ def orthonormal (v : ι → E) : Prop := -(∀ i, ∥v i∥ = 1) ∧ (∀ {i j}, i ≠ j → ⟪v i, v j⟫ = 0) +(∀ i, ‖v i‖ = 1) ∧ (∀ {i j}, i ≠ j → ⟪v i, v j⟫ = 0) omit 𝕜 @@ -708,8 +708,8 @@ begin { intros h, split, { intros i, - have h' : ∥v i∥ ^ 2 = 1 ^ 2 := by simp [norm_sq_eq_inner, h i i], - have h₁ : 0 ≤ ∥v i∥ := norm_nonneg _, + have h' : ‖v i‖ ^ 2 = 1 ^ 2 := by simp [norm_sq_eq_inner, h i i], + have h₁ : 0 ≤ ‖v i‖ := norm_nonneg _, have h₂ : (0:ℝ) ≤ 1 := zero_le_one, rwa sq_eq_sq h₁ h₂ at h' }, { intros i j hij, @@ -915,7 +915,7 @@ end lemma orthonormal.ne_zero {v : ι → E} (hv : orthonormal 𝕜 v) (i : ι) : v i ≠ 0 := begin - have : ∥v i∥ ≠ 0, + have : ‖v i‖ ≠ 0, { rw hv.1 i, norm_num }, simpa using this @@ -938,31 +938,31 @@ end orthonormal_sets section norm -lemma norm_eq_sqrt_inner (x : E) : ∥x∥ = sqrt (re ⟪x, x⟫) := +lemma norm_eq_sqrt_inner (x : E) : ‖x‖ = sqrt (re ⟪x, x⟫) := begin - have h₁ : ∥x∥^2 = re ⟪x, x⟫ := norm_sq_eq_inner x, + have h₁ : ‖x‖^2 = re ⟪x, x⟫ := norm_sq_eq_inner x, have h₂ := congr_arg sqrt h₁, simpa using h₂, end -lemma norm_eq_sqrt_real_inner (x : F) : ∥x∥ = sqrt ⟪x, x⟫_ℝ := +lemma norm_eq_sqrt_real_inner (x : F) : ‖x‖ = sqrt ⟪x, x⟫_ℝ := by { have h := @norm_eq_sqrt_inner ℝ F _ _ x, simpa using h } -lemma inner_self_eq_norm_mul_norm (x : E) : re ⟪x, x⟫ = ∥x∥ * ∥x∥ := +lemma inner_self_eq_norm_mul_norm (x : E) : re ⟪x, x⟫ = ‖x‖ * ‖x‖ := by rw [norm_eq_sqrt_inner, ←sqrt_mul inner_self_nonneg (re ⟪x, x⟫), sqrt_mul_self inner_self_nonneg] -lemma inner_self_eq_norm_sq (x : E) : re ⟪x, x⟫ = ∥x∥^2 := +lemma inner_self_eq_norm_sq (x : E) : re ⟪x, x⟫ = ‖x‖^2 := by rw [pow_two, inner_self_eq_norm_mul_norm] -lemma real_inner_self_eq_norm_mul_norm (x : F) : ⟪x, x⟫_ℝ = ∥x∥ * ∥x∥ := +lemma real_inner_self_eq_norm_mul_norm (x : F) : ⟪x, x⟫_ℝ = ‖x‖ * ‖x‖ := by { have h := @inner_self_eq_norm_mul_norm ℝ F _ _ x, simpa using h } -lemma real_inner_self_eq_norm_sq (x : F) : ⟪x, x⟫_ℝ = ∥x∥^2 := +lemma real_inner_self_eq_norm_sq (x : F) : ⟪x, x⟫_ℝ = ‖x‖^2 := by rw [pow_two, real_inner_self_eq_norm_mul_norm] /-- Expand the square -/ -lemma norm_add_sq {x y : E} : ∥x + y∥^2 = ∥x∥^2 + 2 * (re ⟪x, y⟫) + ∥y∥^2 := +lemma norm_add_sq {x y : E} : ‖x + y‖^2 = ‖x‖^2 + 2 * (re ⟪x, y⟫) + ‖y‖^2 := begin repeat {rw [sq, ←inner_self_eq_norm_mul_norm]}, rw [inner_add_add_self, two_mul], @@ -973,21 +973,21 @@ end alias norm_add_sq ← norm_add_pow_two /-- Expand the square -/ -lemma norm_add_sq_real {x y : F} : ∥x + y∥^2 = ∥x∥^2 + 2 * ⟪x, y⟫_ℝ + ∥y∥^2 := +lemma norm_add_sq_real {x y : F} : ‖x + y‖^2 = ‖x‖^2 + 2 * ⟪x, y⟫_ℝ + ‖y‖^2 := by { have h := @norm_add_sq ℝ F _ _, simpa using h } alias norm_add_sq_real ← norm_add_pow_two_real /-- Expand the square -/ -lemma norm_add_mul_self {x y : E} : ∥x + y∥ * ∥x + y∥ = ∥x∥ * ∥x∥ + 2 * (re ⟪x, y⟫) + ∥y∥ * ∥y∥ := +lemma norm_add_mul_self {x y : E} : ‖x + y‖ * ‖x + y‖ = ‖x‖ * ‖x‖ + 2 * (re ⟪x, y⟫) + ‖y‖ * ‖y‖ := by { repeat {rw [← sq]}, exact norm_add_sq } /-- Expand the square -/ -lemma norm_add_mul_self_real {x y : F} : ∥x + y∥ * ∥x + y∥ = ∥x∥ * ∥x∥ + 2 * ⟪x, y⟫_ℝ + ∥y∥ * ∥y∥ := +lemma norm_add_mul_self_real {x y : F} : ‖x + y‖ * ‖x + y‖ = ‖x‖ * ‖x‖ + 2 * ⟪x, y⟫_ℝ + ‖y‖ * ‖y‖ := by { have h := @norm_add_mul_self ℝ F _ _, simpa using h } /-- Expand the square -/ -lemma norm_sub_sq {x y : E} : ∥x - y∥^2 = ∥x∥^2 - 2 * (re ⟪x, y⟫) + ∥y∥^2 := +lemma norm_sub_sq {x y : E} : ‖x - y‖^2 = ‖x‖^2 - 2 * (re ⟪x, y⟫) + ‖y‖^2 := begin repeat {rw [sq, ←inner_self_eq_norm_mul_norm]}, rw [inner_sub_sub_self], @@ -1003,50 +1003,50 @@ end alias norm_sub_sq ← norm_sub_pow_two /-- Expand the square -/ -lemma norm_sub_sq_real {x y : F} : ∥x - y∥^2 = ∥x∥^2 - 2 * ⟪x, y⟫_ℝ + ∥y∥^2 := +lemma norm_sub_sq_real {x y : F} : ‖x - y‖^2 = ‖x‖^2 - 2 * ⟪x, y⟫_ℝ + ‖y‖^2 := norm_sub_sq alias norm_sub_sq_real ← norm_sub_pow_two_real /-- Expand the square -/ -lemma norm_sub_mul_self {x y : E} : ∥x - y∥ * ∥x - y∥ = ∥x∥ * ∥x∥ - 2 * re ⟪x, y⟫ + ∥y∥ * ∥y∥ := +lemma norm_sub_mul_self {x y : E} : ‖x - y‖ * ‖x - y‖ = ‖x‖ * ‖x‖ - 2 * re ⟪x, y⟫ + ‖y‖ * ‖y‖ := by { repeat {rw [← sq]}, exact norm_sub_sq } /-- Expand the square -/ -lemma norm_sub_mul_self_real {x y : F} : ∥x - y∥ * ∥x - y∥ = ∥x∥ * ∥x∥ - 2 * ⟪x, y⟫_ℝ + ∥y∥ * ∥y∥ := +lemma norm_sub_mul_self_real {x y : F} : ‖x - y‖ * ‖x - y‖ = ‖x‖ * ‖x‖ - 2 * ⟪x, y⟫_ℝ + ‖y‖ * ‖y‖ := by { have h := @norm_sub_mul_self ℝ F _ _, simpa using h } /-- Cauchy–Schwarz inequality with norm -/ -lemma abs_inner_le_norm (x y : E) : abs ⟪x, y⟫ ≤ ∥x∥ * ∥y∥ := +lemma abs_inner_le_norm (x y : E) : abs ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := nonneg_le_nonneg_of_sq_le_sq (mul_nonneg (norm_nonneg _) (norm_nonneg _)) begin - have : ∥x∥ * ∥y∥ * (∥x∥ * ∥y∥) = (re ⟪x, x⟫) * (re ⟪y, y⟫), + have : ‖x‖ * ‖y‖ * (‖x‖ * ‖y‖) = (re ⟪x, x⟫) * (re ⟪y, y⟫), simp only [inner_self_eq_norm_mul_norm], ring, rw this, conv_lhs { congr, skip, rw [inner_abs_conj_sym] }, exact inner_mul_inner_self_le _ _ end -lemma norm_inner_le_norm (x y : E) : ∥⟪x, y⟫∥ ≤ ∥x∥ * ∥y∥ := +lemma norm_inner_le_norm (x y : E) : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := (is_R_or_C.norm_eq_abs _).le.trans (abs_inner_le_norm x y) -lemma nnnorm_inner_le_nnnorm (x y : E) : ∥⟪x, y⟫∥₊ ≤ ∥x∥₊ * ∥y∥₊ := +lemma nnnorm_inner_le_nnnorm (x y : E) : ‖⟪x, y⟫‖₊ ≤ ‖x‖₊ * ‖y‖₊ := norm_inner_le_norm x y -lemma re_inner_le_norm (x y : E) : re ⟪x, y⟫ ≤ ∥x∥ * ∥y∥ := +lemma re_inner_le_norm (x y : E) : re ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := le_trans (re_le_abs (inner x y)) (abs_inner_le_norm x y) /-- Cauchy–Schwarz inequality with norm -/ -lemma abs_real_inner_le_norm (x y : F) : absR ⟪x, y⟫_ℝ ≤ ∥x∥ * ∥y∥ := +lemma abs_real_inner_le_norm (x y : F) : absR ⟪x, y⟫_ℝ ≤ ‖x‖ * ‖y‖ := by { have h := @abs_inner_le_norm ℝ F _ _ x y, simpa using h } /-- Cauchy–Schwarz inequality with norm -/ -lemma real_inner_le_norm (x y : F) : ⟪x, y⟫_ℝ ≤ ∥x∥ * ∥y∥ := +lemma real_inner_le_norm (x y : F) : ⟪x, y⟫_ℝ ≤ ‖x‖ * ‖y‖ := le_trans (le_abs_self _) (abs_real_inner_le_norm _ _) include 𝕜 lemma parallelogram_law_with_norm (x y : E) : - ∥x + y∥ * ∥x + y∥ + ∥x - y∥ * ∥x - y∥ = 2 * (∥x∥ * ∥x∥ + ∥y∥ * ∥y∥) := + ‖x + y‖ * ‖x + y‖ + ‖x - y‖ * ‖x - y‖ = 2 * (‖x‖ * ‖x‖ + ‖y‖ * ‖y‖) := begin simp only [← inner_self_eq_norm_mul_norm], rw [← re.map_add, parallelogram_law, two_mul, two_mul], @@ -1054,34 +1054,34 @@ begin end lemma parallelogram_law_with_nnnorm (x y : E) : - ∥x + y∥₊ * ∥x + y∥₊ + ∥x - y∥₊ * ∥x - y∥₊ = 2 * (∥x∥₊ * ∥x∥₊ + ∥y∥₊ * ∥y∥₊) := + ‖x + y‖₊ * ‖x + y‖₊ + ‖x - y‖₊ * ‖x - y‖₊ = 2 * (‖x‖₊ * ‖x‖₊ + ‖y‖₊ * ‖y‖₊) := subtype.ext $ parallelogram_law_with_norm x y omit 𝕜 /-- Polarization identity: The real part of the inner product, in terms of the norm. -/ lemma re_inner_eq_norm_add_mul_self_sub_norm_mul_self_sub_norm_mul_self_div_two (x y : E) : - re ⟪x, y⟫ = (∥x + y∥ * ∥x + y∥ - ∥x∥ * ∥x∥ - ∥y∥ * ∥y∥) / 2 := + re ⟪x, y⟫ = (‖x + y‖ * ‖x + y‖ - ‖x‖ * ‖x‖ - ‖y‖ * ‖y‖) / 2 := by { rw norm_add_mul_self, ring } /-- Polarization identity: The real part of the inner product, in terms of the norm. -/ lemma re_inner_eq_norm_mul_self_add_norm_mul_self_sub_norm_sub_mul_self_div_two (x y : E) : - re ⟪x, y⟫ = (∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ - ∥x - y∥ * ∥x - y∥) / 2 := + re ⟪x, y⟫ = (‖x‖ * ‖x‖ + ‖y‖ * ‖y‖ - ‖x - y‖ * ‖x - y‖) / 2 := by { rw [norm_sub_mul_self], ring } /-- Polarization identity: The real part of the inner product, in terms of the norm. -/ lemma re_inner_eq_norm_add_mul_self_sub_norm_sub_mul_self_div_four (x y : E) : - re ⟪x, y⟫ = (∥x + y∥ * ∥x + y∥ - ∥x - y∥ * ∥x - y∥) / 4 := + re ⟪x, y⟫ = (‖x + y‖ * ‖x + y‖ - ‖x - y‖ * ‖x - y‖) / 4 := by { rw [norm_add_mul_self, norm_sub_mul_self], ring } /-- Polarization identity: The imaginary part of the inner product, in terms of the norm. -/ lemma im_inner_eq_norm_sub_I_smul_mul_self_sub_norm_add_I_smul_mul_self_div_four (x y : E) : - im ⟪x, y⟫ = (∥x - IK • y∥ * ∥x - IK • y∥ - ∥x + IK • y∥ * ∥x + IK • y∥) / 4 := + im ⟪x, y⟫ = (‖x - IK • y‖ * ‖x - IK • y‖ - ‖x + IK • y‖ * ‖x + IK • y‖) / 4 := by { simp only [norm_add_mul_self, norm_sub_mul_self, inner_smul_right, I_mul_re], ring } /-- Polarization identity: The inner product, in terms of the norm. -/ lemma inner_eq_sum_norm_sq_div_four (x y : E) : - ⟪x, y⟫ = (∥x + y∥ ^ 2 - ∥x - y∥ ^ 2 + (∥x - IK • y∥ ^ 2 - ∥x + IK • y∥ ^ 2) * IK) / 4 := + ⟪x, y⟫ = (‖x + y‖ ^ 2 - ‖x - y‖ ^ 2 + (‖x - IK • y‖ ^ 2 - ‖x + IK • y‖ ^ 2) * IK) / 4 := begin rw [← re_add_im ⟪x, y⟫, re_inner_eq_norm_add_mul_self_sub_norm_sub_mul_self_div_four, im_inner_eq_norm_sub_I_smul_mul_self_sub_norm_add_I_smul_mul_self_div_four], @@ -1093,16 +1093,16 @@ end zero. See also `euclidean_geometry.dist_inversion_inversion` for inversions around a general point. -/ lemma dist_div_norm_sq_smul {x y : F} (hx : x ≠ 0) (hy : y ≠ 0) (R : ℝ) : - dist ((R / ∥x∥) ^ 2 • x) ((R / ∥y∥) ^ 2 • y) = (R ^ 2 / (∥x∥ * ∥y∥)) * dist x y := -have hx' : ∥x∥ ≠ 0, from norm_ne_zero_iff.2 hx, -have hy' : ∥y∥ ≠ 0, from norm_ne_zero_iff.2 hy, -calc dist ((R / ∥x∥) ^ 2 • x) ((R / ∥y∥) ^ 2 • y) - = sqrt (∥(R / ∥x∥) ^ 2 • x - (R / ∥y∥) ^ 2 • y∥^2) : + dist ((R / ‖x‖) ^ 2 • x) ((R / ‖y‖) ^ 2 • y) = (R ^ 2 / (‖x‖ * ‖y‖)) * dist x y := +have hx' : ‖x‖ ≠ 0, from norm_ne_zero_iff.2 hx, +have hy' : ‖y‖ ≠ 0, from norm_ne_zero_iff.2 hy, +calc dist ((R / ‖x‖) ^ 2 • x) ((R / ‖y‖) ^ 2 • y) + = sqrt (‖(R / ‖x‖) ^ 2 • x - (R / ‖y‖) ^ 2 • y‖^2) : by rw [dist_eq_norm, sqrt_sq (norm_nonneg _)] -... = sqrt ((R ^ 2 / (∥x∥ * ∥y∥)) ^ 2 * ∥x - y∥ ^ 2) : +... = sqrt ((R ^ 2 / (‖x‖ * ‖y‖)) ^ 2 * ‖x - y‖ ^ 2) : congr_arg sqrt $ by { field_simp [sq, norm_sub_mul_self_real, norm_smul, real_inner_smul_left, inner_smul_right, real.norm_of_nonneg (mul_self_nonneg _)], ring } -... = (R ^ 2 / (∥x∥ * ∥y∥)) * dist x y : +... = (R ^ 2 / (‖x‖ * ‖y‖)) * dist x y : by rw [sqrt_mul (sq_nonneg _), sqrt_sq (norm_nonneg _), sqrt_sq (div_nonneg (sq_nonneg _) (mul_nonneg (norm_nonneg _) (norm_nonneg _))), dist_eq_norm] @@ -1114,7 +1114,7 @@ instance inner_product_space.to_uniform_convex_space : uniform_convex_space F := exact pow_pos hε _ }, rw sub_sub_cancel, refine le_sqrt_of_sq_le _, - rw [sq, eq_sub_iff_add_eq.2 (parallelogram_law_with_norm x y), ←sq (∥x - y∥), hx, hy], + rw [sq, eq_sub_iff_add_eq.2 (parallelogram_law_with_norm x y), ←sq (‖x - y‖), hx, hy], norm_num, exact pow_le_pow_of_le_left hε.le hxy _, end⟩ @@ -1320,19 +1320,19 @@ end /-- Polarization identity: The real inner product, in terms of the norm. -/ lemma real_inner_eq_norm_add_mul_self_sub_norm_mul_self_sub_norm_mul_self_div_two (x y : F) : - ⟪x, y⟫_ℝ = (∥x + y∥ * ∥x + y∥ - ∥x∥ * ∥x∥ - ∥y∥ * ∥y∥) / 2 := + ⟪x, y⟫_ℝ = (‖x + y‖ * ‖x + y‖ - ‖x‖ * ‖x‖ - ‖y‖ * ‖y‖) / 2 := re_to_real.symm.trans $ re_inner_eq_norm_add_mul_self_sub_norm_mul_self_sub_norm_mul_self_div_two x y /-- Polarization identity: The real inner product, in terms of the norm. -/ lemma real_inner_eq_norm_mul_self_add_norm_mul_self_sub_norm_sub_mul_self_div_two (x y : F) : - ⟪x, y⟫_ℝ = (∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ - ∥x - y∥ * ∥x - y∥) / 2 := + ⟪x, y⟫_ℝ = (‖x‖ * ‖x‖ + ‖y‖ * ‖y‖ - ‖x - y‖ * ‖x - y‖) / 2 := re_to_real.symm.trans $ re_inner_eq_norm_mul_self_add_norm_mul_self_sub_norm_sub_mul_self_div_two x y /-- Pythagorean theorem, if-and-only-if vector inner product form. -/ lemma norm_add_sq_eq_norm_sq_add_norm_sq_iff_real_inner_eq_zero (x y : F) : - ∥x + y∥ * ∥x + y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ ↔ ⟪x, y⟫_ℝ = 0 := + ‖x + y‖ * ‖x + y‖ = ‖x‖ * ‖x‖ + ‖y‖ * ‖y‖ ↔ ⟪x, y⟫_ℝ = 0 := begin rw [norm_add_mul_self, add_right_cancel_iff, add_right_eq_self, mul_eq_zero], norm_num @@ -1340,13 +1340,13 @@ end /-- Pythagorean theorem, if-and-if vector inner product form using square roots. -/ lemma norm_add_eq_sqrt_iff_real_inner_eq_zero {x y : F} : - ∥x + y∥ = sqrt (∥x∥ * ∥x∥ + ∥y∥ * ∥y∥) ↔ ⟪x, y⟫_ℝ = 0 := + ‖x + y‖ = sqrt (‖x‖ * ‖x‖ + ‖y‖ * ‖y‖) ↔ ⟪x, y⟫_ℝ = 0 := by rw [←norm_add_sq_eq_norm_sq_add_norm_sq_iff_real_inner_eq_zero, eq_comm, sqrt_eq_iff_mul_self_eq (add_nonneg (mul_self_nonneg _) (mul_self_nonneg _)) (norm_nonneg _)] /-- Pythagorean theorem, vector inner product form. -/ lemma norm_add_sq_eq_norm_sq_add_norm_sq_of_inner_eq_zero (x y : E) (h : ⟪x, y⟫ = 0) : - ∥x + y∥ * ∥x + y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ := + ‖x + y‖ * ‖x + y‖ = ‖x‖ * ‖x‖ + ‖y‖ * ‖y‖ := begin rw [norm_add_mul_self, add_right_cancel_iff, add_right_eq_self, mul_eq_zero], apply or.inr, @@ -1355,13 +1355,13 @@ end /-- Pythagorean theorem, vector inner product form. -/ lemma norm_add_sq_eq_norm_sq_add_norm_sq_real {x y : F} (h : ⟪x, y⟫_ℝ = 0) : - ∥x + y∥ * ∥x + y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ := + ‖x + y‖ * ‖x + y‖ = ‖x‖ * ‖x‖ + ‖y‖ * ‖y‖ := (norm_add_sq_eq_norm_sq_add_norm_sq_iff_real_inner_eq_zero x y).2 h /-- Pythagorean theorem, subtracting vectors, if-and-only-if vector inner product form. -/ lemma norm_sub_sq_eq_norm_sq_add_norm_sq_iff_real_inner_eq_zero (x y : F) : - ∥x - y∥ * ∥x - y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ ↔ ⟪x, y⟫_ℝ = 0 := + ‖x - y‖ * ‖x - y‖ = ‖x‖ * ‖x‖ + ‖y‖ * ‖y‖ ↔ ⟪x, y⟫_ℝ = 0 := begin rw [norm_sub_mul_self, add_right_cancel_iff, sub_eq_add_neg, add_right_eq_self, neg_eq_zero, mul_eq_zero], @@ -1371,19 +1371,19 @@ end /-- Pythagorean theorem, subtracting vectors, if-and-if vector inner product form using square roots. -/ lemma norm_sub_eq_sqrt_iff_real_inner_eq_zero {x y : F} : - ∥x - y∥ = sqrt (∥x∥ * ∥x∥ + ∥y∥ * ∥y∥) ↔ ⟪x, y⟫_ℝ = 0 := + ‖x - y‖ = sqrt (‖x‖ * ‖x‖ + ‖y‖ * ‖y‖) ↔ ⟪x, y⟫_ℝ = 0 := by rw [←norm_sub_sq_eq_norm_sq_add_norm_sq_iff_real_inner_eq_zero, eq_comm, sqrt_eq_iff_mul_self_eq (add_nonneg (mul_self_nonneg _) (mul_self_nonneg _)) (norm_nonneg _)] /-- Pythagorean theorem, subtracting vectors, vector inner product form. -/ lemma norm_sub_sq_eq_norm_sq_add_norm_sq_real {x y : F} (h : ⟪x, y⟫_ℝ = 0) : - ∥x - y∥ * ∥x - y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ := + ‖x - y‖ * ‖x - y‖ = ‖x‖ * ‖x‖ + ‖y‖ * ‖y‖ := (norm_sub_sq_eq_norm_sq_add_norm_sq_iff_real_inner_eq_zero x y).2 h /-- The sum and difference of two vectors are orthogonal if and only if they have the same norm. -/ -lemma real_inner_add_sub_eq_zero_iff (x y : F) : ⟪x + y, x - y⟫_ℝ = 0 ↔ ∥x∥ = ∥y∥ := +lemma real_inner_add_sub_eq_zero_iff (x y : F) : ⟪x + y, x - y⟫_ℝ = 0 ↔ ‖x‖ = ‖y‖ := begin conv_rhs { rw ←mul_self_inj_of_nonneg (norm_nonneg _) (norm_nonneg _) }, simp only [←inner_self_eq_norm_mul_norm, inner_add_left, inner_sub_right, @@ -1397,7 +1397,7 @@ begin end /-- Given two orthogonal vectors, their sum and difference have equal norms. -/ -lemma norm_sub_eq_norm_add {v w : E} (h : ⟪v, w⟫ = 0) : ∥w - v∥ = ∥w + v∥ := +lemma norm_sub_eq_norm_add {v w : E} (h : ⟪v, w⟫ = 0) : ‖w - v‖ = ‖w + v‖ := begin rw ←mul_self_inj_of_nonneg (norm_nonneg _) (norm_nonneg _), simp [h, ←inner_self_eq_norm_mul_norm, inner_add_left, inner_add_right, inner_sub_left, @@ -1406,34 +1406,34 @@ end /-- The real inner product of two vectors, divided by the product of their norms, has absolute value at most 1. -/ -lemma abs_real_inner_div_norm_mul_norm_le_one (x y : F) : absR (⟪x, y⟫_ℝ / (∥x∥ * ∥y∥)) ≤ 1 := +lemma abs_real_inner_div_norm_mul_norm_le_one (x y : F) : absR (⟪x, y⟫_ℝ / (‖x‖ * ‖y‖)) ≤ 1 := begin rw _root_.abs_div, - by_cases h : 0 = absR (∥x∥ * ∥y∥), + by_cases h : 0 = absR (‖x‖ * ‖y‖), { rw [←h, div_zero], norm_num }, - { change 0 ≠ absR (∥x∥ * ∥y∥) at h, - rw div_le_iff' (lt_of_le_of_ne (ge_iff_le.mp (_root_.abs_nonneg (∥x∥ * ∥y∥))) h), + { change 0 ≠ absR (‖x‖ * ‖y‖) at h, + rw div_le_iff' (lt_of_le_of_ne (ge_iff_le.mp (_root_.abs_nonneg (‖x‖ * ‖y‖))) h), convert abs_real_inner_le_norm x y using 1, rw [_root_.abs_mul, _root_.abs_of_nonneg (norm_nonneg x), _root_.abs_of_nonneg (norm_nonneg y), mul_one] } end /-- The inner product of a vector with a multiple of itself. -/ -lemma real_inner_smul_self_left (x : F) (r : ℝ) : ⟪r • x, x⟫_ℝ = r * (∥x∥ * ∥x∥) := +lemma real_inner_smul_self_left (x : F) (r : ℝ) : ⟪r • x, x⟫_ℝ = r * (‖x‖ * ‖x‖) := by rw [real_inner_smul_left, ←real_inner_self_eq_norm_mul_norm] /-- The inner product of a vector with a multiple of itself. -/ -lemma real_inner_smul_self_right (x : F) (r : ℝ) : ⟪x, r • x⟫_ℝ = r * (∥x∥ * ∥x∥) := +lemma real_inner_smul_self_right (x : F) (r : ℝ) : ⟪x, r • x⟫_ℝ = r * (‖x‖ * ‖x‖) := by rw [inner_smul_right, ←real_inner_self_eq_norm_mul_norm] /-- The inner product of a nonzero vector with a nonzero multiple of itself, divided by the product of their norms, has absolute value 1. -/ lemma abs_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul - {x : E} {r : 𝕜} (hx : x ≠ 0) (hr : r ≠ 0) : abs ⟪x, r • x⟫ / (∥x∥ * ∥r • x∥) = 1 := + {x : E} {r : 𝕜} (hx : x ≠ 0) (hr : r ≠ 0) : abs ⟪x, r • x⟫ / (‖x‖ * ‖r • x‖) = 1 := begin - have hx' : ∥x∥ ≠ 0 := by simp [norm_eq_zero, hx], + have hx' : ‖x‖ ≠ 0 := by simp [norm_eq_zero, hx], have hr' : abs r ≠ 0 := by simp [is_R_or_C.abs_eq_zero, hr], rw [inner_smul_right, is_R_or_C.abs_mul, ←inner_self_re_abs, inner_self_eq_norm_mul_norm, norm_smul], @@ -1445,7 +1445,7 @@ end itself, divided by the product of their norms, has absolute value 1. -/ lemma abs_real_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul - {x : F} {r : ℝ} (hx : x ≠ 0) (hr : r ≠ 0) : absR ⟪x, r • x⟫_ℝ / (∥x∥ * ∥r • x∥) = 1 := + {x : F} {r : ℝ} (hx : x ≠ 0) (hr : r ≠ 0) : absR ⟪x, r • x⟫_ℝ / (‖x‖ * ‖r • x‖) = 1 := begin rw ← abs_to_real, exact abs_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul hx hr @@ -1454,9 +1454,9 @@ end /-- The inner product of a nonzero vector with a positive multiple of itself, divided by the product of their norms, has value 1. -/ lemma real_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_pos_mul - {x : F} {r : ℝ} (hx : x ≠ 0) (hr : 0 < r) : ⟪x, r • x⟫_ℝ / (∥x∥ * ∥r • x∥) = 1 := + {x : F} {r : ℝ} (hx : x ≠ 0) (hr : 0 < r) : ⟪x, r • x⟫_ℝ / (‖x‖ * ‖r • x‖) = 1 := begin - rw [real_inner_smul_self_right, norm_smul, real.norm_eq_abs, ←mul_assoc ∥x∥, mul_comm _ (absR r), + rw [real_inner_smul_self_right, norm_smul, real.norm_eq_abs, ←mul_assoc ‖x‖, mul_comm _ (absR r), mul_assoc, _root_.abs_of_nonneg (le_of_lt hr), div_self], exact mul_ne_zero (ne_of_gt hr) (λ h, hx (norm_eq_zero.1 (eq_zero_of_mul_self_eq_zero h))) @@ -1465,9 +1465,9 @@ end /-- The inner product of a nonzero vector with a negative multiple of itself, divided by the product of their norms, has value -1. -/ lemma real_inner_div_norm_mul_norm_eq_neg_one_of_ne_zero_of_neg_mul - {x : F} {r : ℝ} (hx : x ≠ 0) (hr : r < 0) : ⟪x, r • x⟫_ℝ / (∥x∥ * ∥r • x∥) = -1 := + {x : F} {r : ℝ} (hx : x ≠ 0) (hr : r < 0) : ⟪x, r • x⟫_ℝ / (‖x‖ * ‖r • x‖) = -1 := begin - rw [real_inner_smul_self_right, norm_smul, real.norm_eq_abs, ←mul_assoc ∥x∥, mul_comm _ (absR r), + rw [real_inner_smul_self_right, norm_smul, real.norm_eq_abs, ←mul_assoc ‖x‖, mul_comm _ (absR r), mul_assoc, abs_of_neg hr, neg_mul, div_neg_eq_neg_div, div_self], exact mul_ne_zero (ne_of_lt hr) (λ h, hx (norm_eq_zero.1 (eq_zero_of_mul_self_eq_zero h))) @@ -1477,7 +1477,7 @@ end norms, has absolute value 1 if and only if they are nonzero and one is a multiple of the other. One form of equality case for Cauchy-Schwarz. -/ lemma abs_inner_div_norm_mul_norm_eq_one_iff (x y : E) : - abs (⟪x, y⟫ / (∥x∥ * ∥y∥)) = 1 ↔ (x ≠ 0 ∧ ∃ (r : 𝕜), r ≠ 0 ∧ y = r • x) := + abs (⟪x, y⟫ / (‖x‖ * ‖y‖)) = 1 ↔ (x ≠ 0 ∧ ∃ (r : 𝕜), r ≠ 0 ∧ y = r • x) := begin split, { intro h, @@ -1486,7 +1486,7 @@ begin rw [hx0, inner_zero_left, zero_div] at h, norm_num at h, }, refine and.intro hx0 _, - set r := ⟪x, y⟫ / (∥x∥ * ∥x∥) with hr, + set r := ⟪x, y⟫ / (‖x‖ * ‖x‖) with hr, use r, set t := y - r • x with ht, have ht0 : ⟪x, t⟫ = 0, @@ -1494,7 +1494,7 @@ begin norm_cast, rw [←inner_self_eq_norm_mul_norm, inner_self_re_to_K, div_mul_cancel _ (λ h, hx0 (inner_self_eq_zero.1 h)), sub_self] }, - replace h : ∥r • x∥ / ∥t + r • x∥ = 1, + replace h : ‖r • x‖ / ‖t + r • x‖ = 1, { rw [←sub_add_cancel y (r • x), ←ht, inner_add_right, ht0, zero_add, inner_smul_right, is_R_or_C.abs_div, is_R_or_C.abs_mul, ←inner_self_re_abs, inner_self_eq_norm_mul_norm] at h, @@ -1507,7 +1507,7 @@ begin rw [hr0, zero_smul, norm_zero, zero_div] at h, norm_num at h }, refine and.intro hr0 _, - have h2 : ∥r • x∥ ^ 2 = ∥t + r • x∥ ^ 2, + have h2 : ‖r • x‖ ^ 2 = ‖t + r • x‖ ^ 2, { rw [eq_of_div_eq_one h] }, replace h2 : ⟪r • x, r • x⟫ = ⟪t, t⟫ + ⟪t, r • x⟫ + ⟪r • x, t⟫ + ⟪r • x, r • x⟫, { rw [sq, sq, ←inner_self_eq_norm_mul_norm, ←inner_self_eq_norm_mul_norm ] at h2, @@ -1532,7 +1532,7 @@ end norms, has absolute value 1 if and only if they are nonzero and one is a multiple of the other. One form of equality case for Cauchy-Schwarz. -/ lemma abs_real_inner_div_norm_mul_norm_eq_one_iff (x y : F) : - absR (⟪x, y⟫_ℝ / (∥x∥ * ∥y∥)) = 1 ↔ (x ≠ 0 ∧ ∃ (r : ℝ), r ≠ 0 ∧ y = r • x) := + absR (⟪x, y⟫_ℝ / (‖x‖ * ‖y‖)) = 1 ↔ (x ≠ 0 ∧ ∃ (r : ℝ), r ≠ 0 ∧ y = r • x) := begin have := @abs_inner_div_norm_mul_norm_eq_one_iff ℝ F _ _ x y, simpa [coe_real_eq_id] using this, @@ -1541,14 +1541,14 @@ end /-- If the inner product of two vectors is equal to the product of their norms, then the two vectors are multiples of each other. One form of the equality case for Cauchy-Schwarz. -Compare `inner_eq_norm_mul_iff`, which takes the stronger hypothesis `⟪x, y⟫ = ∥x∥ * ∥y∥`. -/ +Compare `inner_eq_norm_mul_iff`, which takes the stronger hypothesis `⟪x, y⟫ = ‖x‖ * ‖y‖`. -/ lemma abs_inner_eq_norm_iff (x y : E) (hx0 : x ≠ 0) (hy0 : y ≠ 0): - abs ⟪x, y⟫ = ∥x∥ * ∥y∥ ↔ ∃ (r : 𝕜), r ≠ 0 ∧ y = r • x := + abs ⟪x, y⟫ = ‖x‖ * ‖y‖ ↔ ∃ (r : 𝕜), r ≠ 0 ∧ y = r • x := begin - have hx0' : ∥x∥ ≠ 0 := by simp [norm_eq_zero, hx0], - have hy0' : ∥y∥ ≠ 0 := by simp [norm_eq_zero, hy0], - have hxy0 : ∥x∥ * ∥y∥ ≠ 0 := by simp [hx0', hy0'], - have h₁ : abs ⟪x, y⟫ = ∥x∥ * ∥y∥ ↔ abs (⟪x, y⟫ / (∥x∥ * ∥y∥)) = 1, + have hx0' : ‖x‖ ≠ 0 := by simp [norm_eq_zero, hx0], + have hy0' : ‖y‖ ≠ 0 := by simp [norm_eq_zero, hy0], + have hxy0 : ‖x‖ * ‖y‖ ≠ 0 := by simp [hx0', hy0'], + have h₁ : abs ⟪x, y⟫ = ‖x‖ * ‖y‖ ↔ abs (⟪x, y⟫ / (‖x‖ * ‖y‖)) = 1, { refine ⟨_ ,_⟩, { intro h, norm_cast, @@ -1566,7 +1566,7 @@ end norms, has value 1 if and only if they are nonzero and one is a positive multiple of the other. -/ lemma real_inner_div_norm_mul_norm_eq_one_iff (x y : F) : - ⟪x, y⟫_ℝ / (∥x∥ * ∥y∥) = 1 ↔ (x ≠ 0 ∧ ∃ (r : ℝ), 0 < r ∧ y = r • x) := + ⟪x, y⟫_ℝ / (‖x‖ * ‖y‖) = 1 ↔ (x ≠ 0 ∧ ∃ (r : ℝ), 0 < r ∧ y = r • x) := begin split, { intro h, @@ -1591,7 +1591,7 @@ end norms, has value -1 if and only if they are nonzero and one is a negative multiple of the other. -/ lemma real_inner_div_norm_mul_norm_eq_neg_one_iff (x y : F) : - ⟪x, y⟫_ℝ / (∥x∥ * ∥y∥) = -1 ↔ (x ≠ 0 ∧ ∃ (r : ℝ), r < 0 ∧ y = r • x) := + ⟪x, y⟫_ℝ / (‖x‖ * ‖y‖) = -1 ↔ (x ≠ 0 ∧ ∃ (r : ℝ), r < 0 ∧ y = r • x) := begin split, { intro h, @@ -1613,15 +1613,15 @@ begin end /-- If the inner product of two vectors is equal to the product of their norms (i.e., -`⟪x, y⟫ = ∥x∥ * ∥y∥`), then the two vectors are nonnegative real multiples of each other. One form +`⟪x, y⟫ = ‖x‖ * ‖y‖`), then the two vectors are nonnegative real multiples of each other. One form of the equality case for Cauchy-Schwarz. -Compare `abs_inner_eq_norm_iff`, which takes the weaker hypothesis `abs ⟪x, y⟫ = ∥x∥ * ∥y∥`. -/ +Compare `abs_inner_eq_norm_iff`, which takes the weaker hypothesis `abs ⟪x, y⟫ = ‖x‖ * ‖y‖`. -/ lemma inner_eq_norm_mul_iff {x y : E} : - ⟪x, y⟫ = (∥x∥ : 𝕜) * ∥y∥ ↔ (∥y∥ : 𝕜) • x = (∥x∥ : 𝕜) • y := + ⟪x, y⟫ = (‖x‖ : 𝕜) * ‖y‖ ↔ (‖y‖ : 𝕜) • x = (‖x‖ : 𝕜) • y := begin by_cases h : (x = 0 ∨ y = 0), -- WLOG `x` and `y` are nonzero { cases h; simp [h] }, - calc ⟪x, y⟫ = (∥x∥ : 𝕜) * ∥y∥ ↔ ∥x∥ * ∥y∥ = re ⟪x, y⟫ : + calc ⟪x, y⟫ = (‖x‖ : 𝕜) * ‖y‖ ↔ ‖x‖ * ‖y‖ = re ⟪x, y⟫ : begin norm_cast, split, @@ -1632,40 +1632,40 @@ begin rw h' at ⊢ cauchy_schwarz, rwa re_eq_self_of_le } end - ... ↔ 2 * ∥x∥ * ∥y∥ * (∥x∥ * ∥y∥ - re ⟪x, y⟫) = 0 : + ... ↔ 2 * ‖x‖ * ‖y‖ * (‖x‖ * ‖y‖ - re ⟪x, y⟫) = 0 : by simp [h, show (2:ℝ) ≠ 0, by norm_num, sub_eq_zero] - ... ↔ ∥(∥y∥:𝕜) • x - (∥x∥:𝕜) • y∥ * ∥(∥y∥:𝕜) • x - (∥x∥:𝕜) • y∥ = 0 : + ... ↔ ‖(‖y‖:𝕜) • x - (‖x‖:𝕜) • y‖ * ‖(‖y‖:𝕜) • x - (‖x‖:𝕜) • y‖ = 0 : begin simp only [norm_sub_mul_self, inner_smul_left, inner_smul_right, norm_smul, conj_of_real, is_R_or_C.norm_eq_abs, abs_of_real, of_real_im, of_real_re, mul_re, abs_norm_eq_norm], refine eq.congr _ rfl, ring end - ... ↔ (∥y∥ : 𝕜) • x = (∥x∥ : 𝕜) • y : by simp [norm_sub_eq_zero_iff] + ... ↔ (‖y‖ : 𝕜) • x = (‖x‖ : 𝕜) • y : by simp [norm_sub_eq_zero_iff] end /-- If the inner product of two vectors is equal to the product of their norms (i.e., -`⟪x, y⟫ = ∥x∥ * ∥y∥`), then the two vectors are nonnegative real multiples of each other. One form +`⟪x, y⟫ = ‖x‖ * ‖y‖`), then the two vectors are nonnegative real multiples of each other. One form of the equality case for Cauchy-Schwarz. -Compare `abs_inner_eq_norm_iff`, which takes the weaker hypothesis `abs ⟪x, y⟫ = ∥x∥ * ∥y∥`. -/ -lemma inner_eq_norm_mul_iff_real {x y : F} : ⟪x, y⟫_ℝ = ∥x∥ * ∥y∥ ↔ ∥y∥ • x = ∥x∥ • y := +Compare `abs_inner_eq_norm_iff`, which takes the weaker hypothesis `abs ⟪x, y⟫ = ‖x‖ * ‖y‖`. -/ +lemma inner_eq_norm_mul_iff_real {x y : F} : ⟪x, y⟫_ℝ = ‖x‖ * ‖y‖ ↔ ‖y‖ • x = ‖x‖ • y := inner_eq_norm_mul_iff /-- If the inner product of two unit vectors is `1`, then the two vectors are equal. One form of the equality case for Cauchy-Schwarz. -/ -lemma inner_eq_norm_mul_iff_of_norm_one {x y : E} (hx : ∥x∥ = 1) (hy : ∥y∥ = 1) : +lemma inner_eq_norm_mul_iff_of_norm_one {x y : E} (hx : ‖x‖ = 1) (hy : ‖y‖ = 1) : ⟪x, y⟫ = 1 ↔ x = y := by { convert inner_eq_norm_mul_iff using 2; simp [hx, hy] } lemma inner_lt_norm_mul_iff_real {x y : F} : - ⟪x, y⟫_ℝ < ∥x∥ * ∥y∥ ↔ ∥y∥ • x ≠ ∥x∥ • y := -calc ⟪x, y⟫_ℝ < ∥x∥ * ∥y∥ - ↔ ⟪x, y⟫_ℝ ≠ ∥x∥ * ∥y∥ : ⟨ne_of_lt, lt_of_le_of_ne (real_inner_le_norm _ _)⟩ -... ↔ ∥y∥ • x ≠ ∥x∥ • y : not_congr inner_eq_norm_mul_iff_real + ⟪x, y⟫_ℝ < ‖x‖ * ‖y‖ ↔ ‖y‖ • x ≠ ‖x‖ • y := +calc ⟪x, y⟫_ℝ < ‖x‖ * ‖y‖ + ↔ ⟪x, y⟫_ℝ ≠ ‖x‖ * ‖y‖ : ⟨ne_of_lt, lt_of_le_of_ne (real_inner_le_norm _ _)⟩ +... ↔ ‖y‖ • x ≠ ‖x‖ • y : not_congr inner_eq_norm_mul_iff_real /-- If the inner product of two unit vectors is strictly less than `1`, then the two vectors are distinct. One form of the equality case for Cauchy-Schwarz. -/ -lemma inner_lt_one_iff_real_of_norm_one {x y : F} (hx : ∥x∥ = 1) (hy : ∥y∥ = 1) : +lemma inner_lt_one_iff_real_of_norm_one {x y : F} (hx : ‖x‖ = 1) (hy : ‖y‖ = 1) : ⟪x, y⟫_ℝ < 1 ↔ x ≠ y := by { convert inner_lt_norm_mul_iff_real; simp [hx, hy] } @@ -1675,7 +1675,7 @@ lemma inner_sum_smul_sum_smul_of_sum_eq_zero {ι₁ : Type*} {s₁ : finset ι (v₁ : ι₁ → F) (h₁ : ∑ i in s₁, w₁ i = 0) {ι₂ : Type*} {s₂ : finset ι₂} {w₂ : ι₂ → ℝ} (v₂ : ι₂ → F) (h₂ : ∑ i in s₂, w₂ i = 0) : ⟪(∑ i₁ in s₁, w₁ i₁ • v₁ i₁), (∑ i₂ in s₂, w₂ i₂ • v₂ i₂)⟫_ℝ = - (-∑ i₁ in s₁, ∑ i₂ in s₂, w₁ i₁ * w₂ i₂ * (∥v₁ i₁ - v₂ i₂∥ * ∥v₁ i₁ - v₂ i₂∥)) / 2 := + (-∑ i₁ in s₁, ∑ i₂ in s₂, w₁ i₁ * w₂ i₂ * (‖v₁ i₁ - v₂ i₂‖ * ‖v₁ i₁ - v₂ i₂‖)) / 2 := by simp_rw [sum_inner, inner_sum, real_inner_smul_left, real_inner_smul_right, real_inner_eq_norm_mul_self_add_norm_mul_self_sub_norm_sub_mul_self_div_two, ←div_sub_div_same, ←div_add_div_same, mul_sub_left_distrib, left_distrib, @@ -1705,7 +1705,7 @@ linear_map.mk_continuous₂ innerₛₗ 1 /-- `innerSL` is an isometry. Note that the associated `linear_isometry` is defined in `inner_product_space.dual` as `to_dual_map`. -/ -@[simp] lemma innerSL_apply_norm {x : E} : ∥(innerSL x : E →L[𝕜] 𝕜)∥ = ∥x∥ := +@[simp] lemma innerSL_apply_norm {x : E} : ‖(innerSL x : E →L[𝕜] 𝕜)‖ = ‖x‖ := begin refine le_antisymm ((innerSL x : E →L[𝕜] 𝕜).op_norm_le_bound (norm_nonneg _) (λ y, norm_inner_le_norm _ _)) _, @@ -1713,11 +1713,11 @@ begin { have : x = 0 := norm_eq_zero.mp (eq.symm h), simp [this] }, { refine (mul_le_mul_right h).mp _, - calc ∥x∥ * ∥x∥ = ∥x∥ ^ 2 : by ring + calc ‖x‖ * ‖x‖ = ‖x‖ ^ 2 : by ring ... = re ⟪x, x⟫ : norm_sq_eq_inner _ ... ≤ abs ⟪x, x⟫ : re_le_abs _ - ... = ∥innerSL x x∥ : by { rw [←is_R_or_C.norm_eq_abs], refl } - ... ≤ ∥innerSL x∥ * ∥x∥ : (innerSL x : E →L[𝕜] 𝕜).le_op_norm _ } + ... = ‖innerSL x x‖ : by { rw [←is_R_or_C.norm_eq_abs], refl } + ... ≤ ‖innerSL x‖ * ‖x‖ : (innerSL x : E →L[𝕜] 𝕜).le_op_norm _ } end /-- The inner product as a continuous sesquilinear map, with the two arguments flipped. -/ @@ -1741,15 +1741,15 @@ def to_sesq_form : (E →L[𝕜] E') →L[𝕜] E' →L⋆[𝕜] E →L[𝕜] @[simp] lemma to_sesq_form_apply_coe (f : E →L[𝕜] E') (x : E') : to_sesq_form f x = (innerSL x).comp f := rfl -lemma to_sesq_form_apply_norm_le {f : E →L[𝕜] E'} {v : E'} : ∥to_sesq_form f v∥ ≤ ∥f∥ * ∥v∥ := +lemma to_sesq_form_apply_norm_le {f : E →L[𝕜] E'} {v : E'} : ‖to_sesq_form f v‖ ≤ ‖f‖ * ‖v‖ := begin refine op_norm_le_bound _ (mul_nonneg (norm_nonneg _) (norm_nonneg _)) _, intro x, - have h₁ : ∥f x∥ ≤ ∥f∥ * ∥x∥ := le_op_norm _ _, + have h₁ : ‖f x‖ ≤ ‖f‖ * ‖x‖ := le_op_norm _ _, have h₂ := @norm_inner_le_norm 𝕜 E' _ _ v (f x), - calc ∥⟪v, f x⟫∥ ≤ ∥v∥ * ∥f x∥ : h₂ - ... ≤ ∥v∥ * (∥f∥ * ∥x∥) : mul_le_mul_of_nonneg_left h₁ (norm_nonneg v) - ... = ∥f∥ * ∥v∥ * ∥x∥ : by ring, + calc ‖⟪v, f x⟫‖ ≤ ‖v‖ * ‖f x‖ : h₂ + ... ≤ ‖v‖ * (‖f‖ * ‖x‖) : mul_le_mul_of_nonneg_left h₁ (norm_nonneg v) + ... = ‖f‖ * ‖v‖ * ‖x‖ : by ring, end end continuous_linear_map @@ -1781,16 +1781,16 @@ variables {ι: Type*} (x : E) {v : ι → E} /-- Bessel's inequality for finite sums. -/ lemma orthonormal.sum_inner_products_le {s : finset ι} (hv : orthonormal 𝕜 v) : - ∑ i in s, ∥⟪v i, x⟫∥ ^ 2 ≤ ∥x∥ ^ 2 := + ∑ i in s, ‖⟪v i, x⟫‖ ^ 2 ≤ ‖x‖ ^ 2 := begin have h₂ : ∑ i in s, ∑ j in s, ⟪v i, x⟫ * ⟪x, v j⟫ * ⟪v j, v i⟫ = (∑ k in s, (⟪v k, x⟫ * ⟪x, v k⟫) : 𝕜), { exact hv.inner_left_right_finset }, - have h₃ : ∀ z : 𝕜, re (z * conj (z)) = ∥z∥ ^ 2, + have h₃ : ∀ z : 𝕜, re (z * conj (z)) = ‖z‖ ^ 2, { intro z, simp only [mul_conj, norm_sq_eq_def'], norm_cast, }, - suffices hbf: ∥x - ∑ i in s, ⟪v i, x⟫ • (v i)∥ ^ 2 = ∥x∥ ^ 2 - ∑ i in s, ∥⟪v i, x⟫∥ ^ 2, + suffices hbf: ‖x - ∑ i in s, ⟪v i, x⟫ • (v i)‖ ^ 2 = ‖x‖ ^ 2 - ∑ i in s, ‖⟪v i, x⟫‖ ^ 2, { rw [←sub_nonneg, ←hbf], simp only [norm_nonneg, pow_nonneg], }, rw [norm_sub_sq, sub_add], @@ -1802,21 +1802,21 @@ end /-- Bessel's inequality. -/ lemma orthonormal.tsum_inner_products_le (hv : orthonormal 𝕜 v) : - ∑' i, ∥⟪v i, x⟫∥ ^ 2 ≤ ∥x∥ ^ 2 := + ∑' i, ‖⟪v i, x⟫‖ ^ 2 ≤ ‖x‖ ^ 2 := begin refine tsum_le_of_sum_le' _ (λ s, hv.sum_inner_products_le x), simp only [norm_nonneg, pow_nonneg] end /-- The sum defined in Bessel's inequality is summable. -/ -lemma orthonormal.inner_products_summable (hv : orthonormal 𝕜 v) : summable (λ i, ∥⟪v i, x⟫∥ ^ 2) := +lemma orthonormal.inner_products_summable (hv : orthonormal 𝕜 v) : summable (λ i, ‖⟪v i, x⟫‖ ^ 2) := begin - use ⨆ s : finset ι, ∑ i in s, ∥⟪v i, x⟫∥ ^ 2, + use ⨆ s : finset ι, ∑ i in s, ‖⟪v i, x⟫‖ ^ 2, apply has_sum_of_is_lub_of_nonneg, { intro b, simp only [norm_nonneg, pow_nonneg], }, { refine is_lub_csupr _, - use ∥x∥ ^ 2, + use ‖x‖ ^ 2, rintro y ⟨s, rfl⟩, exact hv.sum_inner_products_le x } end @@ -1939,9 +1939,9 @@ end ... = ∑ i in s, ⟪l₁ i, l₂ i⟫ : by simp [finset.sum_ite_of_true] lemma orthogonal_family.norm_sum (l : Π i, G i) (s : finset ι) : - ∥∑ i in s, V i (l i)∥ ^ 2 = ∑ i in s, ∥l i∥ ^ 2 := + ‖∑ i in s, V i (l i)‖ ^ 2 = ∑ i in s, ‖l i‖ ^ 2 := begin - have : (∥∑ i in s, V i (l i)∥ ^ 2 : 𝕜) = ∑ i in s, ∥l i∥ ^ 2, + have : (‖∑ i in s, V i (l i)‖ ^ 2 : 𝕜) = ∑ i in s, ‖l i‖ ^ 2, { simp [← inner_self_eq_norm_sq_to_K, hV.inner_sum] }, exact_mod_cast this, end @@ -1969,20 +1969,20 @@ end include dec_ι lemma orthogonal_family.norm_sq_diff_sum (f : Π i, G i) (s₁ s₂ : finset ι) : - ∥∑ i in s₁, V i (f i) - ∑ i in s₂, V i (f i)∥ ^ 2 - = ∑ i in s₁ \ s₂, ∥f i∥ ^ 2 + ∑ i in s₂ \ s₁, ∥f i∥ ^ 2 := + ‖∑ i in s₁, V i (f i) - ∑ i in s₂, V i (f i)‖ ^ 2 + = ∑ i in s₁ \ s₂, ‖f i‖ ^ 2 + ∑ i in s₂ \ s₁, ‖f i‖ ^ 2 := begin rw [← finset.sum_sdiff_sub_sum_sdiff, sub_eq_add_neg, ← finset.sum_neg_distrib], let F : Π i, G i := λ i, if i ∈ s₁ then f i else - (f i), have hF₁ : ∀ i ∈ s₁ \ s₂, F i = f i := λ i hi, if_pos (finset.sdiff_subset _ _ hi), have hF₂ : ∀ i ∈ s₂ \ s₁, F i = - f i := λ i hi, if_neg (finset.mem_sdiff.mp hi).2, - have hF : ∀ i, ∥F i∥ = ∥f i∥, + have hF : ∀ i, ‖F i‖ = ‖f i‖, { intros i, dsimp [F], split_ifs; simp, }, - have : ∥∑ i in s₁ \ s₂, V i (F i) + ∑ i in s₂ \ s₁, V i (F i)∥ ^ 2 = - ∑ i in s₁ \ s₂, ∥F i∥ ^ 2 + ∑ i in s₂ \ s₁, ∥F i∥ ^ 2, + have : ‖∑ i in s₁ \ s₂, V i (F i) + ∑ i in s₂ \ s₁, V i (F i)‖ ^ 2 = + ∑ i in s₁ \ s₂, ‖F i‖ ^ 2 + ∑ i in s₂ \ s₁, ‖F i‖ ^ 2, { have hs : disjoint (s₁ \ s₂) (s₂ \ s₁) := disjoint_sdiff_sdiff, simpa only [finset.sum_union hs] using hV.norm_sum F (s₁ \ s₂ ∪ s₂ \ s₁) }, convert this using 4, @@ -1997,9 +1997,9 @@ end omit dec_ι /-- A family `f` of mutually-orthogonal elements of `E` is summable, if and only if -`(λ i, ∥f i∥ ^ 2)` is summable. -/ +`(λ i, ‖f i‖ ^ 2)` is summable. -/ lemma orthogonal_family.summable_iff_norm_sq_summable [complete_space E] (f : Π i, G i) : - summable (λ i, V i (f i)) ↔ summable (λ i, ∥f i∥ ^ 2) := + summable (λ i, V i (f i)) ↔ summable (λ i, ‖f i‖ ^ 2) := begin classical, simp only [summable_iff_cauchy_seq_finset, normed_add_comm_group.cauchy_seq_iff, @@ -2011,9 +2011,9 @@ begin intros s₁ hs₁ s₂ hs₂, rw ← finset.sum_sdiff_sub_sum_sdiff, refine (_root_.abs_sub _ _).trans_lt _, - have : ∀ i, 0 ≤ ∥f i∥ ^ 2 := λ i : ι, sq_nonneg _, + have : ∀ i, 0 ≤ ‖f i‖ ^ 2 := λ i : ι, sq_nonneg _, simp only [finset.abs_sum_of_nonneg' this], - have : ∑ i in s₁ \ s₂, ∥f i∥ ^ 2 + ∑ i in s₂ \ s₁, ∥f i∥ ^ 2 < (sqrt ε) ^ 2, + have : ∑ i in s₁ \ s₂, ‖f i‖ ^ 2 + ∑ i in s₂ \ s₁, ‖f i‖ ^ 2 < (sqrt ε) ^ 2, { rw [← hV.norm_sq_diff_sum, sq_lt_sq, _root_.abs_of_nonneg (sqrt_nonneg _), _root_.abs_of_nonneg (norm_nonneg _)], exact H s₁ hs₁ s₂ hs₂ }, @@ -2027,13 +2027,13 @@ begin refine (abs_lt_of_sq_lt_sq' _ (le_of_lt hε)).2, have has : a ≤ s₁ ⊓ s₂ := le_inf hs₁ hs₂, rw hV.norm_sq_diff_sum, - have Hs₁ : ∑ (x : ι) in s₁ \ s₂, ∥f x∥ ^ 2 < ε ^ 2 / 2, + have Hs₁ : ∑ (x : ι) in s₁ \ s₂, ‖f x‖ ^ 2 < ε ^ 2 / 2, { convert H _ hs₁ _ has, have : s₁ ⊓ s₂ ⊆ s₁ := finset.inter_subset_left _ _, rw [← finset.sum_sdiff this, add_tsub_cancel_right, finset.abs_sum_of_nonneg'], { simp }, { exact λ i, sq_nonneg _ } }, - have Hs₂ : ∑ (x : ι) in s₂ \ s₁, ∥f x∥ ^ 2 < ε ^ 2 /2, + have Hs₂ : ∑ (x : ι) in s₂ \ s₁, ‖f x‖ ^ 2 < ε ^ 2 /2, { convert H _ hs₂ _ has, have : s₁ ⊓ s₂ ⊆ s₂ := finset.inter_subset_right _ _, rw [← finset.sum_sdiff this, add_tsub_cancel_right, finset.abs_sum_of_nonneg'], @@ -2187,10 +2187,10 @@ lemma continuous_linear_map.re_apply_inner_self_continuous (T : E →L[𝕜] E) re_clm.continuous.comp $ T.continuous.inner continuous_id lemma continuous_linear_map.re_apply_inner_self_smul (T : E →L[𝕜] E) (x : E) {c : 𝕜} : - T.re_apply_inner_self (c • x) = ∥c∥ ^ 2 * T.re_apply_inner_self x := + T.re_apply_inner_self (c • x) = ‖c‖ ^ 2 * T.re_apply_inner_self x := by simp only [continuous_linear_map.map_smul, continuous_linear_map.re_apply_inner_self_apply, inner_smul_left, inner_smul_right, ← mul_assoc, mul_conj, norm_sq_eq_def', ← smul_re, - algebra.smul_def (∥c∥ ^ 2) ⟪T x, x⟫, algebra_map_eq_of_real] + algebra.smul_def (‖c‖ ^ 2) ⟪T x, x⟫, algebra_map_eq_of_real] end re_apply_inner_self diff --git a/src/analysis/inner_product_space/calculus.lean b/src/analysis/inner_product_space/calculus.lean index 9f95e4b6a1992..99a5cbfa25401 100644 --- a/src/analysis/inner_product_space/calculus.lean +++ b/src/analysis/inner_product_space/calculus.lean @@ -126,30 +126,30 @@ lemma deriv_inner_apply {f g : ℝ → E} {x : ℝ} (hf : differentiable_at ℝ deriv (λ t, ⟪f t, g t⟫) x = ⟪f x, deriv g x⟫ + ⟪deriv f x, g x⟫ := (hf.has_deriv_at.inner hg.has_deriv_at).deriv -lemma cont_diff_norm_sq : cont_diff ℝ n (λ x : E, ∥x∥ ^ 2) := +lemma cont_diff_norm_sq : cont_diff ℝ n (λ x : E, ‖x‖ ^ 2) := begin simp only [sq, ← inner_self_eq_norm_mul_norm], exact (re_clm : 𝕜 →L[ℝ] ℝ).cont_diff.comp (cont_diff_id.inner cont_diff_id) end lemma cont_diff.norm_sq (hf : cont_diff ℝ n f) : - cont_diff ℝ n (λ x, ∥f x∥ ^ 2) := + cont_diff ℝ n (λ x, ‖f x‖ ^ 2) := cont_diff_norm_sq.comp hf lemma cont_diff_within_at.norm_sq (hf : cont_diff_within_at ℝ n f s x) : - cont_diff_within_at ℝ n (λ y, ∥f y∥ ^ 2) s x := + cont_diff_within_at ℝ n (λ y, ‖f y‖ ^ 2) s x := cont_diff_norm_sq.cont_diff_at.comp_cont_diff_within_at x hf lemma cont_diff_at.norm_sq (hf : cont_diff_at ℝ n f x) : - cont_diff_at ℝ n (λ y, ∥f y∥ ^ 2) x := + cont_diff_at ℝ n (λ y, ‖f y‖ ^ 2) x := hf.norm_sq lemma cont_diff_at_norm {x : E} (hx : x ≠ 0) : cont_diff_at ℝ n norm x := -have ∥id x∥ ^ 2 ≠ 0, from pow_ne_zero _ (norm_pos_iff.2 hx).ne', +have ‖id x‖ ^ 2 ≠ 0, from pow_ne_zero _ (norm_pos_iff.2 hx).ne', by simpa only [id, sqrt_sq, norm_nonneg] using cont_diff_at_id.norm_sq.sqrt this lemma cont_diff_at.norm (hf : cont_diff_at ℝ n f x) (h0 : f x ≠ 0) : - cont_diff_at ℝ n (λ y, ∥f y∥) x := + cont_diff_at ℝ n (λ y, ‖f y‖) x := (cont_diff_at_norm h0).comp x hf lemma cont_diff_at.dist (hf : cont_diff_at ℝ n f x) (hg : cont_diff_at ℝ n g x) @@ -158,7 +158,7 @@ lemma cont_diff_at.dist (hf : cont_diff_at ℝ n f x) (hg : cont_diff_at ℝ n g by { simp only [dist_eq_norm], exact (hf.sub hg).norm (sub_ne_zero.2 hne) } lemma cont_diff_within_at.norm (hf : cont_diff_within_at ℝ n f s x) (h0 : f x ≠ 0) : - cont_diff_within_at ℝ n (λ y, ∥f y∥) s x := + cont_diff_within_at ℝ n (λ y, ‖f y‖) s x := (cont_diff_at_norm h0).comp_cont_diff_within_at x hf lemma cont_diff_within_at.dist (hf : cont_diff_within_at ℝ n f s x) @@ -167,11 +167,11 @@ lemma cont_diff_within_at.dist (hf : cont_diff_within_at ℝ n f s x) by { simp only [dist_eq_norm], exact (hf.sub hg).norm (sub_ne_zero.2 hne) } lemma cont_diff_on.norm_sq (hf : cont_diff_on ℝ n f s) : - cont_diff_on ℝ n (λ y, ∥f y∥ ^ 2) s := + cont_diff_on ℝ n (λ y, ‖f y‖ ^ 2) s := (λ x hx, (hf x hx).norm_sq) lemma cont_diff_on.norm (hf : cont_diff_on ℝ n f s) (h0 : ∀ x ∈ s, f x ≠ 0) : - cont_diff_on ℝ n (λ y, ∥f y∥) s := + cont_diff_on ℝ n (λ y, ‖f y‖) s := λ x hx, (hf x hx).norm (h0 x hx) lemma cont_diff_on.dist (hf : cont_diff_on ℝ n f s) @@ -180,7 +180,7 @@ lemma cont_diff_on.dist (hf : cont_diff_on ℝ n f s) λ x hx, (hf x hx).dist (hg x hx) (hne x hx) lemma cont_diff.norm (hf : cont_diff ℝ n f) (h0 : ∀ x, f x ≠ 0) : - cont_diff ℝ n (λ y, ∥f y∥) := + cont_diff ℝ n (λ y, ‖f y‖) := cont_diff_iff_cont_diff_at.2 $ λ x, hf.cont_diff_at.norm (h0 x) lemma cont_diff.dist (hf : cont_diff ℝ n f) (hg : cont_diff ℝ n g) @@ -191,7 +191,7 @@ cont_diff_iff_cont_diff_at.2 $ omit 𝕜 lemma has_strict_fderiv_at_norm_sq (x : F) : - has_strict_fderiv_at (λ x, ∥x∥ ^ 2) (bit0 (innerSL x : F →L[ℝ] ℝ)) x := + has_strict_fderiv_at (λ x, ‖x‖ ^ 2) (bit0 (innerSL x : F →L[ℝ] ℝ)) x := begin simp only [sq, ← inner_self_eq_norm_mul_norm], convert (has_strict_fderiv_at_id x).inner (has_strict_fderiv_at_id x), @@ -201,11 +201,11 @@ end include 𝕜 lemma differentiable_at.norm_sq (hf : differentiable_at ℝ f x) : - differentiable_at ℝ (λ y, ∥f y∥ ^ 2) x := + differentiable_at ℝ (λ y, ‖f y‖ ^ 2) x := (cont_diff_at_id.norm_sq.differentiable_at le_rfl).comp x hf lemma differentiable_at.norm (hf : differentiable_at ℝ f x) (h0 : f x ≠ 0) : - differentiable_at ℝ (λ y, ∥f y∥) x := + differentiable_at ℝ (λ y, ‖f y‖) x := ((cont_diff_at_norm h0).differentiable_at le_rfl).comp x hf lemma differentiable_at.dist (hf : differentiable_at ℝ f x) (hg : differentiable_at ℝ g x) @@ -213,11 +213,11 @@ lemma differentiable_at.dist (hf : differentiable_at ℝ f x) (hg : differentiab differentiable_at ℝ (λ y, dist (f y) (g y)) x := by { simp only [dist_eq_norm], exact (hf.sub hg).norm (sub_ne_zero.2 hne) } -lemma differentiable.norm_sq (hf : differentiable ℝ f) : differentiable ℝ (λ y, ∥f y∥ ^ 2) := +lemma differentiable.norm_sq (hf : differentiable ℝ f) : differentiable ℝ (λ y, ‖f y‖ ^ 2) := λ x, (hf x).norm_sq lemma differentiable.norm (hf : differentiable ℝ f) (h0 : ∀ x, f x ≠ 0) : - differentiable ℝ (λ y, ∥f y∥) := + differentiable ℝ (λ y, ‖f y‖) := λ x, (hf x).norm (h0 x) lemma differentiable.dist (hf : differentiable ℝ f) (hg : differentiable ℝ g) @@ -226,11 +226,11 @@ lemma differentiable.dist (hf : differentiable ℝ f) (hg : differentiable ℝ g λ x, (hf x).dist (hg x) (hne x) lemma differentiable_within_at.norm_sq (hf : differentiable_within_at ℝ f s x) : - differentiable_within_at ℝ (λ y, ∥f y∥ ^ 2) s x := + differentiable_within_at ℝ (λ y, ‖f y‖ ^ 2) s x := (cont_diff_at_id.norm_sq.differentiable_at le_rfl).comp_differentiable_within_at x hf lemma differentiable_within_at.norm (hf : differentiable_within_at ℝ f s x) (h0 : f x ≠ 0) : - differentiable_within_at ℝ (λ y, ∥f y∥) s x := + differentiable_within_at ℝ (λ y, ‖f y‖) s x := ((cont_diff_at_id.norm h0).differentiable_at le_rfl).comp_differentiable_within_at x hf lemma differentiable_within_at.dist (hf : differentiable_within_at ℝ f s x) @@ -239,11 +239,11 @@ lemma differentiable_within_at.dist (hf : differentiable_within_at ℝ f s x) by { simp only [dist_eq_norm], exact (hf.sub hg).norm (sub_ne_zero.2 hne) } lemma differentiable_on.norm_sq (hf : differentiable_on ℝ f s) : - differentiable_on ℝ (λ y, ∥f y∥ ^ 2) s := + differentiable_on ℝ (λ y, ‖f y‖ ^ 2) s := λ x hx, (hf x hx).norm_sq lemma differentiable_on.norm (hf : differentiable_on ℝ f s) (h0 : ∀ x ∈ s, f x ≠ 0) : - differentiable_on ℝ (λ y, ∥f y∥) s := + differentiable_on ℝ (λ y, ‖f y‖) s := λ x hx, (hf x hx).norm (h0 x hx) lemma differentiable_on.dist (hf : differentiable_on ℝ f s) (hg : differentiable_on ℝ g s) @@ -343,8 +343,8 @@ variables {n : ℕ∞} {E : Type*} [inner_product_space ℝ E] lemma cont_diff_homeomorph_unit_ball : cont_diff ℝ n $ λ (x : E), (homeomorph_unit_ball x : E) := begin - suffices : cont_diff ℝ n (λ x, (1 + ∥x∥^2).sqrt⁻¹), { exact this.smul cont_diff_id, }, - have h : ∀ (x : E), 0 < 1 + ∥x∥ ^ 2 := λ x, by positivity, + suffices : cont_diff ℝ n (λ x, (1 + ‖x‖^2).sqrt⁻¹), { exact this.smul cont_diff_id, }, + have h : ∀ (x : E), 0 < 1 + ‖x‖ ^ 2 := λ x, by positivity, refine cont_diff.inv _ (λ x, real.sqrt_ne_zero'.mpr (h x)), exact (cont_diff_const.add cont_diff_norm_sq).sqrt (λ x, (h x).ne.symm), end @@ -355,15 +355,15 @@ lemma cont_diff_on_homeomorph_unit_ball_symm begin intros y hy, apply cont_diff_at.cont_diff_within_at, - have hf : f =ᶠ[𝓝 y] λ y, (1 - ∥(y : E)∥^2).sqrt⁻¹ • (y : E), + have hf : f =ᶠ[𝓝 y] λ y, (1 - ‖(y : E)‖^2).sqrt⁻¹ • (y : E), { rw eventually_eq_iff_exists_mem, refine ⟨ball (0 : E) 1, mem_nhds_iff.mpr ⟨ball (0 : E) 1, set.subset.refl _, is_open_ball, hy⟩, λ z hz, _⟩, rw h z hz, refl, }, refine cont_diff_at.congr_of_eventually_eq _ hf, - suffices : cont_diff_at ℝ n (λy, (1 - ∥(y : E)∥^2).sqrt⁻¹) y, { exact this.smul cont_diff_at_id }, - have h : 0 < 1 - ∥(y : E)∥^2, by rwa [mem_ball_zero_iff, ← _root_.abs_one, ← abs_norm_eq_norm, + suffices : cont_diff_at ℝ n (λy, (1 - ‖(y : E)‖^2).sqrt⁻¹) y, { exact this.smul cont_diff_at_id }, + have h : 0 < 1 - ‖(y : E)‖^2, by rwa [mem_ball_zero_iff, ← _root_.abs_one, ← abs_norm_eq_norm, ← sq_lt_sq, one_pow, ← sub_pos] at hy, refine cont_diff_at.inv _ (real.sqrt_ne_zero'.mpr h), refine cont_diff_at.comp _ (cont_diff_at_sqrt h.ne.symm) _, diff --git a/src/analysis/inner_product_space/dual.lean b/src/analysis/inner_product_space/dual.lean index aaa4d09b14a62..54fe63804f2b5 100644 --- a/src/analysis/inner_product_space/dual.lean +++ b/src/analysis/inner_product_space/dual.lean @@ -61,8 +61,8 @@ variables {E} @[simp] lemma to_dual_map_apply {x y : E} : to_dual_map 𝕜 E x y = ⟪x, y⟫ := rfl -lemma innerSL_norm [nontrivial E] : ∥(innerSL : E →L⋆[𝕜] E →L[𝕜] 𝕜)∥ = 1 := -show ∥(to_dual_map 𝕜 E).to_continuous_linear_map∥ = 1, +lemma innerSL_norm [nontrivial E] : ‖(innerSL : E →L⋆[𝕜] E →L[𝕜] 𝕜)‖ = 1 := +show ‖(to_dual_map 𝕜 E).to_continuous_linear_map‖ = 1, from linear_isometry.norm_to_continuous_linear_map _ variable {𝕜} diff --git a/src/analysis/inner_product_space/gram_schmidt_ortho.lean b/src/analysis/inner_product_space/gram_schmidt_ortho.lean index 4bf0c990774be..1b546aeab8cec 100644 --- a/src/analysis/inner_product_space/gram_schmidt_ortho.lean +++ b/src/analysis/inner_product_space/gram_schmidt_ortho.lean @@ -64,7 +64,7 @@ by rw [gram_schmidt_def, sub_add_cancel] lemma gram_schmidt_def'' (f : ι → E) (n : ι): f n = gram_schmidt 𝕜 f n - + ∑ i in Iio n, (⟪gram_schmidt 𝕜 f i, f n⟫ / ∥gram_schmidt 𝕜 f i∥ ^ 2) • gram_schmidt 𝕜 f i := + + ∑ i in Iio n, (⟪gram_schmidt 𝕜 f i, f n⟫ / ‖gram_schmidt 𝕜 f i‖ ^ 2) • gram_schmidt 𝕜 f i := begin convert gram_schmidt_def' 𝕜 f n, ext i, @@ -260,22 +260,22 @@ variables (𝕜) /-- the normalized `gram_schmidt` (i.e each vector in `gram_schmidt_normed` has unit length.) -/ noncomputable def gram_schmidt_normed (f : ι → E) (n : ι) : E := -(∥gram_schmidt 𝕜 f n∥ : 𝕜)⁻¹ • (gram_schmidt 𝕜 f n) +(‖gram_schmidt 𝕜 f n‖ : 𝕜)⁻¹ • (gram_schmidt 𝕜 f n) variables {𝕜} lemma gram_schmidt_normed_unit_length_coe {f : ι → E} (n : ι) (h₀ : linear_independent 𝕜 (f ∘ (coe : set.Iic n → ι))) : - ∥gram_schmidt_normed 𝕜 f n∥ = 1 := + ‖gram_schmidt_normed 𝕜 f n‖ = 1 := by simp only [gram_schmidt_ne_zero_coe n h₀, gram_schmidt_normed, norm_smul_inv_norm, ne.def, not_false_iff] lemma gram_schmidt_normed_unit_length {f : ι → E} (n : ι) (h₀ : linear_independent 𝕜 f) : - ∥gram_schmidt_normed 𝕜 f n∥ = 1 := + ‖gram_schmidt_normed 𝕜 f n‖ = 1 := gram_schmidt_normed_unit_length_coe _ (linear_independent.comp h₀ _ subtype.coe_injective) lemma gram_schmidt_normed_unit_length' {f : ι → E} {n : ι} (hn : gram_schmidt_normed 𝕜 f n ≠ 0) : - ∥gram_schmidt_normed 𝕜 f n∥ = 1 := + ‖gram_schmidt_normed 𝕜 f n‖ = 1 := begin rw gram_schmidt_normed at *, rw [norm_smul_inv_norm], @@ -319,7 +319,7 @@ begin simp only [coe_singleton, set.image_singleton], by_cases h : gram_schmidt 𝕜 f i = 0, { simp [h] }, - { refine mem_span_singleton.2 ⟨∥gram_schmidt 𝕜 f i∥, smul_inv_smul₀ _ _⟩, + { refine mem_span_singleton.2 ⟨‖gram_schmidt 𝕜 f i‖, smul_inv_smul₀ _ _⟩, exact_mod_cast (norm_ne_zero_iff.2 h) } end @@ -345,9 +345,9 @@ lemma gram_schmidt_orthonormal_basis_apply {f : ι → E} {i : ι} lemma gram_schmidt_orthonormal_basis_apply_of_orthogonal {f : ι → E} (hf : pairwise (λ i j, ⟪f i, f j⟫ = 0)) {i : ι} (hi : f i ≠ 0) : - gram_schmidt_orthonormal_basis h f i = (∥f i∥⁻¹ : 𝕜) • f i := + gram_schmidt_orthonormal_basis h f i = (‖f i‖⁻¹ : 𝕜) • f i := begin - have H : gram_schmidt_normed 𝕜 f i = (∥f i∥⁻¹ : 𝕜) • f i, + have H : gram_schmidt_normed 𝕜 f i = (‖f i‖⁻¹ : 𝕜) • f i, { rw [gram_schmidt_normed, gram_schmidt_of_orthogonal 𝕜 hf] }, rw [gram_schmidt_orthonormal_basis_apply h, H], simpa [H] using hi, diff --git a/src/analysis/inner_product_space/l2_space.lean b/src/analysis/inner_product_space/l2_space.lean index bf39b42ae3c3e..4d5d34a61afb9 100644 --- a/src/analysis/inner_product_space/l2_space.lean +++ b/src/analysis/inner_product_space/l2_space.lean @@ -12,7 +12,7 @@ import analysis.inner_product_space.pi_L2 Given a family `(G : ι → Type*) [Π i, inner_product_space 𝕜 (G i)]` of inner product spaces, this file equips `lp G 2` with an inner product space structure, where `lp G 2` consists of those -dependent functions `f : Π i, G i` for which `∑' i, ∥f i∥ ^ 2`, the sum of the norms-squared, is +dependent functions `f : Π i, G i` for which `∑' i, ‖f i‖ ^ 2`, the sum of the norms-squared, is summable. This construction is sometimes called the *Hilbert sum* of the family `G`. By choosing `G` to be `ι → 𝕜`, the Hilbert space `ℓ²(ι, 𝕜)` may be seen as a special case of this construction. @@ -97,8 +97,8 @@ namespace lp lemma summable_inner (f g : lp G 2) : summable (λ i, ⟪f i, g i⟫) := begin - -- Apply the Direct Comparison Test, comparing with ∑' i, ∥f i∥ * ∥g i∥ (summable by Hölder) - refine summable_of_norm_bounded (λ i, ∥f i∥ * ∥g i∥) (lp.summable_mul _ f g) _, + -- Apply the Direct Comparison Test, comparing with ∑' i, ‖f i‖ * ‖g i‖ (summable by Hölder) + refine summable_of_norm_bounded (λ i, ‖f i‖ * ‖g i‖) (lp.summable_mul _ f g) _, { rw real.is_conjugate_exponent_iff; norm_num }, intros i, -- Then apply Cauchy-Schwarz pointwise @@ -108,9 +108,9 @@ end instance : inner_product_space 𝕜 (lp G 2) := { inner := λ f g, ∑' i, ⟪f i, g i⟫, norm_sq_eq_inner := λ f, begin - calc ∥f∥ ^ 2 = ∥f∥ ^ (2:ℝ≥0∞).to_real : by norm_cast - ... = ∑' i, ∥f i∥ ^ (2:ℝ≥0∞).to_real : lp.norm_rpow_eq_tsum _ f - ... = ∑' i, ∥f i∥ ^ 2 : by norm_cast + calc ‖f‖ ^ 2 = ‖f‖ ^ (2:ℝ≥0∞).to_real : by norm_cast + ... = ∑' i, ‖f i‖ ^ (2:ℝ≥0∞).to_real : lp.norm_rpow_eq_tsum _ f + ... = ∑' i, ‖f i‖ ^ 2 : by norm_cast ... = ∑' i, re ⟪f i, f i⟫ : by simp only [norm_sq_eq_inner] ... = re (∑' i, ⟪f i, f i⟫) : (is_R_or_C.re_clm.map_tsum _).symm ... = _ : by congr, @@ -190,7 +190,7 @@ protected def linear_isometry : lp G 2 →ₗᵢ[𝕜] E := norm_map' := λ f, begin classical, -- needed for lattice instance on `finset ι`, for `filter.at_top_ne_bot` have H : 0 < (2:ℝ≥0∞).to_real := by norm_num, - suffices : ∥∑' (i : ι), V i (f i)∥ ^ ((2:ℝ≥0∞).to_real) = ∥f∥ ^ ((2:ℝ≥0∞).to_real), + suffices : ‖∑' (i : ι), V i (f i)‖ ^ ((2:ℝ≥0∞).to_real) = ‖f‖ ^ ((2:ℝ≥0∞).to_real), { exact real.rpow_left_inj_on H.ne' (norm_nonneg _) (norm_nonneg _) this }, refine tendsto_nhds_unique _ (lp.has_sum_norm H f), convert (hV.summable_of_lp f).has_sum.norm.rpow_const (or.inr H.le), diff --git a/src/analysis/inner_product_space/lax_milgram.lean b/src/analysis/inner_product_space/lax_milgram.lean index ef9a303e8bb9b..b7ed91982ec9e 100644 --- a/src/analysis/inner_product_space/lax_milgram.lean +++ b/src/analysis/inner_product_space/lax_milgram.lean @@ -16,7 +16,7 @@ We consider an Hilbert space `V` over `ℝ` equipped with a bounded bilinear form `B : V →L[ℝ] V →L[ℝ] ℝ`. Recall that a bilinear form `B : V →L[ℝ] V →L[ℝ] ℝ` is *coercive* -iff `∃ C, (0 < C) ∧ ∀ u, C * ∥u∥ * ∥u∥ ≤ B u u`. +iff `∃ C, (0 < C) ∧ ∀ u, C * ‖u‖ * ‖u‖ ≤ B u u`. Under the hypothesis that `B` is coercive we prove the Lax-Milgram theorem: that is, the map `inner_product_space.continuous_linear_map_of_bilin` from @@ -46,17 +46,17 @@ variables {B : V →L[ℝ] V →L[ℝ] ℝ} local postfix `♯`:1025 := @continuous_linear_map_of_bilin ℝ V _ _ _ lemma bounded_below (coercive : is_coercive B) : - ∃ C, 0 < C ∧ ∀ v, C * ∥v∥ ≤ ∥B♯ v∥ := + ∃ C, 0 < C ∧ ∀ v, C * ‖v‖ ≤ ‖B♯ v‖ := begin rcases coercive with ⟨C, C_ge_0, coercivity⟩, refine ⟨C, C_ge_0, _⟩, intro v, - by_cases h : 0 < ∥v∥, + by_cases h : 0 < ‖v‖, { refine (mul_le_mul_right h).mp _, - calc C * ∥v∥ * ∥v∥ + calc C * ‖v‖ * ‖v‖ ≤ B v v : coercivity v ... = ⟪B♯ v, v⟫_ℝ : (continuous_linear_map_of_bilin_apply ℝ B v v).symm - ... ≤ ∥B♯ v∥ * ∥v∥ : real_inner_le_norm (B♯ v) v, }, + ... ≤ ‖B♯ v‖ * ‖v‖ : real_inner_le_norm (B♯ v) v, }, { have : v = 0 := by simpa using h, simp [this], } end @@ -96,7 +96,7 @@ begin obtain rfl : w = 0, { rw [←norm_eq_zero, ←mul_self_eq_zero, ←mul_right_inj' C_pos.ne', mul_zero, ←mul_assoc], apply le_antisymm, - { calc C * ∥w∥ * ∥w∥ + { calc C * ‖w‖ * ‖w‖ ≤ B w w : coercivity w ... = ⟪B♯ w, w⟫_ℝ : (continuous_linear_map_of_bilin_apply ℝ B w w).symm ... = 0 : mem_w_orthogonal _ ⟨w, rfl⟩ }, diff --git a/src/analysis/inner_product_space/orientation.lean b/src/analysis/inner_product_space/orientation.lean index 652b3d743b051..605f4775837ae 100644 --- a/src/analysis/inner_product_space/orientation.lean +++ b/src/analysis/inner_product_space/orientation.lean @@ -260,7 +260,7 @@ end /-- Let `v` be an indexed family of `n` vectors in an oriented `n`-dimensional real inner product space `E`. The output of the volume form of `E` when evaluated on `v` is bounded in absolute value by the product of the norms of the vectors `v i`. -/ -lemma abs_volume_form_apply_le (v : fin n → E) : |o.volume_form v| ≤ ∏ i : fin n, ∥v i∥ := +lemma abs_volume_form_apply_le (v : fin n → E) : |o.volume_form v| ≤ ∏ i : fin n, ‖v i‖ := begin unfreezingI { cases n }, { refine o.eq_or_eq_neg_of_is_empty.by_cases _ _; rintros rfl; simp }, @@ -277,7 +277,7 @@ begin simp [b.orthonormal.1 i], end -lemma volume_form_apply_le (v : fin n → E) : o.volume_form v ≤ ∏ i : fin n, ∥v i∥ := +lemma volume_form_apply_le (v : fin n → E) : o.volume_form v ≤ ∏ i : fin n, ‖v i‖ := (le_abs_self _).trans (o.abs_volume_form_apply_le v) /-- Let `v` be an indexed family of `n` orthogonal vectors in an oriented `n`-dimensional @@ -285,7 +285,7 @@ real inner product space `E`. The output of the volume form of `E` when evaluate sign, the product of the norms of the vectors `v i`. -/ lemma abs_volume_form_apply_of_pairwise_orthogonal {v : fin n → E} (hv : pairwise (λ i j, ⟪v i, v j⟫ = 0)) : - |o.volume_form v| = ∏ i : fin n, ∥v i∥ := + |o.volume_form v| = ∏ i : fin n, ‖v i‖ := begin unfreezingI { cases n }, { refine o.eq_or_eq_neg_of_is_empty.by_cases _ _; rintros rfl; simp }, @@ -301,10 +301,10 @@ begin push_neg at h, congr, ext i, - have hb : b i = ∥v i∥⁻¹ • v i := gram_schmidt_orthonormal_basis_apply_of_orthogonal hdim hv (h i), + have hb : b i = ‖v i‖⁻¹ • v i := gram_schmidt_orthonormal_basis_apply_of_orthogonal hdim hv (h i), simp only [hb, inner_smul_left, real_inner_self_eq_norm_mul_norm, is_R_or_C.conj_to_real], rw abs_of_nonneg, - { have : ∥v i∥ ≠ 0 := by simpa using h i, + { have : ‖v i‖ ≠ 0 := by simpa using h i, field_simp }, { positivity }, end diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index b711d17671bc9..de3e93e839bf2 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -102,11 +102,11 @@ def euclidean_space (𝕜 : Type*) [is_R_or_C 𝕜] (n : Type*) [fintype n] : Type* := pi_Lp 2 (λ (i : n), 𝕜) lemma euclidean_space.nnnorm_eq {𝕜 : Type*} [is_R_or_C 𝕜] {n : Type*} [fintype n] - (x : euclidean_space 𝕜 n) : ∥x∥₊ = nnreal.sqrt (∑ i, ∥x i∥₊ ^ 2) := + (x : euclidean_space 𝕜 n) : ‖x‖₊ = nnreal.sqrt (∑ i, ‖x i‖₊ ^ 2) := pi_Lp.nnnorm_eq_of_L2 x lemma euclidean_space.norm_eq {𝕜 : Type*} [is_R_or_C 𝕜] {n : Type*} [fintype n] - (x : euclidean_space 𝕜 n) : ∥x∥ = real.sqrt (∑ i, ∥x i∥ ^ 2) := + (x : euclidean_space 𝕜 n) : ‖x‖ = real.sqrt (∑ i, ‖x i‖ ^ 2) := by simpa only [real.coe_sqrt, nnreal.coe_sum] using congr_arg (coe : ℝ≥0 → ℝ) x.nnnorm_eq lemma euclidean_space.dist_eq {𝕜 : Type*} [is_R_or_C 𝕜] {n : Type*} [fintype n] @@ -547,7 +547,7 @@ end /-- The determinant of the change-of-basis matrix between two orthonormal bases `a`, `b` has unit length. -/ @[simp] lemma orthonormal_basis.det_to_matrix_orthonormal_basis : - ∥a.to_basis.det b∥ = 1 := + ‖a.to_basis.det b‖ = 1 := begin have : (norm_sq (a.to_basis.det b) : 𝕜) = 1, { simpa [is_R_or_C.mul_conj] @@ -757,7 +757,7 @@ begin -- Build a linear map from the isometries on S and Sᗮ let M := L.to_linear_map.comp p1 + L3.to_linear_map.comp p2, -- Prove that M is an isometry - have M_norm_map : ∀ (x : V), ∥M x∥ = ∥x∥, + have M_norm_map : ∀ (x : V), ‖M x‖ = ‖x‖, { intro x, -- Apply M to the orthogonal decomposition of x have Mx_decomp : M x = L (p1 x) + L3 (p2 x), diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index 55f4a9056003a..defc6a6a1fa51 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -13,7 +13,7 @@ import analysis.normed_space.is_R_or_C Given a nonempty complete subspace `K` of an inner product space `E`, this file constructs `orthogonal_projection K : E →L[𝕜] K`, the orthogonal projection of `E` onto `K`. This map satisfies: for any point `u` in `E`, the point `v = orthogonal_projection K u` in `K` minimizes the -distance `∥u - v∥` to `u`. +distance `‖u - v‖` to `u`. Also a linear isometry equivalence `reflection K : E ≃ₗᵢ[𝕜] E` is constructed, by choosing, for each `u : E`, the point `reflection K u` to satisfy @@ -51,31 +51,31 @@ local notation `absR` := has_abs.abs /-- Existence of minimizers Let `u` be a point in a real inner product space, and let `K` be a nonempty complete convex subset. -Then there exists a (unique) `v` in `K` that minimizes the distance `∥u - v∥` to `u`. +Then there exists a (unique) `v` in `K` that minimizes the distance `‖u - v‖` to `u`. -/ -- FIXME this monolithic proof causes a deterministic timeout with `-T50000` -- It should be broken in a sequence of more manageable pieces, -- perhaps with individual statements for the three steps below. theorem exists_norm_eq_infi_of_complete_convex {K : set F} (ne : K.nonempty) (h₁ : is_complete K) - (h₂ : convex ℝ K) : ∀ u : F, ∃ v ∈ K, ∥u - v∥ = ⨅ w : K, ∥u - w∥ := assume u, + (h₂ : convex ℝ K) : ∀ u : F, ∃ v ∈ K, ‖u - v‖ = ⨅ w : K, ‖u - w‖ := assume u, begin - let δ := ⨅ w : K, ∥u - w∥, + let δ := ⨅ w : K, ‖u - w‖, letI : nonempty K := ne.to_subtype, have zero_le_δ : 0 ≤ δ := le_cinfi (λ _, norm_nonneg _), - have δ_le : ∀ w : K, δ ≤ ∥u - w∥, + have δ_le : ∀ w : K, δ ≤ ‖u - w‖, from cinfi_le ⟨0, set.forall_range_iff.2 $ λ _, norm_nonneg _⟩, - have δ_le' : ∀ w ∈ K, δ ≤ ∥u - w∥ := assume w hw, δ_le ⟨w, hw⟩, + have δ_le' : ∀ w ∈ K, δ ≤ ‖u - w‖ := assume w hw, δ_le ⟨w, hw⟩, -- Step 1: since `δ` is the infimum, can find a sequence `w : ℕ → K` in `K` - -- such that `∥u - w n∥ < δ + 1 / (n + 1)` (which implies `∥u - w n∥ --> δ`); + -- such that `‖u - w n‖ < δ + 1 / (n + 1)` (which implies `‖u - w n‖ --> δ`); -- maybe this should be a separate lemma - have exists_seq : ∃ w : ℕ → K, ∀ n, ∥u - w n∥ < δ + 1 / (n + 1), + have exists_seq : ∃ w : ℕ → K, ∀ n, ‖u - w n‖ < δ + 1 / (n + 1), { have hδ : ∀n:ℕ, δ < δ + 1 / (n + 1), from λ n, lt_add_of_le_of_pos le_rfl nat.one_div_pos_of_nat, have h := λ n, exists_lt_of_cinfi_lt (hδ n), let w : ℕ → K := λ n, classical.some (h n), exact ⟨w, λ n, classical.some_spec (h n)⟩ }, rcases exists_seq with ⟨w, hw⟩, - have norm_tendsto : tendsto (λ n, ∥u - w n∥) at_top (nhds δ), + have norm_tendsto : tendsto (λ n, ‖u - w n‖) at_top (nhds δ), { have h : tendsto (λ n:ℕ, δ) at_top (nhds δ) := tendsto_const_nhds, have h' : tendsto (λ n:ℕ, δ + 1 / (n + 1)) at_top (nhds δ), { convert h.add tendsto_one_div_add_at_top_nhds_0_nat, simp only [add_zero] }, @@ -95,18 +95,18 @@ begin let wp := ((w p):F), let wq := ((w q):F), let a := u - wq, let b := u - wp, let half := 1 / (2:ℝ), let div := 1 / ((N:ℝ) + 1), - have : 4 * ∥u - half • (wq + wp)∥ * ∥u - half • (wq + wp)∥ + ∥wp - wq∥ * ∥wp - wq∥ = - 2 * (∥a∥ * ∥a∥ + ∥b∥ * ∥b∥) := + have : 4 * ‖u - half • (wq + wp)‖ * ‖u - half • (wq + wp)‖ + ‖wp - wq‖ * ‖wp - wq‖ = + 2 * (‖a‖ * ‖a‖ + ‖b‖ * ‖b‖) := calc - 4 * ∥u - half•(wq + wp)∥ * ∥u - half•(wq + wp)∥ + ∥wp - wq∥ * ∥wp - wq∥ - = (2*∥u - half•(wq + wp)∥) * (2 * ∥u - half•(wq + wp)∥) + ∥wp-wq∥*∥wp-wq∥ : by ring - ... = (absR ((2:ℝ)) * ∥u - half•(wq + wp)∥) * (absR ((2:ℝ)) * ∥u - half•(wq+wp)∥) + - ∥wp-wq∥*∥wp-wq∥ : + 4 * ‖u - half•(wq + wp)‖ * ‖u - half•(wq + wp)‖ + ‖wp - wq‖ * ‖wp - wq‖ + = (2*‖u - half•(wq + wp)‖) * (2 * ‖u - half•(wq + wp)‖) + ‖wp-wq‖*‖wp-wq‖ : by ring + ... = (absR ((2:ℝ)) * ‖u - half•(wq + wp)‖) * (absR ((2:ℝ)) * ‖u - half•(wq+wp)‖) + + ‖wp-wq‖*‖wp-wq‖ : by { rw _root_.abs_of_nonneg, exact zero_le_two } - ... = ∥(2:ℝ) • (u - half • (wq + wp))∥ * ∥(2:ℝ) • (u - half • (wq + wp))∥ + - ∥wp-wq∥ * ∥wp-wq∥ : + ... = ‖(2:ℝ) • (u - half • (wq + wp))‖ * ‖(2:ℝ) • (u - half • (wq + wp))‖ + + ‖wp-wq‖ * ‖wp-wq‖ : by simp [norm_smul] - ... = ∥a + b∥ * ∥a + b∥ + ∥a - b∥ * ∥a - b∥ : + ... = ‖a + b‖ * ‖a + b‖ + ‖a - b‖ * ‖a - b‖ : begin rw [smul_sub, smul_smul, mul_one_div_cancel (_root_.two_ne_zero : (2 : ℝ) ≠ 0), ← one_add_one_eq_two, add_smul], @@ -115,29 +115,29 @@ begin have eq₂ : u + u - (wq + wp) = a + b, show u + u - (wq + wp) = (u - wq) + (u - wp), abel, rw [eq₁, eq₂], end - ... = 2 * (∥a∥ * ∥a∥ + ∥b∥ * ∥b∥) : parallelogram_law_with_norm _ _, - have eq : δ ≤ ∥u - half • (wq + wp)∥, + ... = 2 * (‖a‖ * ‖a‖ + ‖b‖ * ‖b‖) : parallelogram_law_with_norm _ _, + have eq : δ ≤ ‖u - half • (wq + wp)‖, { rw smul_add, apply δ_le', apply h₂, repeat {exact subtype.mem _}, repeat {exact le_of_lt one_half_pos}, exact add_halves 1 }, - have eq₁ : 4 * δ * δ ≤ 4 * ∥u - half • (wq + wp)∥ * ∥u - half • (wq + wp)∥, + have eq₁ : 4 * δ * δ ≤ 4 * ‖u - half • (wq + wp)‖ * ‖u - half • (wq + wp)‖, { simp_rw mul_assoc, exact mul_le_mul_of_nonneg_left (mul_self_le_mul_self zero_le_δ eq) zero_le_four }, - have eq₂ : ∥a∥ * ∥a∥ ≤ (δ + div) * (δ + div) := + have eq₂ : ‖a‖ * ‖a‖ ≤ (δ + div) * (δ + div) := mul_self_le_mul_self (norm_nonneg _) (le_trans (le_of_lt $ hw q) (add_le_add_left (nat.one_div_le_one_div hq) _)), - have eq₂' : ∥b∥ * ∥b∥ ≤ (δ + div) * (δ + div) := + have eq₂' : ‖b‖ * ‖b‖ ≤ (δ + div) * (δ + div) := mul_self_le_mul_self (norm_nonneg _) (le_trans (le_of_lt $ hw p) (add_le_add_left (nat.one_div_le_one_div hp) _)), rw dist_eq_norm, apply nonneg_le_nonneg_of_sq_le_sq, { exact sqrt_nonneg _ }, rw mul_self_sqrt, calc - ∥wp - wq∥ * ∥wp - wq∥ = 2 * (∥a∥*∥a∥ + ∥b∥*∥b∥) - - 4 * ∥u - half • (wq+wp)∥ * ∥u - half • (wq+wp)∥ : by { rw ← this, simp } - ... ≤ 2 * (∥a∥ * ∥a∥ + ∥b∥ * ∥b∥) - 4 * δ * δ : sub_le_sub_left eq₁ _ + ‖wp - wq‖ * ‖wp - wq‖ = 2 * (‖a‖*‖a‖ + ‖b‖*‖b‖) - + 4 * ‖u - half • (wq+wp)‖ * ‖u - half • (wq+wp)‖ : by { rw ← this, simp } + ... ≤ 2 * (‖a‖ * ‖a‖ + ‖b‖ * ‖b‖) - 4 * δ * δ : sub_le_sub_left eq₁ _ ... ≤ 2 * ((δ + div) * (δ + div) + (δ + div) * (δ + div)) - 4 * δ * δ : sub_le_sub_right (mul_le_mul_of_nonneg_left (add_le_add eq₂ eq₂') (by norm_num)) _ ... = 8 * δ * div + 4 * div * div : by ring, @@ -161,9 +161,9 @@ begin -- Prove that it satisfies all requirements. rcases cauchy_seq_tendsto_of_is_complete h₁ (λ n, _) seq_is_cauchy with ⟨v, hv, w_tendsto⟩, use v, use hv, - have h_cont : continuous (λ v, ∥u - v∥) := + have h_cont : continuous (λ v, ‖u - v‖) := continuous.comp continuous_norm (continuous.sub continuous_const continuous_id), - have : tendsto (λ n, ∥u - w n∥) at_top (nhds ∥u - v∥), + have : tendsto (λ n, ‖u - w n‖) at_top (nhds ‖u - v‖), convert (tendsto.comp h_cont.continuous_at w_tendsto), exact tendsto_nhds_unique this norm_tendsto, exact subtype.mem _ @@ -172,47 +172,47 @@ end /-- Characterization of minimizers for the projection on a convex set in a real inner product space. -/ theorem norm_eq_infi_iff_real_inner_le_zero {K : set F} (h : convex ℝ K) {u : F} {v : F} - (hv : v ∈ K) : ∥u - v∥ = (⨅ w : K, ∥u - w∥) ↔ ∀ w ∈ K, ⟪u - v, w - v⟫_ℝ ≤ 0 := + (hv : v ∈ K) : ‖u - v‖ = (⨅ w : K, ‖u - w‖) ↔ ∀ w ∈ K, ⟪u - v, w - v⟫_ℝ ≤ 0 := iff.intro begin assume eq w hw, - let δ := ⨅ w : K, ∥u - w∥, let p := ⟪u - v, w - v⟫_ℝ, let q := ∥w - v∥^2, + let δ := ⨅ w : K, ‖u - w‖, let p := ⟪u - v, w - v⟫_ℝ, let q := ‖w - v‖^2, letI : nonempty K := ⟨⟨v, hv⟩⟩, have zero_le_δ : 0 ≤ δ, apply le_cinfi, intro, exact norm_nonneg _, - have δ_le : ∀ w : K, δ ≤ ∥u - w∥, + have δ_le : ∀ w : K, δ ≤ ‖u - w‖, assume w, apply cinfi_le, use (0:ℝ), rintros _ ⟨_, rfl⟩, exact norm_nonneg _, - have δ_le' : ∀ w ∈ K, δ ≤ ∥u - w∥ := assume w hw, δ_le ⟨w, hw⟩, + have δ_le' : ∀ w ∈ K, δ ≤ ‖u - w‖ := assume w hw, δ_le ⟨w, hw⟩, have : ∀θ:ℝ, 0 < θ → θ ≤ 1 → 2 * p ≤ θ * q, assume θ hθ₁ hθ₂, - have : ∥u - v∥^2 ≤ ∥u - v∥^2 - 2 * θ * ⟪u - v, w - v⟫_ℝ + θ*θ*∥w - v∥^2 := + have : ‖u - v‖^2 ≤ ‖u - v‖^2 - 2 * θ * ⟪u - v, w - v⟫_ℝ + θ*θ*‖w - v‖^2 := calc - ∥u - v∥^2 ≤ ∥u - (θ•w + (1-θ)•v)∥^2 : + ‖u - v‖^2 ≤ ‖u - (θ•w + (1-θ)•v)‖^2 : begin simp only [sq], apply mul_self_le_mul_self (norm_nonneg _), rw [eq], apply δ_le', apply h hw hv, exacts [le_of_lt hθ₁, sub_nonneg.2 hθ₂, add_sub_cancel'_right _ _], end - ... = ∥(u - v) - θ • (w - v)∥^2 : + ... = ‖(u - v) - θ • (w - v)‖^2 : begin have : u - (θ•w + (1-θ)•v) = (u - v) - θ • (w - v), { rw [smul_sub, sub_smul, one_smul], simp only [sub_eq_add_neg, add_comm, add_left_comm, add_assoc, neg_add_rev] }, rw this end - ... = ∥u - v∥^2 - 2 * θ * inner (u - v) (w - v) + θ*θ*∥w - v∥^2 : + ... = ‖u - v‖^2 - 2 * θ * inner (u - v) (w - v) + θ*θ*‖w - v‖^2 : begin rw [norm_sub_sq, inner_smul_right, norm_smul], simp only [sq], - show ∥u-v∥*∥u-v∥-2*(θ*inner(u-v)(w-v))+absR (θ)*∥w-v∥*(absR (θ)*∥w-v∥)= - ∥u-v∥*∥u-v∥-2*θ*inner(u-v)(w-v)+θ*θ*(∥w-v∥*∥w-v∥), + show ‖u-v‖*‖u-v‖-2*(θ*inner(u-v)(w-v))+absR (θ)*‖w-v‖*(absR (θ)*‖w-v‖)= + ‖u-v‖*‖u-v‖-2*θ*inner(u-v)(w-v)+θ*θ*(‖w-v‖*‖w-v‖), rw abs_of_pos hθ₁, ring end, - have eq₁ : ∥u-v∥^2-2*θ*inner(u-v)(w-v)+θ*θ*∥w-v∥^2=∥u-v∥^2+(θ*θ*∥w-v∥^2-2*θ*inner(u-v)(w-v)), + have eq₁ : ‖u-v‖^2-2*θ*inner(u-v)(w-v)+θ*θ*‖w-v‖^2=‖u-v‖^2+(θ*θ*‖w-v‖^2-2*θ*inner(u-v)(w-v)), by abel, rw [eq₁, le_add_iff_nonneg_right] at this, - have eq₂ : θ*θ*∥w-v∥^2-2*θ*inner(u-v)(w-v)=θ*(θ*∥w-v∥^2-2*inner(u-v)(w-v)), ring, + have eq₂ : θ*θ*‖w-v‖^2-2*θ*inner(u-v)(w-v)=θ*(θ*‖w-v‖^2-2*inner(u-v)(w-v)), ring, rw eq₂ at this, have := le_of_sub_nonneg (nonneg_of_mul_nonneg_right this hθ₁), exact this, @@ -242,13 +242,13 @@ begin apply nonneg_le_nonneg_of_sq_le_sq (norm_nonneg _), have := h w w.2, calc - ∥u - v∥ * ∥u - v∥ ≤ ∥u - v∥ * ∥u - v∥ - 2 * inner (u - v) ((w:F) - v) : by linarith - ... ≤ ∥u - v∥^2 - 2 * inner (u - v) ((w:F) - v) + ∥(w:F) - v∥^2 : + ‖u - v‖ * ‖u - v‖ ≤ ‖u - v‖ * ‖u - v‖ - 2 * inner (u - v) ((w:F) - v) : by linarith + ... ≤ ‖u - v‖^2 - 2 * inner (u - v) ((w:F) - v) + ‖(w:F) - v‖^2 : by { rw sq, refine le_add_of_nonneg_right _, exact sq_nonneg _ } - ... = ∥(u - v) - (w - v)∥^2 : norm_sub_sq.symm - ... = ∥u - w∥ * ∥u - w∥ : + ... = ‖(u - v) - (w - v)‖^2 : norm_sub_sq.symm + ... = ‖u - w‖ * ‖u - w‖ : by { have : (u - v) - (w - v) = u - w, abel, rw [this, sq] } }, - { show (⨅ (w : K), ∥u - w∥) ≤ (λw:K, ∥u - w∥) ⟨v, hv⟩, + { show (⨅ (w : K), ‖u - w‖) ≤ (λw:K, ‖u - w‖) ⟨v, hv⟩, apply cinfi_le, use 0, rintros y ⟨z, rfl⟩, exact norm_nonneg _ } end @@ -257,11 +257,11 @@ variables (K : submodule 𝕜 E) /-- Existence of projections on complete subspaces. Let `u` be a point in an inner product space, and let `K` be a nonempty complete subspace. -Then there exists a (unique) `v` in `K` that minimizes the distance `∥u - v∥` to `u`. +Then there exists a (unique) `v` in `K` that minimizes the distance `‖u - v‖` to `u`. This point `v` is usually called the orthogonal projection of `u` onto `K`. -/ theorem exists_norm_eq_infi_of_complete_subspace - (h : is_complete (↑K : set E)) : ∀ u : E, ∃ v ∈ K, ∥u - v∥ = ⨅ w : (K : set E), ∥u - w∥ := + (h : is_complete (↑K : set E)) : ∀ u : E, ∃ v ∈ K, ‖u - v‖ = ⨅ w : (K : set E), ‖u - w‖ := begin letI : inner_product_space ℝ E := inner_product_space.is_R_or_C_to_real 𝕜 E, letI : module ℝ E := restrict_scalars.module ℝ 𝕜 E, @@ -272,13 +272,13 @@ end /-- Characterization of minimizers in the projection on a subspace, in the real case. Let `u` be a point in a real inner product space, and let `K` be a nonempty subspace. -Then point `v` minimizes the distance `∥u - v∥` over points in `K` if and only if +Then point `v` minimizes the distance `‖u - v‖` over points in `K` if and only if for all `w ∈ K`, `⟪u - v, w⟫ = 0` (i.e., `u - v` is orthogonal to the subspace `K`). This is superceded by `norm_eq_infi_iff_inner_eq_zero` that gives the same conclusion over any `is_R_or_C` field. -/ theorem norm_eq_infi_iff_real_inner_eq_zero (K : submodule ℝ F) {u : F} {v : F} - (hv : v ∈ K) : ∥u - v∥ = (⨅ w : (↑K : set F), ∥u - w∥) ↔ ∀ w ∈ K, ⟪u - v, w⟫_ℝ = 0 := + (hv : v ∈ K) : ‖u - v‖ = (⨅ w : (↑K : set F), ‖u - w‖) ↔ ∀ w ∈ K, ⟪u - v, w⟫_ℝ = 0 := iff.intro begin assume h, @@ -315,11 +315,11 @@ end /-- Characterization of minimizers in the projection on a subspace. Let `u` be a point in an inner product space, and let `K` be a nonempty subspace. -Then point `v` minimizes the distance `∥u - v∥` over points in `K` if and only if +Then point `v` minimizes the distance `‖u - v‖` over points in `K` if and only if for all `w ∈ K`, `⟪u - v, w⟫ = 0` (i.e., `u - v` is orthogonal to the subspace `K`) -/ theorem norm_eq_infi_iff_inner_eq_zero {u : E} {v : E} - (hv : v ∈ K) : ∥u - v∥ = (⨅ w : K, ∥u - w∥) ↔ ∀ w ∈ K, ⟪u - v, w⟫ = 0 := + (hv : v ∈ K) : ‖u - v‖ = (⨅ w : K, ‖u - w‖) ↔ ∀ w ∈ K, ⟪u - v, w⟫ = 0 := begin letI : inner_product_space ℝ E := inner_product_space.is_R_or_C_to_real 𝕜 E, letI : module ℝ E := restrict_scalars.module ℝ 𝕜 E, @@ -395,8 +395,8 @@ end variables (K) lemma orthogonal_projection_fn_norm_sq (v : E) : - ∥v∥ * ∥v∥ = ∥v - (orthogonal_projection_fn K v)∥ * ∥v - (orthogonal_projection_fn K v)∥ - + ∥orthogonal_projection_fn K v∥ * ∥orthogonal_projection_fn K v∥ := + ‖v‖ * ‖v‖ = ‖v - (orthogonal_projection_fn K v)‖ * ‖v - (orthogonal_projection_fn K v)‖ + + ‖orthogonal_projection_fn K v‖ * ‖orthogonal_projection_fn K v‖ := begin set p := orthogonal_projection_fn K v, have h' : ⟪v - p, p⟫ = 0, @@ -433,7 +433,7 @@ linear_map.mk_continuous (λ x, begin simp only [one_mul, linear_map.coe_mk], refine le_of_pow_le_pow 2 (norm_nonneg _) (by norm_num) _, - change ∥orthogonal_projection_fn K x∥ ^ 2 ≤ ∥x∥ ^ 2, + change ‖orthogonal_projection_fn K x‖ ^ 2 ≤ ‖x‖ ^ 2, nlinarith [orthogonal_projection_fn_norm_sq K x] end) @@ -466,9 +466,9 @@ lemma eq_orthogonal_projection_of_mem_of_inner_eq_zero (orthogonal_projection K u : E) = v := eq_orthogonal_projection_fn_of_mem_of_inner_eq_zero hvm hvo -/-- The orthogonal projection of `y` on `U` minimizes the distance `∥y - x∥` for `x ∈ U`. -/ +/-- The orthogonal projection of `y` on `U` minimizes the distance `‖y - x‖` for `x ∈ U`. -/ lemma orthogonal_projection_minimal {U : submodule 𝕜 E} [complete_space U] (y : E) : - ∥y - orthogonal_projection U y∥ = ⨅ x : U, ∥y - x∥ := + ‖y - orthogonal_projection U y‖ = ⨅ x : U, ‖y - x‖ := begin rw norm_eq_infi_iff_inner_eq_zero _ (submodule.coe_mem _), exact orthogonal_projection_inner_eq_zero _ @@ -538,44 +538,44 @@ by ext variables (K) /-- The orthogonal projection has norm `≤ 1`. -/ -lemma orthogonal_projection_norm_le : ∥orthogonal_projection K∥ ≤ 1 := +lemma orthogonal_projection_norm_le : ‖orthogonal_projection K‖ ≤ 1 := linear_map.mk_continuous_norm_le _ (by norm_num) _ variables (𝕜) lemma smul_orthogonal_projection_singleton {v : E} (w : E) : - (∥v∥ ^ 2 : 𝕜) • (orthogonal_projection (𝕜 ∙ v) w : E) = ⟪v, w⟫ • v := + (‖v‖ ^ 2 : 𝕜) • (orthogonal_projection (𝕜 ∙ v) w : E) = ⟪v, w⟫ • v := begin - suffices : ↑(orthogonal_projection (𝕜 ∙ v) ((∥v∥ ^ 2 : 𝕜) • w)) = ⟪v, w⟫ • v, + suffices : ↑(orthogonal_projection (𝕜 ∙ v) ((‖v‖ ^ 2 : 𝕜) • w)) = ⟪v, w⟫ • v, { simpa using this }, apply eq_orthogonal_projection_of_mem_of_inner_eq_zero, { rw submodule.mem_span_singleton, use ⟪v, w⟫ }, { intros x hx, obtain ⟨c, rfl⟩ := submodule.mem_span_singleton.mp hx, - have hv : ↑∥v∥ ^ 2 = ⟪v, v⟫ := by { norm_cast, simp [norm_sq_eq_inner] }, + have hv : ↑‖v‖ ^ 2 = ⟪v, v⟫ := by { norm_cast, simp [norm_sq_eq_inner] }, simp [inner_sub_left, inner_smul_left, inner_smul_right, map_div₀, mul_comm, hv, inner_product_space.conj_sym, hv] } end /-- Formula for orthogonal projection onto a single vector. -/ lemma orthogonal_projection_singleton {v : E} (w : E) : - (orthogonal_projection (𝕜 ∙ v) w : E) = (⟪v, w⟫ / ∥v∥ ^ 2) • v := + (orthogonal_projection (𝕜 ∙ v) w : E) = (⟪v, w⟫ / ‖v‖ ^ 2) • v := begin by_cases hv : v = 0, { rw [hv, eq_orthogonal_projection_of_eq_submodule (submodule.span_zero_singleton 𝕜)], { simp }, { apply_instance } }, - have hv' : ∥v∥ ≠ 0 := ne_of_gt (norm_pos_iff.mpr hv), - have key : ((∥v∥ ^ 2 : 𝕜)⁻¹ * ∥v∥ ^ 2) • ↑(orthogonal_projection (𝕜 ∙ v) w) - = ((∥v∥ ^ 2 : 𝕜)⁻¹ * ⟪v, w⟫) • v, + have hv' : ‖v‖ ≠ 0 := ne_of_gt (norm_pos_iff.mpr hv), + have key : ((‖v‖ ^ 2 : 𝕜)⁻¹ * ‖v‖ ^ 2) • ↑(orthogonal_projection (𝕜 ∙ v) w) + = ((‖v‖ ^ 2 : 𝕜)⁻¹ * ⟪v, w⟫) • v, { simp [mul_smul, smul_orthogonal_projection_singleton 𝕜 w] }, convert key; field_simp [hv'] end /-- Formula for orthogonal projection onto a single unit vector. -/ -lemma orthogonal_projection_unit_singleton {v : E} (hv : ∥v∥ = 1) (w : E) : +lemma orthogonal_projection_unit_singleton {v : E} (hv : ‖v‖ = 1) (w : E) : (orthogonal_projection (𝕜 ∙ v) w : E) = ⟪v, w⟫ • v := by { rw ← smul_orthogonal_projection_singleton 𝕜 w, simp [hv] } @@ -808,7 +808,7 @@ begin have proj_x : ∀ i, orthogonal_projection (U i) x = orthogonal_projection (U i) y := λ i, (orthogonal_projection_orthogonal_projection_of_le ((le_supr U i).trans (supr U).submodule_topological_closure) _).symm, - suffices : ∀ ε > 0, ∃ I, ∀ i ≥ I, ∥(orthogonal_projection (U i) y : E) - y∥ < ε, + suffices : ∀ ε > 0, ∃ I, ∀ i ≥ I, ‖(orthogonal_projection (U i) y : E) - y‖ < ε, { simpa only [proj_x, normed_add_comm_group.tendsto_at_top] using this }, intros ε hε, obtain ⟨a, ha, hay⟩ : ∃ a ∈ ⨆ i, U i, dist y a < ε, @@ -821,7 +821,7 @@ begin refine ⟨I, λ i (hi : I ≤ i), _⟩, rw [norm_sub_rev, orthogonal_projection_minimal], refine lt_of_le_of_lt _ hay, - change _ ≤ ∥y - (⟨a, hU hi hI⟩ : U i)∥, + change _ ≤ ‖y - (⟨a, hU hi hI⟩ : U i)‖, exact cinfi_le ⟨0, set.forall_range_iff.mpr $ λ _, norm_nonneg _⟩ _, end @@ -905,7 +905,7 @@ lemma reflection_orthogonal_complement_singleton_eq_neg [complete_space E] (v : reflection (𝕜 ∙ v)ᗮ v = -v := reflection_mem_subspace_orthogonal_precomplement_eq_neg (submodule.mem_span_singleton_self v) -lemma reflection_sub [complete_space F] {v w : F} (h : ∥v∥ = ∥w∥) : +lemma reflection_sub [complete_space F] {v w : F} (h : ‖v‖ = ‖w‖) : reflection (ℝ ∙ (v - w))ᗮ v = w := begin set R : F ≃ₗᵢ[ℝ] F := reflection (ℝ ∙ (v - w))ᗮ, @@ -942,7 +942,7 @@ end /-- The Pythagorean theorem, for an orthogonal projection.-/ lemma norm_sq_eq_add_norm_sq_projection (x : E) (S : submodule 𝕜 E) [complete_space E] [complete_space S] : - ∥x∥^2 = ∥orthogonal_projection S x∥^2 + ∥orthogonal_projection Sᗮ x∥^2 := + ‖x‖^2 = ‖orthogonal_projection S x‖^2 + ‖orthogonal_projection Sᗮ x‖^2 := begin let p1 := orthogonal_projection S, let p2 := orthogonal_projection Sᗮ, @@ -1194,8 +1194,8 @@ begin -- ** direction 1: nonempty orthogonal complement implies nonmaximal rintros ⟨x, hx', hx⟩, -- take a nonzero vector and normalize it - let e := (∥x∥⁻¹ : 𝕜) • x, - have he : ∥e∥ = 1 := by simp [e, norm_smul_inv_norm hx], + let e := (‖x‖⁻¹ : 𝕜) • x, + have he : ‖e‖ = 1 := by simp [e, norm_smul_inv_norm hx], have he' : e ∈ (span 𝕜 v)ᗮ := smul_mem' _ _ hx', have he'' : e ∉ v, { intros hev, diff --git a/src/analysis/inner_product_space/rayleigh.lean b/src/analysis/inner_product_space/rayleigh.lean index 66fc2437ad046..57359dd8d9072 100644 --- a/src/analysis/inner_product_space/rayleigh.lean +++ b/src/analysis/inner_product_space/rayleigh.lean @@ -13,24 +13,24 @@ import linear_algebra.eigenspace # The Rayleigh quotient The Rayleigh quotient of a self-adjoint operator `T` on an inner product space `E` is the function -`λ x, ⟪T x, x⟫ / ∥x∥ ^ 2`. +`λ x, ⟪T x, x⟫ / ‖x‖ ^ 2`. The main results of this file are `is_self_adjoint.has_eigenvector_of_is_max_on` and `is_self_adjoint.has_eigenvector_of_is_min_on`, which state that if `E` is complete, and if the Rayleigh quotient attains its global maximum/minimum over some sphere at the point `x₀`, then `x₀` -is an eigenvector of `T`, and the `supr`/`infi` of `λ x, ⟪T x, x⟫ / ∥x∥ ^ 2` is the corresponding +is an eigenvector of `T`, and the `supr`/`infi` of `λ x, ⟪T x, x⟫ / ‖x‖ ^ 2` is the corresponding eigenvalue. The corollaries `is_self_adjoint.has_eigenvalue_supr_of_finite_dimensional` and `is_self_adjoint.has_eigenvalue_supr_of_finite_dimensional` state that if `E` is finite-dimensional and nontrivial, then `T` has some (nonzero) eigenvectors with eigenvalue the `supr`/`infi` of -`λ x, ⟪T x, x⟫ / ∥x∥ ^ 2`. +`λ x, ⟪T x, x⟫ / ‖x‖ ^ 2`. ## TODO A slightly more elaborate corollary is that if `E` is complete and `T` is a compact operator, then -`T` has some (nonzero) eigenvector with eigenvalue either `⨆ x, ⟪T x, x⟫ / ∥x∥ ^ 2` or -`⨅ x, ⟪T x, x⟫ / ∥x∥ ^ 2` (not necessarily both). +`T` has some (nonzero) eigenvector with eigenvalue either `⨆ x, ⟪T x, x⟫ / ‖x‖ ^ 2` or +`⨅ x, ⟪T x, x⟫ / ‖x‖ ^ 2` (not necessarily both). -/ @@ -44,15 +44,15 @@ open module.End metric namespace continuous_linear_map variables (T : E →L[𝕜] E) -local notation `rayleigh_quotient` := λ x : E, T.re_apply_inner_self x / ∥(x:E)∥ ^ 2 +local notation `rayleigh_quotient` := λ x : E, T.re_apply_inner_self x / ‖(x:E)‖ ^ 2 lemma rayleigh_smul (x : E) {c : 𝕜} (hc : c ≠ 0) : rayleigh_quotient (c • x) = rayleigh_quotient x := begin by_cases hx : x = 0, { simp [hx] }, - have : ∥c∥ ≠ 0 := by simp [hc], - have : ∥x∥ ≠ 0 := by simp [hx], + have : ‖c‖ ≠ 0 := by simp [hc], + have : ‖x‖ ≠ 0 := by simp [hx], field_simp [norm_smul, T.re_apply_inner_self_smul], ring end @@ -63,8 +63,8 @@ begin ext a, split, { rintros ⟨x, (hx : x ≠ 0), hxT⟩, - have : ∥x∥ ≠ 0 := by simp [hx], - let c : 𝕜 := ↑∥x∥⁻¹ * r, + have : ‖x‖ ≠ 0 := by simp [hx], + let c : 𝕜 := ↑‖x‖⁻¹ * r, have : c ≠ 0 := by simp [c, hx, hr.ne'], refine ⟨c • x, _, _⟩, { field_simp [norm_smul, is_R_or_C.norm_eq_abs, abs_of_nonneg hr.le] }, @@ -101,18 +101,18 @@ begin end variables [complete_space F] {T : F →L[ℝ] F} -local notation `rayleigh_quotient` := λ x : F, T.re_apply_inner_self x / ∥(x:F)∥ ^ 2 +local notation `rayleigh_quotient` := λ x : F, T.re_apply_inner_self x / ‖(x:F)‖ ^ 2 lemma linearly_dependent_of_is_local_extr_on (hT : is_self_adjoint T) - {x₀ : F} (hextr : is_local_extr_on T.re_apply_inner_self (sphere (0:F) ∥x₀∥) x₀) : + {x₀ : F} (hextr : is_local_extr_on T.re_apply_inner_self (sphere (0:F) ‖x₀‖) x₀) : ∃ a b : ℝ, (a, b) ≠ 0 ∧ a • x₀ + b • T x₀ = 0 := begin - have H : is_local_extr_on T.re_apply_inner_self {x : F | ∥x∥ ^ 2 = ∥x₀∥ ^ 2} x₀, + have H : is_local_extr_on T.re_apply_inner_self {x : F | ‖x‖ ^ 2 = ‖x₀‖ ^ 2} x₀, { convert hextr, ext x, simp [dist_eq_norm] }, -- find Lagrange multipliers for the function `T.re_apply_inner_self` and the - -- hypersurface-defining function `λ x, ∥x∥ ^ 2` + -- hypersurface-defining function `λ x, ‖x‖ ^ 2` obtain ⟨a, b, h₁, h₂⟩ := is_local_extr_on.exists_multipliers_of_has_strict_fderiv_at_1d H (has_strict_fderiv_at_norm_sq x₀) (hT.is_symmetric.has_strict_fderiv_at_re_apply_inner_self x₀), refine ⟨a, b, h₁, _⟩, @@ -124,7 +124,7 @@ begin end lemma eq_smul_self_of_is_local_extr_on_real (hT : is_self_adjoint T) - {x₀ : F} (hextr : is_local_extr_on T.re_apply_inner_self (sphere (0:F) ∥x₀∥) x₀) : + {x₀ : F} (hextr : is_local_extr_on T.re_apply_inner_self (sphere (0:F) ‖x₀‖) x₀) : T x₀ = (rayleigh_quotient x₀) • x₀ := begin obtain ⟨a, b, h₁, h₂⟩ := hT.linearly_dependent_of_is_local_extr_on hextr, @@ -141,7 +141,7 @@ begin apply smul_right_injective F hb, simp [c, eq_neg_of_add_eq_zero_left h₂, ← mul_smul, this] }, convert hc, - have : ∥x₀∥ ≠ 0 := by simp [hx₀], + have : ‖x₀‖ ≠ 0 := by simp [hx₀], field_simp, simpa [inner_smul_left, real_inner_self_eq_norm_mul_norm, sq] using congr_arg (λ x, ⟪x, x₀⟫_ℝ) hc, end @@ -150,10 +150,10 @@ end real section complete_space variables [complete_space E] {T : E →L[𝕜] E} -local notation `rayleigh_quotient` := λ x : E, T.re_apply_inner_self x / ∥(x:E)∥ ^ 2 +local notation `rayleigh_quotient` := λ x : E, T.re_apply_inner_self x / ‖(x:E)‖ ^ 2 lemma eq_smul_self_of_is_local_extr_on (hT : is_self_adjoint T) {x₀ : E} - (hextr : is_local_extr_on T.re_apply_inner_self (sphere (0:E) ∥x₀∥) x₀) : + (hextr : is_local_extr_on T.re_apply_inner_self (sphere (0:E) ‖x₀‖) x₀) : T x₀ = (↑(rayleigh_quotient x₀) : 𝕜) • x₀ := begin letI := inner_product_space.is_R_or_C_to_real 𝕜 E, @@ -164,7 +164,7 @@ end /-- For a self-adjoint operator `T`, a local extremum of the Rayleigh quotient of `T` on a sphere centred at the origin is an eigenvector of `T`. -/ lemma has_eigenvector_of_is_local_extr_on (hT : is_self_adjoint T) {x₀ : E} - (hx₀ : x₀ ≠ 0) (hextr : is_local_extr_on T.re_apply_inner_self (sphere (0:E) ∥x₀∥) x₀) : + (hx₀ : x₀ ≠ 0) (hextr : is_local_extr_on T.re_apply_inner_self (sphere (0:E) ‖x₀‖) x₀) : has_eigenvector (T : E →ₗ[𝕜] E) ↑(rayleigh_quotient x₀) x₀ := begin refine ⟨_, hx₀⟩, @@ -176,38 +176,38 @@ end at the origin is an eigenvector of `T`, with eigenvalue the global supremum of the Rayleigh quotient. -/ lemma has_eigenvector_of_is_max_on (hT : is_self_adjoint T) {x₀ : E} - (hx₀ : x₀ ≠ 0) (hextr : is_max_on T.re_apply_inner_self (sphere (0:E) ∥x₀∥) x₀) : + (hx₀ : x₀ ≠ 0) (hextr : is_max_on T.re_apply_inner_self (sphere (0:E) ‖x₀‖) x₀) : has_eigenvector (T : E →ₗ[𝕜] E) ↑(⨆ x : {x : E // x ≠ 0}, rayleigh_quotient x) x₀ := begin convert hT.has_eigenvector_of_is_local_extr_on hx₀ (or.inr hextr.localize), - have hx₀' : 0 < ∥x₀∥ := by simp [hx₀], - have hx₀'' : x₀ ∈ sphere (0:E) (∥x₀∥) := by simp, + have hx₀' : 0 < ‖x₀‖ := by simp [hx₀], + have hx₀'' : x₀ ∈ sphere (0:E) (‖x₀‖) := by simp, rw T.supr_rayleigh_eq_supr_rayleigh_sphere hx₀', refine is_max_on.supr_eq hx₀'' _, intros x hx, dsimp, - have : ∥x∥ = ∥x₀∥ := by simpa using hx, + have : ‖x‖ = ‖x₀‖ := by simpa using hx, rw this, - exact div_le_div_of_le (sq_nonneg ∥x₀∥) (hextr hx) + exact div_le_div_of_le (sq_nonneg ‖x₀‖) (hextr hx) end /-- For a self-adjoint operator `T`, a minimum of the Rayleigh quotient of `T` on a sphere centred at the origin is an eigenvector of `T`, with eigenvalue the global infimum of the Rayleigh quotient. -/ lemma has_eigenvector_of_is_min_on (hT : is_self_adjoint T) {x₀ : E} - (hx₀ : x₀ ≠ 0) (hextr : is_min_on T.re_apply_inner_self (sphere (0:E) ∥x₀∥) x₀) : + (hx₀ : x₀ ≠ 0) (hextr : is_min_on T.re_apply_inner_self (sphere (0:E) ‖x₀‖) x₀) : has_eigenvector (T : E →ₗ[𝕜] E) ↑(⨅ x : {x : E // x ≠ 0}, rayleigh_quotient x) x₀ := begin convert hT.has_eigenvector_of_is_local_extr_on hx₀ (or.inl hextr.localize), - have hx₀' : 0 < ∥x₀∥ := by simp [hx₀], - have hx₀'' : x₀ ∈ sphere (0:E) (∥x₀∥) := by simp, + have hx₀' : 0 < ‖x₀‖ := by simp [hx₀], + have hx₀'' : x₀ ∈ sphere (0:E) (‖x₀‖) := by simp, rw T.infi_rayleigh_eq_infi_rayleigh_sphere hx₀', refine is_min_on.infi_eq hx₀'' _, intros x hx, dsimp, - have : ∥x∥ = ∥x₀∥ := by simpa using hx, + have : ‖x‖ = ‖x₀‖ := by simpa using hx, rw this, - exact div_le_div_of_le (sq_nonneg ∥x₀∥) (hextr hx) + exact div_le_div_of_le (sq_nonneg ‖x₀‖) (hextr hx) end end complete_space @@ -226,21 +226,21 @@ include _i /-- The supremum of the Rayleigh quotient of a symmetric operator `T` on a nontrivial finite-dimensional vector space is an eigenvalue for that operator. -/ lemma has_eigenvalue_supr_of_finite_dimensional (hT : T.is_symmetric) : - has_eigenvalue T ↑(⨆ x : {x : E // x ≠ 0}, is_R_or_C.re ⟪T x, x⟫ / ∥(x:E)∥ ^ 2) := + has_eigenvalue T ↑(⨆ x : {x : E // x ≠ 0}, is_R_or_C.re ⟪T x, x⟫ / ‖(x:E)‖ ^ 2) := begin haveI := finite_dimensional.proper_is_R_or_C 𝕜 E, let T' := hT.to_self_adjoint, obtain ⟨x, hx⟩ : ∃ x : E, x ≠ 0 := exists_ne 0, - have H₁ : is_compact (sphere (0:E) ∥x∥) := is_compact_sphere _ _, - have H₂ : (sphere (0:E) ∥x∥).nonempty := ⟨x, by simp⟩, + have H₁ : is_compact (sphere (0:E) ‖x‖) := is_compact_sphere _ _, + have H₂ : (sphere (0:E) ‖x‖).nonempty := ⟨x, by simp⟩, -- key point: in finite dimension, a continuous function on the sphere has a max obtain ⟨x₀, hx₀', hTx₀⟩ := H₁.exists_forall_ge H₂ T'.val.re_apply_inner_self_continuous.continuous_on, - have hx₀ : ∥x₀∥ = ∥x∥ := by simpa using hx₀', - have : is_max_on T'.val.re_apply_inner_self (sphere 0 ∥x₀∥) x₀, + have hx₀ : ‖x₀‖ = ‖x‖ := by simpa using hx₀', + have : is_max_on T'.val.re_apply_inner_self (sphere 0 ‖x₀‖) x₀, { simpa only [← hx₀] using hTx₀ }, have hx₀_ne : x₀ ≠ 0, - { have : ∥x₀∥ ≠ 0 := by simp only [hx₀, norm_eq_zero, hx, ne.def, not_false_iff], + { have : ‖x₀‖ ≠ 0 := by simp only [hx₀, norm_eq_zero, hx, ne.def, not_false_iff], simpa [← norm_eq_zero, ne.def] }, exact has_eigenvalue_of_has_eigenvector (T'.prop.has_eigenvector_of_is_max_on hx₀_ne this) end @@ -248,21 +248,21 @@ end /-- The infimum of the Rayleigh quotient of a symmetric operator `T` on a nontrivial finite-dimensional vector space is an eigenvalue for that operator. -/ lemma has_eigenvalue_infi_of_finite_dimensional (hT : T.is_symmetric) : - has_eigenvalue T ↑(⨅ x : {x : E // x ≠ 0}, is_R_or_C.re ⟪T x, x⟫ / ∥(x:E)∥ ^ 2) := + has_eigenvalue T ↑(⨅ x : {x : E // x ≠ 0}, is_R_or_C.re ⟪T x, x⟫ / ‖(x:E)‖ ^ 2) := begin haveI := finite_dimensional.proper_is_R_or_C 𝕜 E, let T' := hT.to_self_adjoint, obtain ⟨x, hx⟩ : ∃ x : E, x ≠ 0 := exists_ne 0, - have H₁ : is_compact (sphere (0:E) ∥x∥) := is_compact_sphere _ _, - have H₂ : (sphere (0:E) ∥x∥).nonempty := ⟨x, by simp⟩, + have H₁ : is_compact (sphere (0:E) ‖x‖) := is_compact_sphere _ _, + have H₂ : (sphere (0:E) ‖x‖).nonempty := ⟨x, by simp⟩, -- key point: in finite dimension, a continuous function on the sphere has a min obtain ⟨x₀, hx₀', hTx₀⟩ := H₁.exists_forall_le H₂ T'.val.re_apply_inner_self_continuous.continuous_on, - have hx₀ : ∥x₀∥ = ∥x∥ := by simpa using hx₀', - have : is_min_on T'.val.re_apply_inner_self (sphere 0 ∥x₀∥) x₀, + have hx₀ : ‖x₀‖ = ‖x‖ := by simpa using hx₀', + have : is_min_on T'.val.re_apply_inner_self (sphere 0 ‖x₀‖) x₀, { simpa only [← hx₀] using hTx₀ }, have hx₀_ne : x₀ ≠ 0, - { have : ∥x₀∥ ≠ 0 := by simp only [hx₀, norm_eq_zero, hx, ne.def, not_false_iff], + { have : ‖x₀‖ ≠ 0 := by simp only [hx₀, norm_eq_zero, hx, ne.def, not_false_iff], simpa [← norm_eq_zero, ne.def] }, exact has_eigenvalue_of_has_eigenvector (T'.prop.has_eigenvector_of_is_min_on hx₀_ne this) end diff --git a/src/analysis/inner_product_space/spectrum.lean b/src/analysis/inner_product_space/spectrum.lean index df1dc284dbe4a..f6d03066c8430 100644 --- a/src/analysis/inner_product_space/spectrum.lean +++ b/src/analysis/inner_product_space/spectrum.lean @@ -250,15 +250,15 @@ section nonneg @[simp] lemma inner_product_apply_eigenvector {μ : 𝕜} {v : E} {T : E →ₗ[𝕜] E} - (h : v ∈ module.End.eigenspace T μ) : ⟪v, T v⟫ = μ * ∥v∥ ^ 2 := + (h : v ∈ module.End.eigenspace T μ) : ⟪v, T v⟫ = μ * ‖v‖ ^ 2 := by simp only [mem_eigenspace_iff.mp h, inner_smul_right, inner_self_eq_norm_sq_to_K] lemma eigenvalue_nonneg_of_nonneg {μ : ℝ} {T : E →ₗ[𝕜] E} (hμ : has_eigenvalue T μ) (hnn : ∀ (x : E), 0 ≤ is_R_or_C.re ⟪x, T x⟫) : 0 ≤ μ := begin obtain ⟨v, hv⟩ := hμ.exists_has_eigenvector, - have hpos : 0 < ∥v∥ ^ 2, by simpa only [sq_pos_iff, norm_ne_zero_iff] using hv.2, - have : is_R_or_C.re ⟪v, T v⟫ = μ * ∥v∥ ^ 2, + have hpos : 0 < ‖v‖ ^ 2, by simpa only [sq_pos_iff, norm_ne_zero_iff] using hv.2, + have : is_R_or_C.re ⟪v, T v⟫ = μ * ‖v‖ ^ 2, { exact_mod_cast congr_arg is_R_or_C.re (inner_product_apply_eigenvector hv.1) }, exact (zero_le_mul_right hpos).mp (this ▸ hnn v), end @@ -267,8 +267,8 @@ lemma eigenvalue_pos_of_pos {μ : ℝ} {T : E →ₗ[𝕜] E} (hμ : has_eigenva (hnn : ∀ (x : E), 0 < is_R_or_C.re ⟪x, T x⟫) : 0 < μ := begin obtain ⟨v, hv⟩ := hμ.exists_has_eigenvector, - have hpos : 0 < ∥v∥ ^ 2, by simpa only [sq_pos_iff, norm_ne_zero_iff] using hv.2, - have : is_R_or_C.re ⟪v, T v⟫ = μ * ∥v∥ ^ 2, + have hpos : 0 < ‖v‖ ^ 2, by simpa only [sq_pos_iff, norm_ne_zero_iff] using hv.2, + have : is_R_or_C.re ⟪v, T v⟫ = μ * ‖v‖ ^ 2, { exact_mod_cast congr_arg is_R_or_C.re (inner_product_apply_eigenvector hv.1) }, exact (zero_lt_mul_right hpos).mp (this ▸ hnn v), end diff --git a/src/analysis/inner_product_space/two_dim.lean b/src/analysis/inner_product_space/two_dim.lean index b5ebbced1efce..3134642c5bd40 100644 --- a/src/analysis/inner_product_space/two_dim.lean +++ b/src/analysis/inner_product_space/two_dim.lean @@ -29,7 +29,7 @@ product space `E`. * `orientation.kahler`: a complex-valued real-bilinear map `E →ₗ[ℝ] E →ₗ[ℝ] ℂ`. Its real part is the inner product and its imaginary part is `orientation.area_form`. For vectors `x` and `y` in `E`, - the complex number `o.kahler x y` has modulus `∥x∥ * ∥y∥`. In a later file, oriented angles + the complex number `o.kahler x y` has modulus `‖x‖ * ‖y‖`. In a later file, oriented angles (`orientation.oangle`) are defined, in such a way that the argument of `o.kahler x y` is the oriented angle from `x` to `y`. @@ -40,7 +40,7 @@ product space `E`. * `orientation.nonneg_inner_and_area_form_eq_zero_iff_same_ray`: `x`, `y` are in the same ray, if and only if `0 ≤ ⟪x, y⟫` and `ω x y = 0` -* `orientation.kahler_mul`: the identity `o.kahler x a * o.kahler a y = ∥a∥ ^ 2 * o.kahler x y` +* `orientation.kahler_mul`: the identity `o.kahler x a * o.kahler a y = ‖a‖ ^ 2 * o.kahler x y` * `complex.area_form`, `complex.right_angle_rotation`, `complex.kahler`: the concrete interpretations of `area_form`, `right_angle_rotation`, `kahler` for the oriented real inner @@ -129,13 +129,13 @@ def area_form' : E →L[ℝ] (E →L[ℝ] ℝ) := o.area_form' x = (o.area_form x).to_continuous_linear_map := rfl -lemma abs_area_form_le (x y : E) : |ω x y| ≤ ∥x∥ * ∥y∥ := +lemma abs_area_form_le (x y : E) : |ω x y| ≤ ‖x‖ * ‖y‖ := by simpa [area_form_to_volume_form, fin.prod_univ_succ] using o.abs_volume_form_apply_le ![x, y] -lemma area_form_le (x y : E) : ω x y ≤ ∥x∥ * ∥y∥ := +lemma area_form_le (x y : E) : ω x y ≤ ‖x‖ * ‖y‖ := by simpa [area_form_to_volume_form, fin.prod_univ_succ] using o.volume_form_apply_le ![x, y] -lemma abs_area_form_of_orthogonal {x y : E} (h : ⟪x, y⟫ = 0) : |ω x y| = ∥x∥ * ∥y∥ := +lemma abs_area_form_of_orthogonal {x y : E} (h : ⟪x, y⟫ = 0) : |ω x y| = ‖x‖ * ‖y‖ := begin rw [o.area_form_to_volume_form, o.abs_volume_form_apply_of_pairwise_orthogonal], { simp [fin.prod_univ_succ] }, @@ -217,7 +217,7 @@ def right_angle_rotation_aux₂ : E →ₗᵢ[ℝ] E := obtain ⟨w, hw₀⟩ : ∃ w : Kᗮ, w ≠ 0 := exists_ne 0, have hw' : ⟪x, (w:E)⟫ = 0 := inner_right_of_mem_orthogonal_singleton x w.2, -- hw'₀, have hw : (w:E) ≠ 0 := λ h, hw₀ (submodule.coe_eq_zero.mp h), - refine le_of_mul_le_mul_right _ (by rwa norm_pos_iff : 0 < ∥(w:E)∥), + refine le_of_mul_le_mul_right _ (by rwa norm_pos_iff : 0 < ‖(w:E)‖), rw ← o.abs_area_form_of_orthogonal hw', rw ← o.inner_right_angle_rotation_aux₁_left x w, exact abs_real_inner_le_norm (o.right_angle_rotation_aux₁ x) w }, @@ -366,10 +366,10 @@ def basis_right_angle_rotation (x : E) (hx : x ≠ 0) : basis (fin 2) ℝ E := ⇑(o.basis_right_angle_rotation x hx) = ![x, J x] := coe_basis_of_linear_independent_of_card_eq_finrank _ _ -/-- For vectors `a x y : E`, the identity `⟪a, x⟫ * ⟪a, y⟫ + ω a x * ω a y = ∥a∥ ^ 2 * ⟪x, y⟫`. (See +/-- For vectors `a x y : E`, the identity `⟪a, x⟫ * ⟪a, y⟫ + ω a x * ω a y = ‖a‖ ^ 2 * ⟪x, y⟫`. (See `orientation.inner_mul_inner_add_area_form_mul_area_form` for the "applied" form.)-/ lemma inner_mul_inner_add_area_form_mul_area_form' (a x : E) : - ⟪a, x⟫ • @innerₛₗ ℝ _ _ _ a + ω a x • ω a = ∥a∥ ^ 2 • @innerₛₗ ℝ _ _ _ x := + ⟪a, x⟫ • @innerₛₗ ℝ _ _ _ a + ω a x • ω a = ‖a‖ ^ 2 • @innerₛₗ ℝ _ _ _ x := begin by_cases ha : a = 0, { simp [ha] }, @@ -388,18 +388,18 @@ begin ring, } end -/-- For vectors `a x y : E`, the identity `⟪a, x⟫ * ⟪a, y⟫ + ω a x * ω a y = ∥a∥ ^ 2 * ⟪x, y⟫`. -/ +/-- For vectors `a x y : E`, the identity `⟪a, x⟫ * ⟪a, y⟫ + ω a x * ω a y = ‖a‖ ^ 2 * ⟪x, y⟫`. -/ lemma inner_mul_inner_add_area_form_mul_area_form (a x y : E) : - ⟪a, x⟫ * ⟪a, y⟫ + ω a x * ω a y = ∥a∥ ^ 2 * ⟪x, y⟫ := + ⟪a, x⟫ * ⟪a, y⟫ + ω a x * ω a y = ‖a‖ ^ 2 * ⟪x, y⟫ := congr_arg (λ f : E →ₗ[ℝ] ℝ, f y) (o.inner_mul_inner_add_area_form_mul_area_form' a x) -lemma inner_sq_add_area_form_sq (a b : E) : ⟪a, b⟫ ^ 2 + ω a b ^ 2 = ∥a∥ ^ 2 * ∥b∥ ^ 2 := +lemma inner_sq_add_area_form_sq (a b : E) : ⟪a, b⟫ ^ 2 + ω a b ^ 2 = ‖a‖ ^ 2 * ‖b‖ ^ 2 := by simpa [sq, real_inner_self_eq_norm_sq] using o.inner_mul_inner_add_area_form_mul_area_form a b b -/-- For vectors `a x y : E`, the identity `⟪a, x⟫ * ω a y - ω a x * ⟪a, y⟫ = ∥a∥ ^ 2 * ω x y`. (See +/-- For vectors `a x y : E`, the identity `⟪a, x⟫ * ω a y - ω a x * ⟪a, y⟫ = ‖a‖ ^ 2 * ω x y`. (See `orientation.inner_mul_area_form_sub` for the "applied" form.) -/ lemma inner_mul_area_form_sub' (a x : E) : - ⟪a, x⟫ • ω a - ω a x • @innerₛₗ ℝ _ _ _ a = ∥a∥ ^ 2 • ω x := + ⟪a, x⟫ • ω a - ω a x • @innerₛₗ ℝ _ _ _ a = ‖a‖ ^ 2 • ω x := begin by_cases ha : a = 0, { simp [ha] }, @@ -417,8 +417,8 @@ begin ring}, end -/-- For vectors `a x y : E`, the identity `⟪a, x⟫ * ω a y - ω a x * ⟪a, y⟫ = ∥a∥ ^ 2 * ω x y`. -/ -lemma inner_mul_area_form_sub (a x y : E) : ⟪a, x⟫ * ω a y - ω a x * ⟪a, y⟫ = ∥a∥ ^ 2 * ω x y := +/-- For vectors `a x y : E`, the identity `⟪a, x⟫ * ω a y - ω a x * ⟪a, y⟫ = ‖a‖ ^ 2 * ω x y`. -/ +lemma inner_mul_area_form_sub (a x y : E) : ⟪a, x⟫ * ω a y - ω a x * ⟪a, y⟫ = ‖a‖ ^ 2 * ω x y := congr_arg (λ f : E →ₗ[ℝ] ℝ, f y) (o.inner_mul_area_form_sub' a x) lemma nonneg_inner_and_area_form_eq_zero_iff_same_ray (x y : E) : @@ -429,7 +429,7 @@ begin split, { let a : ℝ := (o.basis_right_angle_rotation x hx).repr y 0, let b : ℝ := (o.basis_right_angle_rotation x hx).repr y 1, - suffices : 0 ≤ a * ∥x∥ ^ 2 ∧ b * ∥x∥ ^ 2 = 0 → same_ray ℝ x (a • x + b • J x), + suffices : 0 ≤ a * ‖x‖ ^ 2 ∧ b * ‖x‖ ^ 2 = 0 → same_ray ℝ x (a • x + b • J x), { -- TODO trace the `dsimp` lemmas in this block to make a single `simp only` rw ← (o.basis_right_angle_rotation x hx).sum_repr y, simp only [fin.sum_univ_succ, coe_basis_right_angle_rotation], @@ -442,7 +442,7 @@ begin o.inner_right_angle_rotation_right, o.area_form_apply_self, real_inner_self_eq_norm_sq], exact this }, rintros ⟨ha, hb⟩, - have hx' : 0 < ∥x∥ := by simpa using hx, + have hx' : 0 < ‖x‖ := by simpa using hx, have ha' : 0 ≤ a := nonneg_of_mul_nonneg_left ha (by positivity), have hb' : b = 0 := eq_zero_of_ne_zero_of_mul_right_eq_zero (pow_ne_zero 2 hx'.ne') hb, simpa [hb'] using same_ray_nonneg_smul_right x ha' }, @@ -470,7 +470,7 @@ begin simp, end -@[simp] lemma kahler_apply_self (x : E) : o.kahler x x = ∥x∥ ^ 2 := +@[simp] lemma kahler_apply_self (x : E) : o.kahler x x = ‖x‖ ^ 2 := by simp [kahler_apply_apply, real_inner_self_eq_norm_sq] @[simp] lemma kahler_right_angle_rotation_left (x y : E) : @@ -498,9 +498,9 @@ end @[simp] lemma kahler_neg_orientation (x y : E) : (-o).kahler x y = conj (o.kahler x y) := by simp [kahler_apply_apply] -lemma kahler_mul (a x y : E) : o.kahler x a * o.kahler a y = ∥a∥ ^ 2 * o.kahler x y := +lemma kahler_mul (a x y : E) : o.kahler x a * o.kahler a y = ‖a‖ ^ 2 * o.kahler x y := begin - transitivity (↑(∥a∥ ^ 2) : ℂ) * o.kahler x y, + transitivity (↑(‖a‖ ^ 2) : ℂ) * o.kahler x y, { ext, { simp only [o.kahler_apply_apply, complex.add_im, complex.add_re, complex.I_im, complex.I_re, complex.mul_im, complex.mul_re, complex.of_real_im, complex.of_real_re, complex.real_smul], @@ -513,10 +513,10 @@ begin { norm_cast }, end -lemma norm_sq_kahler (x y : E) : complex.norm_sq (o.kahler x y) = ∥x∥ ^ 2 * ∥y∥ ^ 2 := +lemma norm_sq_kahler (x y : E) : complex.norm_sq (o.kahler x y) = ‖x‖ ^ 2 * ‖y‖ ^ 2 := by simpa [kahler_apply_apply, complex.norm_sq, sq] using o.inner_sq_add_area_form_sq x y -lemma abs_kahler (x y : E) : complex.abs (o.kahler x y) = ∥x∥ * ∥y∥ := +lemma abs_kahler (x y : E) : complex.abs (o.kahler x y) = ‖x‖ * ‖y‖ := begin rw [← sq_eq_sq, complex.sq_abs], { linear_combination o.norm_sq_kahler x y }, @@ -524,11 +524,11 @@ begin { positivity } end -lemma norm_kahler (x y : E) : ∥o.kahler x y∥ = ∥x∥ * ∥y∥ := by simpa using o.abs_kahler x y +lemma norm_kahler (x y : E) : ‖o.kahler x y‖ = ‖x‖ * ‖y‖ := by simpa using o.abs_kahler x y lemma eq_zero_or_eq_zero_of_kahler_eq_zero {x y : E} (hx : o.kahler x y = 0) : x = 0 ∨ y = 0 := begin - have : ∥x∥ * ∥y∥ = 0 := by simpa [hx] using (o.norm_kahler x y).symm, + have : ‖x‖ * ‖y‖ = 0 := by simpa [hx] using (o.norm_kahler x y).symm, cases eq_zero_or_eq_zero_of_mul_eq_zero this with h h, { left, simpa using h }, diff --git a/src/analysis/locally_convex/abs_convex.lean b/src/analysis/locally_convex/abs_convex.lean index d50eabc26d284..cfefd68a51df4 100644 --- a/src/analysis/locally_convex/abs_convex.lean +++ b/src/analysis/locally_convex/abs_convex.lean @@ -162,7 +162,7 @@ begin balanced_Inter₂ (λ _ _, seminorm.balanced_ball_zero _ _), convex_Inter₂ (λ _ _, seminorm.convex_ball _ _ _)⟩, -- The only nontrivial part is to show that the ball is open - have hr' : r = ∥(r : 𝕜)∥ * 1 := by simp [abs_of_pos hr], + have hr' : r = ‖(r : 𝕜)‖ * 1 := by simp [abs_of_pos hr], have hr'' : (r : 𝕜) ≠ 0 := by simp [ne_of_gt hr], rw hr', rw ←seminorm.smul_ball_zero (norm_pos_iff.mpr hr''), diff --git a/src/analysis/locally_convex/balanced_core_hull.lean b/src/analysis/locally_convex/balanced_core_hull.lean index 1646e67c14c6d..7d363ef2594ed 100644 --- a/src/analysis/locally_convex/balanced_core_hull.lean +++ b/src/analysis/locally_convex/balanced_core_hull.lean @@ -22,7 +22,7 @@ import analysis.locally_convex.basic The balanced core and hull are implemented differently: for the core we take the obvious definition of the union over all balanced sets that are contained in `s`, whereas for the hull, we take the -union over `r • s`, for `r` the scalars with `∥r∥ ≤ 1`. We show that `balanced_hull` has the +union over `r • s`, for `r` the scalars with `‖r‖ ≤ 1`. We show that `balanced_hull` has the defining properties of a hull in `balanced.hull_minimal` and `subset_balanced_hull`. For the core we need slightly stronger assumptions to obtain a characterization as an intersection, this is `balanced_core_eq_Inter`. @@ -55,10 +55,10 @@ variables (𝕜) [has_smul 𝕜 E] {s t : set E} {x : E} def balanced_core (s : set E) := ⋃₀ {t : set E | balanced 𝕜 t ∧ t ⊆ s} /-- Helper definition to prove `balanced_core_eq_Inter`-/ -def balanced_core_aux (s : set E) := ⋂ (r : 𝕜) (hr : 1 ≤ ∥r∥), r • s +def balanced_core_aux (s : set E) := ⋂ (r : 𝕜) (hr : 1 ≤ ‖r‖), r • s /-- The smallest balanced superset of `s`.-/ -def balanced_hull (s : set E) := ⋃ (r : 𝕜) (hr : ∥r∥ ≤ 1), r • s +def balanced_hull (s : set E) := ⋃ (r : 𝕜) (hr : ‖r‖ ≤ 1), r • s variables {𝕜} @@ -70,7 +70,7 @@ eq_empty_of_subset_empty (balanced_core_subset _) lemma mem_balanced_core_iff : x ∈ balanced_core 𝕜 s ↔ ∃ t, balanced 𝕜 t ∧ t ⊆ s ∧ x ∈ t := by simp_rw [balanced_core, mem_sUnion, mem_set_of_eq, exists_prop, and_assoc] -lemma smul_balanced_core_subset (s : set E) {a : 𝕜} (ha : ∥a∥ ≤ 1) : +lemma smul_balanced_core_subset (s : set E) {a : 𝕜} (ha : ‖a‖ ≤ 1) : a • balanced_core 𝕜 s ⊆ balanced_core 𝕜 s := begin rintro x ⟨y, hy, rfl⟩, @@ -87,10 +87,10 @@ lemma balanced_core_balanced (s : set E) : balanced 𝕜 (balanced_core 𝕜 s) lemma balanced.subset_core_of_subset (hs : balanced 𝕜 s) (h : s ⊆ t) : s ⊆ balanced_core 𝕜 t := subset_sUnion_of_mem ⟨hs, h⟩ -lemma mem_balanced_core_aux_iff : x ∈ balanced_core_aux 𝕜 s ↔ ∀ r : 𝕜, 1 ≤ ∥r∥ → x ∈ r • s := +lemma mem_balanced_core_aux_iff : x ∈ balanced_core_aux 𝕜 s ↔ ∀ r : 𝕜, 1 ≤ ‖r‖ → x ∈ r • s := mem_Inter₂ -lemma mem_balanced_hull_iff : x ∈ balanced_hull 𝕜 s ↔ ∃ (r : 𝕜) (hr : ∥r∥ ≤ 1), x ∈ r • s := +lemma mem_balanced_hull_iff : x ∈ balanced_hull 𝕜 s ↔ ∃ (r : 𝕜) (hr : ‖r‖ ≤ 1), x ∈ r • s := mem_Union₂ /-- The balanced hull of `s` is minimal in the sense that it is contained in any balanced superset @@ -151,7 +151,7 @@ begin { rwa zero_smul }, rw mem_balanced_core_aux_iff at ⊢ hy, intros r hr, - have h'' : 1 ≤ ∥a⁻¹ • r∥, + have h'' : 1 ≤ ‖a⁻¹ • r‖, { rw [norm_smul, norm_inv], exact one_le_mul_of_one_le_of_one_le (one_le_inv (norm_pos_iff.mpr h) ha) hr }, have h' := hy (a⁻¹ • r) h'', @@ -171,14 +171,14 @@ lemma balanced_core_subset_balanced_core_aux : balanced_core 𝕜 s ⊆ balanced balanced_core_aux_maximal (balanced_core_subset s) (balanced_core_balanced s) lemma balanced_core_eq_Inter (hs : (0 : E) ∈ s) : - balanced_core 𝕜 s = ⋂ (r : 𝕜) (hr : 1 ≤ ∥r∥), r • s := + balanced_core 𝕜 s = ⋂ (r : 𝕜) (hr : 1 ≤ ‖r‖), r • s := begin refine balanced_core_subset_balanced_core_aux.antisymm _, refine (balanced_core_aux_balanced _).subset_core_of_subset (balanced_core_aux_subset s), exact balanced_core_subset_balanced_core_aux (balanced_core_zero_mem hs), end -lemma subset_balanced_core (ht : (0 : E) ∈ t) (hst : ∀ (a : 𝕜) (ha : ∥a∥ ≤ 1), a • s ⊆ t) : +lemma subset_balanced_core (ht : (0 : E) ∈ t) (hst : ∀ (a : 𝕜) (ha : ‖a‖ ≤ 1), a • s ⊆ t) : s ⊆ balanced_core 𝕜 t := begin rw balanced_core_eq_Inter ht, @@ -218,7 +218,7 @@ lemma balanced_core_mem_nhds_zero (hU : U ∈ 𝓝 (0 : E)) : balanced_core 𝕜 begin -- Getting neighborhoods of the origin for `0 : 𝕜` and `0 : E` obtain ⟨r, V, hr, hV, hrVU⟩ : ∃ (r : ℝ) (V : set E), 0 < r ∧ V ∈ 𝓝 (0 : E) ∧ - ∀ (c : 𝕜) (y : E), ∥c∥ < r → y ∈ V → c • y ∈ U, + ∀ (c : 𝕜) (y : E), ‖c‖ < r → y ∈ V → c • y ∈ U, { have h : filter.tendsto (λ (x : 𝕜 × E), x.fst • x.snd) (𝓝 (0,0)) (𝓝 0), from continuous_smul.tendsto' (0, 0) _ (smul_zero _), simpa only [← prod.exists', ← prod.forall', ← and_imp, ← and.assoc, exists_prop] diff --git a/src/analysis/locally_convex/basic.lean b/src/analysis/locally_convex/basic.lean index b7adfe9c226d8..9ae5564e54d2e 100644 --- a/src/analysis/locally_convex/basic.lean +++ b/src/analysis/locally_convex/basic.lean @@ -49,7 +49,7 @@ variables (𝕜) [has_smul 𝕜 E] /-- A set `A` absorbs another set `B` if `B` is contained in all scalings of `A` by elements of sufficiently large norm. -/ -def absorbs (A B : set E) := ∃ r, 0 < r ∧ ∀ a : 𝕜, r ≤ ∥a∥ → B ⊆ a • A +def absorbs (A B : set E) := ∃ r, 0 < r ∧ ∀ a : 𝕜, r ≤ ‖a‖ → B ⊆ a • A variables {𝕜} {s t u v A B : set E} @@ -99,7 +99,7 @@ end variables (𝕜) /-- A set is absorbent if it absorbs every singleton. -/ -def absorbent (A : set E) := ∀ x, ∃ r, 0 < r ∧ ∀ a : 𝕜, r ≤ ∥a∥ → x ∈ a • A +def absorbent (A : set E) := ∀ x, ∃ r, 0 < r ∧ ∀ a : 𝕜, r ≤ ‖a‖ → x ∈ a • A variables {𝕜} @@ -115,7 +115,7 @@ by simp_rw [absorbs, absorbent, singleton_subset_iff] lemma absorbent.absorbs (hs : absorbent 𝕜 s) {x : E} : absorbs 𝕜 s {x} := absorbent_iff_forall_absorbs_singleton.1 hs _ -lemma absorbent_iff_nonneg_lt : absorbent 𝕜 A ↔ ∀ x, ∃ r, 0 ≤ r ∧ ∀ ⦃a : 𝕜⦄, r < ∥a∥ → x ∈ a • A := +lemma absorbent_iff_nonneg_lt : absorbent 𝕜 A ↔ ∀ x, ∃ r, 0 ≤ r ∧ ∀ ⦃a : 𝕜⦄, r < ‖a‖ → x ∈ a • A := forall_congr $ λ x, ⟨λ ⟨r, hr, hx⟩, ⟨r, hr.le, λ a ha, hx a ha.le⟩, λ ⟨r, hr, hx⟩, ⟨r + 1, add_pos_of_nonneg_of_pos hr zero_lt_one, λ a ha, hx ((lt_add_of_pos_right r zero_lt_one).trans_le ha)⟩⟩ @@ -130,11 +130,11 @@ end variables (𝕜) /-- A set `A` is balanced if `a • A` is contained in `A` whenever `a` has norm at most `1`. -/ -def balanced (A : set E) := ∀ a : 𝕜, ∥a∥ ≤ 1 → a • A ⊆ A +def balanced (A : set E) := ∀ a : 𝕜, ‖a‖ ≤ 1 → a • A ⊆ A variables {𝕜} -lemma balanced_iff_smul_mem : balanced 𝕜 s ↔ ∀ ⦃a : 𝕜⦄, ∥a∥ ≤ 1 → ∀ ⦃x : E⦄, x ∈ s → a • x ∈ s := +lemma balanced_iff_smul_mem : balanced 𝕜 s ↔ ∀ ⦃a : 𝕜⦄, ‖a‖ ≤ 1 → ∀ ⦃x : E⦄, x ∈ s → a • x ∈ s := forall₂_congr $ λ a ha, smul_set_subset_iff alias balanced_iff_smul_mem ↔ balanced.smul_mem _ @@ -204,7 +204,7 @@ variables [normed_field 𝕜] [normed_ring 𝕝] [normed_space 𝕜 𝕝] [add_c [smul_with_zero 𝕝 E] [is_scalar_tower 𝕜 𝕝 E] {s t u v A B : set E} {x : E} {a b : 𝕜} /-- Scalar multiplication (by possibly different types) of a balanced set is monotone. -/ -lemma balanced.smul_mono (hs : balanced 𝕝 s) {a : 𝕝} {b : 𝕜} (h : ∥a∥ ≤ ∥b∥) : a • s ⊆ b • s := +lemma balanced.smul_mono (hs : balanced 𝕝 s) {a : 𝕝} {b : 𝕜} (h : ‖a‖ ≤ ‖b‖) : a • s ⊆ b • s := begin obtain rfl | hb := eq_or_ne b 0, { rw norm_zero at h, @@ -230,7 +230,7 @@ begin exact inv_le_one ha, end -lemma balanced.subset_smul (hA : balanced 𝕜 A) (ha : 1 ≤ ∥a∥) : A ⊆ a • A := +lemma balanced.subset_smul (hA : balanced 𝕜 A) (ha : 1 ≤ ‖a‖) : A ⊆ a • A := begin refine (subset_set_smul_iff₀ _).2 (hA (a⁻¹) _), { rintro rfl, @@ -240,10 +240,10 @@ begin exact inv_le_one ha } end -lemma balanced.smul_eq (hA : balanced 𝕜 A) (ha : ∥a∥ = 1) : a • A = A := +lemma balanced.smul_eq (hA : balanced 𝕜 A) (ha : ‖a‖ = 1) : a • A = A := (hA _ ha.le).antisymm $ hA.subset_smul ha.ge -lemma balanced.mem_smul_iff (hs : balanced 𝕜 s) (h : ∥a∥ = ∥b∥) : a • x ∈ s ↔ b • x ∈ s := +lemma balanced.mem_smul_iff (hs : balanced 𝕜 s) (h : ‖a‖ = ‖b‖) : a • x ∈ s ↔ b • x ∈ s := begin obtain rfl | hb := eq_or_ne b 0, { rw [norm_zero, norm_eq_zero] at h, @@ -291,10 +291,10 @@ begin (by rwa [mem_preimage, zero_smul]), have hr₃ := inv_pos.mpr (half_pos hr₁), refine ⟨(r / 2)⁻¹, hr₃, λ a ha₁, _⟩, - have ha₂ : 0 < ∥a∥ := hr₃.trans_le ha₁, + have ha₂ : 0 < ‖a‖ := hr₃.trans_le ha₁, refine (mem_smul_set_iff_inv_smul_mem₀ (norm_pos_iff.mp ha₂) _ _).2 (hw₁ $ hr₂ _), rw [metric.mem_ball, dist_zero_right, norm_inv], - calc ∥a∥⁻¹ ≤ r/2 : (inv_le (half_pos hr₁) ha₂).mp ha₁ + calc ‖a‖⁻¹ ≤ r/2 : (inv_le (half_pos hr₁) ha₂).mp ha₁ ... < r : half_lt_self hr₁, end @@ -346,7 +346,7 @@ variables [module ℝ E] [smul_comm_class ℝ 𝕜 E] lemma balanced_convex_hull_of_balanced (hs : balanced 𝕜 s) : balanced 𝕜 (convex_hull ℝ s) := begin - suffices : convex ℝ {x | ∀ a : 𝕜, ∥a∥ ≤ 1 → a • x ∈ convex_hull ℝ s}, + suffices : convex ℝ {x | ∀ a : 𝕜, ‖a‖ ≤ 1 → a • x ∈ convex_hull ℝ s}, { rw balanced_iff_smul_mem at hs ⊢, refine λ a ha x hx, convex_hull_min _ this hx a ha, exact λ y hy a ha, subset_convex_hull ℝ s (hs ha hy) }, diff --git a/src/analysis/locally_convex/bounded.lean b/src/analysis/locally_convex/bounded.lean index 84c2c9eb289b3..d26f8bcfcdcaa 100644 --- a/src/analysis/locally_convex/bounded.lean +++ b/src/analysis/locally_convex/bounded.lean @@ -222,16 +222,16 @@ begin specialize hρball a ha.le, rw [← ball_norm_seminorm 𝕜 E, seminorm.smul_ball_zero (hρ.trans ha), ball_norm_seminorm, mul_one] at hρball, - exact ⟨∥a∥, hρball.trans metric.ball_subset_closed_ball⟩ }, + exact ⟨‖a‖, hρball.trans metric.ball_subset_closed_ball⟩ }, { exact λ ⟨C, hC⟩, (is_vonN_bounded_closed_ball 𝕜 E C).subset hC } end lemma is_vonN_bounded_iff' (s : set E) : - bornology.is_vonN_bounded 𝕜 s ↔ ∃ r : ℝ, ∀ (x : E) (hx : x ∈ s), ∥x∥ ≤ r := + bornology.is_vonN_bounded 𝕜 s ↔ ∃ r : ℝ, ∀ (x : E) (hx : x ∈ s), ‖x‖ ≤ r := by rw [normed_space.is_vonN_bounded_iff, ←metric.bounded_iff_is_bounded, bounded_iff_forall_norm_le] lemma image_is_vonN_bounded_iff (f : E' → E) (s : set E') : - bornology.is_vonN_bounded 𝕜 (f '' s) ↔ ∃ r : ℝ, ∀ (x : E') (hx : x ∈ s), ∥f x∥ ≤ r := + bornology.is_vonN_bounded 𝕜 (f '' s) ↔ ∃ r : ℝ, ∀ (x : E') (hx : x ∈ s), ‖f x‖ ≤ r := by simp_rw [is_vonN_bounded_iff', set.ball_image_iff] /-- In a normed space, the von Neumann bornology (`bornology.vonN_bornology`) is equal to the diff --git a/src/analysis/locally_convex/continuous_of_bounded.lean b/src/analysis/locally_convex/continuous_of_bounded.lean index b2c2be6389a3a..0bfc937e9b80f 100644 --- a/src/analysis/locally_convex/continuous_of_bounded.lean +++ b/src/analysis/locally_convex/continuous_of_bounded.lean @@ -148,9 +148,9 @@ begin rcases hf _ h_bounded hV with ⟨r, hr, h'⟩, cases exists_nat_gt r with n hn, -- We now find a contradiction between `f (u n) ∉ V` and the absorbing property - have h1 : r ≤ ∥(n : 𝕜')∥ := + have h1 : r ≤ ‖(n : 𝕜')‖ := by { rw [is_R_or_C.norm_eq_abs, is_R_or_C.abs_cast_nat], exact hn.le }, - have hn' : 0 < ∥(n : 𝕜')∥ := lt_of_lt_of_le hr h1, + have hn' : 0 < ‖(n : 𝕜')‖ := lt_of_lt_of_le hr h1, rw [norm_pos_iff, ne.def, nat.cast_eq_zero] at hn', have h'' : f (u n) ∈ V := begin diff --git a/src/analysis/locally_convex/polar.lean b/src/analysis/locally_convex/polar.lean index 51cff6e2838b4..1fdba000214d7 100644 --- a/src/analysis/locally_convex/polar.lean +++ b/src/analysis/locally_convex/polar.lean @@ -46,23 +46,23 @@ variables [normed_comm_ring 𝕜] [add_comm_monoid E] [add_comm_monoid F] variables [module 𝕜 E] [module 𝕜 F] variables (B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜) -/-- The (absolute) polar of `s : set E` is given by the set of all `y : F` such that `∥B x y∥ ≤ 1` +/-- The (absolute) polar of `s : set E` is given by the set of all `y : F` such that `‖B x y‖ ≤ 1` for all `x ∈ s`.-/ def polar (s : set E) : set F := - {y : F | ∀ x ∈ s, ∥B x y∥ ≤ 1 } + {y : F | ∀ x ∈ s, ‖B x y‖ ≤ 1 } lemma polar_mem_iff (s : set E) (y : F) : - y ∈ B.polar s ↔ ∀ x ∈ s, ∥B x y∥ ≤ 1 := iff.rfl + y ∈ B.polar s ↔ ∀ x ∈ s, ‖B x y‖ ≤ 1 := iff.rfl lemma polar_mem (s : set E) (y : F) (hy : y ∈ B.polar s) : - ∀ x ∈ s, ∥B x y∥ ≤ 1 := hy + ∀ x ∈ s, ‖B x y‖ ≤ 1 := hy @[simp] lemma zero_mem_polar (s : set E) : (0 : F) ∈ B.polar s := λ _ _, by simp only [map_zero, norm_zero, zero_le_one] lemma polar_eq_Inter {s : set E} : - B.polar s = ⋂ x ∈ s, {y : F | ∥B x y∥ ≤ 1} := + B.polar s = ⋂ x ∈ s, {y : F | ‖B x y‖ ≤ 1} := by { ext, simp only [polar_mem_iff, set.mem_Inter, set.mem_set_of_eq] } /-- The map `B.polar : set E → set F` forms an order-reversing Galois connection with @@ -123,7 +123,7 @@ begin refine ⟨by simp only [zero_mem_polar], λ y hy, h _ (λ x, _)⟩, refine norm_le_zero_iff.mp (le_of_forall_le_of_dense $ λ ε hε, _), rcases normed_field.exists_norm_lt 𝕜 hε with ⟨c, hc, hcε⟩, - calc ∥B x y∥ = ∥c∥ * ∥B (c⁻¹ • x) y∥ : + calc ‖B x y‖ = ‖c‖ * ‖B (c⁻¹ • x) y‖ : by rw [B.map_smul, linear_map.smul_apply, algebra.id.smul_eq_mul, norm_mul, norm_inv, mul_inv_cancel_left₀ hc.ne'] ... ≤ ε * 1 : mul_le_mul hcε.le (hy _ trivial) (norm_nonneg _) hε.le diff --git a/src/analysis/locally_convex/weak_dual.lean b/src/analysis/locally_convex/weak_dual.lean index 257662dfbd58f..df9abacbce362 100644 --- a/src/analysis/locally_convex/weak_dual.lean +++ b/src/analysis/locally_convex/weak_dual.lean @@ -11,12 +11,12 @@ import analysis.locally_convex.with_seminorms # Weak Dual in Topological Vector Spaces We prove that the weak topology induced by a bilinear form `B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜` is locally -convex and we explicit give a neighborhood basis in terms of the family of seminorms `λ x, ∥B x y∥` +convex and we explicit give a neighborhood basis in terms of the family of seminorms `λ x, ‖B x y‖` for `y : F`. ## Main definitions -* `linear_map.to_seminorm`: turn a linear form `f : E →ₗ[𝕜] 𝕜` into a seminorm `λ x, ∥f x∥`. +* `linear_map.to_seminorm`: turn a linear form `f : E →ₗ[𝕜] 𝕜` into a seminorm `λ x, ‖f x‖`. * `linear_map.to_seminorm_family`: turn a bilinear form `B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜` into a map `F → seminorm 𝕜 E`. @@ -48,18 +48,18 @@ namespace linear_map variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] [add_comm_group F] [module 𝕜 F] /-- Construct a seminorm from a linear form `f : E →ₗ[𝕜] 𝕜` over a normed field `𝕜` by -`λ x, ∥f x∥` -/ +`λ x, ‖f x‖` -/ def to_seminorm (f : E →ₗ[𝕜] 𝕜) : seminorm 𝕜 E := (norm_seminorm 𝕜 𝕜).comp f lemma coe_to_seminorm {f : E →ₗ[𝕜] 𝕜} : - ⇑f.to_seminorm = λ x, ∥f x∥ := rfl + ⇑f.to_seminorm = λ x, ‖f x‖ := rfl @[simp] lemma to_seminorm_apply {f : E →ₗ[𝕜] 𝕜} {x : E} : - f.to_seminorm x = ∥f x∥ := rfl + f.to_seminorm x = ‖f x‖ := rfl lemma to_seminorm_ball_zero {f : E →ₗ[𝕜] 𝕜} {r : ℝ} : - seminorm.ball f.to_seminorm 0 r = { x : E | ∥f x∥ < r} := + seminorm.ball f.to_seminorm 0 r = { x : E | ‖f x‖ < r} := by simp only [seminorm.ball_zero_eq, to_seminorm_apply] lemma to_seminorm_comp (f : F →ₗ[𝕜] 𝕜) (g : E →ₗ[𝕜] F) : @@ -71,7 +71,7 @@ def to_seminorm_family (B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜) : seminorm_famil λ y, (B.flip y).to_seminorm @[simp] lemma to_seminorm_family_apply {B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜} {x y} : - (B.to_seminorm_family y) x = ∥B x y∥ := rfl + (B.to_seminorm_family y) x = ‖B x y‖ := rfl end linear_map diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index 9e80fc9dc8d35..6fb325e9d33b3 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -172,7 +172,7 @@ begin rw hU, by_cases h : x ≠ 0, { rw [(s.sup p).smul_ball_preimage 0 r x h, smul_zero], - use (s.sup p).ball 0 (r / ∥x∥), + use (s.sup p).ball 0 (r / ‖x‖), exact ⟨p.basis_sets_mem s (div_pos hr (norm_pos_iff.mpr h)), subset.rfl⟩ }, refine ⟨(s.sup p).ball 0 r, p.basis_sets_mem s hr, _⟩, simp only [not_ne_iff.mp h, subset_def, mem_ball_zero, hr, mem_univ, map_zero, @@ -409,7 +409,7 @@ begin cases normed_field.exists_lt_norm 𝕜 r with a ha, specialize h a (le_of_lt ha), rw [seminorm.smul_ball_zero (lt_trans hr ha), mul_one] at h, - refine ⟨∥a∥, lt_trans hr ha, _⟩, + refine ⟨‖a‖, lt_trans hr ha, _⟩, intros x hx, specialize h hx, exact (finset.sup I p).mem_ball_zero.mp h }, diff --git a/src/analysis/matrix.lean b/src/analysis/matrix.lean index 2068ce405c88f..cfe6a970c3425 100644 --- a/src/analysis/matrix.lean +++ b/src/analysis/matrix.lean @@ -65,57 +65,57 @@ pi.seminormed_add_comm_group local attribute [instance] matrix.seminormed_add_comm_group lemma norm_le_iff {r : ℝ} (hr : 0 ≤ r) {A : matrix m n α} : - ∥A∥ ≤ r ↔ ∀ i j, ∥A i j∥ ≤ r := + ‖A‖ ≤ r ↔ ∀ i j, ‖A i j‖ ≤ r := by simp [pi_norm_le_iff_of_nonneg hr] lemma nnnorm_le_iff {r : ℝ≥0} {A : matrix m n α} : - ∥A∥₊ ≤ r ↔ ∀ i j, ∥A i j∥₊ ≤ r := + ‖A‖₊ ≤ r ↔ ∀ i j, ‖A i j‖₊ ≤ r := by simp [pi_nnnorm_le_iff] lemma norm_lt_iff {r : ℝ} (hr : 0 < r) {A : matrix m n α} : - ∥A∥ < r ↔ ∀ i j, ∥A i j∥ < r := + ‖A‖ < r ↔ ∀ i j, ‖A i j‖ < r := by simp [pi_norm_lt_iff hr] lemma nnnorm_lt_iff {r : ℝ≥0} (hr : 0 < r) {A : matrix m n α} : - ∥A∥₊ < r ↔ ∀ i j, ∥A i j∥₊ < r := + ‖A‖₊ < r ↔ ∀ i j, ‖A i j‖₊ < r := by simp [pi_nnnorm_lt_iff hr] lemma norm_entry_le_entrywise_sup_norm (A : matrix m n α) {i : m} {j : n} : - ∥A i j∥ ≤ ∥A∥ := + ‖A i j‖ ≤ ‖A‖ := (norm_le_pi_norm (A i) j).trans (norm_le_pi_norm A i) lemma nnnorm_entry_le_entrywise_sup_nnnorm (A : matrix m n α) {i : m} {j : n} : - ∥A i j∥₊ ≤ ∥A∥₊ := + ‖A i j‖₊ ≤ ‖A‖₊ := (nnnorm_le_pi_nnnorm (A i) j).trans (nnnorm_le_pi_nnnorm A i) -@[simp] lemma nnnorm_map_eq (A : matrix m n α) (f : α → β) (hf : ∀ a, ∥f a∥₊ = ∥a∥₊) : - ∥A.map f∥₊ = ∥A∥₊ := +@[simp] lemma nnnorm_map_eq (A : matrix m n α) (f : α → β) (hf : ∀ a, ‖f a‖₊ = ‖a‖₊) : + ‖A.map f‖₊ = ‖A‖₊ := by simp_rw [pi.nnnorm_def, matrix.map_apply, hf] -@[simp] lemma norm_map_eq (A : matrix m n α) (f : α → β) (hf : ∀ a, ∥f a∥ = ∥a∥) : - ∥A.map f∥ = ∥A∥ := +@[simp] lemma norm_map_eq (A : matrix m n α) (f : α → β) (hf : ∀ a, ‖f a‖ = ‖a‖) : + ‖A.map f‖ = ‖A‖ := (congr_arg (coe : ℝ≥0 → ℝ) $ nnnorm_map_eq A f $ λ a, subtype.ext $ hf a : _) -@[simp] lemma nnnorm_transpose (A : matrix m n α) : ∥Aᵀ∥₊ = ∥A∥₊ := +@[simp] lemma nnnorm_transpose (A : matrix m n α) : ‖Aᵀ‖₊ = ‖A‖₊ := by { simp_rw [pi.nnnorm_def], exact finset.sup_comm _ _ _ } -@[simp] lemma norm_transpose (A : matrix m n α) : ∥Aᵀ∥ = ∥A∥ := congr_arg coe $ nnnorm_transpose A +@[simp] lemma norm_transpose (A : matrix m n α) : ‖Aᵀ‖ = ‖A‖ := congr_arg coe $ nnnorm_transpose A @[simp] lemma nnnorm_conj_transpose [star_add_monoid α] [normed_star_group α] (A : matrix m n α) : - ∥Aᴴ∥₊ = ∥A∥₊ := + ‖Aᴴ‖₊ = ‖A‖₊ := (nnnorm_map_eq _ _ nnnorm_star).trans A.nnnorm_transpose @[simp] lemma norm_conj_transpose [star_add_monoid α] [normed_star_group α] (A : matrix m n α) : - ∥Aᴴ∥ = ∥A∥ := + ‖Aᴴ‖ = ‖A‖ := congr_arg coe $ nnnorm_conj_transpose A instance [star_add_monoid α] [normed_star_group α] : normed_star_group (matrix m m α) := ⟨norm_conj_transpose⟩ -@[simp] lemma nnnorm_col (v : m → α) : ∥col v∥₊ = ∥v∥₊ := by simp [pi.nnnorm_def] -@[simp] lemma norm_col (v : m → α) : ∥col v∥ = ∥v∥ := congr_arg coe $ nnnorm_col v +@[simp] lemma nnnorm_col (v : m → α) : ‖col v‖₊ = ‖v‖₊ := by simp [pi.nnnorm_def] +@[simp] lemma norm_col (v : m → α) : ‖col v‖ = ‖v‖ := congr_arg coe $ nnnorm_col v -@[simp] lemma nnnorm_row (v : n → α) : ∥row v∥₊ = ∥v∥₊ := by simp [pi.nnnorm_def] -@[simp] lemma norm_row (v : n → α) : ∥row v∥ = ∥v∥ := congr_arg coe $ nnnorm_row v +@[simp] lemma nnnorm_row (v : n → α) : ‖row v‖₊ = ‖v‖₊ := by simp [pi.nnnorm_def] +@[simp] lemma norm_row (v : n → α) : ‖row v‖ = ‖v‖ := congr_arg coe $ nnnorm_row v -@[simp] lemma nnnorm_diagonal [decidable_eq n] (v : n → α) : ∥diagonal v∥₊ = ∥v∥₊ := +@[simp] lemma nnnorm_diagonal [decidable_eq n] (v : n → α) : ‖diagonal v‖₊ = ‖v‖₊ := begin simp_rw pi.nnnorm_def, congr' 1 with i : 1, @@ -128,7 +128,7 @@ begin rw diagonal_apply_eq } end -@[simp] lemma norm_diagonal [decidable_eq n] (v : n → α) : ∥diagonal v∥ = ∥v∥ := +@[simp] lemma norm_diagonal [decidable_eq n] (v : n → α) : ‖diagonal v‖ = ‖v‖ := congr_arg coe $ nnnorm_diagonal v /-- Note this is safe as an instance as it carries no data. -/ @@ -199,35 +199,35 @@ section seminormed_add_comm_group variables [seminormed_add_comm_group α] lemma linfty_op_norm_def (A : matrix m n α) : - ∥A∥ = ((finset.univ : finset m).sup (λ i : m, ∑ j : n, ∥A i j∥₊) : ℝ≥0) := + ‖A‖ = ((finset.univ : finset m).sup (λ i : m, ∑ j : n, ‖A i j‖₊) : ℝ≥0) := by simp [pi.norm_def, pi_Lp.nnnorm_eq_sum ennreal.one_ne_top] lemma linfty_op_nnnorm_def (A : matrix m n α) : - ∥A∥₊ = (finset.univ : finset m).sup (λ i : m, ∑ j : n, ∥A i j∥₊) := + ‖A‖₊ = (finset.univ : finset m).sup (λ i : m, ∑ j : n, ‖A i j‖₊) := subtype.ext $ linfty_op_norm_def A @[simp] lemma linfty_op_nnnorm_col (v : m → α) : - ∥col v∥₊ = ∥v∥₊ := + ‖col v‖₊ = ‖v‖₊ := begin rw [linfty_op_nnnorm_def, pi.nnnorm_def], simp, end @[simp] lemma linfty_op_norm_col (v : m → α) : - ∥col v∥ = ∥v∥ := + ‖col v‖ = ‖v‖ := congr_arg coe $ linfty_op_nnnorm_col v @[simp] lemma linfty_op_nnnorm_row (v : n → α) : - ∥row v∥₊ = ∑ i, ∥v i∥₊ := + ‖row v‖₊ = ∑ i, ‖v i‖₊ := by simp [linfty_op_nnnorm_def] @[simp] lemma linfty_op_norm_row (v : n → α) : - ∥row v∥ = ∑ i, ∥v i∥ := + ‖row v‖ = ∑ i, ‖v i‖ := (congr_arg coe $ linfty_op_nnnorm_row v).trans $ by simp [nnreal.coe_sum] @[simp] lemma linfty_op_nnnorm_diagonal [decidable_eq m] (v : m → α) : - ∥diagonal v∥₊ = ∥v∥₊ := + ‖diagonal v‖₊ = ‖v‖₊ := begin rw [linfty_op_nnnorm_def, pi.nnnorm_def], congr' 1 with i : 1, @@ -238,7 +238,7 @@ end @[simp] lemma linfty_op_norm_diagonal [decidable_eq m] (v : m → α) : - ∥diagonal v∥ = ∥v∥ := + ‖diagonal v‖ = ‖v‖ := congr_arg coe $ linfty_op_nnnorm_diagonal v end seminormed_add_comm_group @@ -246,32 +246,32 @@ end seminormed_add_comm_group section non_unital_semi_normed_ring variables [non_unital_semi_normed_ring α] -lemma linfty_op_nnnorm_mul (A : matrix l m α) (B : matrix m n α) : ∥A ⬝ B∥₊ ≤ ∥A∥₊ * ∥B∥₊ := +lemma linfty_op_nnnorm_mul (A : matrix l m α) (B : matrix m n α) : ‖A ⬝ B‖₊ ≤ ‖A‖₊ * ‖B‖₊ := begin simp_rw [linfty_op_nnnorm_def, matrix.mul_apply], - calc finset.univ.sup (λ i, ∑ k, ∥∑ j, A i j * B j k∥₊) - ≤ finset.univ.sup (λ i, ∑ k j, ∥A i j∥₊ * ∥B j k∥₊) : + calc finset.univ.sup (λ i, ∑ k, ‖∑ j, A i j * B j k‖₊) + ≤ finset.univ.sup (λ i, ∑ k j, ‖A i j‖₊ * ‖B j k‖₊) : finset.sup_mono_fun $ λ i hi, finset.sum_le_sum $ λ k hk, nnnorm_sum_le_of_le _ $ λ j hj, nnnorm_mul_le _ _ - ... = finset.univ.sup (λ i, ∑ j, (∥A i j∥₊ * ∑ k, ∥B j k∥₊)) : + ... = finset.univ.sup (λ i, ∑ j, (‖A i j‖₊ * ∑ k, ‖B j k‖₊)) : by simp_rw [@finset.sum_comm _ m n, finset.mul_sum] - ... ≤ finset.univ.sup (λ i, ∑ j, ∥A i j∥₊ * finset.univ.sup (λ i, ∑ j, ∥B i j∥₊)) : + ... ≤ finset.univ.sup (λ i, ∑ j, ‖A i j‖₊ * finset.univ.sup (λ i, ∑ j, ‖B i j‖₊)) : finset.sup_mono_fun $ λ i hi, finset.sum_le_sum $ λ j hj, mul_le_mul_of_nonneg_left (finset.le_sup hj) (zero_le _) - ... ≤ finset.univ.sup (λ i, ∑ j, ∥A i j∥₊) * finset.univ.sup (λ i, ∑ j, ∥B i j∥₊) : + ... ≤ finset.univ.sup (λ i, ∑ j, ‖A i j‖₊) * finset.univ.sup (λ i, ∑ j, ‖B i j‖₊) : by simp_rw [←finset.sum_mul, ←nnreal.finset_sup_mul], end -lemma linfty_op_norm_mul (A : matrix l m α) (B : matrix m n α) : ∥A ⬝ B∥ ≤ ∥A∥ * ∥B∥ := +lemma linfty_op_norm_mul (A : matrix l m α) (B : matrix m n α) : ‖A ⬝ B‖ ≤ ‖A‖ * ‖B‖ := linfty_op_nnnorm_mul _ _ -lemma linfty_op_nnnorm_mul_vec (A : matrix l m α) (v : m → α) : ∥A.mul_vec v∥₊ ≤ ∥A∥₊ * ∥v∥₊ := +lemma linfty_op_nnnorm_mul_vec (A : matrix l m α) (v : m → α) : ‖A.mul_vec v‖₊ ≤ ‖A‖₊ * ‖v‖₊ := begin rw [←linfty_op_nnnorm_col (A.mul_vec v), ←linfty_op_nnnorm_col v], exact linfty_op_nnnorm_mul A (col v), end -lemma linfty_op_norm_mul_vec (A : matrix l m α) (v : m → α) : ∥matrix.mul_vec A v∥ ≤ ∥A∥ * ∥v∥ := +lemma linfty_op_norm_mul_vec (A : matrix l m α) (v : m → α) : ‖matrix.mul_vec A v‖ ≤ ‖A‖ * ‖v‖ := linfty_op_nnnorm_mul_vec _ _ end non_unital_semi_normed_ring @@ -366,54 +366,54 @@ section seminormed_add_comm_group variables [seminormed_add_comm_group α] [seminormed_add_comm_group β] lemma frobenius_nnnorm_def (A : matrix m n α) : - ∥A∥₊ = (∑ i j, ∥A i j∥₊ ^ (2 : ℝ)) ^ (1/2 : ℝ) := + ‖A‖₊ = (∑ i j, ‖A i j‖₊ ^ (2 : ℝ)) ^ (1/2 : ℝ) := by simp_rw [pi_Lp.nnnorm_eq_of_L2, nnreal.sq_sqrt, nnreal.sqrt_eq_rpow, nnreal.rpow_two] lemma frobenius_norm_def (A : matrix m n α) : - ∥A∥ = (∑ i j, ∥A i j∥ ^ (2 : ℝ)) ^ (1/2 : ℝ) := + ‖A‖ = (∑ i j, ‖A i j‖ ^ (2 : ℝ)) ^ (1/2 : ℝ) := (congr_arg coe (frobenius_nnnorm_def A)).trans $ by simp [nnreal.coe_sum] -@[simp] lemma frobenius_nnnorm_map_eq (A : matrix m n α) (f : α → β) (hf : ∀ a, ∥f a∥₊ = ∥a∥₊) : - ∥A.map f∥₊ = ∥A∥₊ := +@[simp] lemma frobenius_nnnorm_map_eq (A : matrix m n α) (f : α → β) (hf : ∀ a, ‖f a‖₊ = ‖a‖₊) : + ‖A.map f‖₊ = ‖A‖₊ := by simp_rw [frobenius_nnnorm_def, matrix.map_apply, hf] -@[simp] lemma frobenius_norm_map_eq (A : matrix m n α) (f : α → β) (hf : ∀ a, ∥f a∥ = ∥a∥) : - ∥A.map f∥ = ∥A∥ := +@[simp] lemma frobenius_norm_map_eq (A : matrix m n α) (f : α → β) (hf : ∀ a, ‖f a‖ = ‖a‖) : + ‖A.map f‖ = ‖A‖ := (congr_arg (coe : ℝ≥0 → ℝ) $ frobenius_nnnorm_map_eq A f $ λ a, subtype.ext $ hf a : _) -@[simp] lemma frobenius_nnnorm_transpose (A : matrix m n α) : ∥Aᵀ∥₊ = ∥A∥₊ := +@[simp] lemma frobenius_nnnorm_transpose (A : matrix m n α) : ‖Aᵀ‖₊ = ‖A‖₊ := by { rw [frobenius_nnnorm_def, frobenius_nnnorm_def, finset.sum_comm], refl } -@[simp] lemma frobenius_norm_transpose (A : matrix m n α) : ∥Aᵀ∥ = ∥A∥ := +@[simp] lemma frobenius_norm_transpose (A : matrix m n α) : ‖Aᵀ‖ = ‖A‖ := congr_arg coe $ frobenius_nnnorm_transpose A @[simp] lemma frobenius_nnnorm_conj_transpose [star_add_monoid α] [normed_star_group α] - (A : matrix m n α) : ∥Aᴴ∥₊ = ∥A∥₊ := + (A : matrix m n α) : ‖Aᴴ‖₊ = ‖A‖₊ := (frobenius_nnnorm_map_eq _ _ nnnorm_star).trans A.frobenius_nnnorm_transpose @[simp] lemma frobenius_norm_conj_transpose [star_add_monoid α] [normed_star_group α] - (A : matrix m n α) : ∥Aᴴ∥ = ∥A∥ := + (A : matrix m n α) : ‖Aᴴ‖ = ‖A‖ := congr_arg coe $ frobenius_nnnorm_conj_transpose A instance frobenius_normed_star_group [star_add_monoid α] [normed_star_group α] : normed_star_group (matrix m m α) := ⟨frobenius_norm_conj_transpose⟩ -@[simp] lemma frobenius_norm_row (v : m → α) : ∥row v∥ = ∥(pi_Lp.equiv 2 _).symm v∥ := +@[simp] lemma frobenius_norm_row (v : m → α) : ‖row v‖ = ‖(pi_Lp.equiv 2 _).symm v‖ := begin rw [frobenius_norm_def, fintype.sum_unique, pi_Lp.norm_eq_of_L2, real.sqrt_eq_rpow], simp only [row_apply, real.rpow_two, pi_Lp.equiv_symm_apply], end -@[simp] lemma frobenius_nnnorm_row (v : m → α) : ∥row v∥₊ = ∥(pi_Lp.equiv 2 _).symm v∥₊ := +@[simp] lemma frobenius_nnnorm_row (v : m → α) : ‖row v‖₊ = ‖(pi_Lp.equiv 2 _).symm v‖₊ := subtype.ext $ frobenius_norm_row v -@[simp] lemma frobenius_norm_col (v : n → α) : ∥col v∥ = ∥(pi_Lp.equiv 2 _).symm v∥ := +@[simp] lemma frobenius_norm_col (v : n → α) : ‖col v‖ = ‖(pi_Lp.equiv 2 _).symm v‖ := begin simp_rw [frobenius_norm_def, fintype.sum_unique, pi_Lp.norm_eq_of_L2, real.sqrt_eq_rpow], simp only [col_apply, real.rpow_two, pi_Lp.equiv_symm_apply] end -@[simp] lemma frobenius_nnnorm_col (v : n → α) : ∥col v∥₊ = ∥(pi_Lp.equiv 2 _).symm v∥₊ := +@[simp] lemma frobenius_nnnorm_col (v : n → α) : ‖col v‖₊ = ‖(pi_Lp.equiv 2 _).symm v‖₊ := subtype.ext $ frobenius_norm_col v @[simp] lemma frobenius_nnnorm_diagonal [decidable_eq n] (v : n → α) : - ∥diagonal v∥₊ = ∥(pi_Lp.equiv 2 _).symm v∥₊ := + ‖diagonal v‖₊ = ‖(pi_Lp.equiv 2 _).symm v‖₊ := begin simp_rw [frobenius_nnnorm_def, ←finset.sum_product', finset.univ_product_univ, pi_Lp.nnnorm_eq_of_L2], @@ -428,13 +428,13 @@ begin exact finset.mem_map.not.mp his ⟨i.1, finset.mem_univ _, prod.ext rfl h⟩ } end @[simp] lemma frobenius_norm_diagonal [decidable_eq n] (v : n → α) : - ∥diagonal v∥ = ∥(pi_Lp.equiv 2 _).symm v∥ := + ‖diagonal v‖ = ‖(pi_Lp.equiv 2 _).symm v‖ := (congr_arg coe $ frobenius_nnnorm_diagonal v : _).trans rfl end seminormed_add_comm_group lemma frobenius_nnnorm_one [decidable_eq n] [seminormed_add_comm_group α] [has_one α] : - ∥(1 : matrix n n α)∥₊ = nnreal.sqrt (fintype.card n) * ∥(1 : α)∥₊:= + ‖(1 : matrix n n α)‖₊ = nnreal.sqrt (fintype.card n) * ‖(1 : α)‖₊:= begin refine (frobenius_nnnorm_diagonal _).trans _, simp_rw [pi_Lp.nnnorm_equiv_symm_const ennreal.two_ne_top, nnreal.sqrt_eq_rpow], @@ -444,7 +444,7 @@ end section is_R_or_C variables [is_R_or_C α] -lemma frobenius_nnnorm_mul (A : matrix l m α) (B : matrix m n α) : ∥A ⬝ B∥₊ ≤ ∥A∥₊ * ∥B∥₊ := +lemma frobenius_nnnorm_mul (A : matrix l m α) (B : matrix m n α) : ‖A ⬝ B‖₊ ≤ ‖A‖₊ * ‖B‖₊ := begin simp_rw [frobenius_nnnorm_def, matrix.mul_apply], rw [←nnreal.mul_rpow, @finset.sum_comm _ n m, finset.sum_mul_sum, finset.sum_product], @@ -461,7 +461,7 @@ begin star_star, nnnorm_star, nnreal.sqrt_eq_rpow, nnreal.rpow_two] using this, end -lemma frobenius_norm_mul (A : matrix l m α) (B : matrix m n α) : ∥A ⬝ B∥ ≤ ∥A∥ * ∥B∥ := +lemma frobenius_norm_mul (A : matrix l m α) (B : matrix m n α) : ‖A ⬝ B‖ ≤ ‖A‖ * ‖B‖ := frobenius_nnnorm_mul A B /-- Normed ring instance (using frobenius norm) for matrices over `ℝ` or `ℂ`. Not diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index afbbf848cbd37..ea3f874e13211 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -22,14 +22,14 @@ open filter metric open_locale topological_space big_operators nnreal ennreal uniformity pointwise /-- A non-unital seminormed ring is a not-necessarily-unital ring -endowed with a seminorm which satisfies the inequality `∥x y∥ ≤ ∥x∥ ∥y∥`. -/ +endowed with a seminorm which satisfies the inequality `‖x y‖ ≤ ‖x‖ ‖y‖`. -/ class non_unital_semi_normed_ring (α : Type*) extends has_norm α, non_unital_ring α, pseudo_metric_space α := (dist_eq : ∀ x y, dist x y = norm (x - y)) (norm_mul : ∀ a b, norm (a * b) ≤ norm a * norm b) /-- A seminormed ring is a ring endowed with a seminorm which satisfies the inequality -`∥x y∥ ≤ ∥x∥ ∥y∥`. -/ +`‖x y‖ ≤ ‖x‖ ‖y‖`. -/ class semi_normed_ring (α : Type*) extends has_norm α, ring α, pseudo_metric_space α := (dist_eq : ∀ x y, dist x y = norm (x - y)) (norm_mul : ∀ a b, norm (a * b) ≤ norm a * norm b) @@ -41,7 +41,7 @@ instance semi_normed_ring.to_non_unital_semi_normed_ring [β : semi_normed_ring { ..β } /-- A non-unital normed ring is a not-necessarily-unital ring -endowed with a norm which satisfies the inequality `∥x y∥ ≤ ∥x∥ ∥y∥`. -/ +endowed with a norm which satisfies the inequality `‖x y‖ ≤ ‖x‖ ‖y‖`. -/ class non_unital_normed_ring (α : Type*) extends has_norm α, non_unital_ring α, metric_space α := (dist_eq : ∀ x y, dist x y = norm (x - y)) (norm_mul : ∀ a b, norm (a * b) ≤ norm a * norm b) @@ -52,13 +52,13 @@ instance non_unital_normed_ring.to_non_unital_semi_normed_ring [β : non_unital_ non_unital_semi_normed_ring α := { ..β } -/-- A normed ring is a ring endowed with a norm which satisfies the inequality `∥x y∥ ≤ ∥x∥ ∥y∥`. -/ +/-- A normed ring is a ring endowed with a norm which satisfies the inequality `‖x y‖ ≤ ‖x‖ ‖y‖`. -/ class normed_ring (α : Type*) extends has_norm α, ring α, metric_space α := (dist_eq : ∀ x y, dist x y = norm (x - y)) (norm_mul : ∀ a b, norm (a * b) ≤ norm a * norm b) /-- A normed division ring is a division ring endowed with a seminorm which satisfies the equality -`∥x y∥ = ∥x∥ ∥y∥`. -/ +`‖x y‖ = ‖x‖ ‖y‖`. -/ class normed_division_ring (α : Type*) extends has_norm α, division_ring α, metric_space α := (dist_eq : ∀ x y, dist x y = norm (x - y)) (norm_mul' : ∀ a b, norm (a * b) = norm a * norm b) @@ -80,12 +80,12 @@ instance normed_ring.to_non_unital_normed_ring [β : normed_ring α] : non_unita { ..β } /-- A seminormed commutative ring is a commutative ring endowed with a seminorm which satisfies -the inequality `∥x y∥ ≤ ∥x∥ ∥y∥`. -/ +the inequality `‖x y‖ ≤ ‖x‖ ‖y‖`. -/ class semi_normed_comm_ring (α : Type*) extends semi_normed_ring α := (mul_comm : ∀ x y : α, x * y = y * x) /-- A normed commutative ring is a commutative ring endowed with a norm which satisfies -the inequality `∥x y∥ ≤ ∥x∥ ∥y∥`. -/ +the inequality `‖x y‖ ≤ ‖x‖ ‖y‖`. -/ class normed_comm_ring (α : Type*) extends normed_ring α := (mul_comm : ∀ x y : α, x * y = y * x) @@ -99,17 +99,17 @@ instance : normed_comm_ring punit := ..punit.normed_add_comm_group, ..punit.comm_ring, } -/-- A mixin class with the axiom `∥1∥ = 1`. Many `normed_ring`s and all `normed_field`s satisfy this +/-- A mixin class with the axiom `‖1‖ = 1`. Many `normed_ring`s and all `normed_field`s satisfy this axiom. -/ class norm_one_class (α : Type*) [has_norm α] [has_one α] : Prop := -(norm_one : ∥(1:α)∥ = 1) +(norm_one : ‖(1:α)‖ = 1) export norm_one_class (norm_one) attribute [simp] norm_one @[simp] lemma nnnorm_one [seminormed_add_comm_group α] [has_one α] [norm_one_class α] : - ∥(1 : α)∥₊ = 1 := + ‖(1 : α)‖₊ = 1 := nnreal.eq norm_one lemma norm_one_class.nontrivial (α : Type*) [seminormed_add_comm_group α] [has_one α] @@ -145,18 +145,18 @@ instance pi.norm_one_class {ι : Type*} {α : ι → Type*} [nonempty ι] [finty section non_unital_semi_normed_ring variables [non_unital_semi_normed_ring α] -lemma norm_mul_le (a b : α) : (∥a*b∥) ≤ (∥a∥) * (∥b∥) := +lemma norm_mul_le (a b : α) : (‖a*b‖) ≤ (‖a‖) * (‖b‖) := non_unital_semi_normed_ring.norm_mul _ _ -lemma nnnorm_mul_le (a b : α) : ∥a * b∥₊ ≤ ∥a∥₊ * ∥b∥₊ := +lemma nnnorm_mul_le (a b : α) : ‖a * b‖₊ ≤ ‖a‖₊ * ‖b‖₊ := by simpa only [←norm_to_nnreal, ←real.to_nnreal_mul (norm_nonneg _)] using real.to_nnreal_mono (norm_mul_le _ _) -lemma one_le_norm_one (β) [normed_ring β] [nontrivial β] : 1 ≤ ∥(1 : β)∥ := +lemma one_le_norm_one (β) [normed_ring β] [nontrivial β] : 1 ≤ ‖(1 : β)‖ := (le_mul_iff_one_le_left $ norm_pos_iff.mpr (one_ne_zero : (1 : β) ≠ 0)).mp (by simpa only [mul_one] using norm_mul_le (1 : β) 1) -lemma one_le_nnnorm_one (β) [normed_ring β] [nontrivial β] : 1 ≤ ∥(1 : β)∥₊ := +lemma one_le_nnnorm_one (β) [normed_ring β] [nontrivial β] : 1 ≤ ‖(1 : β)‖₊ := one_le_norm_one β lemma filter.tendsto.zero_mul_is_bounded_under_le {f g : ι → α} {l : filter ι} @@ -171,12 +171,12 @@ hg.op_zero_is_bounded_under_le hf (flip (*)) (λ x y, ((norm_mul_le y x).trans_e /-- In a seminormed ring, the left-multiplication `add_monoid_hom` is bounded. -/ lemma mul_left_bound (x : α) : - ∀ (y:α), ∥add_monoid_hom.mul_left x y∥ ≤ ∥x∥ * ∥y∥ := + ∀ (y:α), ‖add_monoid_hom.mul_left x y‖ ≤ ‖x‖ * ‖y‖ := norm_mul_le x /-- In a seminormed ring, the right-multiplication `add_monoid_hom` is bounded. -/ lemma mul_right_bound (x : α) : - ∀ (y:α), ∥add_monoid_hom.mul_right x y∥ ≤ ∥x∥ * ∥y∥ := + ∀ (y:α), ‖add_monoid_hom.mul_right x y‖ ≤ ‖x‖ * ‖y‖ := λ y, by {rw mul_comm, convert norm_mul_le y x} instance : non_unital_semi_normed_ring (ulift α) := @@ -189,15 +189,15 @@ instance prod.non_unital_semi_normed_ring [non_unital_semi_normed_ring β] : non_unital_semi_normed_ring (α × β) := { norm_mul := assume x y, calc - ∥x * y∥ = ∥(x.1*y.1, x.2*y.2)∥ : rfl - ... = (max ∥x.1*y.1∥ ∥x.2*y.2∥) : rfl - ... ≤ (max (∥x.1∥*∥y.1∥) (∥x.2∥*∥y.2∥)) : + ‖x * y‖ = ‖(x.1*y.1, x.2*y.2)‖ : rfl + ... = (max ‖x.1*y.1‖ ‖x.2*y.2‖) : rfl + ... ≤ (max (‖x.1‖*‖y.1‖) (‖x.2‖*‖y.2‖)) : max_le_max (norm_mul_le (x.1) (y.1)) (norm_mul_le (x.2) (y.2)) - ... = (max (∥x.1∥*∥y.1∥) (∥y.2∥*∥x.2∥)) : by simp[mul_comm] - ... ≤ (max (∥x.1∥) (∥x.2∥)) * (max (∥y.2∥) (∥y.1∥)) : + ... = (max (‖x.1‖*‖y.1‖) (‖y.2‖*‖x.2‖)) : by simp[mul_comm] + ... ≤ (max (‖x.1‖) (‖x.2‖)) * (max (‖y.2‖) (‖y.1‖)) : by apply max_mul_mul_le_max_mul_max; simp [norm_nonneg] - ... = (max (∥x.1∥) (∥x.2∥)) * (max (∥y.1∥) (∥y.2∥)) : by simp [max_comm] - ... = (∥x∥*∥y∥) : rfl, + ... = (max (‖x.1‖) (‖x.2‖)) * (max (‖y.1‖) (‖y.2‖)) : by simp [max_comm] + ... = (‖x‖*‖y‖) : rfl, ..prod.seminormed_add_comm_group } /-- Non-unital seminormed ring structure on the product of finitely many non-unital seminormed @@ -206,10 +206,10 @@ instance pi.non_unital_semi_normed_ring {π : ι → Type*} [fintype ι] [Π i, non_unital_semi_normed_ring (π i)] : non_unital_semi_normed_ring (Π i, π i) := { norm_mul := λ x y, nnreal.coe_mono $ - calc finset.univ.sup (λ i, ∥x i * y i∥₊) - ≤ finset.univ.sup ((λ i, ∥x i∥₊) * (λ i, ∥y i∥₊)) : + calc finset.univ.sup (λ i, ‖x i * y i‖₊) + ≤ finset.univ.sup ((λ i, ‖x i‖₊) * (λ i, ‖y i‖₊)) : finset.sup_mono_fun $ λ b hb, norm_mul_le _ _ - ... ≤ finset.univ.sup (λ i, ∥x i∥₊) * finset.univ.sup (λ i, ∥y i∥₊) : + ... ≤ finset.univ.sup (λ i, ‖x i‖₊) * finset.univ.sup (λ i, ‖y i‖₊) : finset.sup_mul_le_mul_sup_of_nonneg _ (λ i _, zero_le _) (λ i _, zero_le _), ..pi.seminormed_add_comm_group } @@ -234,34 +234,34 @@ instance subalgebra.normed_ring {𝕜 : Type*} {_ : comm_ring 𝕜} {E : Type*} [normed_ring E] {_ : algebra 𝕜 E} (s : subalgebra 𝕜 E) : normed_ring s := { ..s.semi_normed_ring } -lemma nat.norm_cast_le : ∀ n : ℕ, ∥(n : α)∥ ≤ n * ∥(1 : α)∥ +lemma nat.norm_cast_le : ∀ n : ℕ, ‖(n : α)‖ ≤ n * ‖(1 : α)‖ | 0 := by simp | (n + 1) := by { rw [n.cast_succ, n.cast_succ, add_mul, one_mul], exact norm_add_le_of_le (nat.norm_cast_le n) le_rfl } -lemma list.norm_prod_le' : ∀ {l : list α}, l ≠ [] → ∥l.prod∥ ≤ (l.map norm).prod +lemma list.norm_prod_le' : ∀ {l : list α}, l ≠ [] → ‖l.prod‖ ≤ (l.map norm).prod | [] h := (h rfl).elim | [a] _ := by simp | (a :: b :: l) _ := begin - rw [list.map_cons, list.prod_cons, @list.prod_cons _ _ _ ∥a∥], + rw [list.map_cons, list.prod_cons, @list.prod_cons _ _ _ ‖a‖], refine le_trans (norm_mul_le _ _) (mul_le_mul_of_nonneg_left _ (norm_nonneg _)), exact list.norm_prod_le' (list.cons_ne_nil b l) end -lemma list.nnnorm_prod_le' {l : list α} (hl : l ≠ []) : ∥l.prod∥₊ ≤ (l.map nnnorm).prod := +lemma list.nnnorm_prod_le' {l : list α} (hl : l ≠ []) : ‖l.prod‖₊ ≤ (l.map nnnorm).prod := (list.norm_prod_le' hl).trans_eq $ by simp [nnreal.coe_list_prod, list.map_map] -lemma list.norm_prod_le [norm_one_class α] : ∀ l : list α, ∥l.prod∥ ≤ (l.map norm).prod +lemma list.norm_prod_le [norm_one_class α] : ∀ l : list α, ‖l.prod‖ ≤ (l.map norm).prod | [] := by simp | (a::l) := list.norm_prod_le' (list.cons_ne_nil a l) -lemma list.nnnorm_prod_le [norm_one_class α] (l : list α) : ∥l.prod∥₊ ≤ (l.map nnnorm).prod := +lemma list.nnnorm_prod_le [norm_one_class α] (l : list α) : ‖l.prod‖₊ ≤ (l.map nnnorm).prod := l.norm_prod_le.trans_eq $ by simp [nnreal.coe_list_prod, list.map_map] lemma finset.norm_prod_le' {α : Type*} [normed_comm_ring α] (s : finset ι) (hs : s.nonempty) (f : ι → α) : - ∥∏ i in s, f i∥ ≤ ∏ i in s, ∥f i∥ := + ‖∏ i in s, f i‖ ≤ ∏ i in s, ‖f i‖ := begin rcases s with ⟨⟨l⟩, hl⟩, have : l.map f ≠ [], by simpa using hs, @@ -270,12 +270,12 @@ end lemma finset.nnnorm_prod_le' {α : Type*} [normed_comm_ring α] (s : finset ι) (hs : s.nonempty) (f : ι → α) : - ∥∏ i in s, f i∥₊ ≤ ∏ i in s, ∥f i∥₊ := + ‖∏ i in s, f i‖₊ ≤ ∏ i in s, ‖f i‖₊ := (s.norm_prod_le' hs f).trans_eq $ by simp [nnreal.coe_prod] lemma finset.norm_prod_le {α : Type*} [normed_comm_ring α] [norm_one_class α] (s : finset ι) (f : ι → α) : - ∥∏ i in s, f i∥ ≤ ∏ i in s, ∥f i∥ := + ‖∏ i in s, f i‖ ≤ ∏ i in s, ‖f i‖ := begin rcases s with ⟨⟨l⟩, hl⟩, simpa using (l.map f).norm_prod_le @@ -283,30 +283,30 @@ end lemma finset.nnnorm_prod_le {α : Type*} [normed_comm_ring α] [norm_one_class α] (s : finset ι) (f : ι → α) : - ∥∏ i in s, f i∥₊ ≤ ∏ i in s, ∥f i∥₊ := + ‖∏ i in s, f i‖₊ ≤ ∏ i in s, ‖f i‖₊ := (s.norm_prod_le f).trans_eq $ by simp [nnreal.coe_prod] -/-- If `α` is a seminormed ring, then `∥a ^ n∥₊ ≤ ∥a∥₊ ^ n` for `n > 0`. +/-- If `α` is a seminormed ring, then `‖a ^ n‖₊ ≤ ‖a‖₊ ^ n` for `n > 0`. See also `nnnorm_pow_le`. -/ -lemma nnnorm_pow_le' (a : α) : ∀ {n : ℕ}, 0 < n → ∥a ^ n∥₊ ≤ ∥a∥₊ ^ n +lemma nnnorm_pow_le' (a : α) : ∀ {n : ℕ}, 0 < n → ‖a ^ n‖₊ ≤ ‖a‖₊ ^ n | 1 h := by simp only [pow_one] | (n + 2) h := by simpa only [pow_succ _ (n + 1)] using le_trans (nnnorm_mul_le _ _) (mul_le_mul_left' (nnnorm_pow_le' n.succ_pos) _) -/-- If `α` is a seminormed ring with `∥1∥₊ = 1`, then `∥a ^ n∥₊ ≤ ∥a∥₊ ^ n`. +/-- If `α` is a seminormed ring with `‖1‖₊ = 1`, then `‖a ^ n‖₊ ≤ ‖a‖₊ ^ n`. See also `nnnorm_pow_le'`.-/ -lemma nnnorm_pow_le [norm_one_class α] (a : α) (n : ℕ) : ∥a ^ n∥₊ ≤ ∥a∥₊ ^ n := +lemma nnnorm_pow_le [norm_one_class α] (a : α) (n : ℕ) : ‖a ^ n‖₊ ≤ ‖a‖₊ ^ n := nat.rec_on n (by simp only [pow_zero, nnnorm_one]) (λ k hk, nnnorm_pow_le' a k.succ_pos) -/-- If `α` is a seminormed ring, then `∥a ^ n∥ ≤ ∥a∥ ^ n` for `n > 0`. See also `norm_pow_le`. -/ -lemma norm_pow_le' (a : α) {n : ℕ} (h : 0 < n) : ∥a ^ n∥ ≤ ∥a∥ ^ n := +/-- If `α` is a seminormed ring, then `‖a ^ n‖ ≤ ‖a‖ ^ n` for `n > 0`. See also `norm_pow_le`. -/ +lemma norm_pow_le' (a : α) {n : ℕ} (h : 0 < n) : ‖a ^ n‖ ≤ ‖a‖ ^ n := by simpa only [nnreal.coe_pow, coe_nnnorm] using nnreal.coe_mono (nnnorm_pow_le' a h) -/-- If `α` is a seminormed ring with `∥1∥ = 1`, then `∥a ^ n∥ ≤ ∥a∥ ^ n`. See also `norm_pow_le'`.-/ -lemma norm_pow_le [norm_one_class α] (a : α) (n : ℕ) : ∥a ^ n∥ ≤ ∥a∥ ^ n := +/-- If `α` is a seminormed ring with `‖1‖ = 1`, then `‖a ^ n‖ ≤ ‖a‖ ^ n`. See also `norm_pow_le'`.-/ +lemma norm_pow_le [norm_one_class α] (a : α) (n : ℕ) : ‖a ^ n‖ ≤ ‖a‖ ^ n := nat.rec_on n (by simp only [pow_zero, norm_one]) (λ n hn, norm_pow_le' a n.succ_pos) -lemma eventually_norm_pow_le (a : α) : ∀ᶠ (n:ℕ) in at_top, ∥a ^ n∥ ≤ ∥a∥ ^ n := +lemma eventually_norm_pow_le (a : α) : ∀ᶠ (n:ℕ) in at_top, ‖a ^ n‖ ≤ ‖a‖ ^ n := eventually_at_top.mpr ⟨1, λ b h, norm_pow_le' a (nat.succ_le_iff.mp h)⟩ instance : semi_normed_ring (ulift α) := @@ -357,10 +357,10 @@ section normed_ring variables [normed_ring α] -lemma units.norm_pos [nontrivial α] (x : αˣ) : 0 < ∥(x:α)∥ := +lemma units.norm_pos [nontrivial α] (x : αˣ) : 0 < ‖(x:α)‖ := norm_pos_iff.mpr (units.ne_zero x) -lemma units.nnnorm_pos [nontrivial α] (x : αˣ) : 0 < ∥(x:α)∥₊ := +lemma units.nnnorm_pos [nontrivial α] (x : αˣ) : 0 < ‖(x:α)‖₊ := x.norm_pos instance : normed_ring (ulift α) := @@ -384,11 +384,11 @@ end normed_ring instance semi_normed_ring_top_monoid [non_unital_semi_normed_ring α] : has_continuous_mul α := ⟨ continuous_iff_continuous_at.2 $ λ x, tendsto_iff_norm_tendsto_zero.2 $ begin - have : ∀ e : α × α, ∥e.1 * e.2 - x.1 * x.2∥ ≤ ∥e.1∥ * ∥e.2 - x.2∥ + ∥e.1 - x.1∥ * ∥x.2∥, + have : ∀ e : α × α, ‖e.1 * e.2 - x.1 * x.2‖ ≤ ‖e.1‖ * ‖e.2 - x.2‖ + ‖e.1 - x.1‖ * ‖x.2‖, { intro e, - calc ∥e.1 * e.2 - x.1 * x.2∥ ≤ ∥e.1 * (e.2 - x.2) + (e.1 - x.1) * x.2∥ : + calc ‖e.1 * e.2 - x.1 * x.2‖ ≤ ‖e.1 * (e.2 - x.2) + (e.1 - x.1) * x.2‖ : by rw [mul_sub, sub_mul, sub_add_sub_cancel] - ... ≤ ∥e.1∥ * ∥e.2 - x.2∥ + ∥e.1 - x.1∥ * ∥x.2∥ : + ... ≤ ‖e.1‖ * ‖e.2 - x.2‖ + ‖e.1 - x.1‖ * ‖x.2‖ : norm_add_le_of_le (norm_mul_le _ _) (norm_mul_le _ _) }, refine squeeze_zero (λ e, norm_nonneg _) this _, convert ((continuous_fst.tendsto x).norm.mul ((continuous_snd.tendsto x).sub @@ -406,7 +406,7 @@ section normed_division_ring variables [normed_division_ring α] -@[simp] lemma norm_mul (a b : α) : ∥a * b∥ = ∥a∥ * ∥b∥ := +@[simp] lemma norm_mul (a b : α) : ‖a * b‖ = ‖a‖ * ‖b‖ := normed_division_ring.norm_mul' a b @[priority 900] @@ -414,7 +414,7 @@ instance normed_division_ring.to_norm_one_class : norm_one_class α := ⟨mul_left_cancel₀ (mt norm_eq_zero.1 (one_ne_zero' α)) $ by rw [← norm_mul, mul_one, mul_one]⟩ -@[simp] lemma nnnorm_mul (a b : α) : ∥a * b∥₊ = ∥a∥₊ * ∥b∥₊ := +@[simp] lemma nnnorm_mul (a b : α) : ‖a * b‖₊ = ‖a‖₊ * ‖b‖₊ := nnreal.eq $ norm_mul a b /-- `norm` as a `monoid_with_zero_hom`. -/ @@ -423,30 +423,30 @@ nnreal.eq $ norm_mul a b /-- `nnnorm` as a `monoid_with_zero_hom`. -/ @[simps] def nnnorm_hom : α →*₀ ℝ≥0 := ⟨nnnorm, nnnorm_zero, nnnorm_one, nnnorm_mul⟩ -@[simp] lemma norm_pow (a : α) : ∀ (n : ℕ), ∥a ^ n∥ = ∥a∥ ^ n := +@[simp] lemma norm_pow (a : α) : ∀ (n : ℕ), ‖a ^ n‖ = ‖a‖ ^ n := (norm_hom.to_monoid_hom : α →* ℝ).map_pow a -@[simp] lemma nnnorm_pow (a : α) (n : ℕ) : ∥a ^ n∥₊ = ∥a∥₊ ^ n := +@[simp] lemma nnnorm_pow (a : α) (n : ℕ) : ‖a ^ n‖₊ = ‖a‖₊ ^ n := (nnnorm_hom.to_monoid_hom : α →* ℝ≥0).map_pow a n -protected lemma list.norm_prod (l : list α) : ∥l.prod∥ = (l.map norm).prod := +protected lemma list.norm_prod (l : list α) : ‖l.prod‖ = (l.map norm).prod := (norm_hom.to_monoid_hom : α →* ℝ).map_list_prod _ -protected lemma list.nnnorm_prod (l : list α) : ∥l.prod∥₊ = (l.map nnnorm).prod := +protected lemma list.nnnorm_prod (l : list α) : ‖l.prod‖₊ = (l.map nnnorm).prod := (nnnorm_hom.to_monoid_hom : α →* ℝ≥0).map_list_prod _ -@[simp] lemma norm_div (a b : α) : ∥a / b∥ = ∥a∥ / ∥b∥ := map_div₀ (norm_hom : α →*₀ ℝ) a b +@[simp] lemma norm_div (a b : α) : ‖a / b‖ = ‖a‖ / ‖b‖ := map_div₀ (norm_hom : α →*₀ ℝ) a b -@[simp] lemma nnnorm_div (a b : α) : ∥a / b∥₊ = ∥a∥₊ / ∥b∥₊ := map_div₀ (nnnorm_hom : α →*₀ ℝ≥0) a b +@[simp] lemma nnnorm_div (a b : α) : ‖a / b‖₊ = ‖a‖₊ / ‖b‖₊ := map_div₀ (nnnorm_hom : α →*₀ ℝ≥0) a b -@[simp] lemma norm_inv (a : α) : ∥a⁻¹∥ = ∥a∥⁻¹ := map_inv₀ (norm_hom : α →*₀ ℝ) a +@[simp] lemma norm_inv (a : α) : ‖a⁻¹‖ = ‖a‖⁻¹ := map_inv₀ (norm_hom : α →*₀ ℝ) a -@[simp] lemma nnnorm_inv (a : α) : ∥a⁻¹∥₊ = ∥a∥₊⁻¹ := +@[simp] lemma nnnorm_inv (a : α) : ‖a⁻¹‖₊ = ‖a‖₊⁻¹ := nnreal.eq $ by simp -@[simp] lemma norm_zpow : ∀ (a : α) (n : ℤ), ∥a^n∥ = ∥a∥^n := map_zpow₀ (norm_hom : α →*₀ ℝ) +@[simp] lemma norm_zpow : ∀ (a : α) (n : ℤ), ‖a^n‖ = ‖a‖^n := map_zpow₀ (norm_hom : α →*₀ ℝ) -@[simp] lemma nnnorm_zpow : ∀ (a : α) (n : ℤ), ∥a ^ n∥₊ = ∥a∥₊ ^ n := +@[simp] lemma nnnorm_zpow : ∀ (a : α) (n : ℤ), ‖a ^ n‖₊ = ‖a‖₊ ^ n := map_zpow₀ (nnnorm_hom : α →*₀ ℝ≥0) /-- Multiplication on the left by a nonzero element of a normed division ring tends to infinity at @@ -467,16 +467,16 @@ by simpa only [tendsto_comap_iff, (∘), norm_mul] instance normed_division_ring.to_has_continuous_inv₀ : has_continuous_inv₀ α := begin refine ⟨λ r r0, tendsto_iff_norm_tendsto_zero.2 _⟩, - have r0' : 0 < ∥r∥ := norm_pos_iff.2 r0, + have r0' : 0 < ‖r‖ := norm_pos_iff.2 r0, rcases exists_between r0' with ⟨ε, ε0, εr⟩, - have : ∀ᶠ e in 𝓝 r, ∥e⁻¹ - r⁻¹∥ ≤ ∥r - e∥ / ∥r∥ / ε, + have : ∀ᶠ e in 𝓝 r, ‖e⁻¹ - r⁻¹‖ ≤ ‖r - e‖ / ‖r‖ / ε, { filter_upwards [(is_open_lt continuous_const continuous_norm).eventually_mem εr] with e he, have e0 : e ≠ 0 := norm_pos_iff.1 (ε0.trans he), - calc ∥e⁻¹ - r⁻¹∥ = ∥r∥⁻¹ * ∥r - e∥ * ∥e∥⁻¹ : by + calc ‖e⁻¹ - r⁻¹‖ = ‖r‖⁻¹ * ‖r - e‖ * ‖e‖⁻¹ : by { rw [←norm_inv, ←norm_inv, ←norm_mul, ←norm_mul, mul_sub, sub_mul, mul_assoc _ e, inv_mul_cancel r0, mul_inv_cancel e0, one_mul, mul_one] } - ... = ∥r - e∥ / ∥r∥ / ∥e∥ : by field_simp [mul_comm] - ... ≤ ∥r - e∥ / ∥r∥ / ε : + ... = ‖r - e‖ / ‖r‖ / ‖e‖ : by field_simp [mul_comm] + ... ≤ ‖r - e‖ / ‖r‖ / ε : div_le_div_of_le_left (div_nonneg (norm_nonneg _) (norm_nonneg _)) ε0 he.le }, refine squeeze_zero' (eventually_of_forall $ λ _, norm_nonneg _) this _, refine (continuous_const.sub continuous_id).norm.div_const.div_const.tendsto' _ _ _, @@ -484,9 +484,9 @@ begin end lemma norm_one_of_pow_eq_one {x : α} {k : ℕ+} (h : x ^ (k : ℕ) = 1) : - ∥x∥ = 1 := + ‖x‖ = 1 := begin - rw ( _ : ∥x∥ = 1 ↔ ∥x∥₊ = 1), + rw ( _ : ‖x‖ = 1 ↔ ‖x‖₊ = 1), apply (@pow_left_inj nnreal _ _ _ ↑k zero_le' zero_le' (pnat.pos k)).mp, { rw [← nnnorm_pow, one_pow, h, nnnorm_one], }, { exact subtype.mk_eq_mk.symm, }, @@ -494,7 +494,7 @@ end lemma norm_map_one_of_pow_eq_one [comm_monoid β] (φ : β →* α) {x : β} {k : ℕ+} (h : x ^ (k : ℕ) = 1) : - ∥φ x∥ = 1 := + ‖φ x‖ = 1 := begin have : (φ x) ^ (k : ℕ) = 1 := by rw [← monoid_hom.map_pow, h, monoid_hom.map_one], exact norm_one_of_pow_eq_one this, @@ -502,7 +502,7 @@ end end normed_division_ring -/-- A normed field is a field with a norm satisfying ∥x y∥ = ∥x∥ ∥y∥. -/ +/-- A normed field is a field with a norm satisfying ‖x y‖ = ‖x‖ ‖y‖. -/ class normed_field (α : Type*) extends has_norm α, field α, metric_space α := (dist_eq : ∀ x y, dist x y = norm (x - y)) (norm_mul' : ∀ a b, norm (a * b) = norm a * norm b) @@ -511,13 +511,13 @@ class normed_field (α : Type*) extends has_norm α, field α, metric_space α : from `0` and `1`. This makes it possible to bring any element arbitrarily close to `0` by multiplication by the powers of any element, and thus to relate algebra and topology. -/ class nontrivially_normed_field (α : Type*) extends normed_field α := -(non_trivial : ∃ x : α, 1 < ∥x∥) +(non_trivial : ∃ x : α, 1 < ‖x‖) /-- A densely normed field is a normed field for which the image of the norm is dense in `ℝ≥0`, which means it is also nontrivially normed. However, not all nontrivally normed fields are densely normed; in particular, the `padic`s exhibit this fact. -/ class densely_normed_field (α : Type*) extends normed_field α := -(lt_norm_lt : ∀ x y : ℝ, 0 ≤ x → x < y → ∃ a : α, x < ∥a∥ ∧ ∥a∥ < y) +(lt_norm_lt : ∀ x y : ℝ, 0 ≤ x → x < y → ∃ a : α, x < ‖a‖ ∧ ‖a‖ < y) section normed_field @@ -540,11 +540,11 @@ instance normed_field.to_normed_comm_ring : normed_comm_ring α := { norm_mul := λ a b, (norm_mul a b).le, ..‹normed_field α› } @[simp] lemma norm_prod (s : finset β) (f : β → α) : - ∥∏ b in s, f b∥ = ∏ b in s, ∥f b∥ := + ‖∏ b in s, f b‖ = ∏ b in s, ‖f b‖ := (norm_hom.to_monoid_hom : α →* ℝ).map_prod f s @[simp] lemma nnnorm_prod (s : finset β) (f : β → α) : - ∥∏ b in s, f b∥₊ = ∏ b in s, ∥f b∥₊ := + ‖∏ b in s, f b‖₊ = ∏ b in s, ‖f b‖₊ := (nnnorm_hom.to_monoid_hom : α →* ℝ≥0).map_prod f s end normed_field @@ -555,18 +555,18 @@ section nontrivially variables (α) [nontrivially_normed_field α] -lemma exists_one_lt_norm : ∃x : α, 1 < ∥x∥ := ‹nontrivially_normed_field α›.non_trivial +lemma exists_one_lt_norm : ∃x : α, 1 < ‖x‖ := ‹nontrivially_normed_field α›.non_trivial -lemma exists_lt_norm (r : ℝ) : ∃ x : α, r < ∥x∥ := +lemma exists_lt_norm (r : ℝ) : ∃ x : α, r < ‖x‖ := let ⟨w, hw⟩ := exists_one_lt_norm α in let ⟨n, hn⟩ := pow_unbounded_of_one_lt r hw in ⟨w^n, by rwa norm_pow⟩ -lemma exists_norm_lt {r : ℝ} (hr : 0 < r) : ∃ x : α, 0 < ∥x∥ ∧ ∥x∥ < r := +lemma exists_norm_lt {r : ℝ} (hr : 0 < r) : ∃ x : α, 0 < ‖x‖ ∧ ‖x‖ < r := let ⟨w, hw⟩ := exists_lt_norm α r⁻¹ in ⟨w⁻¹, by rwa [← set.mem_Ioo, norm_inv, ← set.mem_inv, set.inv_Ioo_0_left hr]⟩ -lemma exists_norm_lt_one : ∃x : α, 0 < ∥x∥ ∧ ∥x∥ < 1 := +lemma exists_norm_lt_one : ∃x : α, 0 < ‖x‖ ∧ ‖x‖ < 1 := exists_norm_lt α one_pos variable {α} @@ -591,28 +591,28 @@ section densely variables (α) [densely_normed_field α] -lemma exists_lt_norm_lt {r₁ r₂ : ℝ} (h₀ : 0 ≤ r₁) (h : r₁ < r₂) : ∃ x : α, r₁ < ∥x∥ ∧ ∥x∥ < r₂ := +lemma exists_lt_norm_lt {r₁ r₂ : ℝ} (h₀ : 0 ≤ r₁) (h : r₁ < r₂) : ∃ x : α, r₁ < ‖x‖ ∧ ‖x‖ < r₂ := densely_normed_field.lt_norm_lt r₁ r₂ h₀ h -lemma exists_lt_nnnorm_lt {r₁ r₂ : ℝ≥0} (h : r₁ < r₂) : ∃ x : α, r₁ < ∥x∥₊ ∧ ∥x∥₊ < r₂ := +lemma exists_lt_nnnorm_lt {r₁ r₂ : ℝ≥0} (h : r₁ < r₂) : ∃ x : α, r₁ < ‖x‖₊ ∧ ‖x‖₊ < r₂ := by exact_mod_cast exists_lt_norm_lt α r₁.prop h instance densely_ordered_range_norm : densely_ordered (set.range (norm : α → ℝ)) := { dense := begin rintro ⟨-, x, rfl⟩ ⟨-, y, rfl⟩ hxy, - exact let ⟨z, h⟩ := exists_lt_norm_lt α (norm_nonneg _) hxy in ⟨⟨∥z∥, z, rfl⟩, h⟩, + exact let ⟨z, h⟩ := exists_lt_norm_lt α (norm_nonneg _) hxy in ⟨⟨‖z‖, z, rfl⟩, h⟩, end } instance densely_ordered_range_nnnorm : densely_ordered (set.range (nnnorm : α → ℝ≥0)) := { dense := begin rintro ⟨-, x, rfl⟩ ⟨-, y, rfl⟩ hxy, - exact let ⟨z, h⟩ := exists_lt_nnnorm_lt α hxy in ⟨⟨∥z∥₊, z, rfl⟩, h⟩, + exact let ⟨z, h⟩ := exists_lt_nnnorm_lt α hxy in ⟨⟨‖z‖₊, z, rfl⟩, h⟩, end } lemma dense_range_nnnorm : dense_range (nnnorm : α → ℝ≥0) := -dense_of_exists_between $ λ _ _ hr, let ⟨x, h⟩ := exists_lt_nnnorm_lt α hr in ⟨∥x∥₊, ⟨x, rfl⟩, h⟩ +dense_of_exists_between $ λ _ _ hr, let ⟨x, h⟩ := exists_lt_nnnorm_lt α hr in ⟨‖x‖₊, ⟨x, rfl⟩, h⟩ end densely @@ -633,10 +633,10 @@ noncomputable instance : densely_normed_field ℝ := namespace real -lemma to_nnreal_mul_nnnorm {x : ℝ} (y : ℝ) (hx : 0 ≤ x) : x.to_nnreal * ∥y∥₊ = ∥x * y∥₊ := +lemma to_nnreal_mul_nnnorm {x : ℝ} (y : ℝ) (hx : 0 ≤ x) : x.to_nnreal * ‖y‖₊ = ‖x * y‖₊ := by simp [real.to_nnreal_of_nonneg, nnnorm, norm_of_nonneg, hx] -lemma nnnorm_mul_to_nnreal (x : ℝ) {y : ℝ} (hy : 0 ≤ y) : ∥x∥₊ * y.to_nnreal = ∥x * y∥₊ := +lemma nnnorm_mul_to_nnreal (x : ℝ) {y : ℝ} (hy : 0 ≤ y) : ‖x‖₊ * y.to_nnreal = ‖x * y‖₊ := by simp [real.to_nnreal_of_nonneg, nnnorm, norm_of_nonneg, hy] /-- If `E` is a nontrivial topological module over `ℝ`, then `E` has no isolated points. @@ -653,24 +653,24 @@ namespace nnreal open_locale nnreal -@[simp] lemma norm_eq (x : ℝ≥0) : ∥(x : ℝ)∥ = x := +@[simp] lemma norm_eq (x : ℝ≥0) : ‖(x : ℝ)‖ = x := by rw [real.norm_eq_abs, x.abs_eq] -@[simp] lemma nnnorm_eq (x : ℝ≥0) : ∥(x : ℝ)∥₊ = x := +@[simp] lemma nnnorm_eq (x : ℝ≥0) : ‖(x : ℝ)‖₊ = x := nnreal.eq $ real.norm_of_nonneg x.2 end nnreal -@[simp] lemma norm_norm [seminormed_add_comm_group α] (x : α) : ∥∥x∥∥ = ∥x∥ := +@[simp] lemma norm_norm [seminormed_add_comm_group α] (x : α) : ‖‖x‖‖ = ‖x‖ := real.norm_of_nonneg (norm_nonneg _) -@[simp] lemma nnnorm_norm [seminormed_add_comm_group α] (a : α) : ∥∥a∥∥₊ = ∥a∥₊ := +@[simp] lemma nnnorm_norm [seminormed_add_comm_group α] (a : α) : ‖‖a‖‖₊ = ‖a‖₊ := by simpa [real.nnnorm_of_nonneg (norm_nonneg a)] /-- A restatement of `metric_space.tendsto_at_top` in terms of the norm. -/ lemma normed_add_comm_group.tendsto_at_top [nonempty α] [semilattice_sup α] {β : Type*} [seminormed_add_comm_group β] {f : α → β} {b : β} : - tendsto f at_top (𝓝 b) ↔ ∀ ε, 0 < ε → ∃ N, ∀ n, N ≤ n → ∥f n - b∥ < ε := + tendsto f at_top (𝓝 b) ↔ ∀ ε, 0 < ε → ∃ N, ∀ n, N ≤ n → ‖f n - b‖ < ε := (at_top_basis.tendsto_iff metric.nhds_basis_ball).trans (by simp [dist_eq_norm]) /-- @@ -680,26 +680,26 @@ uses `∃ N, ∀ n > N, ...` rather than `∃ N, ∀ n ≥ N, ...` lemma normed_add_comm_group.tendsto_at_top' [nonempty α] [semilattice_sup α] [no_max_order α] {β : Type*} [seminormed_add_comm_group β] {f : α → β} {b : β} : - tendsto f at_top (𝓝 b) ↔ ∀ ε, 0 < ε → ∃ N, ∀ n, N < n → ∥f n - b∥ < ε := + tendsto f at_top (𝓝 b) ↔ ∀ ε, 0 < ε → ∃ N, ∀ n, N < n → ‖f n - b‖ < ε := (at_top_basis_Ioi.tendsto_iff metric.nhds_basis_ball).trans (by simp [dist_eq_norm]) instance : normed_comm_ring ℤ := -{ norm := λ n, ∥(n : ℝ)∥, +{ norm := λ n, ‖(n : ℝ)‖, norm_mul := λ m n, le_of_eq $ by simp only [norm, int.cast_mul, abs_mul], dist_eq := λ m n, by simp only [int.dist_eq, norm, int.cast_sub], mul_comm := mul_comm } -@[norm_cast] lemma int.norm_cast_real (m : ℤ) : ∥(m : ℝ)∥ = ∥m∥ := rfl +@[norm_cast] lemma int.norm_cast_real (m : ℤ) : ‖(m : ℝ)‖ = ‖m‖ := rfl -lemma int.norm_eq_abs (n : ℤ) : ∥n∥ = |n| := rfl +lemma int.norm_eq_abs (n : ℤ) : ‖n‖ = |n| := rfl -lemma nnreal.coe_nat_abs (n : ℤ) : (n.nat_abs : ℝ≥0) = ∥n∥₊ := +lemma nnreal.coe_nat_abs (n : ℤ) : (n.nat_abs : ℝ≥0) = ‖n‖₊ := nnreal.eq $ calc ((n.nat_abs : ℝ≥0) : ℝ) = (n.nat_abs : ℤ) : by simp only [int.cast_coe_nat, nnreal.coe_nat_cast] ... = |n| : by simp only [← int.abs_eq_nat_abs, int.cast_abs] - ... = ∥n∥ : rfl + ... = ‖n‖ : rfl -lemma int.abs_le_floor_nnreal_iff (z : ℤ) (c : ℝ≥0) : |z| ≤ ⌊c⌋₊ ↔ ∥z∥₊ ≤ c := +lemma int.abs_le_floor_nnreal_iff (z : ℤ) (c : ℝ≥0) : |z| ≤ ⌊c⌋₊ ↔ ‖z‖₊ ≤ c := begin rw [int.abs_eq_nat_abs, int.coe_nat_le, nat.le_floor_iff (zero_le c)], congr', @@ -710,7 +710,7 @@ instance : norm_one_class ℤ := ⟨by simp [← int.norm_cast_real]⟩ instance : normed_field ℚ := -{ norm := λ r, ∥(r : ℝ)∥, +{ norm := λ r, ‖(r : ℝ)‖, norm_mul' := λ r₁ r₂, by simp only [norm, rat.cast_mul, abs_mul], dist_eq := λ r₁ r₂, by simp only [rat.dist_eq, norm, rat.cast_sub] } @@ -718,9 +718,9 @@ instance : densely_normed_field ℚ := { lt_norm_lt := λ r₁ r₂ h₀ hr, let ⟨q, h⟩ := exists_rat_btwn hr in ⟨q, by { unfold norm, rwa abs_of_pos (h₀.trans_lt h.1) } ⟩ } -@[norm_cast, simp] lemma rat.norm_cast_real (r : ℚ) : ∥(r : ℝ)∥ = ∥r∥ := rfl +@[norm_cast, simp] lemma rat.norm_cast_real (r : ℚ) : ‖(r : ℝ)‖ = ‖r‖ := rfl -@[norm_cast, simp] lemma int.norm_cast_rat (m : ℤ) : ∥(m : ℚ)∥ = ∥m∥ := +@[norm_cast, simp] lemma int.norm_cast_rat (m : ℤ) : ‖(m : ℚ)‖ = ‖m‖ := by rw [← rat.norm_cast_real, ← int.norm_cast_real]; congr' 1; norm_cast -- Now that we've installed the norm on `ℤ`, @@ -728,7 +728,7 @@ by rw [← rat.norm_cast_real, ← int.norm_cast_real]; congr' 1; norm_cast section variables [seminormed_add_comm_group α] -lemma norm_nsmul_le (n : ℕ) (a : α) : ∥n • a∥ ≤ n * ∥a∥ := +lemma norm_nsmul_le (n : ℕ) (a : α) : ‖n • a‖ ≤ n * ‖a‖ := begin induction n with n ih, { simp only [norm_zero, nat.cast_zero, zero_mul, zero_smul] }, @@ -737,7 +737,7 @@ begin exact norm_add_le_of_le ih le_rfl end -lemma norm_zsmul_le (n : ℤ) (a : α) : ∥n • a∥ ≤ ∥n∥ * ∥a∥ := +lemma norm_zsmul_le (n : ℤ) (a : α) : ‖n • a‖ ≤ ‖n‖ * ‖a‖ := begin induction n with n n, { simp only [int.of_nat_eq_coe, coe_nat_zsmul], @@ -748,11 +748,11 @@ begin exact nat.abs_cast n.succ, } end -lemma nnnorm_nsmul_le (n : ℕ) (a : α) : ∥n • a∥₊ ≤ n * ∥a∥₊ := +lemma nnnorm_nsmul_le (n : ℕ) (a : α) : ‖n • a‖₊ ≤ n * ‖a‖₊ := by simpa only [←nnreal.coe_le_coe, nnreal.coe_mul, nnreal.coe_nat_cast] using norm_nsmul_le n a -lemma nnnorm_zsmul_le (n : ℤ) (a : α) : ∥n • a∥₊ ≤ ∥n∥₊ * ∥a∥₊ := +lemma nnnorm_zsmul_le (n : ℤ) (a : α) : ‖n • a‖₊ ≤ ‖n‖₊ * ‖a‖₊ := by simpa only [←nnreal.coe_le_coe, nnreal.coe_mul] using norm_zsmul_le n a end @@ -799,20 +799,20 @@ calc ∑ x in u, f x.1 * g x.2 mul_le_mul_of_nonneg_right (sum_le_has_sum _ (λ _ _, hf' _) hf) (hg.nonneg $ λ _, hg' _) lemma summable.mul_norm {f : ι → α} {g : ι' → α} - (hf : summable (λ x, ∥f x∥)) (hg : summable (λ x, ∥g x∥)) : - summable (λ (x : ι × ι'), ∥f x.1 * g x.2∥) := + (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : + summable (λ (x : ι × ι'), ‖f x.1 * g x.2‖) := summable_of_nonneg_of_le (λ x, norm_nonneg (f x.1 * g x.2)) (λ x, norm_mul_le (f x.1) (g x.2)) (hf.mul_of_nonneg hg (λ x, norm_nonneg $ f x) (λ x, norm_nonneg $ g x) : _) lemma summable_mul_of_summable_norm [complete_space α] {f : ι → α} {g : ι' → α} - (hf : summable (λ x, ∥f x∥)) (hg : summable (λ x, ∥g x∥)) : + (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : summable (λ (x : ι × ι'), f x.1 * g x.2) := summable_of_summable_norm (hf.mul_norm hg) /-- Product of two infinites sums indexed by arbitrary types. See also `tsum_mul_tsum` if `f` and `g` are *not* absolutely summable. -/ lemma tsum_mul_tsum_of_summable_norm [complete_space α] {f : ι → α} {g : ι' → α} - (hf : summable (λ x, ∥f x∥)) (hg : summable (λ x, ∥g x∥)) : + (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : (∑' x, f x) * (∑' y, g y) = (∑' z : ι × ι', f z.1 * g z.2) := tsum_mul_tsum (summable_of_summable_norm hf) (summable_of_summable_norm hg) (summable_mul_of_summable_norm hf hg) @@ -832,16 +832,16 @@ section nat open finset.nat lemma summable_norm_sum_mul_antidiagonal_of_summable_norm {f g : ℕ → α} - (hf : summable (λ x, ∥f x∥)) (hg : summable (λ x, ∥g x∥)) : - summable (λ n, ∥∑ kl in antidiagonal n, f kl.1 * g kl.2∥) := + (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : + summable (λ n, ‖∑ kl in antidiagonal n, f kl.1 * g kl.2‖) := begin have := summable_sum_mul_antidiagonal_of_summable_mul (summable.mul_of_nonneg hf hg (λ _, norm_nonneg _) (λ _, norm_nonneg _)), refine summable_of_nonneg_of_le (λ _, norm_nonneg _) _ this, intros n, - calc ∥∑ kl in antidiagonal n, f kl.1 * g kl.2∥ - ≤ ∑ kl in antidiagonal n, ∥f kl.1 * g kl.2∥ : norm_sum_le _ _ - ... ≤ ∑ kl in antidiagonal n, ∥f kl.1∥ * ∥g kl.2∥ : sum_le_sum (λ i _, norm_mul_le _ _) + calc ‖∑ kl in antidiagonal n, f kl.1 * g kl.2‖ + ≤ ∑ kl in antidiagonal n, ‖f kl.1 * g kl.2‖ : norm_sum_le _ _ + ... ≤ ∑ kl in antidiagonal n, ‖f kl.1‖ * ‖g kl.2‖ : sum_le_sum (λ i _, norm_mul_le _ _) end /-- The Cauchy product formula for the product of two infinite sums indexed by `ℕ`, @@ -849,14 +849,14 @@ end See also `tsum_mul_tsum_eq_tsum_sum_antidiagonal` if `f` and `g` are *not* absolutely summable. -/ lemma tsum_mul_tsum_eq_tsum_sum_antidiagonal_of_summable_norm [complete_space α] {f g : ℕ → α} - (hf : summable (λ x, ∥f x∥)) (hg : summable (λ x, ∥g x∥)) : + (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : (∑' n, f n) * (∑' n, g n) = ∑' n, ∑ kl in antidiagonal n, f kl.1 * g kl.2 := tsum_mul_tsum_eq_tsum_sum_antidiagonal (summable_of_summable_norm hf) (summable_of_summable_norm hg) (summable_mul_of_summable_norm hf hg) lemma summable_norm_sum_mul_range_of_summable_norm {f g : ℕ → α} - (hf : summable (λ x, ∥f x∥)) (hg : summable (λ x, ∥g x∥)) : - summable (λ n, ∥∑ k in range (n+1), f k * g (n - k)∥) := + (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : + summable (λ n, ‖∑ k in range (n+1), f k * g (n - k)‖) := begin simp_rw ← sum_antidiagonal_eq_sum_range_succ (λ k l, f k * g l), exact summable_norm_sum_mul_antidiagonal_of_summable_norm hf hg @@ -867,7 +867,7 @@ end See also `tsum_mul_tsum_eq_tsum_sum_range` if `f` and `g` are *not* absolutely summable. -/ lemma tsum_mul_tsum_eq_tsum_sum_range_of_summable_norm [complete_space α] {f g : ℕ → α} - (hf : summable (λ x, ∥f x∥)) (hg : summable (λ x, ∥g x∥)) : + (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : (∑' n, f n) * (∑' n, g n) = ∑' n, ∑ k in range (n+1), f k * g (n - k) := begin simp_rw ← sum_antidiagonal_eq_sum_range_succ (λ k l, f k * g l), @@ -886,7 +886,7 @@ variables {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} for a continuous semilinear map to be bounded and this is the main use for this typeclass. -/ class ring_hom_isometric [semiring R₁] [semiring R₂] [has_norm R₁] [has_norm R₂] (σ : R₁ →+* R₂) : Prop := -(is_iso : ∀ {x : R₁}, ∥σ x∥ = ∥x∥) +(is_iso : ∀ {x : R₁}, ‖σ x‖ = ‖x‖) attribute [simp] ring_hom_isometric.is_iso @@ -985,7 +985,7 @@ def normed_field.induced [field R] [normed_field S] { .. normed_division_ring.induced R S f hf } /-- A ring homomorphism from a `ring R` to a `semi_normed_ring S` which induces the norm structure -`semi_normed_ring.induced` makes `R` satisfy `∥(1 : R)∥ = 1` whenever `∥(1 : S)∥ = 1`. -/ +`semi_normed_ring.induced` makes `R` satisfy `‖(1 : R)‖ = 1` whenever `‖(1 : S)‖ = 1`. -/ lemma norm_one_class.induced {F : Type*} (R S : Type*) [ring R] [semi_normed_ring S] [norm_one_class S] [ring_hom_class F R S] (f : F) : @norm_one_class R (semi_normed_ring.induced R S f).to_has_norm _ := diff --git a/src/analysis/normed/group/SemiNormedGroup.lean b/src/analysis/normed/group/SemiNormedGroup.lean index cb87e2c2ba58f..c06355e1fe803 100644 --- a/src/analysis/normed/group/SemiNormedGroup.lean +++ b/src/analysis/normed/group/SemiNormedGroup.lean @@ -72,8 +72,8 @@ begin apply add_monoid_hom_class.isometry_of_norm, intro v, apply le_antisymm (h1 v), - calc ∥v∥ = ∥i.inv (i.hom v)∥ : by rw [iso.hom_inv_id_apply] - ... ≤ ∥i.hom v∥ : h2 _, + calc ‖v‖ = ‖i.inv (i.hom v)‖ : by rw [iso.hom_inv_id_apply] + ... ≤ ‖i.hom v‖ : h2 _, end end SemiNormedGroup @@ -171,8 +171,8 @@ begin refine add_monoid_hom_class.isometry_of_norm i.hom _, intro v, apply le_antisymm (i.hom.2 v), - calc ∥v∥ = ∥i.inv (i.hom v)∥ : by rw [iso.hom_inv_id_apply] - ... ≤ ∥i.hom v∥ : i.inv.2 _, + calc ‖v‖ = ‖i.inv (i.hom v)‖ : by rw [iso.hom_inv_id_apply] + ... ≤ ‖i.hom v‖ : i.inv.2 _, end end SemiNormedGroup₁ diff --git a/src/analysis/normed/group/SemiNormedGroup/completion.lean b/src/analysis/normed/group/SemiNormedGroup/completion.lean index 0baa630b179a1..27a0a2315624c 100644 --- a/src/analysis/normed/group/SemiNormedGroup/completion.lean +++ b/src/analysis/normed/group/SemiNormedGroup/completion.lean @@ -56,7 +56,7 @@ def Completion.incl {V : SemiNormedGroup} : V ⟶ Completion.obj V := map_add' := completion.coe_add, bound' := ⟨1, λ v, by simp⟩ } -lemma Completion.norm_incl_eq {V : SemiNormedGroup} {v : V} : ∥Completion.incl v∥ = ∥v∥ := by simp +lemma Completion.norm_incl_eq {V : SemiNormedGroup} {v : V} : ‖Completion.incl v‖ = ‖v‖ := by simp lemma Completion.map_norm_noninc {V W : SemiNormedGroup} {f : V ⟶ W} (hf : f.norm_noninc) : (Completion.map f).norm_noninc := diff --git a/src/analysis/normed/group/SemiNormedGroup/kernels.lean b/src/analysis/normed/group/SemiNormedGroup/kernels.lean index 7d9c71e257445..0dc24df7e08d0 100644 --- a/src/analysis/normed/group/SemiNormedGroup/kernels.lean +++ b/src/analysis/normed/group/SemiNormedGroup/kernels.lean @@ -262,8 +262,8 @@ lemma norm_noninc_explicit_cokernel_π {X Y : SemiNormedGroup.{u}} (f : X ⟶ Y) open_locale nnreal lemma explicit_cokernel_desc_norm_le_of_norm_le {X Y Z : SemiNormedGroup.{u}} - {f : X ⟶ Y} {g : Y ⟶ Z} (w : f ≫ g = 0) (c : ℝ≥0) (h : ∥ g ∥ ≤ c) : - ∥ explicit_cokernel_desc w ∥ ≤ c := + {f : X ⟶ Y} {g : Y ⟶ Z} (w : f ≫ g = 0) (c : ℝ≥0) (h : ‖ g ‖ ≤ c) : + ‖ explicit_cokernel_desc w ‖ ≤ c := normed_add_group_hom.lift_norm_le _ _ _ h lemma explicit_cokernel_desc_norm_noninc {X Y Z : SemiNormedGroup.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} @@ -285,8 +285,8 @@ begin end lemma explicit_cokernel_desc_norm_le {X Y Z : SemiNormedGroup.{u}} - {f : X ⟶ Y} {g : Y ⟶ Z} (w : f ≫ g = 0) : ∥ explicit_cokernel_desc w ∥ ≤ ∥ g ∥ := -explicit_cokernel_desc_norm_le_of_norm_le w ∥ g ∥₊ le_rfl + {f : X ⟶ Y} {g : Y ⟶ Z} (w : f ≫ g = 0) : ‖ explicit_cokernel_desc w ‖ ≤ ‖ g ‖ := +explicit_cokernel_desc_norm_le_of_norm_le w ‖ g ‖₊ le_rfl /-- The explicit cokernel is isomorphic to the usual cokernel. -/ def explicit_cokernel_iso {X Y : SemiNormedGroup.{u}} (f : X ⟶ Y) : diff --git a/src/analysis/normed/group/add_circle.lean b/src/analysis/normed/group/add_circle.lean index 2118d402c0a81..45f17aa4bf13e 100644 --- a/src/analysis/normed/group/add_circle.lean +++ b/src/analysis/normed/group/add_circle.lean @@ -10,7 +10,7 @@ import topology.instances.add_circle # The additive circle as a normed group We define the normed group structure on `add_circle p`, for `p : ℝ`. For example if `p = 1` then: -`∥(x : add_circle 1)∥ = |x - round x|` for any `x : ℝ` (see `unit_add_circle.norm_eq`). +`‖(x : add_circle 1)‖ = |x - round x|` for any `x : ℝ` (see `unit_add_circle.norm_eq`). ## Main definitions: @@ -18,7 +18,7 @@ We define the normed group structure on `add_circle p`, for `p : ℝ`. For examp ## TODO - * The fact `inner_product_geometry.angle (real.cos θ) (real.sin θ) = ∥(θ : real.angle)∥` + * The fact `inner_product_geometry.angle (real.cos θ) (real.sin θ) = ‖(θ : real.angle)‖` -/ @@ -33,7 +33,7 @@ variables (p : ℝ) instance : normed_add_comm_group (add_circle p) := add_subgroup.normed_add_comm_group_quotient _ @[simp] lemma norm_coe_mul (x : ℝ) (t : ℝ) : - ∥(↑(t * x) : add_circle (t * p))∥ = |t| * ∥(x : add_circle p)∥ := + ‖(↑(t * x) : add_circle (t * p))‖ = |t| * ‖(x : add_circle p)‖ := begin have aux : ∀ {a b c : ℝ}, a ∈ zmultiples b → c * a ∈ zmultiples (c * b) := λ a b c h, by { simp only [mem_zmultiples_iff] at ⊢ h, @@ -60,14 +60,14 @@ begin end lemma norm_neg_period (x : ℝ) : - ∥(x : add_circle (-p))∥ = ∥(x : add_circle p)∥ := + ‖(x : add_circle (-p))‖ = ‖(x : add_circle p)‖ := begin - suffices : ∥(↑(-1 * x) : add_circle (-1 * p))∥ = ∥(x : add_circle p)∥, + suffices : ‖(↑(-1 * x) : add_circle (-1 * p))‖ = ‖(x : add_circle p)‖, { rw [← this, neg_one_mul], simp, }, simp only [norm_coe_mul, abs_neg, abs_one, one_mul], end -@[simp] lemma norm_eq_of_zero {x : ℝ} : ∥(x : add_circle (0 : ℝ))∥ = |x| := +@[simp] lemma norm_eq_of_zero {x : ℝ} : ‖(x : add_circle (0 : ℝ))‖ = |x| := begin suffices : {y : ℝ | (y : add_circle (0 : ℝ)) = (x : add_circle (0 : ℝ)) } = { x }, { rw [quotient_norm_eq, this, image_singleton, real.norm_eq_abs, cInf_singleton], }, @@ -75,9 +75,9 @@ begin simp [quotient_add_group.eq_iff_sub_mem, mem_zmultiples_iff, sub_eq_zero], end -lemma norm_eq {x : ℝ} : ∥(x : add_circle p)∥ = |x - round (p⁻¹ * x) * p| := +lemma norm_eq {x : ℝ} : ‖(x : add_circle p)‖ = |x - round (p⁻¹ * x) * p| := begin - suffices : ∀ (x : ℝ), ∥(x : add_circle (1 : ℝ))∥ = |x - round x|, + suffices : ∀ (x : ℝ), ‖(x : add_circle (1 : ℝ))‖ = |x - round x|, { rcases eq_or_ne p 0 with rfl | hp, { simp, }, intros, have hx := norm_coe_mul p x p⁻¹, @@ -109,26 +109,26 @@ begin simp, }, end -lemma norm_le_half_period {x : add_circle p} (hp : p ≠ 0) : ∥x∥ ≤ |p|/2 := +lemma norm_le_half_period {x : add_circle p} (hp : p ≠ 0) : ‖x‖ ≤ |p|/2 := begin obtain ⟨x⟩ := x, - change ∥(x : add_circle p)∥ ≤ |p|/2, + change ‖(x : add_circle p)‖ ≤ |p|/2, rw [norm_eq, ← mul_le_mul_left (abs_pos.mpr (inv_ne_zero hp)), ← abs_mul, mul_sub, mul_left_comm, ← mul_div_assoc, ← abs_mul, inv_mul_cancel hp, mul_one, abs_one], exact abs_sub_round (p⁻¹ * x), end -@[simp] lemma norm_half_period_eq : ∥(↑(p/2) : add_circle p)∥ = |p|/2 := +@[simp] lemma norm_half_period_eq : ‖(↑(p/2) : add_circle p)‖ = |p|/2 := begin rcases eq_or_ne p 0 with rfl | hp, { simp, }, rw [norm_eq, ← mul_div_assoc, inv_mul_cancel hp, one_div, round_two_inv, algebra_map.coe_one, one_mul, (by linarith : p / 2 - p = -(p / 2)), abs_neg, abs_div, abs_two], end -lemma norm_coe_eq_abs_iff {x : ℝ} (hp : p ≠ 0) : ∥(x : add_circle p)∥ = |x| ↔ |x| ≤ |p|/2 := +lemma norm_coe_eq_abs_iff {x : ℝ} (hp : p ≠ 0) : ‖(x : add_circle p)‖ = |x| ↔ |x| ≤ |p|/2 := begin refine ⟨λ hx, hx ▸ norm_le_half_period p hp, λ hx, _⟩, - suffices : ∀ (p : ℝ), 0 < p → |x| ≤ p/2 → ∥(x : add_circle p)∥ = |x|, + suffices : ∀ (p : ℝ), 0 < p → |x| ≤ p/2 → ‖(x : add_circle p)‖ = |x|, { rcases lt_trichotomy 0 p with hp | rfl | hp, { rw abs_eq_self.mpr hp.le at hx, exact this p hp hx, }, @@ -215,6 +215,6 @@ end add_circle namespace unit_add_circle -lemma norm_eq {x : ℝ} : ∥(x : unit_add_circle)∥ = |x - round x| := by simp [add_circle.norm_eq] +lemma norm_eq {x : ℝ} : ‖(x : unit_add_circle)‖ = |x - round x| := by simp [add_circle.norm_eq] end unit_add_circle diff --git a/src/analysis/normed/group/add_torsor.lean b/src/analysis/normed/group/add_torsor.lean index fc06853e8b08b..0e1f3e2a565d9 100644 --- a/src/analysis/normed/group/add_torsor.lean +++ b/src/analysis/normed/group/add_torsor.lean @@ -28,7 +28,7 @@ results in type class problems). -/ class normed_add_torsor (V : out_param $ Type*) (P : Type*) [out_param $ seminormed_add_comm_group V] [pseudo_metric_space P] extends add_torsor V P := -(dist_eq_norm' : ∀ (x y : P), dist x y = ∥(x -ᵥ y : V)∥) +(dist_eq_norm' : ∀ (x y : P), dist x y = ‖(x -ᵥ y : V)‖) /-- Shortcut instance to help typeclass inference out. -/ @[priority 100] @@ -59,12 +59,12 @@ variables (V W) /-- The distance equals the norm of subtracting two points. In this lemma, it is necessary to have `V` as an explicit argument; otherwise `rw dist_eq_norm_vsub` sometimes doesn't work. -/ -lemma dist_eq_norm_vsub (x y : P) : dist x y = ∥x -ᵥ y∥ := normed_add_torsor.dist_eq_norm' x y +lemma dist_eq_norm_vsub (x y : P) : dist x y = ‖x -ᵥ y‖ := normed_add_torsor.dist_eq_norm' x y /-- The distance equals the norm of subtracting two points. In this lemma, it is necessary to have `V` as an explicit argument; otherwise `rw dist_eq_norm_vsub'` sometimes doesn't work. -/ -lemma dist_eq_norm_vsub' (x y : P) : dist x y = ∥y -ᵥ x∥ := +lemma dist_eq_norm_vsub' (x y : P) : dist x y = ‖y -ᵥ x‖ := (dist_comm _ _).trans (dist_eq_norm_vsub _ _ _) end @@ -77,10 +77,10 @@ by rw [dist_eq_norm_vsub V, dist_eq_norm_vsub V, vadd_vsub_vadd_cancel_left] dist (v₁ +ᵥ x) (v₂ +ᵥ x) = dist v₁ v₂ := by rw [dist_eq_norm_vsub V, dist_eq_norm, vadd_vsub_vadd_cancel_right] -@[simp] lemma dist_vadd_left (v : V) (x : P) : dist (v +ᵥ x) x = ∥v∥ := +@[simp] lemma dist_vadd_left (v : V) (x : P) : dist (v +ᵥ x) x = ‖v‖ := by simp [dist_eq_norm_vsub V _ x] -@[simp] lemma dist_vadd_right (v : V) (x : P) : dist x (v +ᵥ x) = ∥v∥ := +@[simp] lemma dist_vadd_right (v : V) (x : P) : dist x (v +ᵥ x) = ‖v‖ := by rw [dist_comm, dist_vadd_left] /-- Isometry between the tangent space `V` of a (semi)normed add torsor `P` and `P` given by @@ -162,12 +162,12 @@ is not an instance because it depends on `V` to define a `metric_space P`. -/ def pseudo_metric_space_of_normed_add_comm_group_of_add_torsor (V P : Type*) [seminormed_add_comm_group V] [add_torsor V P] : pseudo_metric_space P := -{ dist := λ x y, ∥(x -ᵥ y : V)∥, +{ dist := λ x y, ‖(x -ᵥ y : V)‖, dist_self := λ x, by simp, dist_comm := λ x y, by simp only [←neg_vsub_eq_vsub_rev y x, norm_neg], dist_triangle := begin intros x y z, - change ∥x -ᵥ z∥ ≤ ∥x -ᵥ y∥ + ∥y -ᵥ z∥, + change ‖x -ᵥ z‖ ≤ ‖x -ᵥ y‖ + ‖y -ᵥ z‖, rw ←vsub_add_vsub_cancel, apply norm_add_le end } @@ -178,13 +178,13 @@ P`. -/ def metric_space_of_normed_add_comm_group_of_add_torsor (V P : Type*) [normed_add_comm_group V] [add_torsor V P] : metric_space P := -{ dist := λ x y, ∥(x -ᵥ y : V)∥, +{ dist := λ x y, ‖(x -ᵥ y : V)‖, dist_self := λ x, by simp, eq_of_dist_eq_zero := λ x y h, by simpa using h, dist_comm := λ x y, by simp only [←neg_vsub_eq_vsub_rev y x, norm_neg], dist_triangle := begin intros x y z, - change ∥x -ᵥ z∥ ≤ ∥x -ᵥ y∥ + ∥y -ᵥ z∥, + change ‖x -ᵥ z‖ ≤ ‖x -ᵥ y‖ + ‖y -ᵥ z‖, rw ←vsub_add_vsub_cancel, apply norm_add_le end } diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 69630645b2da5..53654d2e0850f 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -16,10 +16,10 @@ import topology.sequences In this file we define 10 classes: * `has_norm`, `has_nnnorm`: auxiliary classes endowing a type `α` with a function `norm : α → ℝ` - (notation: `∥x∥`) and `nnnorm : α → ℝ≥0` (notation: `∥x∥₊`), respectively; + (notation: `‖x‖`) and `nnnorm : α → ℝ≥0` (notation: `‖x‖₊`), respectively; * `seminormed_..._group`: A seminormed (additive) (commutative) group is an (additive) (commutative) group with a norm and a compatible pseudometric space structure: - `∀ x y, dist x y = ∥x / y∥` or `∀ x y, dist x y = ∥x - y∥`, depending on the group operation. + `∀ x y, dist x y = ‖x / y‖` or `∀ x y, dist x y = ‖x - y‖`, depending on the group operation. * `normed_..._group`: A normed (additive) (commutative) group is an (additive) (commutative) group with a norm and a compatible metric space structure. @@ -27,9 +27,9 @@ We also prove basic properties of (semi)normed groups and provide some instances ## Notes -The current convention `dist x y = ∥x - y∥` means that the distance is invariant under right +The current convention `dist x y = ‖x - y‖` means that the distance is invariant under right addition, but actions in mathlib are usually from the left. This means we might want to change it to -`dist x y = ∥-x + y∥`. +`dist x y = ‖-x + y‖`. The normed group hierarchy would lend itself well to a mixin design (that is, having `seminormed_group` and `seminormed_add_group` not extend `group` and `add_group`), but we choose not @@ -45,73 +45,73 @@ variables {𝓕 𝕜 α ι κ E F G : Type*} open filter function metric open_locale big_operators ennreal filter nnreal uniformity pointwise topological_space -/-- Auxiliary class, endowing a type `E` with a function `norm : E → ℝ` with notation `∥x∥`. This +/-- Auxiliary class, endowing a type `E` with a function `norm : E → ℝ` with notation `‖x‖`. This class is designed to be extended in more interesting classes specifying the properties of the norm. -/ @[notation_class] class has_norm (E : Type*) := (norm : E → ℝ) -/-- Auxiliary class, endowing a type `α` with a function `nnnorm : α → ℝ≥0` with notation `∥x∥₊`. -/ +/-- Auxiliary class, endowing a type `α` with a function `nnnorm : α → ℝ≥0` with notation `‖x‖₊`. -/ @[notation_class] class has_nnnorm (E : Type*) := (nnnorm : E → ℝ≥0) export has_norm (norm) export has_nnnorm (nnnorm) -notation `∥` e `∥` := norm e -notation `∥` e `∥₊` := nnnorm e +notation `‖` e `‖` := norm e +notation `‖` e `‖₊` := nnnorm e -/-- A seminormed group is an additive group endowed with a norm for which `dist x y = ∥x - y∥` +/-- A seminormed group is an additive group endowed with a norm for which `dist x y = ‖x - y‖` defines a pseudometric space structure. -/ class seminormed_add_group (E : Type*) extends has_norm E, add_group E, pseudo_metric_space E := -(dist := λ x y, ∥x - y∥) -(dist_eq : ∀ x y, dist x y = ∥x - y∥ . obviously) +(dist := λ x y, ‖x - y‖) +(dist_eq : ∀ x y, dist x y = ‖x - y‖ . obviously) -/-- A seminormed group is a group endowed with a norm for which `dist x y = ∥x / y∥` defines a +/-- A seminormed group is a group endowed with a norm for which `dist x y = ‖x / y‖` defines a pseudometric space structure. -/ @[to_additive] class seminormed_group (E : Type*) extends has_norm E, group E, pseudo_metric_space E := -(dist := λ x y, ∥x / y∥) -(dist_eq : ∀ x y, dist x y = ∥x / y∥ . obviously) +(dist := λ x y, ‖x / y‖) +(dist_eq : ∀ x y, dist x y = ‖x / y‖ . obviously) -/-- A normed group is an additive group endowed with a norm for which `dist x y = ∥x - y∥` defines a +/-- A normed group is an additive group endowed with a norm for which `dist x y = ‖x - y‖` defines a metric space structure. -/ class normed_add_group (E : Type*) extends has_norm E, add_group E, metric_space E := -(dist := λ x y, ∥x - y∥) -(dist_eq : ∀ x y, dist x y = ∥x - y∥ . obviously) +(dist := λ x y, ‖x - y‖) +(dist_eq : ∀ x y, dist x y = ‖x - y‖ . obviously) -/-- A normed group is a group endowed with a norm for which `dist x y = ∥x / y∥` defines a metric +/-- A normed group is a group endowed with a norm for which `dist x y = ‖x / y‖` defines a metric space structure. -/ @[to_additive] class normed_group (E : Type*) extends has_norm E, group E, metric_space E := -(dist := λ x y, ∥x / y∥) -(dist_eq : ∀ x y, dist x y = ∥x / y∥ . obviously) +(dist := λ x y, ‖x / y‖) +(dist_eq : ∀ x y, dist x y = ‖x / y‖ . obviously) -/-- A seminormed group is an additive group endowed with a norm for which `dist x y = ∥x - y∥` +/-- A seminormed group is an additive group endowed with a norm for which `dist x y = ‖x - y‖` defines a pseudometric space structure. -/ class seminormed_add_comm_group (E : Type*) extends has_norm E, add_comm_group E, pseudo_metric_space E := -(dist := λ x y, ∥x - y∥) -(dist_eq : ∀ x y, dist x y = ∥x - y∥ . obviously) +(dist := λ x y, ‖x - y‖) +(dist_eq : ∀ x y, dist x y = ‖x - y‖ . obviously) -/-- A seminormed group is a group endowed with a norm for which `dist x y = ∥x / y∥` +/-- A seminormed group is a group endowed with a norm for which `dist x y = ‖x / y‖` defines a pseudometric space structure. -/ @[to_additive] class seminormed_comm_group (E : Type*) extends has_norm E, comm_group E, pseudo_metric_space E := -(dist := λ x y, ∥x / y∥) -(dist_eq : ∀ x y, dist x y = ∥x / y∥ . obviously) +(dist := λ x y, ‖x / y‖) +(dist_eq : ∀ x y, dist x y = ‖x / y‖ . obviously) -/-- A normed group is an additive group endowed with a norm for which `dist x y = ∥x - y∥` defines a +/-- A normed group is an additive group endowed with a norm for which `dist x y = ‖x - y‖` defines a metric space structure. -/ class normed_add_comm_group (E : Type*) extends has_norm E, add_comm_group E, metric_space E := -(dist := λ x y, ∥x - y∥) -(dist_eq : ∀ x y, dist x y = ∥x - y∥ . obviously) +(dist := λ x y, ‖x - y‖) +(dist_eq : ∀ x y, dist x y = ‖x - y‖ . obviously) -/-- A normed group is a group endowed with a norm for which `dist x y = ∥x / y∥` defines a metric +/-- A normed group is a group endowed with a norm for which `dist x y = ‖x / y‖` defines a metric space structure. -/ @[to_additive] class normed_comm_group (E : Type*) extends has_norm E, comm_group E, metric_space E := -(dist := λ x y, ∥x / y∥) -(dist_eq : ∀ x y, dist x y = ∥x / y∥ . obviously) +(dist := λ x y, ‖x / y‖) +(dist_eq : ∀ x y, dist x y = ‖x / y‖ . obviously) @[priority 100, to_additive] -- See note [lower instance priority] instance normed_group.to_seminormed_group [normed_group E] : seminormed_group E := @@ -130,35 +130,35 @@ instance seminormed_comm_group.to_seminormed_group [seminormed_comm_group E] : s instance normed_comm_group.to_normed_group [normed_comm_group E] : normed_group E := { ..‹normed_comm_group E› } -/-- Construct a `normed_group` from a `seminormed_group` satisfying `∀ x, ∥x∥ = 0 → x = 1`. This +/-- Construct a `normed_group` from a `seminormed_group` satisfying `∀ x, ‖x‖ = 0 → x = 1`. This avoids having to go back to the `(pseudo_)metric_space` level when declaring a `normed_group` instance as a special case of a more general `seminormed_group` instance. -/ @[to_additive "Construct a `normed_add_group` from a `seminormed_add_group` satisfying -`∀ x, ∥x∥ = 0 → x = 0`. This avoids having to go back to the `(pseudo_)metric_space` level when +`∀ x, ‖x‖ = 0 → x = 0`. This avoids having to go back to the `(pseudo_)metric_space` level when declaring a `normed_add_group` instance as a special case of a more general `seminormed_add_group` instance.", reducible] -- See note [reducible non-instances] -def normed_group.of_separation [seminormed_group E] (h : ∀ x : E, ∥x∥ = 0 → x = 1) : +def normed_group.of_separation [seminormed_group E] (h : ∀ x : E, ‖x‖ = 0 → x = 1) : normed_group E := { to_metric_space := { eq_of_dist_eq_zero := λ x y hxy, div_eq_one.1 $ h _ $ by rwa ←‹seminormed_group E›.dist_eq }, ..‹seminormed_group E› } /-- Construct a `normed_comm_group` from a `seminormed_comm_group` satisfying -`∀ x, ∥x∥ = 0 → x = 1`. This avoids having to go back to the `(pseudo_)metric_space` level when +`∀ x, ‖x‖ = 0 → x = 1`. This avoids having to go back to the `(pseudo_)metric_space` level when declaring a `normed_comm_group` instance as a special case of a more general `seminormed_comm_group` instance. -/ @[to_additive "Construct a `normed_add_comm_group` from a `seminormed_add_comm_group` satisfying -`∀ x, ∥x∥ = 0 → x = 0`. This avoids having to go back to the `(pseudo_)metric_space` level when +`∀ x, ‖x‖ = 0 → x = 0`. This avoids having to go back to the `(pseudo_)metric_space` level when declaring a `normed_add_comm_group` instance as a special case of a more general `seminormed_add_comm_group` instance.", reducible] -- See note [reducible non-instances] -def normed_comm_group.of_separation [seminormed_comm_group E] (h : ∀ x : E, ∥x∥ = 0 → x = 1) : +def normed_comm_group.of_separation [seminormed_comm_group E] (h : ∀ x : E, ‖x‖ = 0 → x = 1) : normed_comm_group E := { ..‹seminormed_comm_group E›, ..normed_group.of_separation h } /-- Construct a seminormed group from a multiplication-invariant distance. -/ @[to_additive "Construct a seminormed group from a translation-invariant distance."] def seminormed_group.of_mul_dist [has_norm E] [group E] [pseudo_metric_space E] - (h₁ : ∀ x : E, ∥x∥ = dist x 1) (h₂ : ∀ x y z : E, dist x y ≤ dist (x * z) (y * z)) : + (h₁ : ∀ x : E, ‖x‖ = dist x 1) (h₂ : ∀ x y z : E, dist x y ≤ dist (x * z) (y * z)) : seminormed_group E := { dist_eq := λ x y, begin rw h₁, apply le_antisymm, @@ -169,7 +169,7 @@ def seminormed_group.of_mul_dist [has_norm E] [group E] [pseudo_metric_space E] /-- Construct a seminormed group from a multiplication-invariant pseudodistance. -/ @[to_additive "Construct a seminormed group from a translation-invariant pseudodistance."] def seminormed_group.of_mul_dist' [has_norm E] [group E] [pseudo_metric_space E] - (h₁ : ∀ x : E, ∥x∥ = dist x 1) (h₂ : ∀ x y z : E, dist (x * z) (y * z) ≤ dist x y) : + (h₁ : ∀ x : E, ‖x‖ = dist x 1) (h₂ : ∀ x y z : E, dist (x * z) (y * z) ≤ dist x y) : seminormed_group E := { dist_eq := λ x y, begin rw h₁, apply le_antisymm, @@ -180,42 +180,42 @@ def seminormed_group.of_mul_dist' [has_norm E] [group E] [pseudo_metric_space E] /-- Construct a seminormed group from a multiplication-invariant pseudodistance. -/ @[to_additive "Construct a seminormed group from a translation-invariant pseudodistance."] def seminormed_comm_group.of_mul_dist [has_norm E] [comm_group E] [pseudo_metric_space E] - (h₁ : ∀ x : E, ∥x∥ = dist x 1) (h₂ : ∀ x y z : E, dist x y ≤ dist (x * z) (y * z)) : + (h₁ : ∀ x : E, ‖x‖ = dist x 1) (h₂ : ∀ x y z : E, dist x y ≤ dist (x * z) (y * z)) : seminormed_comm_group E := { ..seminormed_group.of_mul_dist h₁ h₂ } /-- Construct a seminormed group from a multiplication-invariant pseudodistance. -/ @[to_additive "Construct a seminormed group from a translation-invariant pseudodistance."] def seminormed_comm_group.of_mul_dist' [has_norm E] [comm_group E] [pseudo_metric_space E] - (h₁ : ∀ x : E, ∥x∥ = dist x 1) (h₂ : ∀ x y z : E, dist (x * z) (y * z) ≤ dist x y) : + (h₁ : ∀ x : E, ‖x‖ = dist x 1) (h₂ : ∀ x y z : E, dist (x * z) (y * z) ≤ dist x y) : seminormed_comm_group E := { ..seminormed_group.of_mul_dist' h₁ h₂ } /-- Construct a normed group from a multiplication-invariant distance. -/ @[to_additive "Construct a normed group from a translation-invariant distance."] def normed_group.of_mul_dist [has_norm E] [group E] [metric_space E] - (h₁ : ∀ x : E, ∥x∥ = dist x 1) (h₂ : ∀ x y z : E, dist x y ≤ dist (x * z) (y * z)) : + (h₁ : ∀ x : E, ‖x‖ = dist x 1) (h₂ : ∀ x y z : E, dist x y ≤ dist (x * z) (y * z)) : normed_group E := { ..seminormed_group.of_mul_dist h₁ h₂ } /-- Construct a normed group from a multiplication-invariant pseudodistance. -/ @[to_additive "Construct a normed group from a translation-invariant pseudodistance."] def normed_group.of_mul_dist' [has_norm E] [group E] [metric_space E] - (h₁ : ∀ x : E, ∥x∥ = dist x 1) (h₂ : ∀ x y z : E, dist (x * z) (y * z) ≤ dist x y) : + (h₁ : ∀ x : E, ‖x‖ = dist x 1) (h₂ : ∀ x y z : E, dist (x * z) (y * z) ≤ dist x y) : normed_group E := { ..seminormed_group.of_mul_dist' h₁ h₂ } /-- Construct a normed group from a multiplication-invariant pseudodistance. -/ @[to_additive "Construct a normed group from a translation-invariant pseudodistance."] def normed_comm_group.of_mul_dist [has_norm E] [comm_group E] [metric_space E] - (h₁ : ∀ x : E, ∥x∥ = dist x 1) (h₂ : ∀ x y z : E, dist x y ≤ dist (x * z) (y * z)) : + (h₁ : ∀ x : E, ‖x‖ = dist x 1) (h₂ : ∀ x y z : E, dist x y ≤ dist (x * z) (y * z)) : normed_comm_group E := { ..normed_group.of_mul_dist h₁ h₂ } /-- Construct a normed group from a multiplication-invariant pseudodistance. -/ @[to_additive "Construct a normed group from a translation-invariant pseudodistance."] def normed_comm_group.of_mul_dist' [has_norm E] [comm_group E] [metric_space E] - (h₁ : ∀ x : E, ∥x∥ = dist x 1) (h₂ : ∀ x y z : E, dist (x * z) (y * z) ≤ dist x y) : + (h₁ : ∀ x : E, ‖x‖ = dist x 1) (h₂ : ∀ x y z : E, dist (x * z) (y * z) ≤ dist x y) : normed_comm_group E := { ..normed_group.of_mul_dist' h₁ h₂ } @@ -276,22 +276,22 @@ instance : normed_add_comm_group punit := { norm := function.const _ 0, dist_eq := λ _ _, rfl, } -@[simp] lemma punit.norm_eq_zero (r : punit) : ∥r∥ = 0 := rfl +@[simp] lemma punit.norm_eq_zero (r : punit) : ‖r‖ = 0 := rfl section seminormed_group variables [seminormed_group E] [seminormed_group F] [seminormed_group G] {s : set E} {a a₁ a₂ b b₁ b₂ : E} {r r₁ r₂ : ℝ} @[to_additive] -lemma dist_eq_norm_div (a b : E) : dist a b = ∥a / b∥ := seminormed_group.dist_eq _ _ +lemma dist_eq_norm_div (a b : E) : dist a b = ‖a / b‖ := seminormed_group.dist_eq _ _ @[to_additive] -lemma dist_eq_norm_div' (a b : E) : dist a b = ∥b / a∥ := by rw [dist_comm, dist_eq_norm_div] +lemma dist_eq_norm_div' (a b : E) : dist a b = ‖b / a‖ := by rw [dist_comm, dist_eq_norm_div] alias dist_eq_norm_sub ← dist_eq_norm alias dist_eq_norm_sub' ← dist_eq_norm' -@[simp, to_additive] lemma dist_one_right (a : E) : dist a 1 = ∥a∥ := +@[simp, to_additive] lemma dist_one_right (a : E) : dist a 1 = ‖a‖ := by rw [dist_eq_norm_div, div_one] @[simp, to_additive] lemma dist_one_left : dist (1 : E) = norm := @@ -299,26 +299,26 @@ funext $ λ a, by rw [dist_comm, dist_one_right] @[to_additive] lemma isometry.norm_map_of_map_one {f : E → F} (hi : isometry f) (h₁ : f 1 = 1) (x : E) : - ∥f x∥ = ∥x∥ := + ‖f x‖ = ‖x‖ := by rw [←dist_one_right, ←h₁, hi.dist_eq, dist_one_right] @[to_additive tendsto_norm_cocompact_at_top] lemma tendsto_norm_cocompact_at_top' [proper_space E] : tendsto norm (cocompact E) at_top := by simpa only [dist_one_right] using tendsto_dist_right_cocompact_at_top (1 : E) -@[to_additive] lemma norm_div_rev (a b : E) : ∥a / b∥ = ∥b / a∥ := +@[to_additive] lemma norm_div_rev (a b : E) : ‖a / b‖ = ‖b / a‖ := by simpa only [dist_eq_norm_div] using dist_comm a b @[simp, to_additive norm_neg] -lemma norm_inv' (a : E) : ∥a⁻¹∥ = ∥a∥ := by simpa using norm_div_rev 1 a +lemma norm_inv' (a : E) : ‖a⁻¹‖ = ‖a‖ := by simpa using norm_div_rev 1 a @[simp, to_additive] lemma dist_mul_right (a₁ a₂ b : E) : dist (a₁ * b) (a₂ * b) = dist a₁ a₂ := by simp [dist_eq_norm_div] -@[simp, to_additive] lemma dist_mul_self_right (a b : E) : dist b (a * b) = ∥a∥ := +@[simp, to_additive] lemma dist_mul_self_right (a b : E) : dist b (a * b) = ‖a‖ := by rw [←dist_one_left, ←dist_mul_right 1 a b, one_mul] -@[simp, to_additive] lemma dist_mul_self_left (a b : E) : dist (a * b) b = ∥a∥ := +@[simp, to_additive] lemma dist_mul_self_left (a b : E) : dist (a * b) b = ‖a‖ := by rw [dist_comm, dist_mul_self_right] @[to_additive] lemma dist_div_right (a₁ a₂ b : E) : dist (a₁ / b) (a₂ / b) = dist a₁ a₂ := @@ -342,16 +342,16 @@ by simpa only [norm_inv', tendsto_comap_iff, (∘)] using tendsto_comap /-- **Triangle inequality** for the norm. -/ @[to_additive norm_add_le "**Triangle inequality** for the norm."] -lemma norm_mul_le' (a b : E) : ∥a * b∥ ≤ ∥a∥ + ∥b∥ := +lemma norm_mul_le' (a b : E) : ‖a * b‖ ≤ ‖a‖ + ‖b‖ := by simpa [dist_eq_norm_div] using dist_triangle a 1 b⁻¹ -@[to_additive] lemma norm_mul_le_of_le (h₁ : ∥a₁∥ ≤ r₁) (h₂ : ∥a₂∥ ≤ r₂) : ∥a₁ * a₂∥ ≤ r₁ + r₂ := +@[to_additive] lemma norm_mul_le_of_le (h₁ : ‖a₁‖ ≤ r₁) (h₂ : ‖a₂‖ ≤ r₂) : ‖a₁ * a₂‖ ≤ r₁ + r₂ := (norm_mul_le' a₁ a₂).trans $ add_le_add h₁ h₂ -@[to_additive norm_add₃_le] lemma norm_mul₃_le (a b c : E) : ∥a * b * c∥ ≤ ∥a∥ + ∥b∥ + ∥c∥ := +@[to_additive norm_add₃_le] lemma norm_mul₃_le (a b c : E) : ‖a * b * c‖ ≤ ‖a‖ + ‖b‖ + ‖c‖ := norm_mul_le_of_le (norm_mul_le' _ _) le_rfl -@[simp, to_additive norm_nonneg] lemma norm_nonneg' (a : E) : 0 ≤ ∥a∥ := +@[simp, to_additive norm_nonneg] lemma norm_nonneg' (a : E) : 0 ≤ ‖a‖ := by { rw [←dist_one_right], exact dist_nonneg } section @@ -360,101 +360,101 @@ open tactic tactic.positivity /-- Extension for the `positivity` tactic: norms are nonnegative. -/ @[positivity] meta def _root_.tactic.positivity_norm : expr → tactic strictness -| `(∥%%a∥) := nonnegative <$> mk_app ``norm_nonneg [a] <|> nonnegative <$> mk_app ``norm_nonneg' [a] +| `(‖%%a‖) := nonnegative <$> mk_app ``norm_nonneg [a] <|> nonnegative <$> mk_app ``norm_nonneg' [a] | _ := failed end -@[simp, to_additive norm_zero] lemma norm_one' : ∥(1 : E)∥ = 0 := by rw [←dist_one_right, dist_self] +@[simp, to_additive norm_zero] lemma norm_one' : ‖(1 : E)‖ = 0 := by rw [←dist_one_right, dist_self] -@[to_additive] lemma ne_one_of_norm_ne_zero : ∥a∥ ≠ 0 → a ≠ 1 := +@[to_additive] lemma ne_one_of_norm_ne_zero : ‖a‖ ≠ 0 → a ≠ 1 := mt $ by { rintro rfl, exact norm_one' } @[nontriviality, to_additive norm_of_subsingleton] -lemma norm_of_subsingleton' [subsingleton E] (a : E) : ∥a∥ = 0 := +lemma norm_of_subsingleton' [subsingleton E] (a : E) : ‖a‖ = 0 := by rw [subsingleton.elim a 1, norm_one'] attribute [nontriviality] norm_of_subsingleton @[to_additive zero_lt_one_add_norm_sq] -lemma zero_lt_one_add_norm_sq' (x : E) : 0 < 1 + ∥x∥^2 := by positivity +lemma zero_lt_one_add_norm_sq' (x : E) : 0 < 1 + ‖x‖^2 := by positivity -@[to_additive] lemma norm_div_le (a b : E) : ∥a / b∥ ≤ ∥a∥ + ∥b∥ := +@[to_additive] lemma norm_div_le (a b : E) : ‖a / b‖ ≤ ‖a‖ + ‖b‖ := by simpa [dist_eq_norm_div] using dist_triangle a 1 b -@[to_additive] lemma norm_div_le_of_le {r₁ r₂ : ℝ} (H₁ : ∥a₁∥ ≤ r₁) (H₂ : ∥a₂∥ ≤ r₂) : - ∥a₁ / a₂∥ ≤ r₁ + r₂ := +@[to_additive] lemma norm_div_le_of_le {r₁ r₂ : ℝ} (H₁ : ‖a₁‖ ≤ r₁) (H₂ : ‖a₂‖ ≤ r₂) : + ‖a₁ / a₂‖ ≤ r₁ + r₂ := (norm_div_le a₁ a₂).trans $ add_le_add H₁ H₂ -@[to_additive] lemma dist_le_norm_mul_norm (a b : E) : dist a b ≤ ∥a∥ + ∥b∥ := +@[to_additive] lemma dist_le_norm_mul_norm (a b : E) : dist a b ≤ ‖a‖ + ‖b‖ := by { rw dist_eq_norm_div, apply norm_div_le } -@[to_additive abs_norm_sub_norm_le] lemma abs_norm_sub_norm_le' (a b : E) : |∥a∥ - ∥b∥| ≤ ∥a / b∥ := +@[to_additive abs_norm_sub_norm_le] lemma abs_norm_sub_norm_le' (a b : E) : |‖a‖ - ‖b‖| ≤ ‖a / b‖ := by simpa [dist_eq_norm_div] using abs_dist_sub_le a b 1 -@[to_additive norm_sub_norm_le] lemma norm_sub_norm_le' (a b : E) : ∥a∥ - ∥b∥ ≤ ∥a / b∥ := +@[to_additive norm_sub_norm_le] lemma norm_sub_norm_le' (a b : E) : ‖a‖ - ‖b‖ ≤ ‖a / b‖ := (le_abs_self _).trans (abs_norm_sub_norm_le' a b) -@[to_additive dist_norm_norm_le] lemma dist_norm_norm_le' (a b : E) : dist ∥a∥ ∥b∥ ≤ ∥a / b∥ := +@[to_additive dist_norm_norm_le] lemma dist_norm_norm_le' (a b : E) : dist ‖a‖ ‖b‖ ≤ ‖a / b‖ := abs_norm_sub_norm_le' a b -@[to_additive] lemma norm_le_norm_add_norm_div' (u v : E) : ∥u∥ ≤ ∥v∥ + ∥u / v∥ := +@[to_additive] lemma norm_le_norm_add_norm_div' (u v : E) : ‖u‖ ≤ ‖v‖ + ‖u / v‖ := by { rw add_comm, refine (norm_mul_le' _ _).trans_eq' _, rw div_mul_cancel' } -@[to_additive] lemma norm_le_norm_add_norm_div (u v : E) : ∥v∥ ≤ ∥u∥ + ∥u / v∥ := +@[to_additive] lemma norm_le_norm_add_norm_div (u v : E) : ‖v‖ ≤ ‖u‖ + ‖u / v‖ := by { rw norm_div_rev, exact norm_le_norm_add_norm_div' v u } alias norm_le_norm_add_norm_sub' ← norm_le_insert' alias norm_le_norm_add_norm_sub ← norm_le_insert -@[to_additive] lemma norm_le_mul_norm_add (u v : E) : ∥u∥ ≤ ∥u * v∥ + ∥v∥ := -calc ∥u∥ = ∥u * v / v∥ : by rw mul_div_cancel'' -... ≤ ∥u * v∥ + ∥v∥ : norm_div_le _ _ +@[to_additive] lemma norm_le_mul_norm_add (u v : E) : ‖u‖ ≤ ‖u * v‖ + ‖v‖ := +calc ‖u‖ = ‖u * v / v‖ : by rw mul_div_cancel'' +... ≤ ‖u * v‖ + ‖v‖ : norm_div_le _ _ -@[to_additive ball_eq] lemma ball_eq' (y : E) (ε : ℝ) : ball y ε = {x | ∥x / y∥ < ε} := +@[to_additive ball_eq] lemma ball_eq' (y : E) (ε : ℝ) : ball y ε = {x | ‖x / y‖ < ε} := set.ext $ λ a, by simp [dist_eq_norm_div] -@[to_additive] lemma ball_one_eq (r : ℝ) : ball (1 : E) r = {x | ∥x∥ < r} := +@[to_additive] lemma ball_one_eq (r : ℝ) : ball (1 : E) r = {x | ‖x‖ < r} := set.ext $ assume a, by simp -@[to_additive mem_ball_iff_norm] lemma mem_ball_iff_norm'' : b ∈ ball a r ↔ ∥b / a∥ < r := +@[to_additive mem_ball_iff_norm] lemma mem_ball_iff_norm'' : b ∈ ball a r ↔ ‖b / a‖ < r := by rw [mem_ball, dist_eq_norm_div] -@[to_additive mem_ball_iff_norm'] lemma mem_ball_iff_norm''' : b ∈ ball a r ↔ ∥a / b∥ < r := +@[to_additive mem_ball_iff_norm'] lemma mem_ball_iff_norm''' : b ∈ ball a r ↔ ‖a / b‖ < r := by rw [mem_ball', dist_eq_norm_div] -@[simp, to_additive] lemma mem_ball_one_iff : a ∈ ball (1 : E) r ↔ ∥a∥ < r := +@[simp, to_additive] lemma mem_ball_one_iff : a ∈ ball (1 : E) r ↔ ‖a‖ < r := by rw [mem_ball, dist_one_right] @[to_additive mem_closed_ball_iff_norm] -lemma mem_closed_ball_iff_norm'' : b ∈ closed_ball a r ↔ ∥b / a∥ ≤ r := +lemma mem_closed_ball_iff_norm'' : b ∈ closed_ball a r ↔ ‖b / a‖ ≤ r := by rw [mem_closed_ball, dist_eq_norm_div] -@[simp, to_additive] lemma mem_closed_ball_one_iff : a ∈ closed_ball (1 : E) r ↔ ∥a∥ ≤ r := +@[simp, to_additive] lemma mem_closed_ball_one_iff : a ∈ closed_ball (1 : E) r ↔ ‖a‖ ≤ r := by rw [mem_closed_ball, dist_one_right] @[to_additive mem_closed_ball_iff_norm'] -lemma mem_closed_ball_iff_norm''' : b ∈ closed_ball a r ↔ ∥a / b∥ ≤ r := +lemma mem_closed_ball_iff_norm''' : b ∈ closed_ball a r ↔ ‖a / b‖ ≤ r := by rw [mem_closed_ball', dist_eq_norm_div] @[to_additive norm_le_of_mem_closed_ball] -lemma norm_le_of_mem_closed_ball' (h : b ∈ closed_ball a r) : ∥b∥ ≤ ∥a∥ + r := +lemma norm_le_of_mem_closed_ball' (h : b ∈ closed_ball a r) : ‖b‖ ≤ ‖a‖ + r := (norm_le_norm_add_norm_div' _ _).trans $ add_le_add_left (by rwa ←dist_eq_norm_div) _ @[to_additive norm_le_norm_add_const_of_dist_le] -lemma norm_le_norm_add_const_of_dist_le' : dist a b ≤ r → ∥a∥ ≤ ∥b∥ + r := +lemma norm_le_norm_add_const_of_dist_le' : dist a b ≤ r → ‖a‖ ≤ ‖b‖ + r := norm_le_of_mem_closed_ball' @[to_additive norm_lt_of_mem_ball] -lemma norm_lt_of_mem_ball' (h : b ∈ ball a r) : ∥b∥ < ∥a∥ + r := +lemma norm_lt_of_mem_ball' (h : b ∈ ball a r) : ‖b‖ < ‖a‖ + r := (norm_le_norm_add_norm_div' _ _).trans_lt $ add_lt_add_left (by rwa ←dist_eq_norm_div) _ @[to_additive] -lemma norm_div_sub_norm_div_le_norm_div (u v w : E) : ∥u / w∥ - ∥v / w∥ ≤ ∥u / v∥ := +lemma norm_div_sub_norm_div_le_norm_div (u v w : E) : ‖u / w‖ - ‖v / w‖ ≤ ‖u / v‖ := by simpa only [div_div_div_cancel_right'] using norm_sub_norm_le' (u / w) (v / w) @[to_additive bounded_iff_forall_norm_le] -lemma bounded_iff_forall_norm_le' : bounded s ↔ ∃ C, ∀ x ∈ s, ∥x∥ ≤ C := +lemma bounded_iff_forall_norm_le' : bounded s ↔ ∃ C, ∀ x ∈ s, ‖x‖ ≤ C := by simpa only [set.subset_def, mem_closed_ball_one_iff] using bounded_iff_subset_ball (1 : E) alias bounded_iff_forall_norm_le' ↔ metric.bounded.exists_norm_le' _ @@ -463,18 +463,18 @@ alias bounded_iff_forall_norm_le ↔ metric.bounded.exists_norm_le _ attribute [to_additive metric.bounded.exists_norm_le] metric.bounded.exists_norm_le' @[to_additive metric.bounded.exists_pos_norm_le] -lemma metric.bounded.exists_pos_norm_le' (hs : metric.bounded s) : ∃ R > 0, ∀ x ∈ s, ∥x∥ ≤ R := +lemma metric.bounded.exists_pos_norm_le' (hs : metric.bounded s) : ∃ R > 0, ∀ x ∈ s, ‖x‖ ≤ R := let ⟨R₀, hR₀⟩ := hs.exists_norm_le' in ⟨max R₀ 1, by positivity, λ x hx, (hR₀ x hx).trans $ le_max_left _ _⟩ @[simp, to_additive mem_sphere_iff_norm] -lemma mem_sphere_iff_norm' : b ∈ sphere a r ↔ ∥b / a∥ = r := by simp [dist_eq_norm_div] +lemma mem_sphere_iff_norm' : b ∈ sphere a r ↔ ‖b / a‖ = r := by simp [dist_eq_norm_div] -@[simp, to_additive] lemma mem_sphere_one_iff_norm : a ∈ sphere (1 : E) r ↔ ∥a∥ = r := +@[simp, to_additive] lemma mem_sphere_one_iff_norm : a ∈ sphere (1 : E) r ↔ ‖a‖ = r := by simp [dist_eq_norm_div] @[simp, to_additive norm_eq_of_mem_sphere] -lemma norm_eq_of_mem_sphere' (x : sphere (1:E) r) : ∥(x : E)∥ = r := mem_sphere_one_iff_norm.mp x.2 +lemma norm_eq_of_mem_sphere' (x : sphere (1:E) r) : ‖(x : E)‖ = r := mem_sphere_one_iff_norm.mp x.2 @[to_additive] lemma ne_one_of_mem_sphere (hr : r ≠ 0) (x : sphere (1 : E) r) : (x : E) ≠ 1 := ne_one_of_norm_ne_zero $ by rwa norm_eq_of_mem_sphere' x @@ -518,43 +518,43 @@ ext $ λ y, rfl end isometric @[to_additive] lemma normed_comm_group.tendsto_nhds_one {f : α → E} {l : filter α} : - tendsto f l (𝓝 1) ↔ ∀ ε > 0, ∀ᶠ x in l, ∥ f x ∥ < ε := + tendsto f l (𝓝 1) ↔ ∀ ε > 0, ∀ᶠ x in l, ‖ f x ‖ < ε := metric.tendsto_nhds.trans $ by simp only [dist_one_right] @[to_additive] lemma normed_comm_group.tendsto_nhds_nhds {f : E → F} {x : E} {y : F} : - tendsto f (𝓝 x) (𝓝 y) ↔ ∀ ε > 0, ∃ δ > 0, ∀ x', ∥x' / x∥ < δ → ∥f x' / y∥ < ε := + tendsto f (𝓝 x) (𝓝 y) ↔ ∀ ε > 0, ∃ δ > 0, ∀ x', ‖x' / x‖ < δ → ‖f x' / y‖ < ε := by simp_rw [metric.tendsto_nhds_nhds, dist_eq_norm_div] @[to_additive] lemma normed_comm_group.cauchy_seq_iff [nonempty α] [semilattice_sup α] {u : α → E} : - cauchy_seq u ↔ ∀ ε > 0, ∃ N, ∀ m, N ≤ m → ∀ n, N ≤ n → ∥u m / u n∥ < ε := + cauchy_seq u ↔ ∀ ε > 0, ∃ N, ∀ m, N ≤ m → ∀ n, N ≤ n → ‖u m / u n‖ < ε := by simp [metric.cauchy_seq_iff, dist_eq_norm_div] @[to_additive] lemma normed_comm_group.nhds_basis_norm_lt (x : E) : - (𝓝 x).has_basis (λ ε : ℝ, 0 < ε) (λ ε, {y | ∥y / x∥ < ε}) := + (𝓝 x).has_basis (λ ε : ℝ, 0 < ε) (λ ε, {y | ‖y / x‖ < ε}) := by { simp_rw ← ball_eq', exact metric.nhds_basis_ball } @[to_additive] lemma normed_comm_group.nhds_one_basis_norm_lt : - (𝓝 (1 : E)).has_basis (λ ε : ℝ, 0 < ε) (λ ε, {y | ∥y∥ < ε}) := + (𝓝 (1 : E)).has_basis (λ ε : ℝ, 0 < ε) (λ ε, {y | ‖y‖ < ε}) := by { convert normed_comm_group.nhds_basis_norm_lt (1 : E), simp } @[to_additive] lemma normed_comm_group.uniformity_basis_dist : - (𝓤 E).has_basis (λ ε : ℝ, 0 < ε) (λ ε, {p : E × E | ∥p.fst / p.snd∥ < ε}) := + (𝓤 E).has_basis (λ ε : ℝ, 0 < ε) (λ ε, {p : E × E | ‖p.fst / p.snd‖ < ε}) := by { convert metric.uniformity_basis_dist, simp [dist_eq_norm_div] } open finset /-- A homomorphism `f` of seminormed groups is Lipschitz, if there exists a constant `C` such that -for all `x`, one has `∥f x∥ ≤ C * ∥x∥`. The analogous condition for a linear map of +for all `x`, one has `‖f x‖ ≤ C * ‖x‖`. The analogous condition for a linear map of (semi)normed spaces is in `normed_space.operator_norm`. -/ @[to_additive "A homomorphism `f` of seminormed groups is Lipschitz, if there exists a constant `C` -such that for all `x`, one has `∥f x∥ ≤ C * ∥x∥`. The analogous condition for a linear map of +such that for all `x`, one has `‖f x‖ ≤ C * ‖x‖`. The analogous condition for a linear map of (semi)normed spaces is in `normed_space.operator_norm`."] lemma monoid_hom_class.lipschitz_of_bound [monoid_hom_class 𝓕 E F] (f : 𝓕) (C : ℝ) - (h : ∀ x, ∥f x∥ ≤ C * ∥x∥) : lipschitz_with (real.to_nnreal C) f := + (h : ∀ x, ‖f x‖ ≤ C * ‖x‖) : lipschitz_with (real.to_nnreal C) f := lipschitz_with.of_dist_le' $ λ x y, by simpa only [dist_eq_norm_div, map_div] using h (x / y) @[to_additive] lemma lipschitz_on_with_iff_norm_div_le {f : E → F} {C : ℝ≥0} : - lipschitz_on_with C f s ↔ ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, y ∈ s → ∥f x / f y∥ ≤ C * ∥x / y∥ := + lipschitz_on_with C f s ↔ ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, y ∈ s → ‖f x / f y‖ ≤ C * ‖x / y‖ := by simp only [lipschitz_on_with_iff_dist_le_mul, dist_eq_norm_div] alias lipschitz_on_with_iff_norm_div_le ↔ lipschitz_on_with.norm_div_le _ @@ -562,12 +562,12 @@ alias lipschitz_on_with_iff_norm_div_le ↔ lipschitz_on_with.norm_div_le _ attribute [to_additive] lipschitz_on_with.norm_div_le @[to_additive] lemma lipschitz_on_with.norm_div_le_of_le {f : E → F} {C : ℝ≥0} - (h : lipschitz_on_with C f s) (ha : a ∈ s) (hb : b ∈ s) (hr : ∥a / b∥ ≤ r) : - ∥f a / f b∥ ≤ C * r := + (h : lipschitz_on_with C f s) (ha : a ∈ s) (hb : b ∈ s) (hr : ‖a / b‖ ≤ r) : + ‖f a / f b‖ ≤ C * r := (h.norm_div_le ha hb).trans $ mul_le_mul_of_nonneg_left hr C.2 @[to_additive] lemma lipschitz_with_iff_norm_div_le {f : E → F} {C : ℝ≥0} : - lipschitz_with C f ↔ ∀ x y, ∥f x / f y∥ ≤ C * ∥x / y∥ := + lipschitz_with C f ↔ ∀ x y, ‖f x / f y‖ ≤ C * ‖x / y‖ := by simp only [lipschitz_with_iff_dist_le_mul, dist_eq_norm_div] alias lipschitz_with_iff_norm_div_le ↔ lipschitz_with.norm_div_le _ @@ -575,30 +575,30 @@ alias lipschitz_with_iff_norm_div_le ↔ lipschitz_with.norm_div_le _ attribute [to_additive] lipschitz_with.norm_div_le @[to_additive] lemma lipschitz_with.norm_div_le_of_le {f : E → F} {C : ℝ≥0} (h : lipschitz_with C f) - (hr : ∥a / b∥ ≤ r) : ∥f a / f b∥ ≤ C * r := + (hr : ‖a / b‖ ≤ r) : ‖f a / f b‖ ≤ C * r := (h.norm_div_le _ _).trans $ mul_le_mul_of_nonneg_left hr C.2 /-- A homomorphism `f` of seminormed groups is continuous, if there exists a constant `C` such that -for all `x`, one has `∥f x∥ ≤ C * ∥x∥`. -/ +for all `x`, one has `‖f x‖ ≤ C * ‖x‖`. -/ @[to_additive "A homomorphism `f` of seminormed groups is continuous, if there exists a constant `C` -such that for all `x`, one has `∥f x∥ ≤ C * ∥x∥`"] +such that for all `x`, one has `‖f x‖ ≤ C * ‖x‖`"] lemma monoid_hom_class.continuous_of_bound [monoid_hom_class 𝓕 E F] (f : 𝓕) (C : ℝ) - (h : ∀ x, ∥f x∥ ≤ C * ∥x∥) : continuous f := + (h : ∀ x, ‖f x‖ ≤ C * ‖x‖) : continuous f := (monoid_hom_class.lipschitz_of_bound f C h).continuous @[to_additive] lemma monoid_hom_class.uniform_continuous_of_bound [monoid_hom_class 𝓕 E F] - (f : 𝓕) (C : ℝ) (h : ∀x, ∥f x∥ ≤ C * ∥x∥) : uniform_continuous f := + (f : 𝓕) (C : ℝ) (h : ∀x, ‖f x‖ ≤ C * ‖x‖) : uniform_continuous f := (monoid_hom_class.lipschitz_of_bound f C h).uniform_continuous @[to_additive is_compact.exists_bound_of_continuous_on] lemma is_compact.exists_bound_of_continuous_on' [topological_space α] {s : set α} (hs : is_compact s) {f : α → E} (hf : continuous_on f s) : - ∃ C, ∀ x ∈ s, ∥f x∥ ≤ C := + ∃ C, ∀ x ∈ s, ‖f x‖ ≤ C := (bounded_iff_forall_norm_le'.1 (hs.image_of_continuous_on hf).bounded).imp $ λ C hC x hx, hC _ $ set.mem_image_of_mem _ hx @[to_additive] lemma monoid_hom_class.isometry_iff_norm [monoid_hom_class 𝓕 E F] (f : 𝓕) : - isometry f ↔ ∀ x, ∥f x∥ = ∥x∥ := + isometry f ↔ ∀ x, ‖f x‖ = ‖x‖ := begin simp only [isometry_iff_dist_eq, dist_eq_norm_div, ←map_div], refine ⟨λ h x, _, λ h x y, h _⟩, @@ -612,62 +612,62 @@ attribute [to_additive] monoid_hom_class.isometry_of_norm section nnnorm @[priority 100, to_additive] -- See note [lower instance priority] -instance seminormed_group.to_has_nnnorm : has_nnnorm E := ⟨λ a, ⟨∥a∥, norm_nonneg' a⟩⟩ +instance seminormed_group.to_has_nnnorm : has_nnnorm E := ⟨λ a, ⟨‖a‖, norm_nonneg' a⟩⟩ -@[simp, norm_cast, to_additive coe_nnnorm] lemma coe_nnnorm' (a : E) : (∥a∥₊ : ℝ) = ∥a∥ := rfl +@[simp, norm_cast, to_additive coe_nnnorm] lemma coe_nnnorm' (a : E) : (‖a‖₊ : ℝ) = ‖a‖ := rfl @[simp, to_additive coe_comp_nnnorm] lemma coe_comp_nnnorm' : (coe : ℝ≥0 → ℝ) ∘ (nnnorm : E → ℝ≥0) = norm := rfl @[to_additive norm_to_nnreal] -lemma norm_to_nnreal' : ∥a∥.to_nnreal = ∥a∥₊ := @real.to_nnreal_coe ∥a∥₊ +lemma norm_to_nnreal' : ‖a‖.to_nnreal = ‖a‖₊ := @real.to_nnreal_coe ‖a‖₊ @[to_additive] -lemma nndist_eq_nnnorm_div (a b : E) : nndist a b = ∥a / b∥₊ := nnreal.eq $ dist_eq_norm_div _ _ +lemma nndist_eq_nnnorm_div (a b : E) : nndist a b = ‖a / b‖₊ := nnreal.eq $ dist_eq_norm_div _ _ alias nndist_eq_nnnorm_sub ← nndist_eq_nnnorm -@[simp, to_additive nnnorm_zero] lemma nnnorm_one' : ∥(1 : E)∥₊ = 0 := nnreal.eq norm_one' +@[simp, to_additive nnnorm_zero] lemma nnnorm_one' : ‖(1 : E)‖₊ = 0 := nnreal.eq norm_one' -@[to_additive] lemma ne_one_of_nnnorm_ne_zero {a : E} : ∥a∥₊ ≠ 0 → a ≠ 1 := +@[to_additive] lemma ne_one_of_nnnorm_ne_zero {a : E} : ‖a‖₊ ≠ 0 → a ≠ 1 := mt $ by { rintro rfl, exact nnnorm_one' } @[to_additive nnnorm_add_le] -lemma nnnorm_mul_le' (a b : E) : ∥a * b∥₊ ≤ ∥a∥₊ + ∥b∥₊ := nnreal.coe_le_coe.1 $ norm_mul_le' a b +lemma nnnorm_mul_le' (a b : E) : ‖a * b‖₊ ≤ ‖a‖₊ + ‖b‖₊ := nnreal.coe_le_coe.1 $ norm_mul_le' a b -@[simp, to_additive nnnorm_neg] lemma nnnorm_inv' (a : E) : ∥a⁻¹∥₊ = ∥a∥₊ := nnreal.eq $ norm_inv' a +@[simp, to_additive nnnorm_neg] lemma nnnorm_inv' (a : E) : ‖a⁻¹‖₊ = ‖a‖₊ := nnreal.eq $ norm_inv' a -@[to_additive] lemma nnnorm_div_le (a b : E) : ∥a / b∥₊ ≤ ∥a∥₊ + ∥b∥₊ := +@[to_additive] lemma nnnorm_div_le (a b : E) : ‖a / b‖₊ ≤ ‖a‖₊ + ‖b‖₊ := nnreal.coe_le_coe.1 $ norm_div_le _ _ @[to_additive nndist_nnnorm_nnnorm_le] -lemma nndist_nnnorm_nnnorm_le' (a b : E) : nndist ∥a∥₊ ∥b∥₊ ≤ ∥a / b∥₊ := +lemma nndist_nnnorm_nnnorm_le' (a b : E) : nndist ‖a‖₊ ‖b‖₊ ≤ ‖a / b‖₊ := nnreal.coe_le_coe.1 $ dist_norm_norm_le' a b -@[to_additive] lemma nnnorm_le_nnnorm_add_nnnorm_div (a b : E) : ∥b∥₊ ≤ ∥a∥₊ + ∥a / b∥₊ := +@[to_additive] lemma nnnorm_le_nnnorm_add_nnnorm_div (a b : E) : ‖b‖₊ ≤ ‖a‖₊ + ‖a / b‖₊ := norm_le_norm_add_norm_div _ _ -@[to_additive] lemma nnnorm_le_nnnorm_add_nnnorm_div' (a b : E) : ∥a∥₊ ≤ ∥b∥₊ + ∥a / b∥₊ := +@[to_additive] lemma nnnorm_le_nnnorm_add_nnnorm_div' (a b : E) : ‖a‖₊ ≤ ‖b‖₊ + ‖a / b‖₊ := norm_le_norm_add_norm_div' _ _ alias nnnorm_le_nnnorm_add_nnnorm_sub' ← nnnorm_le_insert' alias nnnorm_le_nnnorm_add_nnnorm_sub ← nnnorm_le_insert @[to_additive] -lemma nnnorm_le_mul_nnnorm_add (a b : E) : ∥a∥₊ ≤ ∥a * b∥₊ + ∥b∥₊ := norm_le_mul_norm_add _ _ +lemma nnnorm_le_mul_nnnorm_add (a b : E) : ‖a‖₊ ≤ ‖a * b‖₊ + ‖b‖₊ := norm_le_mul_norm_add _ _ @[to_additive of_real_norm_eq_coe_nnnorm] -lemma of_real_norm_eq_coe_nnnorm' (a : E) : ennreal.of_real ∥a∥ = ∥a∥₊ := +lemma of_real_norm_eq_coe_nnnorm' (a : E) : ennreal.of_real ‖a‖ = ‖a‖₊ := ennreal.of_real_eq_coe_nnreal _ -@[to_additive] lemma edist_eq_coe_nnnorm_div (a b : E) : edist a b = ∥a / b∥₊ := +@[to_additive] lemma edist_eq_coe_nnnorm_div (a b : E) : edist a b = ‖a / b‖₊ := by rw [edist_dist, dist_eq_norm_div, of_real_norm_eq_coe_nnnorm'] -@[to_additive edist_eq_coe_nnnorm] lemma edist_eq_coe_nnnorm' (x : E) : edist x 1 = (∥x∥₊ : ℝ≥0∞) := +@[to_additive edist_eq_coe_nnnorm] lemma edist_eq_coe_nnnorm' (x : E) : edist x 1 = (‖x‖₊ : ℝ≥0∞) := by rw [edist_eq_coe_nnnorm_div, div_one] @[to_additive] -lemma mem_emetric_ball_one_iff {r : ℝ≥0∞} : a ∈ emetric.ball (1 : E) r ↔ ↑∥a∥₊ < r := +lemma mem_emetric_ball_one_iff {r : ℝ≥0∞} : a ∈ emetric.ball (1 : E) r ↔ ↑‖a‖₊ < r := by rw [emetric.mem_ball, edist_eq_coe_nnnorm'] @[simp, to_additive] lemma edist_mul_right (a₁ a₂ b : E) : edist (a₁ * b) (a₂ * b) = edist a₁ a₂ := @@ -677,26 +677,26 @@ by simp [edist_dist] by simpa only [div_eq_mul_inv] using edist_mul_right _ _ _ @[to_additive] lemma monoid_hom_class.lipschitz_of_bound_nnnorm [monoid_hom_class 𝓕 E F] (f : 𝓕) - (C : ℝ≥0) (h : ∀ x, ∥f x∥₊ ≤ C * ∥x∥₊) : lipschitz_with C f := + (C : ℝ≥0) (h : ∀ x, ‖f x‖₊ ≤ C * ‖x‖₊) : lipschitz_with C f := @real.to_nnreal_coe C ▸ monoid_hom_class.lipschitz_of_bound f C h @[to_additive] lemma monoid_hom_class.antilipschitz_of_bound [monoid_hom_class 𝓕 E F] (f : 𝓕) - {K : ℝ≥0} (h : ∀ x, ∥x∥ ≤ K * ∥f x∥) : + {K : ℝ≥0} (h : ∀ x, ‖x‖ ≤ K * ‖f x‖) : antilipschitz_with K f := antilipschitz_with.of_le_mul_dist $ λ x y, by simpa only [dist_eq_norm_div, map_div] using h (x / y) @[to_additive] lemma monoid_hom_class.bound_of_antilipschitz [monoid_hom_class 𝓕 E F] (f : 𝓕) - {K : ℝ≥0} (h : antilipschitz_with K f) (x) : ∥x∥ ≤ K * ∥f x∥ := + {K : ℝ≥0} (h : antilipschitz_with K f) (x) : ‖x‖ ≤ K * ‖f x‖ := by simpa only [dist_one_right, map_one] using h.le_mul_dist x 1 end nnnorm @[to_additive] lemma tendsto_iff_norm_tendsto_one {f : α → E} {a : filter α} {b : E} : - tendsto f a (𝓝 b) ↔ tendsto (λ e, ∥f e / b∥) a (𝓝 0) := + tendsto f a (𝓝 b) ↔ tendsto (λ e, ‖f e / b‖) a (𝓝 0) := by { convert tendsto_iff_dist_tendsto_zero, simp [dist_eq_norm_div] } @[to_additive] lemma tendsto_one_iff_norm_tendsto_one {f : α → E} {a : filter α} : - tendsto f a (𝓝 1) ↔ tendsto (λ e, ∥f e∥) a (𝓝 0) := + tendsto f a (𝓝 1) ↔ tendsto (λ e, ‖f e‖) a (𝓝 0) := by { rw tendsto_iff_norm_tendsto_one, simp only [div_one] } @[to_additive] lemma comap_norm_nhds_one : comap norm (𝓝 0) = 𝓝 (1 : E) := @@ -712,7 +712,7 @@ real function `a` which tends to `0`, then `f` tends to `1`. In this pair of lem (`squeeze_zero_norm'` and `squeeze_zero_norm`), following a convention of similar lemmas in `topology.metric_space.basic` and `topology.algebra.order`, the `'` version is phrased using \"eventually\" and the non-`'` version is phrased absolutely."] -lemma squeeze_one_norm' {f : α → E} {a : α → ℝ} {t₀ : filter α} (h : ∀ᶠ n in t₀, ∥f n∥ ≤ a n) +lemma squeeze_one_norm' {f : α → E} {a : α → ℝ} {t₀ : filter α} (h : ∀ᶠ n in t₀, ‖f n‖ ≤ a n) (h' : tendsto a t₀ (𝓝 0)) : tendsto f t₀ (𝓝 1) := tendsto_one_iff_norm_tendsto_one.2 $ squeeze_zero' (eventually_of_forall $ λ n, norm_nonneg' _) h h' @@ -720,25 +720,25 @@ tendsto_one_iff_norm_tendsto_one.2 $ squeeze_zero' (eventually_of_forall $ λ n, tends to `0`, then `f` tends to `1`. -/ @[to_additive "Special case of the sandwich theorem: if the norm of `f` is bounded by a real function `a` which tends to `0`, then `f` tends to `0`."] -lemma squeeze_one_norm {f : α → E} {a : α → ℝ} {t₀ : filter α} (h : ∀ n, ∥f n∥ ≤ a n) : +lemma squeeze_one_norm {f : α → E} {a : α → ℝ} {t₀ : filter α} (h : ∀ n, ‖f n‖ ≤ a n) : tendsto a t₀ (𝓝 0) → tendsto f t₀ (𝓝 1) := squeeze_one_norm' $ eventually_of_forall h -@[to_additive] lemma tendsto_norm_div_self (x : E) : tendsto (λ a, ∥a / x∥) (𝓝 x) (𝓝 0) := +@[to_additive] lemma tendsto_norm_div_self (x : E) : tendsto (λ a, ‖a / x‖) (𝓝 x) (𝓝 0) := by simpa [dist_eq_norm_div] using tendsto_id.dist (tendsto_const_nhds : tendsto (λ a, (x:E)) (𝓝 x) _) -@[to_additive tendsto_norm]lemma tendsto_norm' {x : E} : tendsto (λ a, ∥a∥) (𝓝 x) (𝓝 ∥x∥) := +@[to_additive tendsto_norm]lemma tendsto_norm' {x : E} : tendsto (λ a, ‖a‖) (𝓝 x) (𝓝 ‖x‖) := by simpa using tendsto_id.dist (tendsto_const_nhds : tendsto (λ a, (1:E)) _ _) -@[to_additive] lemma tendsto_norm_one : tendsto (λ a : E, ∥a∥) (𝓝 1) (𝓝 0) := +@[to_additive] lemma tendsto_norm_one : tendsto (λ a : E, ‖a‖) (𝓝 1) (𝓝 0) := by simpa using tendsto_norm_div_self (1:E) -@[continuity, to_additive continuous_norm] lemma continuous_norm' : continuous (λ a : E, ∥a∥) := +@[continuity, to_additive continuous_norm] lemma continuous_norm' : continuous (λ a : E, ‖a‖) := by simpa using continuous_id.dist (continuous_const : continuous (λ a, (1:E))) @[continuity, to_additive continuous_nnnorm] -lemma continuous_nnnorm' : continuous (λ a : E, ∥a∥₊) := continuous_norm'.subtype_mk _ +lemma continuous_nnnorm' : continuous (λ a : E, ‖a‖₊) := continuous_norm'.subtype_mk _ @[to_additive lipschitz_with_one_norm] lemma lipschitz_with_one_norm' : lipschitz_with 1 (norm : E → ℝ) := @@ -753,26 +753,26 @@ lemma uniform_continuous_norm' : uniform_continuous (norm : E → ℝ) := lipschitz_with_one_norm'.uniform_continuous @[to_additive uniform_continuous_nnnorm] -lemma uniform_continuous_nnnorm' : uniform_continuous (λ (a : E), ∥a∥₊) := +lemma uniform_continuous_nnnorm' : uniform_continuous (λ (a : E), ‖a‖₊) := uniform_continuous_norm'.subtype_mk _ -@[to_additive] lemma mem_closure_one_iff_norm {x : E} : x ∈ closure ({1} : set E) ↔ ∥x∥ = 0 := +@[to_additive] lemma mem_closure_one_iff_norm {x : E} : x ∈ closure ({1} : set E) ↔ ‖x‖ = 0 := by rw [←closed_ball_zero', mem_closed_ball_one_iff, (norm_nonneg' x).le_iff_eq] -@[to_additive] lemma closure_one_eq : closure ({1} : set E) = {x | ∥x∥ = 0} := +@[to_additive] lemma closure_one_eq : closure ({1} : set E) = {x | ‖x‖ = 0} := set.ext (λ x, mem_closure_one_iff_norm) /-- A helper lemma used to prove that the (scalar or usual) product of a function that tends to one and a bounded function tends to one. This lemma is formulated for any binary operation -`op : E → F → G` with an estimate `∥op x y∥ ≤ A * ∥x∥ * ∥y∥` for some constant A instead of +`op : E → F → G` with an estimate `‖op x y‖ ≤ A * ‖x‖ * ‖y‖` for some constant A instead of multiplication so that it can be applied to `(*)`, `flip (*)`, `(•)`, and `flip (•)`. -/ @[to_additive "A helper lemma used to prove that the (scalar or usual) product of a function that tends to zero and a bounded function tends to zero. This lemma is formulated for any binary -operation `op : E → F → G` with an estimate `∥op x y∥ ≤ A * ∥x∥ * ∥y∥` for some constant A instead +operation `op : E → F → G` with an estimate `‖op x y‖ ≤ A * ‖x‖ * ‖y‖` for some constant A instead of multiplication so that it can be applied to `(*)`, `flip (*)`, `(•)`, and `flip (•)`."] lemma filter.tendsto.op_one_is_bounded_under_le' {f : α → E} {g : α → F} {l : filter α} (hf : tendsto f l (𝓝 1)) (hg : is_bounded_under (≤) l (norm ∘ g)) (op : E → F → G) - (h_op : ∃ A, ∀ x y, ∥op x y∥ ≤ A * ∥x∥ * ∥y∥) : + (h_op : ∃ A, ∀ x y, ‖op x y‖ ≤ A * ‖x‖ * ‖y‖) : tendsto (λ x, op (f x) (g x)) l (𝓝 1) := begin cases h_op with A h_op, @@ -785,7 +785,7 @@ begin cases le_total A 0 with hA hA, { exact (mul_nonpos_of_nonpos_of_nonneg (mul_nonpos_of_nonpos_of_nonneg hA $ norm_nonneg' _) $ norm_nonneg' _).trans_lt ε₀ }, - calc A * ∥f i∥ * ∥g i∥ ≤ A * δ * C : + calc A * ‖f i‖ * ‖g i‖ ≤ A * δ * C : mul_le_mul (mul_le_mul_of_nonneg_left hf.le hA) hg (norm_nonneg' _) (mul_nonneg hA δ₀.le) ... = A * C * δ : mul_right_comm _ _ _ ... < ε : hδ, @@ -793,27 +793,27 @@ end /-- A helper lemma used to prove that the (scalar or usual) product of a function that tends to one and a bounded function tends to one. This lemma is formulated for any binary operation -`op : E → F → G` with an estimate `∥op x y∥ ≤ ∥x∥ * ∥y∥` instead of multiplication so that it +`op : E → F → G` with an estimate `‖op x y‖ ≤ ‖x‖ * ‖y‖` instead of multiplication so that it can be applied to `(*)`, `flip (*)`, `(•)`, and `flip (•)`. -/ @[to_additive "A helper lemma used to prove that the (scalar or usual) product of a function that tends to zero and a bounded function tends to zero. This lemma is formulated for any binary -operation `op : E → F → G` with an estimate `∥op x y∥ ≤ ∥x∥ * ∥y∥` instead of multiplication so that +operation `op : E → F → G` with an estimate `‖op x y‖ ≤ ‖x‖ * ‖y‖` instead of multiplication so that it can be applied to `(*)`, `flip (*)`, `(•)`, and `flip (•)`."] lemma filter.tendsto.op_one_is_bounded_under_le {f : α → E} {g : α → F} {l : filter α} (hf : tendsto f l (𝓝 1)) (hg : is_bounded_under (≤) l (norm ∘ g)) (op : E → F → G) - (h_op : ∀ x y, ∥op x y∥ ≤ ∥x∥ * ∥y∥) : + (h_op : ∀ x y, ‖op x y‖ ≤ ‖x‖ * ‖y‖) : tendsto (λ x, op (f x) (g x)) l (𝓝 1) := -hf.op_one_is_bounded_under_le' hg op ⟨1, λ x y, (one_mul (∥x∥)).symm ▸ h_op x y⟩ +hf.op_one_is_bounded_under_le' hg op ⟨1, λ x y, (one_mul (‖x‖)).symm ▸ h_op x y⟩ section variables {l : filter α} {f : α → E} @[to_additive filter.tendsto.norm] lemma filter.tendsto.norm' (h : tendsto f l (𝓝 a)) : - tendsto (λ x, ∥f x∥) l (𝓝 ∥a∥) := + tendsto (λ x, ‖f x‖) l (𝓝 ‖a‖) := tendsto_norm'.comp h @[to_additive filter.tendsto.nnnorm] lemma filter.tendsto.nnnorm' (h : tendsto f l (𝓝 a)) : - tendsto (λ x, ∥f x∥₊) l (𝓝 (∥a∥₊)) := + tendsto (λ x, ‖f x‖₊) l (𝓝 (‖a‖₊)) := tendsto.comp continuous_nnnorm'.continuous_at h end @@ -821,67 +821,67 @@ end section variables [topological_space α] {f : α → E} -@[to_additive continuous.norm] lemma continuous.norm' : continuous f → continuous (λ x, ∥f x∥) := +@[to_additive continuous.norm] lemma continuous.norm' : continuous f → continuous (λ x, ‖f x‖) := continuous_norm'.comp @[to_additive continuous.nnnorm] -lemma continuous.nnnorm' : continuous f → continuous (λ x, ∥f x∥₊) := continuous_nnnorm'.comp +lemma continuous.nnnorm' : continuous f → continuous (λ x, ‖f x‖₊) := continuous_nnnorm'.comp @[to_additive continuous_at.norm] -lemma continuous_at.norm' {a : α} (h : continuous_at f a) : continuous_at (λ x, ∥f x∥) a := h.norm' +lemma continuous_at.norm' {a : α} (h : continuous_at f a) : continuous_at (λ x, ‖f x‖) a := h.norm' @[to_additive continuous_at.nnnorm] -lemma continuous_at.nnnorm' {a : α} (h : continuous_at f a) : continuous_at (λ x, ∥f x∥₊) a := +lemma continuous_at.nnnorm' {a : α} (h : continuous_at f a) : continuous_at (λ x, ‖f x‖₊) a := h.nnnorm' @[to_additive continuous_within_at.norm] lemma continuous_within_at.norm' {s : set α} {a : α} (h : continuous_within_at f s a) : - continuous_within_at (λ x, ∥f x∥) s a := + continuous_within_at (λ x, ‖f x‖) s a := h.norm' @[to_additive continuous_within_at.nnnorm] lemma continuous_within_at.nnnorm' {s : set α} {a : α} (h : continuous_within_at f s a) : - continuous_within_at (λ x, ∥f x∥₊) s a := + continuous_within_at (λ x, ‖f x‖₊) s a := h.nnnorm' @[to_additive continuous_on.norm] -lemma continuous_on.norm' {s : set α} (h : continuous_on f s) : continuous_on (λ x, ∥f x∥) s := +lemma continuous_on.norm' {s : set α} (h : continuous_on f s) : continuous_on (λ x, ‖f x‖) s := λ x hx, (h x hx).norm' @[to_additive continuous_on.nnnorm] -lemma continuous_on.nnnorm' {s : set α} (h : continuous_on f s) : continuous_on (λ x, ∥f x∥₊) s := +lemma continuous_on.nnnorm' {s : set α} (h : continuous_on f s) : continuous_on (λ x, ‖f x‖₊) s := λ x hx, (h x hx).nnnorm' end -/-- If `∥y∥ → ∞`, then we can assume `y ≠ x` for any fixed `x`. -/ -@[to_additive eventually_ne_of_tendsto_norm_at_top "If `∥y∥→∞`, then we can assume `y≠x` for any +/-- If `‖y‖ → ∞`, then we can assume `y ≠ x` for any fixed `x`. -/ +@[to_additive eventually_ne_of_tendsto_norm_at_top "If `‖y‖→∞`, then we can assume `y≠x` for any fixed `x`"] lemma eventually_ne_of_tendsto_norm_at_top' {l : filter α} {f : α → E} - (h : tendsto (λ y, ∥f y∥) l at_top) (x : E) : + (h : tendsto (λ y, ‖f y‖) l at_top) (x : E) : ∀ᶠ y in l, f y ≠ x := (h.eventually_ne_at_top _).mono $ λ x, ne_of_apply_ne norm @[to_additive] lemma seminormed_comm_group.mem_closure_iff : - a ∈ closure s ↔ ∀ ε, 0 < ε → ∃ b ∈ s, ∥a / b∥ < ε := + a ∈ closure s ↔ ∀ ε, 0 < ε → ∃ b ∈ s, ‖a / b‖ < ε := by simp [metric.mem_closure_iff, dist_eq_norm_div] -@[to_additive norm_le_zero_iff'] lemma norm_le_zero_iff''' [t0_space E] {a : E} : ∥a∥ ≤ 0 ↔ a = 1 := +@[to_additive norm_le_zero_iff'] lemma norm_le_zero_iff''' [t0_space E] {a : E} : ‖a‖ ≤ 0 ↔ a = 1 := begin letI : normed_group E := { to_metric_space := metric.of_t0_pseudo_metric_space E, ..‹seminormed_group E› }, rw [←dist_one_right, dist_le_zero], end -@[to_additive norm_eq_zero'] lemma norm_eq_zero''' [t0_space E] {a : E} : ∥a∥ = 0 ↔ a = 1 := +@[to_additive norm_eq_zero'] lemma norm_eq_zero''' [t0_space E] {a : E} : ‖a‖ = 0 ↔ a = 1 := (norm_nonneg' a).le_iff_eq.symm.trans norm_le_zero_iff''' -@[to_additive norm_pos_iff'] lemma norm_pos_iff''' [t0_space E] {a : E} : 0 < ∥a∥ ↔ a ≠ 1 := +@[to_additive norm_pos_iff'] lemma norm_pos_iff''' [t0_space E] {a : E} : 0 < ‖a‖ ↔ a ≠ 1 := by rw [← not_le, norm_le_zero_iff'''] @[to_additive] lemma seminormed_group.tendsto_uniformly_on_one {f : ι → κ → G} {s : set κ} {l : filter ι} : - tendsto_uniformly_on f 1 l s ↔ ∀ ε > 0, ∀ᶠ i in l, ∀ x ∈ s, ∥f i x∥ < ε := + tendsto_uniformly_on f 1 l s ↔ ∀ ε > 0, ∀ᶠ i in l, ∀ x ∈ s, ‖f i x‖ < ε := by simp_rw [tendsto_uniformly_on_iff, pi.one_apply, dist_one_left] @[to_additive] @@ -922,7 +922,7 @@ to_additive "A group homomorphism from an `add_group` to a `seminormed_add_group `seminormed_add_group` structure on the domain."] def seminormed_group.induced [group E] [seminormed_group F] [monoid_hom_class 𝓕 E F] (f : 𝓕) : seminormed_group E := -{ norm := λ x, ∥f x∥, +{ norm := λ x, ‖f x‖, dist_eq := λ x y, by simpa only [map_div, ←dist_eq_norm_div], ..pseudo_metric_space.induced f _ } @@ -970,16 +970,16 @@ by rw [dist_inv, inv_inv] @[simp, to_additive] lemma dist_div_left (a b₁ b₂ : E) : dist (a / b₁) (a / b₂) = dist b₁ b₂ := by simp only [div_eq_mul_inv, dist_mul_left, dist_inv_inv] -@[simp, to_additive] lemma dist_self_mul_right (a b : E) : dist a (a * b) = ∥b∥ := +@[simp, to_additive] lemma dist_self_mul_right (a b : E) : dist a (a * b) = ‖b‖ := by rw [←dist_one_left, ←dist_mul_left a 1 b, mul_one] -@[simp, to_additive] lemma dist_self_mul_left (a b : E) : dist (a * b) a = ∥b∥ := +@[simp, to_additive] lemma dist_self_mul_left (a b : E) : dist (a * b) a = ‖b‖ := by rw [dist_comm, dist_self_mul_right] -@[simp, to_additive] lemma dist_self_div_right (a b : E) : dist a (a / b) = ∥b∥ := +@[simp, to_additive] lemma dist_self_div_right (a b : E) : dist a (a / b) = ‖b‖ := by rw [div_eq_mul_inv, dist_self_mul_right, norm_inv'] -@[simp, to_additive] lemma dist_self_div_left (a b : E) : dist (a / b) a = ∥b∥ := +@[simp, to_additive] lemma dist_self_div_left (a b : E) : dist (a / b) a = ‖b‖ := by rw [dist_comm, dist_self_div_right] @[to_additive] lemma dist_mul_mul_le (a₁ a₂ b₁ b₂ : E) : @@ -1004,11 +1004,11 @@ by simpa only [dist_mul_left, dist_mul_right, dist_comm b₂] using abs_dist_sub_le (a₁ * a₂) (b₁ * b₂) (b₁ * a₂) lemma norm_multiset_sum_le {E} [seminormed_add_comm_group E] (m : multiset E) : - ∥m.sum∥ ≤ (m.map (λ x, ∥x∥)).sum := + ‖m.sum‖ ≤ (m.map (λ x, ‖x‖)).sum := m.le_sum_of_subadditive norm norm_zero norm_add_le @[to_additive] -lemma norm_multiset_prod_le (m : multiset E) : ∥m.prod∥ ≤ (m.map $ λ x, ∥x∥).sum := +lemma norm_multiset_prod_le (m : multiset E) : ‖m.prod‖ ≤ (m.map $ λ x, ‖x‖).sum := begin rw [←multiplicative.of_add_le, of_add_multiset_prod, multiset.map_map], refine multiset.le_prod_of_submultiplicative (multiplicative.of_add ∘ norm) _ (λ x y, _) _, @@ -1017,10 +1017,10 @@ begin end lemma norm_sum_le {E} [seminormed_add_comm_group E] (s : finset ι) (f : ι → E) : - ∥∑ i in s, f i∥ ≤ ∑ i in s, ∥f i∥ := + ‖∑ i in s, f i‖ ≤ ∑ i in s, ‖f i‖ := s.le_sum_of_subadditive norm norm_zero norm_add_le f -@[to_additive] lemma norm_prod_le (s : finset ι) (f : ι → E) : ∥∏ i in s, f i∥ ≤ ∑ i in s, ∥f i∥ := +@[to_additive] lemma norm_prod_le (s : finset ι) (f : ι → E) : ‖∏ i in s, f i‖ ≤ ∑ i in s, ‖f i‖ := begin rw [←multiplicative.of_add_le, of_add_sum], refine finset.le_prod_of_submultiplicative (multiplicative.of_add ∘ norm) _ (λ x y, _) _ _, @@ -1029,8 +1029,8 @@ begin end @[to_additive] -lemma norm_prod_le_of_le (s : finset ι) {f : ι → E} {n : ι → ℝ} (h : ∀ b ∈ s, ∥f b∥ ≤ n b) : - ∥∏ b in s, f b∥ ≤ ∑ b in s, n b := +lemma norm_prod_le_of_le (s : finset ι) {f : ι → E} {n : ι → ℝ} (h : ∀ b ∈ s, ‖f b‖ ≤ n b) : + ‖∏ b in s, f b‖ ≤ ∑ b in s, n b := (norm_prod_le s f).trans $ finset.sum_le_sum h @[to_additive] lemma dist_prod_prod_le_of_le (s : finset ι) {f a : ι → E} {d : ι → ℝ} @@ -1042,10 +1042,10 @@ by { simp only [dist_eq_norm_div, ← finset.prod_div_distrib] at *, exact norm_ dist (∏ b in s, f b) (∏ b in s, a b) ≤ ∑ b in s, dist (f b) (a b) := dist_prod_prod_le_of_le s $ λ _ _, le_rfl -@[to_additive] lemma mul_mem_ball_iff_norm : a * b ∈ ball a r ↔ ∥b∥ < r := +@[to_additive] lemma mul_mem_ball_iff_norm : a * b ∈ ball a r ↔ ‖b‖ < r := by rw [mem_ball_iff_norm'', mul_div_cancel'''] -@[to_additive] lemma mul_mem_closed_ball_iff_norm : a * b ∈ closed_ball a r ↔ ∥b∥ ≤ r := +@[to_additive] lemma mul_mem_closed_ball_iff_norm : a * b ∈ closed_ball a r ↔ ‖b‖ ≤ r := by rw [mem_closed_ball_iff_norm'', mul_div_cancel'''] @[simp, to_additive] lemma preimage_mul_ball (a b : E) (r : ℝ) : @@ -1100,22 +1100,22 @@ open finset ∃ v : ℕ → E, tendsto (λ n, ∏ i in range (n+1), v i) at_top (𝓝 a) ∧ (∀ n, v n ∈ s) ∧ - ∥v 0 / a∥ < b 0 ∧ - ∀ n, 0 < n → ∥v n∥ < b n := + ‖v 0 / a‖ < b 0 ∧ + ∀ n, 0 < n → ‖v n‖ < b n := begin obtain ⟨u : ℕ → E, u_in : ∀ n, u n ∈ s, lim_u : tendsto u at_top (𝓝 a)⟩ := mem_closure_iff_seq_limit.mp hg, - obtain ⟨n₀, hn₀⟩ : ∃ n₀, ∀ n ≥ n₀, ∥u n / a∥ < b 0, - { have : {x | ∥x / a∥ < b 0} ∈ 𝓝 a, + obtain ⟨n₀, hn₀⟩ : ∃ n₀, ∀ n ≥ n₀, ‖u n / a‖ < b 0, + { have : {x | ‖x / a‖ < b 0} ∈ 𝓝 a, { simp_rw ← dist_eq_norm_div, exact metric.ball_mem_nhds _ (b_pos _) }, exact filter.tendsto_at_top'.mp lim_u _ this }, set z : ℕ → E := λ n, u (n + n₀), have lim_z : tendsto z at_top (𝓝 a) := lim_u.comp (tendsto_add_at_top_nat n₀), - have mem_𝓤 : ∀ n, {p : E × E | ∥p.1 / p.2∥ < b (n + 1)} ∈ 𝓤 E := + have mem_𝓤 : ∀ n, {p : E × E | ‖p.1 / p.2‖ < b (n + 1)} ∈ 𝓤 E := λ n, by simpa [← dist_eq_norm_div] using metric.dist_mem_uniformity (b_pos $ n+1), obtain ⟨φ : ℕ → ℕ, φ_extr : strict_mono φ, - hφ : ∀ n, ∥z (φ $ n + 1) / z (φ n)∥ < b (n + 1)⟩ := + hφ : ∀ n, ‖z (φ $ n + 1) / z (φ n)‖ < b (n + 1)⟩ := lim_z.cauchy_seq.subseq_mem mem_𝓤, set w : ℕ → E := z ∘ φ, have hw : tendsto w at_top (𝓝 a), @@ -1136,8 +1136,8 @@ end (hb : b ∈ closure (j.range : set F)) {f : ℕ → ℝ} (b_pos : ∀ n, 0 < f n) : ∃ a : ℕ → E, tendsto (λ n, ∏ i in range (n + 1), j (a i)) at_top (𝓝 b) ∧ - ∥j (a 0) / b∥ < f 0 ∧ - ∀ n, 0 < n → ∥j (a n)∥ < f n := + ‖j (a 0) / b‖ < f 0 ∧ + ∀ n, 0 < n → ‖j (a n)‖ < f n := begin obtain ⟨v, sum_v, v_in, hv₀, hv_pos⟩ := controlled_prod_of_mem_closure hb b_pos, choose g hg using v_in, @@ -1166,50 +1166,50 @@ by rw [edist_inv, inv_inv] by simp only [div_eq_mul_inv, edist_mul_left, edist_inv_inv] @[to_additive] -lemma nnnorm_multiset_prod_le (m : multiset E) : ∥m.prod∥₊ ≤ (m.map (λ x, ∥x∥₊)).sum := +lemma nnnorm_multiset_prod_le (m : multiset E) : ‖m.prod‖₊ ≤ (m.map (λ x, ‖x‖₊)).sum := nnreal.coe_le_coe.1 $ by { push_cast, rw multiset.map_map, exact norm_multiset_prod_le _ } @[to_additive] lemma nnnorm_prod_le (s : finset ι) (f : ι → E) : - ∥∏ a in s, f a∥₊ ≤ ∑ a in s, ∥f a∥₊ := + ‖∏ a in s, f a‖₊ ≤ ∑ a in s, ‖f a‖₊ := nnreal.coe_le_coe.1 $ by { push_cast, exact norm_prod_le _ _ } @[to_additive] -lemma nnnorm_prod_le_of_le (s : finset ι) {f : ι → E} {n : ι → ℝ≥0} (h : ∀ b ∈ s, ∥f b∥₊ ≤ n b) : - ∥∏ b in s, f b∥₊ ≤ ∑ b in s, n b := +lemma nnnorm_prod_le_of_le (s : finset ι) {f : ι → E} {n : ι → ℝ≥0} (h : ∀ b ∈ s, ‖f b‖₊ ≤ n b) : + ‖∏ b in s, f b‖₊ ≤ ∑ b in s, n b := (norm_prod_le_of_le s h).trans_eq nnreal.coe_sum.symm namespace real instance : has_norm ℝ := { norm := λ r, |r| } -@[simp] lemma norm_eq_abs (r : ℝ) : ∥r∥ = |r| := rfl +@[simp] lemma norm_eq_abs (r : ℝ) : ‖r‖ = |r| := rfl instance : normed_add_comm_group ℝ := ⟨λ r y, rfl⟩ -lemma norm_of_nonneg (hr : 0 ≤ r) : ∥r∥ = r := abs_of_nonneg hr -lemma norm_of_nonpos (hr : r ≤ 0) : ∥r∥ = -r := abs_of_nonpos hr -lemma le_norm_self (r : ℝ) : r ≤ ∥r∥ := le_abs_self r +lemma norm_of_nonneg (hr : 0 ≤ r) : ‖r‖ = r := abs_of_nonneg hr +lemma norm_of_nonpos (hr : r ≤ 0) : ‖r‖ = -r := abs_of_nonpos hr +lemma le_norm_self (r : ℝ) : r ≤ ‖r‖ := le_abs_self r -@[simp] lemma norm_coe_nat (n : ℕ) : ∥(n : ℝ)∥ = n := abs_of_nonneg n.cast_nonneg -@[simp] lemma nnnorm_coe_nat (n : ℕ) : ∥(n : ℝ)∥₊ = n := nnreal.eq $ norm_coe_nat _ +@[simp] lemma norm_coe_nat (n : ℕ) : ‖(n : ℝ)‖ = n := abs_of_nonneg n.cast_nonneg +@[simp] lemma nnnorm_coe_nat (n : ℕ) : ‖(n : ℝ)‖₊ = n := nnreal.eq $ norm_coe_nat _ -@[simp] lemma norm_two : ∥(2 : ℝ)∥ = 2 := abs_of_pos (@zero_lt_two ℝ _ _) +@[simp] lemma norm_two : ‖(2 : ℝ)‖ = 2 := abs_of_pos (@zero_lt_two ℝ _ _) -@[simp] lemma nnnorm_two : ∥(2 : ℝ)∥₊ = 2 := nnreal.eq $ by simp +@[simp] lemma nnnorm_two : ‖(2 : ℝ)‖₊ = 2 := nnreal.eq $ by simp -lemma nnnorm_of_nonneg (hr : 0 ≤ r) : ∥r∥₊ = ⟨r, hr⟩ := nnreal.eq $ norm_of_nonneg hr +lemma nnnorm_of_nonneg (hr : 0 ≤ r) : ‖r‖₊ = ⟨r, hr⟩ := nnreal.eq $ norm_of_nonneg hr -lemma ennnorm_eq_of_real (hr : 0 ≤ r) : (∥r∥₊ : ℝ≥0∞) = ennreal.of_real r := +lemma ennnorm_eq_of_real (hr : 0 ≤ r) : (‖r‖₊ : ℝ≥0∞) = ennreal.of_real r := by { rw [← of_real_norm_eq_coe_nnnorm, norm_of_nonneg hr] } -lemma to_nnreal_eq_nnnorm_of_nonneg (hr : 0 ≤ r) : r.to_nnreal = ∥r∥₊ := +lemma to_nnreal_eq_nnnorm_of_nonneg (hr : 0 ≤ r) : r.to_nnreal = ‖r‖₊ := begin rw real.to_nnreal_of_nonneg hr, congr, rw [real.norm_eq_abs, abs_of_nonneg hr], end -lemma of_real_le_ennnorm (r : ℝ) : ennreal.of_real r ≤ ∥r∥₊ := +lemma of_real_le_ennnorm (r : ℝ) : ennreal.of_real r ≤ ‖r‖₊ := begin obtain hr | hr := le_total 0 r, { exact (real.ennnorm_eq_of_real hr).ge }, @@ -1259,7 +1259,7 @@ end by simpa only [pi.div_apply, mul_div_cancel'_right] using hf.mul_lipschitz_with hg hK @[to_additive] lemma le_mul_norm_div {f : E → F} (hf : antilipschitz_with K f) (x y : E) : - ∥x / y∥ ≤ K * ∥f x / f y∥ := + ‖x / y‖ ≤ K * ‖f x / f y‖ := by simp [← dist_eq_norm_div, hf.le_mul_dist x y] end antilipschitz_with @@ -1300,36 +1300,36 @@ end seminormed_comm_group section normed_group variables [normed_group E] [normed_group F] {a b : E} -@[simp, to_additive norm_eq_zero] lemma norm_eq_zero'' : ∥a∥ = 0 ↔ a = 1 := norm_eq_zero''' +@[simp, to_additive norm_eq_zero] lemma norm_eq_zero'' : ‖a‖ = 0 ↔ a = 1 := norm_eq_zero''' -@[to_additive norm_ne_zero_iff] lemma norm_ne_zero_iff' : ∥a∥ ≠ 0 ↔ a ≠ 1 := norm_eq_zero''.not +@[to_additive norm_ne_zero_iff] lemma norm_ne_zero_iff' : ‖a‖ ≠ 0 ↔ a ≠ 1 := norm_eq_zero''.not -@[simp, to_additive norm_pos_iff] lemma norm_pos_iff'' : 0 < ∥a∥ ↔ a ≠ 1 := norm_pos_iff''' +@[simp, to_additive norm_pos_iff] lemma norm_pos_iff'' : 0 < ‖a‖ ↔ a ≠ 1 := norm_pos_iff''' @[simp, to_additive norm_le_zero_iff] -lemma norm_le_zero_iff'' : ∥a∥ ≤ 0 ↔ a = 1 := norm_le_zero_iff''' +lemma norm_le_zero_iff'' : ‖a‖ ≤ 0 ↔ a = 1 := norm_le_zero_iff''' @[to_additive] -lemma norm_div_eq_zero_iff : ∥a / b∥ = 0 ↔ a = b := by rw [norm_eq_zero'', div_eq_one] +lemma norm_div_eq_zero_iff : ‖a / b‖ = 0 ↔ a = b := by rw [norm_eq_zero'', div_eq_one] -@[to_additive] lemma norm_div_pos_iff : 0 < ∥a / b∥ ↔ a ≠ b := +@[to_additive] lemma norm_div_pos_iff : 0 < ‖a / b‖ ↔ a ≠ b := by { rw [(norm_nonneg' _).lt_iff_ne, ne_comm], exact norm_div_eq_zero_iff.not } -@[to_additive] lemma eq_of_norm_div_le_zero (h : ∥a / b∥ ≤ 0) : a = b := +@[to_additive] lemma eq_of_norm_div_le_zero (h : ‖a / b‖ ≤ 0) : a = b := by rwa [←div_eq_one, ← norm_le_zero_iff''] alias norm_div_eq_zero_iff ↔ eq_of_norm_div_eq_zero _ attribute [to_additive] eq_of_norm_div_eq_zero -@[simp, to_additive nnnorm_eq_zero] lemma nnnorm_eq_zero' : ∥a∥₊ = 0 ↔ a = 1 := +@[simp, to_additive nnnorm_eq_zero] lemma nnnorm_eq_zero' : ‖a‖₊ = 0 ↔ a = 1 := by rw [← nnreal.coe_eq_zero, coe_nnnorm', norm_eq_zero''] @[to_additive nnnorm_ne_zero_iff] -lemma nnnorm_ne_zero_iff' : ∥a∥₊ ≠ 0 ↔ a ≠ 1 := nnnorm_eq_zero'.not +lemma nnnorm_ne_zero_iff' : ‖a‖₊ ≠ 0 ↔ a ≠ 1 := nnnorm_eq_zero'.not @[to_additive] -lemma tendsto_norm_div_self_punctured_nhds (a : E) : tendsto (λ x, ∥x / a∥) (𝓝[≠] a) (𝓝[>] 0) := +lemma tendsto_norm_div_self_punctured_nhds (a : E) : tendsto (λ x, ‖x / a‖) (𝓝[≠] a) (𝓝[>] 0) := (tendsto_norm_div_self a).inf $ tendsto_principal_principal.2 $ λ x hx, norm_pos_iff''.2 $ div_ne_one.2 hx @@ -1352,13 +1352,13 @@ variables [normed_add_group E] [topological_space α] {f : α → E} /-! Some relations with `has_compact_support` -/ -lemma has_compact_support_norm_iff : has_compact_support (λ x, ∥f x∥) ↔ has_compact_support f := +lemma has_compact_support_norm_iff : has_compact_support (λ x, ‖f x‖) ↔ has_compact_support f := has_compact_support_comp_left $ λ x, norm_eq_zero alias has_compact_support_norm_iff ↔ _ has_compact_support.norm lemma continuous.bounded_above_of_compact_support (hf : continuous f) (h : has_compact_support f) : - ∃ C, ∀ x, ∥f x∥ ≤ C := + ∃ C, ∀ x, ‖f x‖ ≤ C := by simpa [bdd_above_def] using hf.norm.bdd_above_range_of_has_compact_support h.norm end normed_add_group @@ -1369,22 +1369,22 @@ namespace ulift section has_norm variables [has_norm E] -instance : has_norm (ulift E) := ⟨λ x, ∥x.down∥⟩ +instance : has_norm (ulift E) := ⟨λ x, ‖x.down‖⟩ -lemma norm_def (x : ulift E) : ∥x∥ = ∥x.down∥ := rfl -@[simp] lemma norm_up (x : E) : ∥ulift.up x∥ = ∥x∥ := rfl -@[simp] lemma norm_down (x : ulift E) : ∥x.down∥ = ∥x∥ := rfl +lemma norm_def (x : ulift E) : ‖x‖ = ‖x.down‖ := rfl +@[simp] lemma norm_up (x : E) : ‖ulift.up x‖ = ‖x‖ := rfl +@[simp] lemma norm_down (x : ulift E) : ‖x.down‖ = ‖x‖ := rfl end has_norm section has_nnnorm variables [has_nnnorm E] -instance : has_nnnorm (ulift E) := ⟨λ x, ∥x.down∥₊⟩ +instance : has_nnnorm (ulift E) := ⟨λ x, ‖x.down‖₊⟩ -lemma nnnorm_def (x : ulift E) : ∥x∥₊ = ∥x.down∥₊ := rfl -@[simp] lemma nnnorm_up (x : E) : ∥ulift.up x∥₊ = ∥x∥₊ := rfl -@[simp] lemma nnnorm_down (x : ulift E) : ∥x.down∥₊ = ∥x∥₊ := rfl +lemma nnnorm_def (x : ulift E) : ‖x‖₊ = ‖x.down‖₊ := rfl +@[simp] lemma nnnorm_up (x : E) : ‖ulift.up x‖₊ = ‖x‖₊ := rfl +@[simp] lemma nnnorm_down (x : ulift E) : ‖x.down‖₊ = ‖x‖₊ := rfl end has_nnnorm @@ -1416,10 +1416,10 @@ variables [has_norm E] instance : has_norm (additive E) := ‹has_norm E› instance : has_norm (multiplicative E) := ‹has_norm E› -@[simp] lemma norm_to_mul (x) : ∥(to_mul x : E)∥ = ∥x∥ := rfl -@[simp] lemma norm_of_mul (x : E) : ∥of_mul x∥ = ∥x∥ := rfl -@[simp] lemma norm_to_add (x) : ∥(to_add x : E)∥ = ∥x∥ := rfl -@[simp] lemma norm_of_add (x : E) : ∥of_add x∥ = ∥x∥ := rfl +@[simp] lemma norm_to_mul (x) : ‖(to_mul x : E)‖ = ‖x‖ := rfl +@[simp] lemma norm_of_mul (x : E) : ‖of_mul x‖ = ‖x‖ := rfl +@[simp] lemma norm_to_add (x) : ‖(to_add x : E)‖ = ‖x‖ := rfl +@[simp] lemma norm_of_add (x : E) : ‖of_add x‖ = ‖x‖ := rfl end has_norm @@ -1429,10 +1429,10 @@ variables [has_nnnorm E] instance : has_nnnorm (additive E) := ‹has_nnnorm E› instance : has_nnnorm (multiplicative E) := ‹has_nnnorm E› -@[simp] lemma nnnorm_to_mul (x) : ∥(to_mul x : E)∥₊ = ∥x∥₊ := rfl -@[simp] lemma nnnorm_of_mul (x : E) : ∥of_mul x∥₊ = ∥x∥₊ := rfl -@[simp] lemma nnnorm_to_add (x) : ∥(to_add x : E)∥₊ = ∥x∥₊ := rfl -@[simp] lemma nnnorm_of_add (x : E) : ∥of_add x∥₊ = ∥x∥₊ := rfl +@[simp] lemma nnnorm_to_mul (x) : ‖(to_mul x : E)‖₊ = ‖x‖₊ := rfl +@[simp] lemma nnnorm_of_mul (x : E) : ‖of_mul x‖₊ = ‖x‖₊ := rfl +@[simp] lemma nnnorm_to_add (x) : ‖(to_add x : E)‖₊ = ‖x‖₊ := rfl +@[simp] lemma nnnorm_of_add (x : E) : ‖of_add x‖₊ = ‖x‖₊ := rfl end has_nnnorm @@ -1473,8 +1473,8 @@ variables [has_norm E] instance : has_norm Eᵒᵈ := ‹has_norm E› -@[simp] lemma norm_to_dual (x : E) : ∥to_dual x∥ = ∥x∥ := rfl -@[simp] lemma norm_of_dual (x : Eᵒᵈ) : ∥of_dual x∥ = ∥x∥ := rfl +@[simp] lemma norm_to_dual (x : E) : ‖to_dual x‖ = ‖x‖ := rfl +@[simp] lemma norm_of_dual (x : Eᵒᵈ) : ‖of_dual x‖ = ‖x‖ := rfl end has_norm @@ -1483,8 +1483,8 @@ variables [has_nnnorm E] instance : has_nnnorm Eᵒᵈ := ‹has_nnnorm E› -@[simp] lemma nnnorm_to_dual (x : E) : ∥to_dual x∥₊ = ∥x∥₊ := rfl -@[simp] lemma nnnorm_of_dual (x : Eᵒᵈ) : ∥of_dual x∥₊ = ∥x∥₊ := rfl +@[simp] lemma nnnorm_to_dual (x : E) : ‖to_dual x‖₊ = ‖x‖₊ := rfl +@[simp] lemma nnnorm_of_dual (x : Eᵒᵈ) : ‖of_dual x‖₊ = ‖x‖₊ := rfl end has_nnnorm @@ -1507,13 +1507,13 @@ end order_dual section has_norm variables [has_norm E] [has_norm F] {x : E × F} {r : ℝ} -instance : has_norm (E × F) := ⟨λ x, ∥x.1∥ ⊔ ∥x.2∥⟩ +instance : has_norm (E × F) := ⟨λ x, ‖x.1‖ ⊔ ‖x.2‖⟩ -lemma prod.norm_def (x : E × F) : ∥x∥ = (max ∥x.1∥ ∥x.2∥) := rfl -lemma norm_fst_le (x : E × F) : ∥x.1∥ ≤ ∥x∥ := le_max_left _ _ -lemma norm_snd_le (x : E × F) : ∥x.2∥ ≤ ∥x∥ := le_max_right _ _ +lemma prod.norm_def (x : E × F) : ‖x‖ = (max ‖x.1‖ ‖x.2‖) := rfl +lemma norm_fst_le (x : E × F) : ‖x.1‖ ≤ ‖x‖ := le_max_left _ _ +lemma norm_snd_le (x : E × F) : ‖x.2‖ ≤ ‖x‖ := le_max_right _ _ -lemma norm_prod_le_iff : ∥x∥ ≤ r ↔ ∥x.1∥ ≤ r ∧ ∥x.2∥ ≤ r := max_le_iff +lemma norm_prod_le_iff : ‖x‖ ≤ r ↔ ‖x.1‖ ≤ r ∧ ‖x.2‖ ≤ r := max_le_iff end has_norm @@ -1526,7 +1526,7 @@ instance : seminormed_group (E × F) := ⟨λ x y, by simp only [prod.norm_def, prod.dist_eq, dist_eq_norm_div, prod.fst_div, prod.snd_div]⟩ @[to_additive prod.nnnorm_def'] -lemma prod.nnorm_def (x : E × F) : ∥x∥₊ = (max ∥x.1∥₊ ∥x.2∥₊) := rfl +lemma prod.nnorm_def (x : E × F) : ‖x‖₊ = (max ‖x.1‖₊ ‖x.2‖₊) := rfl end seminormed_group @@ -1556,28 +1556,28 @@ variables [Π i, seminormed_group (π i)] [seminormed_group E] (f : Π i, π i) /-- Finite product of seminormed groups, using the sup norm. -/ @[to_additive "Finite product of seminormed groups, using the sup norm."] instance : seminormed_group (Π i, π i) := -{ norm := λ f, ↑(finset.univ.sup (λ b, ∥f b∥₊)), +{ norm := λ f, ↑(finset.univ.sup (λ b, ‖f b‖₊)), dist_eq := λ x y, congr_arg (coe : ℝ≥0 → ℝ) $ congr_arg (finset.sup finset.univ) $ funext $ λ a, - show nndist (x a) (y a) = ∥x a / y a∥₊, from nndist_eq_nnnorm_div (x a) (y a) } + show nndist (x a) (y a) = ‖x a / y a‖₊, from nndist_eq_nnnorm_div (x a) (y a) } -@[to_additive pi.norm_def] lemma pi.norm_def' : ∥f∥ = ↑(finset.univ.sup (λ b, ∥f b∥₊)) := rfl -@[to_additive pi.nnnorm_def] lemma pi.nnnorm_def' : ∥f∥₊ = finset.univ.sup (λ b, ∥f b∥₊) := +@[to_additive pi.norm_def] lemma pi.norm_def' : ‖f‖ = ↑(finset.univ.sup (λ b, ‖f b‖₊)) := rfl +@[to_additive pi.nnnorm_def] lemma pi.nnnorm_def' : ‖f‖₊ = finset.univ.sup (λ b, ‖f b‖₊) := subtype.eta _ _ /-- The seminorm of an element in a product space is `≤ r` if and only if the norm of each component is. -/ @[to_additive pi_norm_le_iff_of_nonneg "The seminorm of an element in a product space is `≤ r` if and only if the norm of each component is."] -lemma pi_norm_le_iff_of_nonneg' (hr : 0 ≤ r) : ∥x∥ ≤ r ↔ ∀ i, ∥x i∥ ≤ r := +lemma pi_norm_le_iff_of_nonneg' (hr : 0 ≤ r) : ‖x‖ ≤ r ↔ ∀ i, ‖x i‖ ≤ r := by simp only [←dist_one_right, dist_pi_le_iff hr, pi.one_apply] @[to_additive pi_nnnorm_le_iff] -lemma pi_nnnorm_le_iff' {r : ℝ≥0} : ∥x∥₊ ≤ r ↔ ∀ i, ∥x i∥₊ ≤ r := +lemma pi_nnnorm_le_iff' {r : ℝ≥0} : ‖x‖₊ ≤ r ↔ ∀ i, ‖x i‖₊ ≤ r := pi_norm_le_iff_of_nonneg' r.coe_nonneg @[to_additive pi_norm_le_iff_of_nonempty] -lemma pi_norm_le_iff_of_nonempty' [nonempty ι] : ∥f∥ ≤ r ↔ ∀ b, ∥f b∥ ≤ r := +lemma pi_norm_le_iff_of_nonempty' [nonempty ι] : ‖f‖ ≤ r ↔ ∀ b, ‖f b‖ ≤ r := begin by_cases hr : 0 ≤ r, { exact pi_norm_le_iff_of_nonneg' hr }, @@ -1589,43 +1589,43 @@ end component is. -/ @[to_additive pi_norm_lt_iff "The seminorm of an element in a product space is `< r` if and only if the norm of each component is."] -lemma pi_norm_lt_iff' (hr : 0 < r) : ∥x∥ < r ↔ ∀ i, ∥x i∥ < r := +lemma pi_norm_lt_iff' (hr : 0 < r) : ‖x‖ < r ↔ ∀ i, ‖x i‖ < r := by simp only [←dist_one_right, dist_pi_lt_iff hr, pi.one_apply] @[to_additive pi_nnnorm_lt_iff] -lemma pi_nnnorm_lt_iff' {r : ℝ≥0} (hr : 0 < r) : ∥x∥₊ < r ↔ ∀ i, ∥x i∥₊ < r := pi_norm_lt_iff' hr +lemma pi_nnnorm_lt_iff' {r : ℝ≥0} (hr : 0 < r) : ‖x‖₊ < r ↔ ∀ i, ‖x i‖₊ < r := pi_norm_lt_iff' hr @[to_additive norm_le_pi_norm] -lemma norm_le_pi_norm' (i : ι) : ∥f i∥ ≤ ∥f∥ := +lemma norm_le_pi_norm' (i : ι) : ‖f i‖ ≤ ‖f‖ := (pi_norm_le_iff_of_nonneg' $ norm_nonneg' _).1 le_rfl i @[to_additive nnnorm_le_pi_nnnorm] -lemma nnnorm_le_pi_nnnorm' (i : ι) : ∥f i∥₊ ≤ ∥f∥₊ := norm_le_pi_norm' _ i +lemma nnnorm_le_pi_nnnorm' (i : ι) : ‖f i‖₊ ≤ ‖f‖₊ := norm_le_pi_norm' _ i @[to_additive pi_norm_const_le] -lemma pi_norm_const_le' (a : E) : ∥(λ _ : ι, a)∥ ≤ ∥a∥ := +lemma pi_norm_const_le' (a : E) : ‖(λ _ : ι, a)‖ ≤ ‖a‖ := (pi_norm_le_iff_of_nonneg' $ norm_nonneg' _).2 $ λ _, le_rfl @[to_additive pi_nnnorm_const_le] -lemma pi_nnnorm_const_le' (a : E) : ∥(λ _ : ι, a)∥₊ ≤ ∥a∥₊ := pi_norm_const_le' _ +lemma pi_nnnorm_const_le' (a : E) : ‖(λ _ : ι, a)‖₊ ≤ ‖a‖₊ := pi_norm_const_le' _ @[simp, to_additive pi_norm_const] -lemma pi_norm_const' [nonempty ι] (a : E) : ∥(λ i : ι, a)∥ = ∥a∥ := +lemma pi_norm_const' [nonempty ι] (a : E) : ‖(λ i : ι, a)‖ = ‖a‖ := by simpa only [←dist_one_right] using dist_pi_const a 1 @[simp, to_additive pi_nnnorm_const] -lemma pi_nnnorm_const' [nonempty ι] (a : E) : ∥(λ i : ι, a)∥₊ = ∥a∥₊ := nnreal.eq $ pi_norm_const' a +lemma pi_nnnorm_const' [nonempty ι] (a : E) : ‖(λ i : ι, a)‖₊ = ‖a‖₊ := nnreal.eq $ pi_norm_const' a /-- The $L^1$ norm is less than the $L^\infty$ norm scaled by the cardinality. -/ @[to_additive pi.sum_norm_apply_le_norm "The $L^1$ norm is less than the $L^\\infty$ norm scaled by the cardinality."] -lemma pi.sum_norm_apply_le_norm' : ∑ i, ∥f i∥ ≤ fintype.card ι • ∥f∥ := +lemma pi.sum_norm_apply_le_norm' : ∑ i, ‖f i‖ ≤ fintype.card ι • ‖f‖ := finset.sum_le_card_nsmul _ _ _ $ λ i hi, norm_le_pi_norm' _ i /-- The $L^1$ norm is less than the $L^\infty$ norm scaled by the cardinality. -/ @[to_additive pi.sum_nnnorm_apply_le_nnnorm "The $L^1$ norm is less than the $L^\\infty$ norm scaled by the cardinality."] -lemma pi.sum_nnnorm_apply_le_nnnorm' : ∑ i, ∥f i∥₊ ≤ fintype.card ι • ∥f∥₊ := +lemma pi.sum_nnnorm_apply_le_nnnorm' : ∑ i, ‖f i‖₊ ≤ fintype.card ι • ‖f‖₊ := nnreal.coe_sum.trans_le $ pi.sum_norm_apply_le_norm' _ end seminormed_group @@ -1664,7 +1664,7 @@ instance seminormed_group : seminormed_group s := seminormed_group.induced _ _ s its norm in `E`. -/ @[simp, to_additive "If `x` is an element of a subgroup `s` of a seminormed group `E`, its norm in `s` is equal to its norm in `E`."] -lemma coe_norm (x : s) : ∥x∥ = ∥(x : E)∥ := rfl +lemma coe_norm (x : s) : ‖x‖ = ‖(x : E)‖ := rfl /-- If `x` is an element of a subgroup `s` of a seminormed group `E`, its norm in `s` is equal to its norm in `E`. @@ -1674,7 +1674,7 @@ This is a reversed version of the `simp` lemma `subgroup.coe_norm` for use by `n in `s` is equal to its norm in `E`. This is a reversed version of the `simp` lemma `add_subgroup.coe_norm` for use by `norm_cast`."] -lemma norm_coe {s : subgroup E} (x : s) : ∥(x : E)∥ = ∥x∥ := rfl +lemma norm_coe {s : subgroup E} (x : s) : ‖(x : E)‖ = ‖x‖ := rfl end seminormed_group @@ -1708,7 +1708,7 @@ norm in `E`. -/ -- See note [implicit instance arguments]. @[simp] lemma coe_norm {_ : ring 𝕜} [seminormed_add_comm_group E] {_ : module 𝕜 E} {s : submodule 𝕜 E} (x : s) : - ∥x∥ = ∥(x : E)∥ := rfl + ‖x‖ = ‖(x : E)‖ := rfl /-- If `x` is an element of a submodule `s` of a normed group `E`, its norm in `E` is equal to its norm in `s`. @@ -1717,7 +1717,7 @@ This is a reversed version of the `simp` lemma `submodule.coe_norm` for use by ` -- See note [implicit instance arguments]. @[norm_cast] lemma norm_coe {_ : ring 𝕜} [seminormed_add_comm_group E] {_ : module 𝕜 E} {s : submodule 𝕜 E} (x : s) : - ∥(x : E)∥ = ∥x∥ := rfl + ‖(x : E)‖ = ‖x‖ := rfl /-- A submodule of a normed group is also a normed group, with the restriction of the norm. -/ -- See note [implicit instance arguments]. diff --git a/src/analysis/normed/group/completion.lean b/src/analysis/normed/group/completion.lean index de99d3ecfcde7..8e3ba8389dc67 100644 --- a/src/analysis/normed/group/completion.lean +++ b/src/analysis/normed/group/completion.lean @@ -29,7 +29,7 @@ instance [uniform_space E] [has_norm E] : { norm := completion.extension has_norm.norm } @[simp] lemma norm_coe {E} [seminormed_add_comm_group E] (x : E) : - ∥(x : completion E)∥ = ∥x∥ := + ‖(x : completion E)‖ = ‖x‖ := completion.extension_coe uniform_continuous_norm x instance [seminormed_add_comm_group E] : normed_add_comm_group (completion E) := diff --git a/src/analysis/normed/group/hom.lean b/src/analysis/normed/group/hom.lean index 21172be652425..62bd8e1f499b3 100644 --- a/src/analysis/normed/group/hom.lean +++ b/src/analysis/normed/group/hom.lean @@ -20,7 +20,7 @@ simple constructions for normed group homs, like kernel, range and equalizer. Some easy other constructions are related to subgroups of normed groups. -Since a lot of elementary properties don't require `∥x∥ = 0 → x = 0` we start setting up the +Since a lot of elementary properties don't require `‖x‖ = 0 → x = 0` we start setting up the theory of `seminormed_add_group_hom` and we specialize to `normed_add_group_hom` when needed. -/ @@ -32,7 +32,7 @@ structure normed_add_group_hom (V W : Type*) [seminormed_add_comm_group V] [seminormed_add_comm_group W] := (to_fun : V → W) (map_add' : ∀ v₁ v₂, to_fun (v₁ + v₂) = to_fun v₁ + to_fun v₂) -(bound' : ∃ C, ∀ v, ∥to_fun v∥ ≤ C * ∥v∥) +(bound' : ∃ C, ∀ v, ‖to_fun v‖ ≤ C * ‖v‖) namespace add_monoid_hom @@ -43,13 +43,13 @@ variables {V W : Type*} [seminormed_add_comm_group V] [seminormed_add_comm_group See `add_monoid_hom.mk_normed_add_group_hom'` for a version that uses `ℝ≥0` for the bound. -/ def mk_normed_add_group_hom (f : V →+ W) - (C : ℝ) (h : ∀ v, ∥f v∥ ≤ C * ∥v∥) : normed_add_group_hom V W := + (C : ℝ) (h : ∀ v, ‖f v‖ ≤ C * ‖v‖) : normed_add_group_hom V W := { bound' := ⟨C, h⟩, ..f } /-- Associate to a group homomorphism a bounded group homomorphism under a norm control condition. See `add_monoid_hom.mk_normed_add_group_hom` for a version that uses `ℝ` for the bound. -/ -def mk_normed_add_group_hom' (f : V →+ W) (C : ℝ≥0) (hC : ∀ x, ∥f x∥₊ ≤ C * ∥x∥₊) : +def mk_normed_add_group_hom' (f : V →+ W) (C : ℝ≥0) (hC : ∀ x, ‖f x‖₊ ≤ C * ‖x‖₊) : normed_add_group_hom V W := { bound' := ⟨C, hC⟩ .. f} @@ -57,11 +57,11 @@ end add_monoid_hom lemma exists_pos_bound_of_bound {V W : Type*} [seminormed_add_comm_group V] [seminormed_add_comm_group W] - {f : V → W} (M : ℝ) (h : ∀x, ∥f x∥ ≤ M * ∥x∥) : - ∃ N, 0 < N ∧ ∀x, ∥f x∥ ≤ N * ∥x∥ := + {f : V → W} (M : ℝ) (h : ∀x, ‖f x‖ ≤ M * ‖x‖) : + ∃ N, 0 < N ∧ ∀x, ‖f x‖ ≤ N * ‖x‖ := ⟨max M 1, lt_of_lt_of_le zero_lt_one (le_max_right _ _), λx, calc - ∥f x∥ ≤ M * ∥x∥ : h x - ... ≤ max M 1 * ∥x∥ : mul_le_mul_of_nonneg_right (le_max_left _ _) (norm_nonneg _) ⟩ + ‖f x‖ ≤ M * ‖x‖ : h x + ... ≤ max M 1 * ‖x‖ : mul_le_mul_of_nonneg_right (le_max_left _ _) (norm_nonneg _) ⟩ namespace normed_add_group_hom @@ -117,20 +117,20 @@ instance : add_monoid_hom_class (normed_add_group_hom V₁ V₂) V₁ V₂ := map_add := λ f, f.to_add_monoid_hom.map_add, map_zero := λ f, f.to_add_monoid_hom.map_zero } -lemma bound : ∃ C, 0 < C ∧ ∀ x, ∥f x∥ ≤ C * ∥x∥ := +lemma bound : ∃ C, 0 < C ∧ ∀ x, ‖f x‖ ≤ C * ‖x‖ := let ⟨C, hC⟩ := f.bound' in exists_pos_bound_of_bound _ hC -theorem antilipschitz_of_norm_ge {K : ℝ≥0} (h : ∀ x, ∥x∥ ≤ K * ∥f x∥) : +theorem antilipschitz_of_norm_ge {K : ℝ≥0} (h : ∀ x, ‖x‖ ≤ K * ‖f x‖) : antilipschitz_with K f := antilipschitz_with.of_le_mul_dist $ λ x y, by simpa only [dist_eq_norm, map_sub] using h (x - y) /-- A normed group hom is surjective on the subgroup `K` with constant `C` if every element -`x` of `K` has a preimage whose norm is bounded above by `C*∥x∥`. This is a more +`x` of `K` has a preimage whose norm is bounded above by `C*‖x‖`. This is a more abstract version of `f` having a right inverse defined on `K` with operator norm at most `C`. -/ def surjective_on_with (f : normed_add_group_hom V₁ V₂) (K : add_subgroup V₂) (C : ℝ) : Prop := - ∀ h ∈ K, ∃ g, f g = h ∧ ∥g∥ ≤ C*∥h∥ + ∀ h ∈ K, ∃ g, f g = h ∧ ‖g‖ ≤ C*‖h‖ lemma surjective_on_with.mono {f : normed_add_group_hom V₁ V₂} {K : add_subgroup V₂} {C C' : ℝ} (h : f.surjective_on_with K C) (H : C ≤ C') : f.surjective_on_with K C' := @@ -138,7 +138,7 @@ begin intros k k_in, rcases h k k_in with ⟨g, rfl, hg⟩, use [g, rfl], - by_cases Hg : ∥f g∥ = 0, + by_cases Hg : ‖f g‖ = 0, { simpa [Hg] using hg }, { exact hg.trans ((mul_le_mul_right $ (ne.symm Hg).le_iff_lt.mp (norm_nonneg _)).mpr H) } end @@ -159,44 +159,44 @@ lemma surjective_on_with.surj_on {f : normed_add_group_hom V₁ V₂} {K : add_s /-! ### The operator norm -/ /-- The operator norm of a seminormed group homomorphism is the inf of all its bounds. -/ -def op_norm (f : normed_add_group_hom V₁ V₂) := Inf {c | 0 ≤ c ∧ ∀ x, ∥f x∥ ≤ c * ∥x∥} +def op_norm (f : normed_add_group_hom V₁ V₂) := Inf {c | 0 ≤ c ∧ ∀ x, ‖f x‖ ≤ c * ‖x‖} instance has_op_norm : has_norm (normed_add_group_hom V₁ V₂) := ⟨op_norm⟩ -lemma norm_def : ∥f∥ = Inf {c | 0 ≤ c ∧ ∀ x, ∥f x∥ ≤ c * ∥x∥} := rfl +lemma norm_def : ‖f‖ = Inf {c | 0 ≤ c ∧ ∀ x, ‖f x‖ ≤ c * ‖x‖} := rfl -- So that invocations of `le_cInf` make sense: we show that the set of -- bounds is nonempty and bounded below. lemma bounds_nonempty {f : normed_add_group_hom V₁ V₂} : - ∃ c, c ∈ { c | 0 ≤ c ∧ ∀ x, ∥f x∥ ≤ c * ∥x∥ } := + ∃ c, c ∈ { c | 0 ≤ c ∧ ∀ x, ‖f x‖ ≤ c * ‖x‖ } := let ⟨M, hMp, hMb⟩ := f.bound in ⟨M, le_of_lt hMp, hMb⟩ lemma bounds_bdd_below {f : normed_add_group_hom V₁ V₂} : - bdd_below {c | 0 ≤ c ∧ ∀ x, ∥f x∥ ≤ c * ∥x∥} := + bdd_below {c | 0 ≤ c ∧ ∀ x, ‖f x‖ ≤ c * ‖x‖} := ⟨0, λ _ ⟨hn, _⟩, hn⟩ -lemma op_norm_nonneg : 0 ≤ ∥f∥ := +lemma op_norm_nonneg : 0 ≤ ‖f‖ := le_cInf bounds_nonempty (λ _ ⟨hx, _⟩, hx) -/-- The fundamental property of the operator norm: `∥f x∥ ≤ ∥f∥ * ∥x∥`. -/ -theorem le_op_norm (x : V₁) : ∥f x∥ ≤ ∥f∥ * ∥x∥ := +/-- The fundamental property of the operator norm: `‖f x‖ ≤ ‖f‖ * ‖x‖`. -/ +theorem le_op_norm (x : V₁) : ‖f x‖ ≤ ‖f‖ * ‖x‖ := begin obtain ⟨C, Cpos, hC⟩ := f.bound, replace hC := hC x, - by_cases h : ∥x∥ = 0, + by_cases h : ‖x‖ = 0, { rwa [h, mul_zero] at ⊢ hC }, - have hlt : 0 < ∥x∥ := lt_of_le_of_ne (norm_nonneg x) (ne.symm h), + have hlt : 0 < ‖x‖ := lt_of_le_of_ne (norm_nonneg x) (ne.symm h), exact (div_le_iff hlt).mp (le_cInf bounds_nonempty (λ c ⟨_, hc⟩, (div_le_iff hlt).mpr $ by { apply hc })), end -theorem le_op_norm_of_le {c : ℝ} {x} (h : ∥x∥ ≤ c) : ∥f x∥ ≤ ∥f∥ * c := +theorem le_op_norm_of_le {c : ℝ} {x} (h : ‖x‖ ≤ c) : ‖f x‖ ≤ ‖f‖ * c := le_trans (f.le_op_norm x) (mul_le_mul_of_nonneg_left h f.op_norm_nonneg) -theorem le_of_op_norm_le {c : ℝ} (h : ∥f∥ ≤ c) (x : V₁) : ∥f x∥ ≤ c * ∥x∥ := +theorem le_of_op_norm_le {c : ℝ} (h : ‖f‖ ≤ c) (x : V₁) : ‖f x‖ ≤ c * ‖x‖ := (f.le_op_norm x).trans (mul_le_mul_of_nonneg_right h (norm_nonneg x)) /-- continuous linear maps are Lipschitz continuous. -/ -theorem lipschitz : lipschitz_with ⟨∥f∥, op_norm_nonneg f⟩ f := +theorem lipschitz : lipschitz_with ⟨‖f‖, op_norm_nonneg f⟩ f := lipschitz_with.of_dist_le_mul $ λ x y, by { rw [dist_eq_norm, dist_eq_norm, ←map_sub], apply le_op_norm } @@ -207,39 +207,39 @@ protected lemma uniform_continuous (f : normed_add_group_hom V₁ V₂) : protected lemma continuous (f : normed_add_group_hom V₁ V₂) : continuous f := f.uniform_continuous.continuous -lemma ratio_le_op_norm (x : V₁) : ∥f x∥ / ∥x∥ ≤ ∥f∥ := +lemma ratio_le_op_norm (x : V₁) : ‖f x‖ / ‖x‖ ≤ ‖f‖ := div_le_of_nonneg_of_le_mul (norm_nonneg _) f.op_norm_nonneg (le_op_norm _ _) /-- If one controls the norm of every `f x`, then one controls the norm of `f`. -/ -lemma op_norm_le_bound {M : ℝ} (hMp: 0 ≤ M) (hM : ∀ x, ∥f x∥ ≤ M * ∥x∥) : - ∥f∥ ≤ M := +lemma op_norm_le_bound {M : ℝ} (hMp: 0 ≤ M) (hM : ∀ x, ‖f x‖ ≤ M * ‖x‖) : + ‖f‖ ≤ M := cInf_le bounds_bdd_below ⟨hMp, hM⟩ lemma op_norm_eq_of_bounds {M : ℝ} (M_nonneg : 0 ≤ M) - (h_above : ∀ x, ∥f x∥ ≤ M*∥x∥) (h_below : ∀ N ≥ 0, (∀ x, ∥f x∥ ≤ N*∥x∥) → M ≤ N) : - ∥f∥ = M := + (h_above : ∀ x, ‖f x‖ ≤ M*‖x‖) (h_below : ∀ N ≥ 0, (∀ x, ‖f x‖ ≤ N*‖x‖) → M ≤ N) : + ‖f‖ = M := le_antisymm (f.op_norm_le_bound M_nonneg h_above) ((le_cInf_iff normed_add_group_hom.bounds_bdd_below ⟨M, M_nonneg, h_above⟩).mpr $ λ N ⟨N_nonneg, hN⟩, h_below N N_nonneg hN) theorem op_norm_le_of_lipschitz {f : normed_add_group_hom V₁ V₂} {K : ℝ≥0} (hf : lipschitz_with K f) : - ∥f∥ ≤ K := + ‖f‖ ≤ K := f.op_norm_le_bound K.2 $ λ x, by simpa only [dist_zero_right, map_zero] using hf.dist_le_mul x 0 /-- If a bounded group homomorphism map is constructed from a group homomorphism via the constructor `mk_normed_add_group_hom`, then its norm is bounded by the bound given to the constructor if it is nonnegative. -/ lemma mk_normed_add_group_hom_norm_le (f : V₁ →+ V₂) {C : ℝ} (hC : 0 ≤ C) - (h : ∀ x, ∥f x∥ ≤ C * ∥x∥) : - ∥f.mk_normed_add_group_hom C h∥ ≤ C := + (h : ∀ x, ‖f x‖ ≤ C * ‖x‖) : + ‖f.mk_normed_add_group_hom C h‖ ≤ C := op_norm_le_bound _ hC h /-- If a bounded group homomorphism map is constructed from a group homomorphism via the constructor `mk_normed_add_group_hom`, then its norm is bounded by the bound given to the constructor or zero if this bound is negative. -/ -lemma mk_normed_add_group_hom_norm_le' (f : V₁ →+ V₂) {C : ℝ} (h : ∀x, ∥f x∥ ≤ C * ∥x∥) : - ∥f.mk_normed_add_group_hom C h∥ ≤ max C 0 := +lemma mk_normed_add_group_hom_norm_le' (f : V₁ →+ V₂) {C : ℝ} (h : ∀x, ‖f x‖ ≤ C * ‖x‖) : + ‖f.mk_normed_add_group_hom C h‖ ≤ max C 0 := op_norm_le_bound _ (le_max_right _ _) $ λ x, (h x).trans $ mul_le_mul_of_nonneg_right (le_max_left _ _) (norm_nonneg x) @@ -250,14 +250,14 @@ alias mk_normed_add_group_hom_norm_le' ← _root_.add_monoid_hom.mk_normed_add_g /-- Addition of normed group homs. -/ instance : has_add (normed_add_group_hom V₁ V₂) := -⟨λ f g, (f.to_add_monoid_hom + g.to_add_monoid_hom).mk_normed_add_group_hom (∥f∥ + ∥g∥) $ λ v, calc - ∥f v + g v∥ - ≤ ∥f v∥ + ∥g v∥ : norm_add_le _ _ - ... ≤ ∥f∥ * ∥v∥ + ∥g∥ * ∥v∥ : add_le_add (le_op_norm f v) (le_op_norm g v) - ... = (∥f∥ + ∥g∥) * ∥v∥ : by rw add_mul⟩ +⟨λ f g, (f.to_add_monoid_hom + g.to_add_monoid_hom).mk_normed_add_group_hom (‖f‖ + ‖g‖) $ λ v, calc + ‖f v + g v‖ + ≤ ‖f v‖ + ‖g v‖ : norm_add_le _ _ + ... ≤ ‖f‖ * ‖v‖ + ‖g‖ * ‖v‖ : add_le_add (le_op_norm f v) (le_op_norm g v) + ... = (‖f‖ + ‖g‖) * ‖v‖ : by rw add_mul⟩ /-- The operator norm satisfies the triangle inequality. -/ -theorem op_norm_add_le : ∥f + g∥ ≤ ∥f∥ + ∥g∥ := +theorem op_norm_add_le : ‖f + g‖ ≤ ‖f‖ + ‖g‖ := mk_normed_add_group_hom_norm_le _ (add_nonneg (op_norm_nonneg _) (op_norm_nonneg _)) _ /-- @@ -281,17 +281,17 @@ instance : has_zero (normed_add_group_hom V₁ V₂) := instance : inhabited (normed_add_group_hom V₁ V₂) := ⟨0⟩ /-- The norm of the `0` operator is `0`. -/ -theorem op_norm_zero : ∥(0 : normed_add_group_hom V₁ V₂)∥ = 0 := +theorem op_norm_zero : ‖(0 : normed_add_group_hom V₁ V₂)‖ = 0 := le_antisymm (cInf_le bounds_bdd_below ⟨ge_of_eq rfl, λ _, le_of_eq (by { rw [zero_mul], exact norm_zero })⟩) (op_norm_nonneg _) /-- For normed groups, an operator is zero iff its norm vanishes. -/ theorem op_norm_zero_iff {V₁ V₂ : Type*} [normed_add_comm_group V₁] [normed_add_comm_group V₂] - {f : normed_add_group_hom V₁ V₂} : ∥f∥ = 0 ↔ f = 0 := + {f : normed_add_group_hom V₁ V₂} : ‖f‖ = 0 ↔ f = 0 := iff.intro (λ hn, ext (λ x, norm_le_zero_iff.1 - (calc _ ≤ ∥f∥ * ∥x∥ : le_op_norm _ _ + (calc _ ≤ ‖f‖ * ‖x‖ : le_op_norm _ _ ... = _ : by rw [hn, zero_mul]))) (λ hf, by rw [hf, op_norm_zero] ) @@ -313,19 +313,19 @@ def id : normed_add_group_hom V V := /-- The norm of the identity is at most `1`. It is in fact `1`, except when the norm of every element vanishes, where it is `0`. (Since we are working with seminorms this can happen even if the space is non-trivial.) It means that one can not do better than an inequality in general. -/ -lemma norm_id_le : ∥(id V : normed_add_group_hom V V)∥ ≤ 1 := +lemma norm_id_le : ‖(id V : normed_add_group_hom V V)‖ ≤ 1 := op_norm_le_bound _ zero_le_one (λx, by simp) /-- If there is an element with norm different from `0`, then the norm of the identity equals `1`. (Since we are working with seminorms supposing that the space is non-trivial is not enough.) -/ -lemma norm_id_of_nontrivial_seminorm (h : ∃ (x : V), ∥x∥ ≠ 0 ) : - ∥(id V)∥ = 1 := +lemma norm_id_of_nontrivial_seminorm (h : ∃ (x : V), ‖x‖ ≠ 0 ) : + ‖(id V)‖ = 1 := le_antisymm (norm_id_le V) $ let ⟨x, hx⟩ := h in have _ := (id V).ratio_le_op_norm x, by rwa [id_apply, div_self hx] at this /-- If a normed space is non-trivial, then the norm of the identity equals `1`. -/ -lemma norm_id {V : Type*} [normed_add_comm_group V] [nontrivial V] : ∥(id V)∥ = 1 := +lemma norm_id {V : Type*} [normed_add_comm_group V] [nontrivial V] : ‖(id V)‖ = 1 := begin refine norm_id_of_nontrivial_seminorm V _, obtain ⟨x, hx⟩ := exists_ne (0 : V), @@ -338,14 +338,14 @@ lemma coe_id : ((normed_add_group_hom.id V) : V → V) = (_root_.id : V → V) : /-- Opposite of a normed group hom. -/ instance : has_neg (normed_add_group_hom V₁ V₂) := -⟨λ f, (-f.to_add_monoid_hom).mk_normed_add_group_hom (∥f∥) (λ v, by simp [le_op_norm f v])⟩ +⟨λ f, (-f.to_add_monoid_hom).mk_normed_add_group_hom (‖f‖) (λ v, by simp [le_op_norm f v])⟩ -- see Note [addition on function coercions] @[simp] lemma coe_neg (f : normed_add_group_hom V₁ V₂) : ⇑(-f) = (-f : V₁ → V₂) := rfl @[simp] lemma neg_apply (f : normed_add_group_hom V₁ V₂) (v : V₁) : (-f : normed_add_group_hom V₁ V₂) v = - (f v) := rfl -lemma op_norm_neg (f : normed_add_group_hom V₁ V₂) : ∥-f∥ = ∥f∥ := +lemma op_norm_neg (f : normed_add_group_hom V₁ V₂) : ‖-f‖ = ‖f‖ := by simp only [norm_def, coe_neg, norm_neg, pi.neg_apply] /-! ### Subtraction of normed group homs -/ @@ -420,7 +420,7 @@ instance has_int_scalar : has_smul ℤ (normed_add_group_hom V₁ V₂) := { smul := λ z f, { to_fun := z • f, map_add' := (z • f.to_add_monoid_hom).map_add', - bound' := let ⟨b, hb⟩ := f.bound' in ⟨∥z∥ • b, λ v, begin + bound' := let ⟨b, hb⟩ := f.bound' in ⟨‖z‖ • b, λ v, begin rw [pi.smul_apply, smul_eq_mul, mul_assoc], exact (norm_zsmul_le _ _).trans (mul_le_mul_of_nonneg_left (hb _) $ norm_nonneg _), end⟩ } } @@ -488,22 +488,22 @@ function.injective.module _ coe_fn_add_hom coe_injective coe_smul @[simps] protected def comp (g : normed_add_group_hom V₂ V₃) (f : normed_add_group_hom V₁ V₂) : normed_add_group_hom V₁ V₃ := -(g.to_add_monoid_hom.comp f.to_add_monoid_hom).mk_normed_add_group_hom (∥g∥ * ∥f∥) $ λ v, calc -∥g (f v)∥ ≤ ∥g∥ * ∥f v∥ : le_op_norm _ _ -... ≤ ∥g∥ * (∥f∥ * ∥v∥) : mul_le_mul_of_nonneg_left (le_op_norm _ _) (op_norm_nonneg _) -... = ∥g∥ * ∥f∥ * ∥v∥ : by rw mul_assoc +(g.to_add_monoid_hom.comp f.to_add_monoid_hom).mk_normed_add_group_hom (‖g‖ * ‖f‖) $ λ v, calc +‖g (f v)‖ ≤ ‖g‖ * ‖f v‖ : le_op_norm _ _ +... ≤ ‖g‖ * (‖f‖ * ‖v‖) : mul_le_mul_of_nonneg_left (le_op_norm _ _) (op_norm_nonneg _) +... = ‖g‖ * ‖f‖ * ‖v‖ : by rw mul_assoc lemma norm_comp_le (g : normed_add_group_hom V₂ V₃) (f : normed_add_group_hom V₁ V₂) : - ∥g.comp f∥ ≤ ∥g∥ * ∥f∥ := + ‖g.comp f‖ ≤ ‖g‖ * ‖f‖ := mk_normed_add_group_hom_norm_le _ (mul_nonneg (op_norm_nonneg _) (op_norm_nonneg _)) _ -lemma norm_comp_le_of_le {g : normed_add_group_hom V₂ V₃} {C₁ C₂ : ℝ} (hg : ∥g∥ ≤ C₂) - (hf : ∥f∥ ≤ C₁) : - ∥g.comp f∥ ≤ C₂ * C₁ := +lemma norm_comp_le_of_le {g : normed_add_group_hom V₂ V₃} {C₁ C₂ : ℝ} (hg : ‖g‖ ≤ C₂) + (hf : ‖f‖ ≤ C₁) : + ‖g.comp f‖ ≤ C₂ * C₁ := le_trans (norm_comp_le g f) $ mul_le_mul hg hf (norm_nonneg _) (le_trans (norm_nonneg _) hg) lemma norm_comp_le_of_le' {g : normed_add_group_hom V₂ V₃} (C₁ C₂ C₃ : ℝ) (h : C₃ = C₂ * C₁) - (hg : ∥g∥ ≤ C₂) (hf : ∥f∥ ≤ C₁) : ∥g.comp f∥ ≤ C₃ := + (hg : ‖g‖ ≤ C₂) (hf : ‖f‖ ≤ C₁) : ‖g.comp f‖ ≤ C₃ := by { rw h, exact norm_comp_le_of_le hg hf } /-- Composition of normed groups hom as an additive group morphism. -/ @@ -543,7 +543,7 @@ variables {V W V₁ V₂ V₃ : Type*} [seminormed_add_comm_group V] [seminormed map_add' := λ v w, add_subgroup.coe_add _ _ _, bound' := ⟨1, λ v, by { rw [one_mul], refl }⟩ } -lemma norm_incl {V' : add_subgroup V} (x : V') : ∥incl _ x∥ = ∥x∥ := +lemma norm_incl {V' : add_subgroup V} (x : V') : ‖incl _ x‖ = ‖x‖ := rfl /-!### Kernel -/ @@ -611,13 +611,13 @@ end range variables {f : normed_add_group_hom V W} -/-- A `normed_add_group_hom` is *norm-nonincreasing* if `∥f v∥ ≤ ∥v∥` for all `v`. -/ +/-- A `normed_add_group_hom` is *norm-nonincreasing* if `‖f v‖ ≤ ‖v‖` for all `v`. -/ def norm_noninc (f : normed_add_group_hom V W) : Prop := -∀ v, ∥f v∥ ≤ ∥v∥ +∀ v, ‖f v‖ ≤ ‖v‖ namespace norm_noninc -lemma norm_noninc_iff_norm_le_one : f.norm_noninc ↔ ∥f∥ ≤ 1 := +lemma norm_noninc_iff_norm_le_one : f.norm_noninc ↔ ‖f‖ ≤ 1 := begin refine ⟨λ h, _, λ h, λ v, _⟩, { refine op_norm_le_bound _ (zero_le_one) (λ v, _), @@ -644,7 +644,7 @@ end norm_noninc section isometry lemma norm_eq_of_isometry {f : normed_add_group_hom V W} (hf : isometry f) (v : V) : - ∥f v∥ = ∥v∥ := + ‖f v‖ = ‖v‖ := (add_monoid_hom_class.isometry_iff_norm f).mp hf v lemma isometry_id : @isometry V V _ _ (id V) := @@ -740,16 +740,16 @@ lemma lift_norm_noninc (φ : normed_add_group_hom V₁ V) (h : f.comp φ = g.com (lift φ h).norm_noninc := hφ -/-- If `φ` satisfies `∥φ∥ ≤ C`, then the same is true for the lifted morphism. -/ +/-- If `φ` satisfies `‖φ‖ ≤ C`, then the same is true for the lifted morphism. -/ lemma norm_lift_le (φ : normed_add_group_hom V₁ V) (h : f.comp φ = g.comp φ) - (C : ℝ) (hφ : ∥φ∥ ≤ C) : ∥(lift φ h)∥ ≤ C := hφ + (C : ℝ) (hφ : ‖φ‖ ≤ C) : ‖(lift φ h)‖ ≤ C := hφ lemma map_norm_noninc (hf : ψ.comp f₁ = f₂.comp φ) (hg : ψ.comp g₁ = g₂.comp φ) (hφ : φ.norm_noninc) : (map φ ψ hf hg).norm_noninc := lift_norm_noninc _ _ $ hφ.comp ι_norm_noninc lemma norm_map_le (hf : ψ.comp f₁ = f₂.comp φ) (hg : ψ.comp g₁ = g₂.comp φ) - (C : ℝ) (hφ : ∥φ.comp (ι f₁ g₁)∥ ≤ C) : ∥map φ ψ hf hg∥ ≤ C := + (C : ℝ) (hφ : ‖φ.comp (ι f₁ g₁)‖ ≤ C) : ‖map φ ψ hf hg‖ ≤ C := norm_lift_le _ _ _ hφ end equalizer @@ -763,7 +763,7 @@ variables {G : Type*} [normed_add_comm_group G] [complete_space G] variables {H : Type*} [normed_add_comm_group H] /-- Given `f : normed_add_group_hom G H` for some complete `G` and a subgroup `K` of `H`, if every -element `x` of `K` has a preimage under `f` whose norm is at most `C*∥x∥` then the same holds for +element `x` of `K` has a preimage under `f` whose norm is at most `C*‖x‖` then the same holds for elements of the (topological) closure of `K` with constant `C+ε` instead of `C`, for any positive `ε`. -/ @@ -781,18 +781,18 @@ begin the series will be guaranteed by completeness of `G`. We first write `h` as the sum of a sequence `v` of elements of `K` which starts close to `h` and then quickly goes to zero. The sequence `b` below quantifies this. -/ - set b : ℕ → ℝ := λ i, (1/2)^i*(ε*∥h∥/2)/C, + set b : ℕ → ℝ := λ i, (1/2)^i*(ε*‖h‖/2)/C, have b_pos : ∀ i, 0 < b i, { intro i, field_simp [b, hC], exact div_pos (mul_pos hε (norm_pos_iff.mpr hyp_h)) (mul_pos (by norm_num : (0 : ℝ) < 2^i*2) hC) }, obtain ⟨v : ℕ → H, lim_v : tendsto (λ (n : ℕ), ∑ k in range (n + 1), v k) at_top (𝓝 h), - v_in : ∀ n, v n ∈ K, hv₀ : ∥v 0 - h∥ < b 0, hv : ∀ n > 0, ∥v n∥ < b n⟩ := + v_in : ∀ n, v n ∈ K, hv₀ : ‖v 0 - h‖ < b 0, hv : ∀ n > 0, ‖v n‖ < b n⟩ := controlled_sum_of_mem_closure h_in b_pos, /- The controlled surjectivity assumption on `f` allows to build preimages `u n` for all elements `v n` of the `v` sequence.-/ - have : ∀ n, ∃ m' : G, f m' = v n ∧ ∥m'∥ ≤ C * ∥v n∥ := λ (n : ℕ), hyp (v n) (v_in n), + have : ∀ n, ∃ m' : G, f m' = v n ∧ ‖m'‖ ≤ C * ‖v n‖ := λ (n : ℕ), hyp (v n) (v_in n), choose u hu hnorm_u using this, /- The desired series `s` is then obtained by summing `u`. We then check our choice of `b` ensures `s` is Cauchy. -/ @@ -800,10 +800,10 @@ begin have : cauchy_seq s, { apply normed_add_comm_group.cauchy_series_of_le_geometric'' (by norm_num) one_half_lt_one, rintro n (hn : n ≥ 1), - calc ∥u n∥ ≤ C*∥v n∥ : hnorm_u n + calc ‖u n‖ ≤ C*‖v n‖ : hnorm_u n ... ≤ C * b n : mul_le_mul_of_nonneg_left (hv _ $ nat.succ_le_iff.mp hn).le hC.le - ... = (1/2)^n * (ε * ∥h∥/2) : by simp [b, mul_div_cancel' _ hC.ne.symm] - ... = (ε * ∥h∥/2) * (1/2)^n : mul_comm _ _ }, + ... = (1/2)^n * (ε * ‖h‖/2) : by simp [b, mul_div_cancel' _ hC.ne.symm] + ... = (ε * ‖h‖/2) * (1/2)^n : mul_comm _ _ }, /- We now show that the limit `g` of `s` is the desired preimage. -/ obtain ⟨g : G, hg⟩ := cauchy_seq_tendsto_of_complete this, refine ⟨g, _, _⟩, @@ -817,44 +817,44 @@ begin rw ← this at lim_v, exact tendsto_nhds_unique ((f.continuous.tendsto g).comp hg) lim_v }, { /- Then we need to estimate the norm of `g`, using our careful choice of `b`. -/ - suffices : ∀ n, ∥s n∥ ≤ (C + ε) * ∥h∥, + suffices : ∀ n, ‖s n‖ ≤ (C + ε) * ‖h‖, from le_of_tendsto' (continuous_norm.continuous_at.tendsto.comp hg) this, intros n, - have hnorm₀ : ∥u 0∥ ≤ C*b 0 + C*∥h∥, + have hnorm₀ : ‖u 0‖ ≤ C*b 0 + C*‖h‖, { have := calc - ∥v 0∥ ≤ ∥h∥ + ∥v 0 - h∥ : norm_le_insert' _ _ - ... ≤ ∥h∥ + b 0 : by apply add_le_add_left hv₀.le, - calc ∥u 0∥ ≤ C*∥v 0∥ : hnorm_u 0 - ... ≤ C*(∥h∥ + b 0) : mul_le_mul_of_nonneg_left this hC.le - ... = C * b 0 + C * ∥h∥ : by rw [add_comm, mul_add] }, - have : ∑ k in range (n + 1), C * b k ≤ ε * ∥h∥ := calc - ∑ k in range (n + 1), C * b k = (∑ k in range (n + 1), (1 / 2) ^ k) * (ε * ∥h∥ / 2) : + ‖v 0‖ ≤ ‖h‖ + ‖v 0 - h‖ : norm_le_insert' _ _ + ... ≤ ‖h‖ + b 0 : by apply add_le_add_left hv₀.le, + calc ‖u 0‖ ≤ C*‖v 0‖ : hnorm_u 0 + ... ≤ C*(‖h‖ + b 0) : mul_le_mul_of_nonneg_left this hC.le + ... = C * b 0 + C * ‖h‖ : by rw [add_comm, mul_add] }, + have : ∑ k in range (n + 1), C * b k ≤ ε * ‖h‖ := calc + ∑ k in range (n + 1), C * b k = (∑ k in range (n + 1), (1 / 2) ^ k) * (ε * ‖h‖ / 2) : by simp only [b, mul_div_cancel' _ hC.ne.symm, ← sum_mul] - ... ≤ 2 * (ε * ∥h∥ / 2) : mul_le_mul_of_nonneg_right (sum_geometric_two_le _) + ... ≤ 2 * (ε * ‖h‖ / 2) : mul_le_mul_of_nonneg_right (sum_geometric_two_le _) (by nlinarith [hε, norm_nonneg h]) - ... = ε * ∥h∥ : mul_div_cancel' _ two_ne_zero, - calc ∥s n∥ ≤ ∑ k in range (n+1), ∥u k∥ : norm_sum_le _ _ - ... = ∑ k in range n, ∥u (k + 1)∥ + ∥u 0∥ : sum_range_succ' _ _ - ... ≤ ∑ k in range n, C*∥v (k + 1)∥ + ∥u 0∥ : add_le_add_right (sum_le_sum (λ _ _, hnorm_u _)) _ - ... ≤ ∑ k in range n, C*b (k+1) + (C*b 0 + C*∥h∥) : + ... = ε * ‖h‖ : mul_div_cancel' _ two_ne_zero, + calc ‖s n‖ ≤ ∑ k in range (n+1), ‖u k‖ : norm_sum_le _ _ + ... = ∑ k in range n, ‖u (k + 1)‖ + ‖u 0‖ : sum_range_succ' _ _ + ... ≤ ∑ k in range n, C*‖v (k + 1)‖ + ‖u 0‖ : add_le_add_right (sum_le_sum (λ _ _, hnorm_u _)) _ + ... ≤ ∑ k in range n, C*b (k+1) + (C*b 0 + C*‖h‖) : add_le_add (sum_le_sum (λ k _, mul_le_mul_of_nonneg_left (hv _ k.succ_pos).le hC.le)) hnorm₀ - ... = ∑ k in range (n+1), C*b k + C*∥h∥ : by rw [← add_assoc, sum_range_succ'] - ... ≤ (C+ε)*∥h∥ : by { rw [add_comm, add_mul], apply add_le_add_left this } } + ... = ∑ k in range (n+1), C*b k + C*‖h‖ : by rw [← add_assoc, sum_range_succ'] + ... ≤ (C+ε)*‖h‖ : by { rw [add_comm, add_mul], apply add_le_add_left this } } end /-- Given `f : normed_add_group_hom G H` for some complete `G`, if every element `x` of the image of an isometric immersion `j : normed_add_group_hom K H` has a preimage under `f` whose norm is at most -`C*∥x∥` then the same holds for elements of the (topological) closure of this image with constant +`C*‖x‖` then the same holds for elements of the (topological) closure of this image with constant `C+ε` instead of `C`, for any positive `ε`. This is useful in particular if `j` is the inclusion of a normed group into its completion (in this case the closure is the full target group). -/ lemma controlled_closure_range_of_complete {f : normed_add_group_hom G H} - {K : Type*} [seminormed_add_comm_group K] {j : normed_add_group_hom K H} (hj : ∀ x, ∥j x∥ = ∥x∥) - {C ε : ℝ} (hC : 0 < C) (hε : 0 < ε) (hyp : ∀ k, ∃ g, f g = j k ∧ ∥g∥ ≤ C*∥k∥) : + {K : Type*} [seminormed_add_comm_group K] {j : normed_add_group_hom K H} (hj : ∀ x, ‖j x‖ = ‖x‖) + {C ε : ℝ} (hC : 0 < C) (hε : 0 < ε) (hyp : ∀ k, ∃ g, f g = j k ∧ ‖g‖ ≤ C*‖k‖) : f.surjective_on_with j.range.topological_closure (C + ε) := begin - replace hyp : ∀ h ∈ j.range, ∃ g, f g = h ∧ ∥g∥ ≤ C*∥h∥, + replace hyp : ∀ h ∈ j.range, ∃ g, f g = h ∧ ‖g‖ ≤ C*‖h‖, { intros h h_in, rcases (j.mem_range _).mp h_in with ⟨k, rfl⟩, rw hj, diff --git a/src/analysis/normed/group/hom_completion.lean b/src/analysis/normed/group/hom_completion.lean index 2e3227117b81f..ab054fd2d501e 100644 --- a/src/analysis/normed/group/hom_completion.lean +++ b/src/analysis/normed/group/hom_completion.lean @@ -37,7 +37,7 @@ The vertical maps in the above diagrams are also normed group homs constructed i * `normed_add_comm_group.to_compl : normed_add_group_hom G (completion G)`: the canonical map from `G` to its completion, as a normed group hom * `normed_add_group_hom.completion_to_compl`: the above diagram indeed commutes. -* `normed_add_group_hom.norm_completion`: `∥f.completion∥ = ∥f∥` +* `normed_add_group_hom.norm_completion`: `‖f.completion‖ = ‖f‖` * `normed_add_group_hom.ker_le_ker_completion`: the kernel of `f.completion` contains the image of the kernel of `f`. * `normed_add_group_hom.ker_completion`: the kernel of `f.completion` is the closure of the image of @@ -60,14 +60,14 @@ variables {K : Type*} [seminormed_add_comm_group K] def normed_add_group_hom.completion (f : normed_add_group_hom G H) : normed_add_group_hom (completion G) (completion H) := { bound' := begin - use ∥f∥, + use ‖f‖, intro y, apply completion.induction_on y, { exact is_closed_le (continuous_norm.comp $ f.to_add_monoid_hom.continuous_completion f.continuous) (continuous_const.mul continuous_norm) }, { intro x, - change ∥f.to_add_monoid_hom.completion _ ↑x∥ ≤ ∥f∥ * ∥↑x∥, + change ‖f.to_add_monoid_hom.completion _ ↑x‖ ≤ ‖f‖ * ‖↑x‖, rw f.to_add_monoid_hom.completion_coe f.continuous, simp only [completion.norm_coe], exact f.le_op_norm x } @@ -145,7 +145,7 @@ def normed_add_comm_group.to_compl : normed_add_group_hom G (completion G) := open normed_add_comm_group -lemma normed_add_comm_group.norm_to_compl (x : G) : ∥to_compl x∥ = ∥x∥ := +lemma normed_add_comm_group.norm_to_compl (x : G) : ‖to_compl x‖ = ‖x‖ := completion.norm_coe x lemma normed_add_comm_group.dense_range_to_compl : dense_range (to_compl : G → completion G) := @@ -161,7 +161,7 @@ begin end @[simp] lemma normed_add_group_hom.norm_completion (f : normed_add_group_hom G H) : - ∥f.completion∥ = ∥f∥ := + ‖f.completion‖ = ‖f‖ := begin apply f.completion.op_norm_eq_of_bounds (norm_nonneg _), { intro x, @@ -196,40 +196,40 @@ begin { intros hatg hatg_in, rw seminormed_add_comm_group.mem_closure_iff, intros ε ε_pos, - have hCf : 0 ≤ C'*∥f∥ := (zero_le_mul_left C'_pos).mpr (norm_nonneg f), - have ineq : 0 < 1 + C'*∥f∥, by linarith, - set δ := ε/(1 + C'*∥f∥), + have hCf : 0 ≤ C'*‖f‖ := (zero_le_mul_left C'_pos).mpr (norm_nonneg f), + have ineq : 0 < 1 + C'*‖f‖, by linarith, + set δ := ε/(1 + C'*‖f‖), have δ_pos : δ > 0, from div_pos ε_pos ineq, - obtain ⟨_, ⟨g : G, rfl⟩, hg : ∥hatg - g∥ < δ⟩ := seminormed_add_comm_group.mem_closure_iff.mp + obtain ⟨_, ⟨g : G, rfl⟩, hg : ‖hatg - g‖ < δ⟩ := seminormed_add_comm_group.mem_closure_iff.mp (completion.dense_inducing_coe.dense hatg) δ δ_pos, - obtain ⟨g' : G, hgg' : f g' = f g, hfg : ∥g'∥ ≤ C' * ∥f g∥⟩ := + obtain ⟨g' : G, hgg' : f g' = f g, hfg : ‖g'‖ ≤ C' * ‖f g‖⟩ := hC' (f g) (mem_range_self g), have mem_ker : g - g' ∈ f.ker, by rw [f.mem_ker, map_sub, sub_eq_zero.mpr hgg'.symm], - have : ∥f g∥ ≤ ∥f∥*∥hatg - g∥, + have : ‖f g‖ ≤ ‖f‖*‖hatg - g‖, calc - ∥f g∥ = ∥f.completion g∥ : by rw [f.completion_coe, completion.norm_coe] - ... = ∥f.completion g - 0∥ : by rw [sub_zero _] - ... = ∥f.completion g - (f.completion hatg)∥ : by rw [(f.completion.mem_ker _).mp hatg_in] - ... = ∥f.completion (g - hatg)∥ : by rw [map_sub] - ... ≤ ∥f.completion∥ * ∥(g :completion G) - hatg∥ : f.completion.le_op_norm _ - ... = ∥f∥ * ∥hatg - g∥ : by rw [norm_sub_rev, f.norm_completion], - have : ∥(g' : completion G)∥ ≤ C'*∥f∥*∥hatg - g∥, + ‖f g‖ = ‖f.completion g‖ : by rw [f.completion_coe, completion.norm_coe] + ... = ‖f.completion g - 0‖ : by rw [sub_zero _] + ... = ‖f.completion g - (f.completion hatg)‖ : by rw [(f.completion.mem_ker _).mp hatg_in] + ... = ‖f.completion (g - hatg)‖ : by rw [map_sub] + ... ≤ ‖f.completion‖ * ‖(g :completion G) - hatg‖ : f.completion.le_op_norm _ + ... = ‖f‖ * ‖hatg - g‖ : by rw [norm_sub_rev, f.norm_completion], + have : ‖(g' : completion G)‖ ≤ C'*‖f‖*‖hatg - g‖, calc - ∥(g' : completion G)∥ = ∥g'∥ : completion.norm_coe _ - ... ≤ C' * ∥f g∥ : hfg - ... ≤ C' * ∥f∥ * ∥hatg - g∥ : by { rw mul_assoc, + ‖(g' : completion G)‖ = ‖g'‖ : completion.norm_coe _ + ... ≤ C' * ‖f g‖ : hfg + ... ≤ C' * ‖f‖ * ‖hatg - g‖ : by { rw mul_assoc, exact (mul_le_mul_left C'_pos).mpr this }, refine ⟨g - g', _, _⟩, { norm_cast, rw normed_add_group_hom.comp_range, apply add_subgroup.mem_map_of_mem, simp only [incl_range, mem_ker] }, - { calc ∥hatg - (g - g')∥ = ∥hatg - g + g'∥ : by abel - ... ≤ ∥hatg - g∥ + ∥(g' : completion G)∥ : norm_add_le _ _ - ... < δ + C'*∥f∥*∥hatg - g∥ : by linarith - ... ≤ δ + C'*∥f∥*δ : add_le_add_left (mul_le_mul_of_nonneg_left hg.le hCf) δ - ... = (1 + C'*∥f∥)*δ : by ring + { calc ‖hatg - (g - g')‖ = ‖hatg - g + g'‖ : by abel + ... ≤ ‖hatg - g‖ + ‖(g' : completion G)‖ : norm_add_le _ _ + ... < δ + C'*‖f‖*‖hatg - g‖ : by linarith + ... ≤ δ + C'*‖f‖*δ : add_le_add_left (mul_le_mul_of_nonneg_left hg.le hCf) δ + ... = (1 + C'*‖f‖)*δ : by ring ... = ε : mul_div_cancel' _ ineq.ne.symm } }, { rw ← f.completion.is_closed_ker.closure_eq, exact closure_mono f.ker_le_ker_completion } @@ -247,7 +247,7 @@ variables {H : Type*} [seminormed_add_comm_group H] [separated_space H] [complet def normed_add_group_hom.extension (f : normed_add_group_hom G H) : normed_add_group_hom (completion G) H := { bound' := begin - refine ⟨∥f∥, λ v, completion.induction_on v (is_closed_le _ _) (λ a, _)⟩, + refine ⟨‖f‖, λ v, completion.induction_on v (is_closed_le _ _) (λ a, _)⟩, { exact continuous.comp continuous_norm completion.continuous_extension }, { exact continuous.mul continuous_const continuous_norm }, { rw [completion.norm_coe, add_monoid_hom.to_fun_eq_coe, add_monoid_hom.extension_coe], diff --git a/src/analysis/normed/group/infinite_sum.lean b/src/analysis/normed/group/infinite_sum.lean index 79d5aa4ad15c5..29cb794586107 100644 --- a/src/analysis/normed/group/infinite_sum.lean +++ b/src/analysis/normed/group/infinite_sum.lean @@ -15,12 +15,12 @@ In a complete (semi)normed group, there exists a finite set `s` such that the sum `∑ i in t, f i` over any finite set `t` disjoint with `s` has norm less than `ε`; -- `summable_of_norm_bounded`, `summable_of_norm_bounded_eventually`: if `∥f i∥` is bounded above by +- `summable_of_norm_bounded`, `summable_of_norm_bounded_eventually`: if `‖f i‖` is bounded above by a summable series `∑' i, g i`, then `∑' i, f i` is summable as well; the same is true if the inequality hold only off some finite set. -- `tsum_of_norm_bounded`, `has_sum.norm_le_of_bounded`: if `∥f i∥ ≤ g i`, where `∑' i, g i` is a - summable series, then `∥∑' i, f i∥ ≤ ∑' i, g i`. +- `tsum_of_norm_bounded`, `has_sum.norm_le_of_bounded`: if `‖f i‖ ≤ g i`, where `∑' i, g i` is a + summable series, then `‖∑' i, f i‖ ≤ ∑' i, g i`. ## Tags @@ -34,7 +34,7 @@ variables {ι α E F : Type*} [seminormed_add_comm_group E] [seminormed_add_comm lemma cauchy_seq_finset_iff_vanishing_norm {f : ι → E} : cauchy_seq (λ s : finset ι, ∑ i in s, f i) ↔ - ∀ε > (0 : ℝ), ∃s:finset ι, ∀t, disjoint t s → ∥ ∑ i in t, f i ∥ < ε := + ∀ε > (0 : ℝ), ∃s:finset ι, ∀t, disjoint t s → ‖ ∑ i in t, f i ‖ < ε := begin rw [cauchy_seq_finset_iff_vanishing, nhds_basis_ball.forall_iff], { simp only [ball_zero_eq, set.mem_set_of_eq] }, @@ -43,60 +43,60 @@ begin end lemma summable_iff_vanishing_norm [complete_space E] {f : ι → E} : - summable f ↔ ∀ε > (0 : ℝ), ∃s:finset ι, ∀t, disjoint t s → ∥ ∑ i in t, f i ∥ < ε := + summable f ↔ ∀ε > (0 : ℝ), ∃s:finset ι, ∀t, disjoint t s → ‖ ∑ i in t, f i ‖ < ε := by rw [summable_iff_cauchy_seq_finset, cauchy_seq_finset_iff_vanishing_norm] lemma cauchy_seq_finset_of_norm_bounded_eventually {f : ι → E} {g : ι → ℝ} (hg : summable g) - (h : ∀ᶠ i in cofinite, ∥f i∥ ≤ g i) : cauchy_seq (λ s, ∑ i in s, f i) := + (h : ∀ᶠ i in cofinite, ‖f i‖ ≤ g i) : cauchy_seq (λ s, ∑ i in s, f i) := begin refine cauchy_seq_finset_iff_vanishing_norm.2 (λ ε hε, _), rcases summable_iff_vanishing_norm.1 hg ε hε with ⟨s, hs⟩, refine ⟨s ∪ h.to_finset, λ t ht, _⟩, - have : ∀ i ∈ t, ∥f i∥ ≤ g i, + have : ∀ i ∈ t, ‖f i‖ ≤ g i, { intros i hi, simp only [disjoint_left, mem_union, not_or_distrib, h.mem_to_finset, set.mem_compl_iff, not_not] at ht, exact (ht hi).2 }, - calc ∥∑ i in t, f i∥ ≤ ∑ i in t, g i : norm_sum_le_of_le _ this - ... ≤ ∥∑ i in t, g i∥ : le_abs_self _ + calc ‖∑ i in t, f i‖ ≤ ∑ i in t, g i : norm_sum_le_of_le _ this + ... ≤ ‖∑ i in t, g i‖ : le_abs_self _ ... < ε : hs _ (ht.mono_right le_sup_left), end lemma cauchy_seq_finset_of_norm_bounded {f : ι → E} (g : ι → ℝ) (hg : summable g) - (h : ∀i, ∥f i∥ ≤ g i) : cauchy_seq (λ s : finset ι, ∑ i in s, f i) := + (h : ∀i, ‖f i‖ ≤ g i) : cauchy_seq (λ s : finset ι, ∑ i in s, f i) := cauchy_seq_finset_of_norm_bounded_eventually hg $ eventually_of_forall h /-- A version of the **direct comparison test** for conditionally convergent series. See `cauchy_seq_finset_of_norm_bounded` for the same statement about absolutely convergent ones. -/ lemma cauchy_seq_range_of_norm_bounded {f : ℕ → E} (g : ℕ → ℝ) - (hg : cauchy_seq (λ n, ∑ i in range n, g i)) (hf : ∀ i, ∥f i∥ ≤ g i) : + (hg : cauchy_seq (λ n, ∑ i in range n, g i)) (hf : ∀ i, ‖f i‖ ≤ g i) : cauchy_seq (λ n, ∑ i in range n, f i) := begin refine metric.cauchy_seq_iff'.2 (λ ε hε, _), refine (metric.cauchy_seq_iff'.1 hg ε hε).imp (λ N hg n hn, _), specialize hg n hn, rw [dist_eq_norm, ←sum_Ico_eq_sub _ hn] at ⊢ hg, - calc ∥∑ k in Ico N n, f k∥ - ≤ ∑ k in _, ∥f k∥ : norm_sum_le _ _ + calc ‖∑ k in Ico N n, f k‖ + ≤ ∑ k in _, ‖f k‖ : norm_sum_le _ _ ... ≤ ∑ k in _, g k : sum_le_sum (λ x _, hf x) - ... ≤ ∥∑ k in _, g k∥ : le_abs_self _ + ... ≤ ‖∑ k in _, g k‖ : le_abs_self _ ... < ε : hg end -lemma cauchy_seq_finset_of_summable_norm {f : ι → E} (hf : summable (λa, ∥f a∥)) : +lemma cauchy_seq_finset_of_summable_norm {f : ι → E} (hf : summable (λa, ‖f a‖)) : cauchy_seq (λ s : finset ι, ∑ a in s, f a) := cauchy_seq_finset_of_norm_bounded _ hf (assume i, le_rfl) /-- If a function `f` is summable in norm, and along some sequence of finsets exhausting the space its sum is converging to a limit `a`, then this holds along all finsets, i.e., `f` is summable with sum `a`. -/ -lemma has_sum_of_subseq_of_summable {f : ι → E} (hf : summable (λa, ∥f a∥)) +lemma has_sum_of_subseq_of_summable {f : ι → E} (hf : summable (λa, ‖f a‖)) {s : α → finset ι} {p : filter α} [ne_bot p] (hs : tendsto s p at_top) {a : E} (ha : tendsto (λ b, ∑ i in s b, f i) p (𝓝 a)) : has_sum f a := tendsto_nhds_of_cauchy_seq_of_subseq (cauchy_seq_finset_of_summable_norm hf) hs ha -lemma has_sum_iff_tendsto_nat_of_summable_norm {f : ℕ → E} {a : E} (hf : summable (λi, ∥f i∥)) : +lemma has_sum_iff_tendsto_nat_of_summable_norm {f : ℕ → E} {a : E} (hf : summable (λi, ‖f i‖)) : has_sum f a ↔ tendsto (λn:ℕ, ∑ i in range n, f i) at_top (𝓝 a) := ⟨λ h, h.tendsto_sum_nat, λ h, has_sum_of_subseq_of_summable hf tendsto_finset_range h⟩ @@ -104,21 +104,21 @@ lemma has_sum_iff_tendsto_nat_of_summable_norm {f : ℕ → E} {a : E} (hf : sum /-- The direct comparison test for series: if the norm of `f` is bounded by a real function `g` which is summable, then `f` is summable. -/ lemma summable_of_norm_bounded - [complete_space E] {f : ι → E} (g : ι → ℝ) (hg : summable g) (h : ∀i, ∥f i∥ ≤ g i) : + [complete_space E] {f : ι → E} (g : ι → ℝ) (hg : summable g) (h : ∀i, ‖f i‖ ≤ g i) : summable f := by { rw summable_iff_cauchy_seq_finset, exact cauchy_seq_finset_of_norm_bounded g hg h } lemma has_sum.norm_le_of_bounded {f : ι → E} {g : ι → ℝ} {a : E} {b : ℝ} - (hf : has_sum f a) (hg : has_sum g b) (h : ∀ i, ∥f i∥ ≤ g i) : - ∥a∥ ≤ b := + (hf : has_sum f a) (hg : has_sum g b) (h : ∀ i, ‖f i‖ ≤ g i) : + ‖a‖ ≤ b := le_of_tendsto_of_tendsto' hf.norm hg $ λ s, norm_sum_le_of_le _ $ λ i hi, h i /-- Quantitative result associated to the direct comparison test for series: If `∑' i, g i` is -summable, and for all `i`, `∥f i∥ ≤ g i`, then `∥∑' i, f i∥ ≤ ∑' i, g i`. Note that we do not +summable, and for all `i`, `‖f i‖ ≤ g i`, then `‖∑' i, f i‖ ≤ ∑' i, g i`. Note that we do not assume that `∑' i, f i` is summable, and it might not be the case if `α` is not a complete space. -/ lemma tsum_of_norm_bounded {f : ι → E} {g : ι → ℝ} {a : ℝ} (hg : has_sum g a) - (h : ∀ i, ∥f i∥ ≤ g i) : - ∥∑' i : ι, f i∥ ≤ a := + (h : ∀ i, ‖f i‖ ≤ g i) : + ‖∑' i : ι, f i‖ ≤ a := begin by_cases hf : summable f, { exact hf.has_sum.norm_le_of_bounded hg h }, @@ -126,29 +126,29 @@ begin exact ge_of_tendsto' hg (λ s, sum_nonneg $ λ i hi, (norm_nonneg _).trans (h i)) } end -/-- If `∑' i, ∥f i∥` is summable, then `∥∑' i, f i∥ ≤ (∑' i, ∥f i∥)`. Note that we do not assume +/-- If `∑' i, ‖f i‖` is summable, then `‖∑' i, f i‖ ≤ (∑' i, ‖f i‖)`. Note that we do not assume that `∑' i, f i` is summable, and it might not be the case if `α` is not a complete space. -/ -lemma norm_tsum_le_tsum_norm {f : ι → E} (hf : summable (λi, ∥f i∥)) : - ∥∑' i, f i∥ ≤ ∑' i, ∥f i∥ := +lemma norm_tsum_le_tsum_norm {f : ι → E} (hf : summable (λi, ‖f i‖)) : + ‖∑' i, f i‖ ≤ ∑' i, ‖f i‖ := tsum_of_norm_bounded hf.has_sum $ λ i, le_rfl /-- Quantitative result associated to the direct comparison test for series: If `∑' i, g i` is -summable, and for all `i`, `∥f i∥₊ ≤ g i`, then `∥∑' i, f i∥₊ ≤ ∑' i, g i`. Note that we +summable, and for all `i`, `‖f i‖₊ ≤ g i`, then `‖∑' i, f i‖₊ ≤ ∑' i, g i`. Note that we do not assume that `∑' i, f i` is summable, and it might not be the case if `α` is not a complete space. -/ lemma tsum_of_nnnorm_bounded {f : ι → E} {g : ι → ℝ≥0} {a : ℝ≥0} (hg : has_sum g a) - (h : ∀ i, ∥f i∥₊ ≤ g i) : - ∥∑' i : ι, f i∥₊ ≤ a := + (h : ∀ i, ‖f i‖₊ ≤ g i) : + ‖∑' i : ι, f i‖₊ ≤ a := begin simp only [← nnreal.coe_le_coe, ← nnreal.has_sum_coe, coe_nnnorm] at *, exact tsum_of_norm_bounded hg h end -/-- If `∑' i, ∥f i∥₊` is summable, then `∥∑' i, f i∥₊ ≤ ∑' i, ∥f i∥₊`. Note that +/-- If `∑' i, ‖f i‖₊` is summable, then `‖∑' i, f i‖₊ ≤ ∑' i, ‖f i‖₊`. Note that we do not assume that `∑' i, f i` is summable, and it might not be the case if `α` is not a complete space. -/ -lemma nnnorm_tsum_le {f : ι → E} (hf : summable (λi, ∥f i∥₊)) : - ∥∑' i, f i∥₊ ≤ ∑' i, ∥f i∥₊ := +lemma nnnorm_tsum_le {f : ι → E} (hf : summable (λi, ‖f i‖₊)) : + ‖∑' i, f i‖₊ ≤ ∑' i, ‖f i‖₊ := tsum_of_nnnorm_bounded hf.has_sum (λ i, le_rfl) variable [complete_space E] @@ -156,15 +156,15 @@ variable [complete_space E] /-- Variant of the direct comparison test for series: if the norm of `f` is eventually bounded by a real function `g` which is summable, then `f` is summable. -/ lemma summable_of_norm_bounded_eventually {f : ι → E} (g : ι → ℝ) (hg : summable g) - (h : ∀ᶠ i in cofinite, ∥f i∥ ≤ g i) : summable f := + (h : ∀ᶠ i in cofinite, ‖f i‖ ≤ g i) : summable f := summable_iff_cauchy_seq_finset.2 $ cauchy_seq_finset_of_norm_bounded_eventually hg h lemma summable_of_nnnorm_bounded {f : ι → E} (g : ι → ℝ≥0) (hg : summable g) - (h : ∀i, ∥f i∥₊ ≤ g i) : summable f := + (h : ∀i, ‖f i‖₊ ≤ g i) : summable f := summable_of_norm_bounded (λ i, (g i : ℝ)) (nnreal.summable_coe.2 hg) (λ i, by exact_mod_cast h i) -lemma summable_of_summable_norm {f : ι → E} (hf : summable (λa, ∥f a∥)) : summable f := +lemma summable_of_summable_norm {f : ι → E} (hf : summable (λa, ‖f a‖)) : summable f := summable_of_norm_bounded _ hf (assume i, le_rfl) -lemma summable_of_summable_nnnorm {f : ι → E} (hf : summable (λ a, ∥f a∥₊)) : summable f := +lemma summable_of_summable_nnnorm {f : ι → E} (hf : summable (λ a, ‖f a‖₊)) : summable f := summable_of_nnnorm_bounded _ hf (assume i, le_rfl) diff --git a/src/analysis/normed/group/pointwise.lean b/src/analysis/normed/group/pointwise.lean index c8c7590a205e9..8cb7806fbb3da 100644 --- a/src/analysis/normed/group/pointwise.lean +++ b/src/analysis/normed/group/pointwise.lean @@ -23,8 +23,8 @@ variables [seminormed_group E] {ε δ : ℝ} {s t : set E} {x y : E} @[to_additive] lemma metric.bounded.mul (hs : bounded s) (ht : bounded t) : bounded (s * t) := begin - obtain ⟨Rs, hRs⟩ : ∃ R, ∀ x ∈ s, ∥x∥ ≤ R := hs.exists_norm_le', - obtain ⟨Rt, hRt⟩ : ∃ R, ∀ x ∈ t, ∥x∥ ≤ R := ht.exists_norm_le', + obtain ⟨Rs, hRs⟩ : ∃ R, ∀ x ∈ s, ‖x‖ ≤ R := hs.exists_norm_le', + obtain ⟨Rt, hRt⟩ : ∃ R, ∀ x ∈ t, ‖x‖ ≤ R := ht.exists_norm_le', refine bounded_iff_forall_norm_le'.2 ⟨Rs + Rt, _⟩, rintro z ⟨x, y, hx, hy, rfl⟩, exact norm_mul_le_of_le (hRs x hx) (hRt y hy), diff --git a/src/analysis/normed/group/quotient.lean b/src/analysis/normed/group/quotient.lean index 12e5cd58965d7..836809a7fec30 100644 --- a/src/analysis/normed/group/quotient.lean +++ b/src/analysis/normed/group/quotient.lean @@ -56,13 +56,13 @@ All the following definitions are in the `add_subgroup` namespace. Hence we can * `norm_normed_mk` : the operator norm of the projection is `1` if the subspace is not dense. * `is_quotient.norm_lift`: Provided `f : normed_hom M N` satisfies `is_quotient f`, for every - `n : N` and positive `ε`, there exists `m` such that `f m = n ∧ ∥m∥ < ∥n∥ + ε`. + `n : N` and positive `ε`, there exists `m` such that `f m = n ∧ ‖m‖ < ‖n‖ + ε`. ## Implementation details For any `seminormed_add_comm_group M` and any `S : add_subgroup M` we define a norm on `M ⧸ S` by -`∥x∥ = Inf (norm '' {m | mk' S m = x})`. This formula is really an implementation detail, it +`‖x‖ = Inf (norm '' {m | mk' S m = x})`. This formula is really an implementation detail, it shouldn't be needed outside of this file setting up the theory. Since `M ⧸ S` is automatically a topological space (as any quotient of a topological space), @@ -70,7 +70,7 @@ one needs to be careful while defining the `seminormed_add_comm_group` instance different topologies on this quotient. This is not purely a technological issue. Mathematically there is something to prove. The main point is proved in the auxiliary lemma `quotient_nhd_basis` that has no use beyond this verification and states that zero in the quotient -admits as basis of neighborhoods in the quotient topology the sets `{x | ∥x∥ < ε}` for positive `ε`. +admits as basis of neighborhoods in the quotient topology the sets `{x | ‖x‖ < ε}` for positive `ε`. Once this mathematical point it settled, we have two topologies that are propositionaly equal. This is not good enough for the type class system. As usual we ensure *definitional* equality @@ -101,7 +101,7 @@ instance norm_on_quotient (S : add_subgroup M) : has_norm (M ⧸ S) := { norm := λ x, Inf (norm '' {m | mk' S m = x}) } lemma add_subgroup.quotient_norm_eq {S : add_subgroup M} (x : M ⧸ S) : - ∥x∥ = Inf (norm '' {m : M | (m : M ⧸ S) = x}) := + ‖x‖ = Inf (norm '' {m : M | (m : M ⧸ S) = x}) := rfl lemma image_norm_nonempty {S : add_subgroup M} : @@ -121,8 +121,8 @@ begin apply norm_nonneg end -/-- The norm on the quotient satisfies `∥-x∥ = ∥x∥`. -/ -lemma quotient_norm_neg {S : add_subgroup M} (x : M ⧸ S) : ∥-x∥ = ∥x∥ := +/-- The norm on the quotient satisfies `‖-x‖ = ‖x‖`. -/ +lemma quotient_norm_neg {S : add_subgroup M} (x : M ⧸ S) : ‖-x‖ = ‖x‖ := begin suffices : norm '' {m | mk' S m = x} = norm '' {m | mk' S m = -x}, by simp only [this, norm], @@ -135,12 +135,12 @@ begin exact ⟨-m, by simpa [eq_comm] using eq_neg_iff_eq_neg.mp ((mk'_apply _ _).symm.trans hm)⟩ } end -lemma quotient_norm_sub_rev {S : add_subgroup M} (x y : M ⧸ S) : ∥x - y∥ = ∥y - x∥ := +lemma quotient_norm_sub_rev {S : add_subgroup M} (x y : M ⧸ S) : ‖x - y‖ = ‖y - x‖ := by rw [show x - y = -(y - x), by abel, quotient_norm_neg] /-- The norm of the projection is smaller or equal to the norm of the original element. -/ lemma quotient_norm_mk_le (S : add_subgroup M) (m : M) : - ∥mk' S m∥ ≤ ∥m∥ := + ‖mk' S m‖ ≤ ‖m‖ := begin apply cInf_le, use 0, @@ -152,11 +152,11 @@ end /-- The norm of the projection is smaller or equal to the norm of the original element. -/ lemma quotient_norm_mk_le' (S : add_subgroup M) (m : M) : - ∥(m : M ⧸ S)∥ ≤ ∥m∥ := quotient_norm_mk_le S m + ‖(m : M ⧸ S)‖ ≤ ‖m‖ := quotient_norm_mk_le S m /-- The norm of the image under the natural morphism to the quotient. -/ lemma quotient_norm_mk_eq (S : add_subgroup M) (m : M) : - ∥mk' S m∥ = Inf ((λ x, ∥m + x∥) '' S) := + ‖mk' S m‖ = Inf ((λ x, ‖m + x‖) '' S) := begin change Inf _ = _, congr' 1, @@ -172,30 +172,30 @@ begin end /-- The quotient norm is nonnegative. -/ -lemma quotient_norm_nonneg (S : add_subgroup M) : ∀ x : M ⧸ S, 0 ≤ ∥x∥ := +lemma quotient_norm_nonneg (S : add_subgroup M) : ∀ x : M ⧸ S, 0 ≤ ‖x‖ := begin rintros ⟨m⟩, - change 0 ≤ ∥mk' S m∥, + change 0 ≤ ‖mk' S m‖, apply le_cInf (image_norm_nonempty _), rintros _ ⟨n, h, rfl⟩, apply norm_nonneg end /-- The quotient norm is nonnegative. -/ -lemma norm_mk_nonneg (S : add_subgroup M) (m : M) : 0 ≤ ∥mk' S m∥ := +lemma norm_mk_nonneg (S : add_subgroup M) (m : M) : 0 ≤ ‖mk' S m‖ := quotient_norm_nonneg S _ /-- The norm of the image of `m : M` in the quotient by `S` is zero if and only if `m` belongs to the closure of `S`. -/ lemma quotient_norm_eq_zero_iff (S : add_subgroup M) (m : M) : - ∥mk' S m∥ = 0 ↔ m ∈ closure (S : set M) := + ‖mk' S m‖ = 0 ↔ m ∈ closure (S : set M) := begin - have : 0 ≤ ∥mk' S m∥ := norm_mk_nonneg S m, + have : 0 ≤ ‖mk' S m‖ := norm_mk_nonneg S m, rw [← this.le_iff_eq, quotient_norm_mk_eq, real.Inf_le_iff], simp_rw [zero_add], - { calc (∀ ε > (0 : ℝ), ∃ r ∈ (λ x, ∥m + x∥) '' (S : set M), r < ε) ↔ - (∀ ε > 0, (∃ x ∈ S, ∥m + x∥ < ε)) : by simp [set.bex_image_iff] - ... ↔ ∀ ε > 0, (∃ x ∈ S, ∥m + -x∥ < ε) : _ + { calc (∀ ε > (0 : ℝ), ∃ r ∈ (λ x, ‖m + x‖) '' (S : set M), r < ε) ↔ + (∀ ε > 0, (∃ x ∈ S, ‖m + x‖ < ε)) : by simp [set.bex_image_iff] + ... ↔ ∀ ε > 0, (∃ x ∈ S, ‖m + -x‖ < ε) : _ ... ↔ ∀ ε > 0, (∃ x ∈ S, x ∈ metric.ball m ε) : by simp [dist_eq_norm, ← sub_eq_add_neg, norm_sub_rev] ... ↔ m ∈ closure ↑S : by simp [metric.mem_closure_iff, dist_comm], @@ -210,21 +210,21 @@ begin end /-- For any `x : M ⧸ S` and any `0 < ε`, there is `m : M` such that `mk' S m = x` -and `∥m∥ < ∥x∥ + ε`. -/ +and `‖m‖ < ‖x‖ + ε`. -/ lemma norm_mk_lt {S : add_subgroup M} (x : M ⧸ S) {ε : ℝ} (hε : 0 < ε) : - ∃ (m : M), mk' S m = x ∧ ∥m∥ < ∥x∥ + ε := + ∃ (m : M), mk' S m = x ∧ ‖m‖ < ‖x‖ + ε := begin - obtain ⟨_, ⟨m : M, H : mk' S m = x, rfl⟩, hnorm : ∥m∥ < ∥x∥ + ε⟩ := + obtain ⟨_, ⟨m : M, H : mk' S m = x, rfl⟩, hnorm : ‖m‖ < ‖x‖ + ε⟩ := real.lt_Inf_add_pos (image_norm_nonempty x) hε, subst H, exact ⟨m, rfl, hnorm⟩, end -/-- For any `m : M` and any `0 < ε`, there is `s ∈ S` such that `∥m + s∥ < ∥mk' S m∥ + ε`. -/ +/-- For any `m : M` and any `0 < ε`, there is `s ∈ S` such that `‖m + s‖ < ‖mk' S m‖ + ε`. -/ lemma norm_mk_lt' (S : add_subgroup M) (m : M) {ε : ℝ} (hε : 0 < ε) : - ∃ s ∈ S, ∥m + s∥ < ∥mk' S m∥ + ε := + ∃ s ∈ S, ‖m + s‖ < ‖mk' S m‖ + ε := begin - obtain ⟨n : M, hn : mk' S n = mk' S m, hn' : ∥n∥ < ∥mk' S m∥ + ε⟩ := + obtain ⟨n : M, hn : mk' S n = mk' S m, hn' : ‖n‖ < ‖mk' S m‖ + ε⟩ := norm_mk_lt (quotient_add_group.mk' S m) hε, erw [eq_comm, quotient_add_group.eq] at hn, use [- m + n, hn], @@ -232,20 +232,20 @@ begin end /-- The quotient norm satisfies the triangle inequality. -/ -lemma quotient_norm_add_le (S : add_subgroup M) (x y : M ⧸ S) : ∥x + y∥ ≤ ∥x∥ + ∥y∥ := +lemma quotient_norm_add_le (S : add_subgroup M) (x y : M ⧸ S) : ‖x + y‖ ≤ ‖x‖ + ‖y‖ := begin refine le_of_forall_pos_le_add (λ ε hε, _), replace hε := half_pos hε, - obtain ⟨m, rfl, hm : ∥m∥ < ∥mk' S m∥ + ε / 2⟩ := norm_mk_lt x hε, - obtain ⟨n, rfl, hn : ∥n∥ < ∥mk' S n∥ + ε / 2⟩ := norm_mk_lt y hε, - calc ∥mk' S m + mk' S n∥ = ∥mk' S (m + n)∥ : by rw (mk' S).map_add - ... ≤ ∥m + n∥ : quotient_norm_mk_le S (m + n) - ... ≤ ∥m∥ + ∥n∥ : norm_add_le _ _ - ... ≤ ∥mk' S m∥ + ∥mk' S n∥ + ε : by linarith + obtain ⟨m, rfl, hm : ‖m‖ < ‖mk' S m‖ + ε / 2⟩ := norm_mk_lt x hε, + obtain ⟨n, rfl, hn : ‖n‖ < ‖mk' S n‖ + ε / 2⟩ := norm_mk_lt y hε, + calc ‖mk' S m + mk' S n‖ = ‖mk' S (m + n)‖ : by rw (mk' S).map_add + ... ≤ ‖m + n‖ : quotient_norm_mk_le S (m + n) + ... ≤ ‖m‖ + ‖n‖ : norm_add_le _ _ + ... ≤ ‖mk' S m‖ + ‖mk' S n‖ + ε : by linarith end /-- The quotient norm of `0` is `0`. -/ -lemma norm_mk_zero (S : add_subgroup M) : ∥(0 : M ⧸ S)∥ = 0 := +lemma norm_mk_zero (S : add_subgroup M) : ‖(0 : M ⧸ S)‖ = 0 := begin erw quotient_norm_eq_zero_iff, exact subset_closure S.zero_mem @@ -254,11 +254,11 @@ end /-- If `(m : M)` has norm equal to `0` in `M ⧸ S` for a closed subgroup `S` of `M`, then `m ∈ S`. -/ lemma norm_zero_eq_zero (S : add_subgroup M) (hS : is_closed (S : set M)) (m : M) - (h : ∥mk' S m∥ = 0) : m ∈ S := + (h : ‖mk' S m‖ = 0) : m ∈ S := by rwa [quotient_norm_eq_zero_iff, hS.closure_eq] at h lemma quotient_nhd_basis (S : add_subgroup M) : - (𝓝 (0 : M ⧸ S)).has_basis (λ ε : ℝ, 0 < ε) (λ ε, {x | ∥x∥ < ε}) := + (𝓝 (0 : M ⧸ S)).has_basis (λ ε : ℝ, 0 < ε) (λ ε, {x | ‖x‖ < ε}) := ⟨begin intros U, split, @@ -275,7 +275,7 @@ lemma quotient_nhd_basis (S : add_subgroup M) : dsimp, linarith }, { rintros ⟨ε, ε_pos, h⟩, - have : (mk' S) '' (ball (0 : M) ε) ⊆ {x | ∥x∥ < ε}, + have : (mk' S) '' (ball (0 : M) ε) ⊆ {x | ‖x‖ < ε}, { rintros _ ⟨x, x_in, rfl⟩, rw mem_ball_zero_iff at x_in, exact lt_of_le_of_lt (quotient_norm_mk_le S x) x_in }, @@ -294,7 +294,7 @@ end⟩ noncomputable instance add_subgroup.seminormed_add_comm_group_quotient (S : add_subgroup M) : seminormed_add_comm_group (M ⧸ S) := -{ dist := λ x y, ∥x - y∥, +{ dist := λ x y, ‖x - y‖, dist_self := λ x, by simp only [norm_mk_zero, sub_self], dist_comm := quotient_norm_sub_rev, dist_triangle := λ x y z, @@ -311,8 +311,8 @@ instance add_subgroup.seminormed_add_comm_group_quotient (S : add_subgroup M) : rw uniformity_eq_comap_nhds_zero', have := (quotient_nhd_basis S).comap (λ (p : (M ⧸ S) × M ⧸ S), p.2 - p.1), apply this.eq_of_same_basis, - have : ∀ ε : ℝ, (λ (p : (M ⧸ S) × M ⧸ S), p.snd - p.fst) ⁻¹' {x | ∥x∥ < ε} = - {p : (M ⧸ S) × M ⧸ S | ∥p.fst - p.snd∥ < ε}, + have : ∀ ε : ℝ, (λ (p : (M ⧸ S) × M ⧸ S), p.snd - p.fst) ⁻¹' {x | ‖x‖ < ε} = + {p : (M ⧸ S) × M ⧸ S | ‖p.fst - p.snd‖ < ε}, { intro ε, ext x, dsimp, @@ -321,7 +321,7 @@ instance add_subgroup.seminormed_add_comm_group_quotient (S : add_subgroup M) : refine filter.has_basis_binfi_principal _ set.nonempty_Ioi, rintros ε (ε_pos : 0 < ε) η (η_pos : 0 < η), refine ⟨min ε η, lt_min ε_pos η_pos, _, _⟩, - { suffices : ∀ (a b : M ⧸ S), ∥a - b∥ < ε → ∥a - b∥ < η → ∥a - b∥ < ε, by simpa, + { suffices : ∀ (a b : M ⧸ S), ‖a - b‖ < ε → ‖a - b‖ < η → ‖a - b‖ < ε, by simpa, exact λ a b h h', h }, { simp } end } @@ -338,7 +338,7 @@ instance add_subgroup.normed_add_comm_group_quotient (S : add_subgroup M) [is_cl normed_add_comm_group (M ⧸ S) := { eq_of_dist_eq_zero := begin - rintros ⟨m⟩ ⟨m'⟩ (h : ∥mk' S m - mk' S m'∥ = 0), + rintros ⟨m⟩ ⟨m'⟩ (h : ‖mk' S m - mk' S m'‖ = 0), erw [← (mk' S).map_sub, quotient_norm_eq_zero_iff, ‹is_closed _›.closure_eq, ← quotient_add_group.eq_iff_sub_mem] at h, exact h @@ -375,29 +375,29 @@ lemma ker_normed_mk (S : add_subgroup M) : S.normed_mk.ker = S := quotient_add_group.ker_mk _ /-- The operator norm of the projection is at most `1`. -/ -lemma norm_normed_mk_le (S : add_subgroup M) : ∥S.normed_mk∥ ≤ 1 := +lemma norm_normed_mk_le (S : add_subgroup M) : ‖S.normed_mk‖ ≤ 1 := normed_add_group_hom.op_norm_le_bound _ zero_le_one (λ m, by simp [quotient_norm_mk_le']) /-- The operator norm of the projection is `1` if the subspace is not dense. -/ lemma norm_normed_mk (S : add_subgroup M) (h : (S.topological_closure : set M) ≠ univ) : - ∥S.normed_mk∥ = 1 := + ‖S.normed_mk‖ = 1 := begin obtain ⟨x, hx⟩ := set.nonempty_compl.2 h, let y := S.normed_mk x, - have hy : ∥y∥ ≠ 0, + have hy : ‖y‖ ≠ 0, { intro h0, exact set.not_mem_of_mem_compl hx ((quotient_norm_eq_zero_iff S x).1 h0) }, refine le_antisymm (norm_normed_mk_le S) (le_of_forall_pos_le_add (λ ε hε, _)), - suffices : 1 ≤ ∥S.normed_mk∥ + min ε ((1 : ℝ)/2), + suffices : 1 ≤ ‖S.normed_mk‖ + min ε ((1 : ℝ)/2), { exact le_add_of_le_add_left this (min_le_left ε ((1 : ℝ)/2)) }, have hδ := sub_pos.mpr (lt_of_le_of_lt (min_le_right ε ((1 : ℝ)/2)) one_half_lt_one), have hδpos : 0 < min ε ((1 : ℝ)/2) := lt_min hε one_half_pos, have hδnorm := mul_pos (div_pos hδpos hδ) (lt_of_le_of_ne (norm_nonneg y) hy.symm), obtain ⟨m, hm, hlt⟩ := norm_mk_lt y hδnorm, - have hrw : ∥y∥ + min ε (1 / 2) / (1 - min ε (1 / 2)) * ∥y∥ = - ∥y∥ * (1 + min ε (1 / 2) / (1 - min ε (1 / 2))) := by ring, + have hrw : ‖y‖ + min ε (1 / 2) / (1 - min ε (1 / 2)) * ‖y‖ = + ‖y‖ * (1 + min ε (1 / 2) / (1 - min ε (1 / 2))) := by ring, rw [hrw] at hlt, - have hm0 : ∥m∥ ≠ 0, + have hm0 : ‖m‖ ≠ 0, { intro h0, have hnorm := quotient_norm_mk_le S m, rw [h0, hm] at hnorm, @@ -405,21 +405,21 @@ begin simpa [hnorm] using hy }, replace hlt := (div_lt_div_right (lt_of_le_of_ne (norm_nonneg m) hm0.symm)).2 hlt, simp only [hm0, div_self, ne.def, not_false_iff] at hlt, - have hrw₁ : ∥y∥ * (1 + min ε (1 / 2) / (1 - min ε (1 / 2))) / ∥m∥ = - (∥y∥ / ∥m∥) * (1 + min ε (1 / 2) / (1 - min ε (1 / 2))) := by ring, + have hrw₁ : ‖y‖ * (1 + min ε (1 / 2) / (1 - min ε (1 / 2))) / ‖m‖ = + (‖y‖ / ‖m‖) * (1 + min ε (1 / 2) / (1 - min ε (1 / 2))) := by ring, rw [hrw₁] at hlt, replace hlt := (inv_pos_lt_iff_one_lt_mul (lt_trans (div_pos hδpos hδ) (lt_one_add _))).2 hlt, - suffices : ∥S.normed_mk∥ ≥ 1 - min ε (1 / 2), + suffices : ‖S.normed_mk‖ ≥ 1 - min ε (1 / 2), { exact sub_le_iff_le_add.mp this }, - calc ∥S.normed_mk∥ ≥ ∥(S.normed_mk) m∥ / ∥m∥ : ratio_le_op_norm S.normed_mk m - ... = ∥y∥ / ∥m∥ : by rw [normed_mk.apply, hm] + calc ‖S.normed_mk‖ ≥ ‖(S.normed_mk) m‖ / ‖m‖ : ratio_le_op_norm S.normed_mk m + ... = ‖y‖ / ‖m‖ : by rw [normed_mk.apply, hm] ... ≥ (1 + min ε (1 / 2) / (1 - min ε (1 / 2)))⁻¹ : le_of_lt hlt ... = 1 - min ε (1 / 2) : by field_simp [(ne_of_lt hδ).symm] end /-- The operator norm of the projection is `0` if the subspace is dense. -/ lemma norm_trivial_quotient_mk (S : add_subgroup M) - (h : (S.topological_closure : set M) = set.univ) : ∥S.normed_mk∥ = 0 := + (h : (S.topological_closure : set M) = set.univ) : ‖S.normed_mk‖ = 0 := begin refine le_antisymm (op_norm_le_bound _ le_rfl (λ x, _)) (norm_nonneg _), have hker : x ∈ (S.normed_mk).ker.topological_closure, @@ -437,7 +437,7 @@ namespace normed_add_group_hom by the kernel of `f`. -/ structure is_quotient (f : normed_add_group_hom M N) : Prop := (surjective : function.surjective f) -(norm : ∀ x, ∥f x∥ = Inf ((λ m, ∥x + m∥) '' f.ker)) +(norm : ∀ x, ‖f x‖ = Inf ((λ m, ‖x + m‖) '' f.ker)) /-- Given `f : normed_add_group_hom M N` such that `f s = 0` for all `s ∈ S`, where, `S : add_subgroup M` is closed, the induced morphism `normed_add_group_hom (M ⧸ S) N`. -/ @@ -447,13 +447,13 @@ def lift {N : Type*} [seminormed_add_comm_group N] (S : add_subgroup M) normed_add_group_hom (M ⧸ S) N := { bound' := begin - obtain ⟨c : ℝ, hcpos : (0 : ℝ) < c, hc : ∀ x, ∥f x∥ ≤ c * ∥x∥⟩ := f.bound, + obtain ⟨c : ℝ, hcpos : (0 : ℝ) < c, hc : ∀ x, ‖f x‖ ≤ c * ‖x‖⟩ := f.bound, refine ⟨c, λ mbar, le_of_forall_pos_le_add (λ ε hε, _)⟩, - obtain ⟨m : M, rfl : mk' S m = mbar, hmnorm : ∥m∥ < ∥mk' S m∥ + ε/c⟩ := + obtain ⟨m : M, rfl : mk' S m = mbar, hmnorm : ‖m‖ < ‖mk' S m‖ + ε/c⟩ := norm_mk_lt mbar (div_pos hε hcpos), - calc ∥f m∥ ≤ c * ∥m∥ : hc m - ... ≤ c*(∥mk' S m∥ + ε/c) : ((mul_lt_mul_left hcpos).mpr hmnorm).le - ... = c * ∥mk' S m∥ + ε : by rw [mul_add, mul_div_cancel' _ hcpos.ne.symm] + calc ‖f m‖ ≤ c * ‖m‖ : hc m + ... ≤ c*(‖mk' S m‖ + ε/c) : ((mul_lt_mul_left hcpos).mpr hmnorm).le + ... = c * ‖mk' S m‖ + ε : by rw [mul_add, mul_div_cancel' _ hcpos.ne.symm] end, .. quotient_add_group.lift S f.to_add_monoid_hom hf } @@ -478,20 +478,20 @@ lemma is_quotient_quotient (S : add_subgroup M) : is_quotient (S.normed_mk) := ⟨S.surjective_normed_mk, λ m, by simpa [S.ker_normed_mk] using quotient_norm_mk_eq _ m⟩ lemma is_quotient.norm_lift {f : normed_add_group_hom M N} (hquot : is_quotient f) {ε : ℝ} - (hε : 0 < ε) (n : N) : ∃ (m : M), f m = n ∧ ∥m∥ < ∥n∥ + ε := + (hε : 0 < ε) (n : N) : ∃ (m : M), f m = n ∧ ‖m‖ < ‖n‖ + ε := begin obtain ⟨m, rfl⟩ := hquot.surjective n, - have nonemp : ((λ m', ∥m + m'∥) '' f.ker).nonempty, + have nonemp : ((λ m', ‖m + m'‖) '' f.ker).nonempty, { rw set.nonempty_image_iff, exact ⟨0, f.ker.zero_mem⟩ }, rcases real.lt_Inf_add_pos nonemp hε with - ⟨_, ⟨⟨x, hx, rfl⟩, H : ∥m + x∥ < Inf ((λ (m' : M), ∥m + m'∥) '' f.ker) + ε⟩⟩, + ⟨_, ⟨⟨x, hx, rfl⟩, H : ‖m + x‖ < Inf ((λ (m' : M), ‖m + m'‖) '' f.ker) + ε⟩⟩, exact ⟨m+x, by rw [map_add,(normed_add_group_hom.mem_ker f x).mp hx, add_zero], by rwa hquot.norm⟩, end lemma is_quotient.norm_le {f : normed_add_group_hom M N} (hquot : is_quotient f) (m : M) : - ∥f m∥ ≤ ∥m∥ := + ‖f m‖ ≤ ‖m‖ := begin rw hquot.norm, apply cInf_le, @@ -503,26 +503,26 @@ end lemma lift_norm_le {N : Type*} [seminormed_add_comm_group N] (S : add_subgroup M) (f : normed_add_group_hom M N) (hf : ∀ s ∈ S, f s = 0) - {c : ℝ≥0} (fb : ∥f∥ ≤ c) : - ∥lift S f hf∥ ≤ c := + {c : ℝ≥0} (fb : ‖f‖ ≤ c) : + ‖lift S f hf‖ ≤ c := begin apply op_norm_le_bound _ c.coe_nonneg, intros x, by_cases hc : c = 0, { simp only [hc, nnreal.coe_zero, zero_mul] at fb ⊢, obtain ⟨x, rfl⟩ := surjective_quot_mk _ x, - show ∥f x∥ ≤ 0, - calc ∥f x∥ ≤ 0 * ∥x∥ : f.le_of_op_norm_le fb x + show ‖f x‖ ≤ 0, + calc ‖f x‖ ≤ 0 * ‖x‖ : f.le_of_op_norm_le fb x ... = 0 : zero_mul _ }, { replace hc : 0 < c := pos_iff_ne_zero.mpr hc, apply le_of_forall_pos_le_add, intros ε hε, have aux : 0 < (ε / c) := div_pos hε hc, - obtain ⟨x, rfl, Hx⟩ : ∃ x', S.normed_mk x' = x ∧ ∥x'∥ < ∥x∥ + (ε / c) := + obtain ⟨x, rfl, Hx⟩ : ∃ x', S.normed_mk x' = x ∧ ‖x'‖ < ‖x‖ + (ε / c) := (is_quotient_quotient _).norm_lift aux _, rw lift_mk, - calc ∥f x∥ ≤ c * ∥x∥ : f.le_of_op_norm_le fb x - ... ≤ c * (∥S.normed_mk x∥ + ε / c) : (mul_le_mul_left _).mpr Hx.le + calc ‖f x‖ ≤ c * ‖x‖ : f.le_of_op_norm_le fb x + ... ≤ c * (‖S.normed_mk x‖ + ε / c) : (mul_le_mul_left _).mpr Hx.le ... = c * _ + ε : _, { exact_mod_cast hc }, { rw [mul_add, mul_div_cancel'], exact_mod_cast hc.ne' } }, @@ -534,7 +534,7 @@ lemma lift_norm_noninc {N : Type*} [seminormed_add_comm_group N] (S : add_subgro (lift S f hf).norm_noninc := λ x, begin - have fb' : ∥f∥ ≤ (1 : ℝ≥0) := norm_noninc.norm_noninc_iff_norm_le_one.mp fb, + have fb' : ‖f‖ ≤ (1 : ℝ≥0) := norm_noninc.norm_noninc_iff_norm_le_one.mp fb, simpa using le_of_op_norm_le _ (f.lift_norm_le _ _ fb') _, end @@ -570,13 +570,13 @@ instance submodule.quotient.complete_space [complete_space M] : complete_space ( quotient_add_group.complete_space M S.to_add_subgroup /-- For any `x : M ⧸ S` and any `0 < ε`, there is `m : M` such that `submodule.quotient.mk m = x` -and `∥m∥ < ∥x∥ + ε`. -/ +and `‖m‖ < ‖x‖ + ε`. -/ lemma submodule.quotient.norm_mk_lt {S : submodule R M} (x : M ⧸ S) {ε : ℝ} (hε : 0 < ε) : - ∃ m : M, submodule.quotient.mk m = x ∧ ∥m∥ < ∥x∥ + ε := + ∃ m : M, submodule.quotient.mk m = x ∧ ‖m‖ < ‖x‖ + ε := norm_mk_lt x hε lemma submodule.quotient.norm_mk_le (m : M) : - ∥(submodule.quotient.mk m : M ⧸ S)∥ ≤ ∥m∥ := + ‖(submodule.quotient.mk m : M ⧸ S)‖ ≤ ‖m‖ := quotient_norm_mk_le S.to_add_subgroup m instance submodule.quotient.normed_space (𝕜 : Type*) [normed_field 𝕜] [normed_space 𝕜 M] @@ -584,12 +584,12 @@ instance submodule.quotient.normed_space (𝕜 : Type*) [normed_field 𝕜] [nor { norm_smul_le := λ k x, le_of_forall_pos_le_add $ λ ε hε, begin have := (nhds_basis_ball.tendsto_iff nhds_basis_ball).mp - ((@real.uniform_continuous_const_mul (∥k∥)).continuous.tendsto (∥x∥)) ε hε, + ((@real.uniform_continuous_const_mul (‖k‖)).continuous.tendsto (‖x‖)) ε hε, simp only [mem_ball, exists_prop, dist, abs_sub_lt_iff] at this, rcases this with ⟨δ, hδ, h⟩, obtain ⟨a, rfl, ha⟩ := submodule.quotient.norm_mk_lt x hδ, - specialize h (∥a∥) (⟨by linarith, by linarith [submodule.quotient.norm_mk_le S a]⟩), - calc _ ≤ ∥k∥ * ∥a∥ : (quotient_norm_mk_le S.to_add_subgroup (k • a)).trans_eq (norm_smul k a) + specialize h (‖a‖) (⟨by linarith, by linarith [submodule.quotient.norm_mk_le S a]⟩), + calc _ ≤ ‖k‖ * ‖a‖ : (quotient_norm_mk_le S.to_add_subgroup (k • a)).trans_eq (norm_smul k a) ... ≤ _ : (sub_lt_iff_lt_add'.mp h.1).le end, .. submodule.quotient.module' S, } @@ -601,11 +601,11 @@ section ideal variables {R : Type*} [semi_normed_comm_ring R] (I : ideal R) lemma ideal.quotient.norm_mk_lt {I : ideal R} (x : R ⧸ I) {ε : ℝ} (hε : 0 < ε) : - ∃ r : R, ideal.quotient.mk I r = x ∧ ∥r∥ < ∥x∥ + ε := + ∃ r : R, ideal.quotient.mk I r = x ∧ ‖r‖ < ‖x‖ + ε := norm_mk_lt x hε lemma ideal.quotient.norm_mk_le (r : R) : - ∥ideal.quotient.mk I r∥ ≤ ∥r∥ := + ‖ideal.quotient.mk I r‖ ≤ ‖r‖ := quotient_norm_mk_le I.to_add_subgroup r instance ideal.quotient.semi_normed_comm_ring : semi_normed_comm_ring (R ⧸ I) := @@ -613,15 +613,15 @@ instance ideal.quotient.semi_normed_comm_ring : semi_normed_comm_ring (R ⧸ I) norm_mul := λ x y, le_of_forall_pos_le_add $ λ ε hε, begin have := ((nhds_basis_ball.prod_nhds nhds_basis_ball).tendsto_iff nhds_basis_ball).mp - (real.continuous_mul.tendsto (∥x∥, ∥y∥)) ε hε, + (real.continuous_mul.tendsto (‖x‖, ‖y‖)) ε hε, simp only [set.mem_prod, mem_ball, and_imp, prod.forall, exists_prop, prod.exists] at this, rcases this with ⟨ε₁, ε₂, ⟨h₁, h₂⟩, h⟩, obtain ⟨⟨a, rfl, ha⟩, ⟨b, rfl, hb⟩⟩ := ⟨ideal.quotient.norm_mk_lt x h₁, ideal.quotient.norm_mk_lt y h₂⟩, simp only [dist, abs_sub_lt_iff] at h, - specialize h (∥a∥) (∥b∥) (⟨by linarith, by linarith [ideal.quotient.norm_mk_le I a]⟩) + specialize h (‖a‖) (‖b‖) (⟨by linarith, by linarith [ideal.quotient.norm_mk_le I a]⟩) (⟨by linarith, by linarith [ideal.quotient.norm_mk_le I b]⟩), - calc _ ≤ ∥a∥ * ∥b∥ : (ideal.quotient.norm_mk_le I (a * b)).trans (norm_mul_le a b) + calc _ ≤ ‖a‖ * ‖b‖ : (ideal.quotient.norm_mk_le I (a * b)).trans (norm_mul_le a b) ... ≤ _ : (sub_lt_iff_lt_add'.mp h.1).le end, .. submodule.quotient.seminormed_add_comm_group I } diff --git a/src/analysis/normed/order/basic.lean b/src/analysis/normed/order/basic.lean index 555fc9ef08fa5..924b53f36f901 100644 --- a/src/analysis/normed/order/basic.lean +++ b/src/analysis/normed/order/basic.lean @@ -23,21 +23,21 @@ variables {α : Type*} carrying their own group structure. -/ class normed_ordered_add_group (α : Type*) extends ordered_add_comm_group α, has_norm α, metric_space α := -(dist_eq : ∀ x y, dist x y = ∥x - y∥ . obviously) +(dist_eq : ∀ x y, dist x y = ‖x - y‖ . obviously) /-- A `normed_ordered_group` is a group that is both a `normed_comm_group` and an `ordered_comm_group`. This class is necessary to avoid diamonds caused by both classes carrying their own group structure. -/ @[to_additive] class normed_ordered_group (α : Type*) extends ordered_comm_group α, has_norm α, metric_space α := -(dist_eq : ∀ x y, dist x y = ∥x / y∥ . obviously) +(dist_eq : ∀ x y, dist x y = ‖x / y‖ . obviously) /-- A `normed_linear_ordered_add_group` is an additive group that is both a `normed_add_comm_group` and a `linear_ordered_add_comm_group`. This class is necessary to avoid diamonds caused by both classes carrying their own group structure. -/ class normed_linear_ordered_add_group (α : Type*) extends linear_ordered_add_comm_group α, has_norm α, metric_space α := -(dist_eq : ∀ x y, dist x y = ∥x - y∥ . obviously) +(dist_eq : ∀ x y, dist x y = ‖x - y‖ . obviously) /-- A `normed_linear_ordered_group` is a group that is both a `normed_comm_group` and a `linear_ordered_comm_group`. This class is necessary to avoid diamonds caused by both classes @@ -45,14 +45,14 @@ carrying their own group structure. -/ @[to_additive] class normed_linear_ordered_group (α : Type*) extends linear_ordered_comm_group α, has_norm α, metric_space α := -(dist_eq : ∀ x y, dist x y = ∥x / y∥ . obviously) +(dist_eq : ∀ x y, dist x y = ‖x / y‖ . obviously) /-- A `normed_linear_ordered_field` is a field that is both a `normed_field` and a `linear_ordered_field`. This class is necessary to avoid diamonds. -/ class normed_linear_ordered_field (α : Type*) extends linear_ordered_field α, has_norm α, metric_space α := -(dist_eq : ∀ x y, dist x y = ∥x - y∥ . obviously) -(norm_mul' : ∀ x y : α, ∥x * y∥ = ∥x∥ * ∥y∥) +(dist_eq : ∀ x y, dist x y = ‖x - y‖ . obviously) +(norm_mul' : ∀ x y : α, ‖x * y‖ = ‖x‖ * ‖y‖) @[to_additive, priority 100] instance normed_ordered_group.to_normed_comm_group [normed_ordered_group α] : normed_comm_group α := diff --git a/src/analysis/normed/order/lattice.lean b/src/analysis/normed/order/lattice.lean index 8f873933c3b34..e9e8318dbba71 100644 --- a/src/analysis/normed/order/lattice.lean +++ b/src/analysis/normed/order/lattice.lean @@ -38,15 +38,15 @@ local notation (name := abs) `|`a`|` := abs a /-- Let `α` be a normed commutative group equipped with a partial order covariant with addition, with respect which `α` forms a lattice. Suppose that `α` is *solid*, that is to say, for `a` and `b` in -`α`, with absolute values `|a|` and `|b|` respectively, `|a| ≤ |b|` implies `∥a∥ ≤ ∥b∥`. Then `α` is +`α`, with absolute values `|a|` and `|b|` respectively, `|a| ≤ |b|` implies `‖a‖ ≤ ‖b‖`. Then `α` is said to be a normed lattice ordered group. -/ class normed_lattice_add_comm_group (α : Type*) extends normed_add_comm_group α, lattice α := (add_le_add_left : ∀ a b : α, a ≤ b → ∀ c : α, c + a ≤ c + b) -(solid : ∀ a b : α, |a| ≤ |b| → ∥a∥ ≤ ∥b∥) +(solid : ∀ a b : α, |a| ≤ |b| → ‖a‖ ≤ ‖b‖) -lemma solid {α : Type*} [normed_lattice_add_comm_group α] {a b : α} (h : |a| ≤ |b|) : ∥a∥ ≤ ∥b∥ := +lemma solid {α : Type*} [normed_lattice_add_comm_group α] {a b : α} (h : |a| ≤ |b|) : ‖a‖ ≤ ‖b‖ := normed_lattice_add_comm_group.solid a b h instance : normed_lattice_add_comm_group ℝ := @@ -63,7 +63,7 @@ instance normed_lattice_add_comm_group_to_ordered_add_comm_group {α : Type*} variables {α : Type*} [normed_lattice_add_comm_group α] open lattice_ordered_comm_group -lemma dual_solid (a b : α) (h: b⊓-b ≤ a⊓-a) : ∥a∥ ≤ ∥b∥ := +lemma dual_solid (a b : α) (h: b⊓-b ≤ a⊓-a) : ‖a‖ ≤ ‖b‖ := begin apply solid, rw abs_eq_sup_neg, @@ -83,10 +83,10 @@ instance : normed_lattice_add_comm_group αᵒᵈ := { solid := dual_solid, ..order_dual.ordered_add_comm_group, ..order_dual.normed_add_comm_group } -lemma norm_abs_eq_norm (a : α) : ∥|a|∥ = ∥a∥ := +lemma norm_abs_eq_norm (a : α) : ‖|a|‖ = ‖a‖ := (solid (abs_abs a).le).antisymm (solid (abs_abs a).symm.le) -lemma norm_inf_sub_inf_le_add_norm (a b c d : α) : ∥a ⊓ b - c ⊓ d∥ ≤ ∥a - c∥ + ∥b - d∥ := +lemma norm_inf_sub_inf_le_add_norm (a b c d : α) : ‖a ⊓ b - c ⊓ d‖ ≤ ‖a - c‖ + ‖b - d‖ := begin rw [← norm_abs_eq_norm (a - c), ← norm_abs_eq_norm (b - d)], refine le_trans (solid _) (norm_add_le (|a - c|) (|b - d|)), @@ -101,7 +101,7 @@ begin exact abs_inf_sub_inf_le_abs _ _ _, } }, end -lemma norm_sup_sub_sup_le_add_norm (a b c d : α) : ∥a ⊔ b - (c ⊔ d)∥ ≤ ∥a - c∥ + ∥b - d∥ := +lemma norm_sup_sub_sup_le_add_norm (a b c d : α) : ‖a ⊔ b - (c ⊔ d)‖ ≤ ‖a - c‖ + ‖b - d‖ := begin rw [← norm_abs_eq_norm (a - c), ← norm_abs_eq_norm (b - d)], refine le_trans (solid _) (norm_add_le (|a - c|) (|b - d|)), @@ -116,15 +116,15 @@ begin exact abs_sup_sub_sup_le_abs _ _ _, } }, end -lemma norm_inf_le_add (x y : α) : ∥x ⊓ y∥ ≤ ∥x∥ + ∥y∥ := +lemma norm_inf_le_add (x y : α) : ‖x ⊓ y‖ ≤ ‖x‖ + ‖y‖ := begin - have h : ∥x ⊓ y - 0 ⊓ 0∥ ≤ ∥x - 0∥ + ∥y - 0∥ := norm_inf_sub_inf_le_add_norm x y 0 0, + have h : ‖x ⊓ y - 0 ⊓ 0‖ ≤ ‖x - 0‖ + ‖y - 0‖ := norm_inf_sub_inf_le_add_norm x y 0 0, simpa only [inf_idem, sub_zero] using h, end -lemma norm_sup_le_add (x y : α) : ∥x ⊔ y∥ ≤ ∥x∥ + ∥y∥ := +lemma norm_sup_le_add (x y : α) : ‖x ⊔ y‖ ≤ ‖x‖ + ‖y‖ := begin - have h : ∥x ⊔ y - 0 ⊔ 0∥ ≤ ∥x - 0∥ + ∥y - 0∥ := norm_sup_sub_sup_le_add_norm x y 0 0, + have h : ‖x ⊔ y - 0 ⊔ 0‖ ≤ ‖x - 0‖ + ‖y - 0‖ := norm_sup_sub_sup_le_add_norm x y 0 0, simpa only [sup_idem, sub_zero] using h, end @@ -135,7 +135,7 @@ Let `α` be a normed lattice ordered group. Then the infimum is jointly continuo instance normed_lattice_add_comm_group_has_continuous_inf : has_continuous_inf α := begin refine ⟨continuous_iff_continuous_at.2 $ λ q, tendsto_iff_norm_tendsto_zero.2 $ _⟩, - have : ∀ p : α × α, ∥p.1 ⊓ p.2 - q.1 ⊓ q.2∥ ≤ ∥p.1 - q.1∥ + ∥p.2 - q.2∥, + have : ∀ p : α × α, ‖p.1 ⊓ p.2 - q.1 ⊓ q.2‖ ≤ ‖p.1 - q.1‖ + ‖p.2 - q.2‖, from λ _, norm_inf_sub_inf_le_add_norm _ _ _ _, refine squeeze_zero (λ e, norm_nonneg _) this _, convert (((continuous_fst.tendsto q).sub tendsto_const_nhds).norm).add @@ -157,13 +157,13 @@ instance normed_lattice_add_comm_group_topological_lattice : topological_lattice topological_lattice.mk lemma norm_abs_sub_abs (a b : α) : - ∥ |a| - |b| ∥ ≤ ∥a-b∥ := + ‖ |a| - |b| ‖ ≤ ‖a-b‖ := solid (lattice_ordered_comm_group.abs_abs_sub_abs_le _ _) -lemma norm_sup_sub_sup_le_norm (x y z : α) : ∥x ⊔ z - (y ⊔ z)∥ ≤ ∥x - y∥ := +lemma norm_sup_sub_sup_le_norm (x y z : α) : ‖x ⊔ z - (y ⊔ z)‖ ≤ ‖x - y‖ := solid (abs_sup_sub_sup_le_abs x y z) -lemma norm_inf_sub_inf_le_norm (x y z : α) : ∥x ⊓ z - (y ⊓ z)∥ ≤ ∥x - y∥ := +lemma norm_inf_sub_inf_le_norm (x y z : α) : ‖x ⊓ z - (y ⊓ z)‖ ≤ ‖x - y‖ := solid (abs_inf_sub_inf_le_abs x y z) lemma lipschitz_with_sup_right (z : α) : lipschitz_with 1 (λ x, x ⊔ z) := diff --git a/src/analysis/normed_space/M_structure.lean b/src/analysis/normed_space/M_structure.lean index 6bd393b3ba9c0..6f38dbc4af335 100644 --- a/src/analysis/normed_space/M_structure.lean +++ b/src/analysis/normed_space/M_structure.lean @@ -68,7 +68,7 @@ Note that we write `P • x` instead of `P x` for reasons described in the modul -/ structure is_Lprojection (P : M) : Prop := (proj : is_idempotent_elem P) -(Lnorm : ∀ (x : X), ∥x∥ = ∥P • x∥ + ∥(1 - P) • x∥) +(Lnorm : ∀ (x : X), ‖x‖ = ‖P • x‖ + ‖(1 - P) • x‖) /-- A projection on a normed space `X` is said to be an M-projection if, for all `x` in `X`, @@ -78,7 +78,7 @@ Note that we write `P • x` instead of `P x` for reasons described in the modul -/ structure is_Mprojection (P : M) : Prop := (proj : is_idempotent_elem P) -(Mnorm : ∀ (x : X), ∥x∥ = (max ∥P • x∥ ∥(1 - P) • x∥)) +(Mnorm : ∀ (x : X), ‖x‖ = (max ‖P • x‖ ‖(1 - P) • x‖)) variables {X} @@ -97,26 +97,26 @@ begin begin refine @eq_of_smul_eq_smul _ X _ _ _ _ (λ x, _), rw ← norm_sub_eq_zero_iff, - have e1 : ∥R • x∥ ≥ ∥R • x∥ + 2 • ∥ (P * R) • x - (R * P * R) • x∥ := - calc ∥R • x∥ = ∥R • (P • (R • x))∥ + ∥(1 - R) • (P • (R • x))∥ + - (∥(R * R) • x - R • (P • (R • x))∥ + ∥(1 - R) • ((1 - P) • (R • x))∥) : + have e1 : ‖R • x‖ ≥ ‖R • x‖ + 2 • ‖ (P * R) • x - (R * P * R) • x‖ := + calc ‖R • x‖ = ‖R • (P • (R • x))‖ + ‖(1 - R) • (P • (R • x))‖ + + (‖(R * R) • x - R • (P • (R • x))‖ + ‖(1 - R) • ((1 - P) • (R • x))‖) : by rw [h₁.Lnorm, h₃.Lnorm, h₃.Lnorm ((1 - P) • (R • x)), sub_smul 1 P, one_smul, smul_sub, mul_smul] - ... = ∥R • (P • (R • x))∥ + ∥(1 - R) • (P • (R • x))∥ + (∥R • x - R • (P • (R • x))∥ - + ∥((1 - R) * R) • x - (1 - R) • (P • (R • x))∥) : by rw [h₃.proj.eq, + ... = ‖R • (P • (R • x))‖ + ‖(1 - R) • (P • (R • x))‖ + (‖R • x - R • (P • (R • x))‖ + + ‖((1 - R) * R) • x - (1 - R) • (P • (R • x))‖) : by rw [h₃.proj.eq, sub_smul 1 P, one_smul, smul_sub, mul_smul] - ... = ∥R • (P • (R • x))∥ + ∥(1 - R) • (P • (R • x))∥ + - (∥R • x - R • (P • (R • x))∥ + ∥(1 - R) • (P • (R • x))∥) : + ... = ‖R • (P • (R • x))‖ + ‖(1 - R) • (P • (R • x))‖ + + (‖R • x - R • (P • (R • x))‖ + ‖(1 - R) • (P • (R • x))‖) : by rw [sub_mul, h₃.proj.eq, one_mul, sub_self, zero_smul, zero_sub, norm_neg] - ... = ∥R • (P • (R • x))∥ + ∥R • x - R • (P • (R • x))∥ + 2•∥(1 - R) • (P • (R • x))∥ : by abel - ... ≥ ∥R • x∥ + 2 • ∥ (P * R) • x - (R * P * R) • x∥ : by + ... = ‖R • (P • (R • x))‖ + ‖R • x - R • (P • (R • x))‖ + 2•‖(1 - R) • (P • (R • x))‖ : by abel + ... ≥ ‖R • x‖ + 2 • ‖ (P * R) • x - (R * P * R) • x‖ : by { rw ge, have := add_le_add_right - (norm_le_insert' (R • x) (R • (P • (R • x)))) (2•∥(1 - R) • (P • (R • x))∥), + (norm_le_insert' (R • x) (R • (P • (R • x)))) (2•‖(1 - R) • (P • (R • x))‖), simpa only [mul_smul, sub_smul, one_smul] using this }, rw ge at e1, - nth_rewrite_rhs 0 ← add_zero (∥R • x∥) at e1, + nth_rewrite_rhs 0 ← add_zero (‖R • x‖) at e1, rw [add_le_add_iff_left, two_smul, ← two_mul] at e1, rw le_antisymm_iff, refine ⟨_, norm_nonneg _⟩, @@ -136,17 +136,17 @@ begin refine ⟨is_idempotent_elem.mul_of_commute (h₁.commute h₂) h₁.proj h₂.proj, _⟩, intro x, refine le_antisymm _ _, - { calc ∥ x ∥ = ∥(P * Q) • x + (x - (P * Q) • x)∥ : by rw add_sub_cancel'_right ((P * Q) • x) x - ... ≤ ∥(P * Q) • x∥ + ∥ x - (P * Q) • x ∥ : by apply norm_add_le - ... = ∥(P * Q) • x∥ + ∥(1 - P * Q) • x∥ : by rw [sub_smul, + { calc ‖ x ‖ = ‖(P * Q) • x + (x - (P * Q) • x)‖ : by rw add_sub_cancel'_right ((P * Q) • x) x + ... ≤ ‖(P * Q) • x‖ + ‖ x - (P * Q) • x ‖ : by apply norm_add_le + ... = ‖(P * Q) • x‖ + ‖(1 - P * Q) • x‖ : by rw [sub_smul, one_smul] }, - { calc ∥x∥ = ∥P • (Q • x)∥ + (∥Q • x - P • (Q • x)∥ + ∥x - Q • x∥) : by + { calc ‖x‖ = ‖P • (Q • x)‖ + (‖Q • x - P • (Q • x)‖ + ‖x - Q • x‖) : by rw [h₂.Lnorm x, h₁.Lnorm (Q • x), sub_smul, one_smul, sub_smul, one_smul, add_assoc] - ... ≥ ∥P • (Q • x)∥ + ∥(Q • x - P • (Q • x)) + (x - Q • x)∥ : - (add_le_add_iff_left (∥P • (Q • x)∥)).mpr (norm_add_le (Q • x - P • (Q • x)) (x - Q • x)) - ... = ∥(P * Q) • x∥ + ∥(1 - P * Q) • x∥ : by rw [sub_add_sub_cancel', + ... ≥ ‖P • (Q • x)‖ + ‖(Q • x - P • (Q • x)) + (x - Q • x)‖ : + (add_le_add_iff_left (‖P • (Q • x)‖)).mpr (norm_add_le (Q • x - P • (Q • x)) (x - Q • x)) + ... = ‖(P * Q) • x‖ + ‖(1 - P * Q) • x‖ : by rw [sub_add_sub_cancel', sub_smul, one_smul, mul_smul] } end diff --git a/src/analysis/normed_space/add_torsor.lean b/src/analysis/normed_space/add_torsor.lean index 6beb40bab3293..21d02eeb0372b 100644 --- a/src/analysis/normed_space/add_torsor.lean +++ b/src/analysis/normed_space/add_torsor.lean @@ -40,11 +40,11 @@ end include V @[simp] lemma dist_center_homothety (p₁ p₂ : P) (c : 𝕜) : - dist p₁ (homothety p₁ c p₂) = ∥c∥ * dist p₁ p₂ := + dist p₁ (homothety p₁ c p₂) = ‖c‖ * dist p₁ p₂ := by simp [homothety_def, norm_smul, ← dist_eq_norm_vsub, dist_comm] @[simp] lemma dist_homothety_center (p₁ p₂ : P) (c : 𝕜) : - dist (homothety p₁ c p₂) p₁ = ∥c∥ * dist p₁ p₂ := + dist (homothety p₁ c p₂) p₁ = ‖c‖ * dist p₁ p₂ := by rw [dist_comm, dist_center_homothety] @[simp] lemma dist_line_map_line_map (p₁ p₂ : P) (c₁ c₂ : 𝕜) : @@ -61,27 +61,27 @@ lipschitz_with.of_dist_le_mul $ λ c₁ c₂, ((dist_line_map_line_map p₁ p₂ c₁ c₂).trans (mul_comm _ _)).le @[simp] lemma dist_line_map_left (p₁ p₂ : P) (c : 𝕜) : - dist (line_map p₁ p₂ c) p₁ = ∥c∥ * dist p₁ p₂ := + dist (line_map p₁ p₂ c) p₁ = ‖c‖ * dist p₁ p₂ := by simpa only [line_map_apply_zero, dist_zero_right] using dist_line_map_line_map p₁ p₂ c 0 @[simp] lemma dist_left_line_map (p₁ p₂ : P) (c : 𝕜) : - dist p₁ (line_map p₁ p₂ c) = ∥c∥ * dist p₁ p₂ := + dist p₁ (line_map p₁ p₂ c) = ‖c‖ * dist p₁ p₂ := (dist_comm _ _).trans (dist_line_map_left _ _ _) @[simp] lemma dist_line_map_right (p₁ p₂ : P) (c : 𝕜) : - dist (line_map p₁ p₂ c) p₂ = ∥1 - c∥ * dist p₁ p₂ := + dist (line_map p₁ p₂ c) p₂ = ‖1 - c‖ * dist p₁ p₂ := by simpa only [line_map_apply_one, dist_eq_norm'] using dist_line_map_line_map p₁ p₂ c 1 @[simp] lemma dist_right_line_map (p₁ p₂ : P) (c : 𝕜) : - dist p₂ (line_map p₁ p₂ c) = ∥1 - c∥ * dist p₁ p₂ := + dist p₂ (line_map p₁ p₂ c) = ‖1 - c‖ * dist p₁ p₂ := (dist_comm _ _).trans (dist_line_map_right _ _ _) @[simp] lemma dist_homothety_self (p₁ p₂ : P) (c : 𝕜) : - dist (homothety p₁ c p₂) p₂ = ∥1 - c∥ * dist p₁ p₂ := + dist (homothety p₁ c p₂) p₂ = ‖1 - c‖ * dist p₁ p₂ := by rw [homothety_eq_line_map, dist_line_map_right] @[simp] lemma dist_self_homothety (p₁ p₂ : P) (c : 𝕜) : - dist p₂ (homothety p₁ c p₂) = ∥1 - c∥ * dist p₁ p₂ := + dist p₂ (homothety p₁ c p₂) = ‖1 - c‖ * dist p₁ p₂ := by rw [dist_comm, dist_homothety_self] section invertible_two @@ -89,23 +89,23 @@ section invertible_two variables [invertible (2:𝕜)] @[simp] lemma dist_left_midpoint (p₁ p₂ : P) : - dist p₁ (midpoint 𝕜 p₁ p₂) = ∥(2:𝕜)∥⁻¹ * dist p₁ p₂ := + dist p₁ (midpoint 𝕜 p₁ p₂) = ‖(2:𝕜)‖⁻¹ * dist p₁ p₂ := by rw [midpoint, dist_comm, dist_line_map_left, inv_of_eq_inv, ← norm_inv] @[simp] lemma dist_midpoint_left (p₁ p₂ : P) : - dist (midpoint 𝕜 p₁ p₂) p₁ = ∥(2:𝕜)∥⁻¹ * dist p₁ p₂ := + dist (midpoint 𝕜 p₁ p₂) p₁ = ‖(2:𝕜)‖⁻¹ * dist p₁ p₂ := by rw [dist_comm, dist_left_midpoint] @[simp] lemma dist_midpoint_right (p₁ p₂ : P) : - dist (midpoint 𝕜 p₁ p₂) p₂ = ∥(2:𝕜)∥⁻¹ * dist p₁ p₂ := + dist (midpoint 𝕜 p₁ p₂) p₂ = ‖(2:𝕜)‖⁻¹ * dist p₁ p₂ := by rw [midpoint_comm, dist_midpoint_left, dist_comm] @[simp] lemma dist_right_midpoint (p₁ p₂ : P) : - dist p₂ (midpoint 𝕜 p₁ p₂) = ∥(2:𝕜)∥⁻¹ * dist p₁ p₂ := + dist p₂ (midpoint 𝕜 p₁ p₂) = ‖(2:𝕜)‖⁻¹ * dist p₁ p₂ := by rw [dist_comm, dist_midpoint_right] lemma dist_midpoint_midpoint_le' (p₁ p₂ p₃ p₄ : P) : - dist (midpoint 𝕜 p₁ p₂) (midpoint 𝕜 p₃ p₄) ≤ (dist p₁ p₃ + dist p₂ p₄) / ∥(2 : 𝕜)∥ := + dist (midpoint 𝕜 p₁ p₂) (midpoint 𝕜 p₃ p₄) ≤ (dist p₁ p₃ + dist p₂ p₄) / ‖(2 : 𝕜)‖ := begin rw [dist_eq_norm_vsub V, dist_eq_norm_vsub V, dist_eq_norm_vsub V, midpoint_vsub_midpoint]; try { apply_instance }, @@ -130,10 +130,10 @@ lemma eventually_homothety_mem_of_mem_interior (x : Q) {s : set Q} {y : Q} (hy : begin rw (normed_add_comm_group.nhds_basis_norm_lt (1 : 𝕜)).eventually_iff, cases eq_or_ne y x with h h, { use 1, simp [h.symm, interior_subset hy], }, - have hxy : 0 < ∥y -ᵥ x∥, { rwa [norm_pos_iff, vsub_ne_zero], }, + have hxy : 0 < ‖y -ᵥ x‖, { rwa [norm_pos_iff, vsub_ne_zero], }, obtain ⟨u, hu₁, hu₂, hu₃⟩ := mem_interior.mp hy, obtain ⟨ε, hε, hyε⟩ := metric.is_open_iff.mp hu₂ y hu₃, - refine ⟨ε / ∥y -ᵥ x∥, div_pos hε hxy, λ δ (hδ : ∥δ - 1∥ < ε / ∥y -ᵥ x∥), hu₁ (hyε _)⟩, + refine ⟨ε / ‖y -ᵥ x‖, div_pos hε hxy, λ δ (hδ : ‖δ - 1‖ < ε / ‖y -ᵥ x‖), hu₁ (hyε _)⟩, rw [lt_div_iff hxy, ← norm_smul, sub_smul, one_smul] at hδ, rwa [homothety_apply, metric.mem_ball, dist_eq_norm_vsub W, vadd_vsub_eq_sub_vsub], end diff --git a/src/analysis/normed_space/affine_isometry.lean b/src/analysis/normed_space/affine_isometry.lean index 85ba9c7b1241a..d95debabb91d5 100644 --- a/src/analysis/normed_space/affine_isometry.lean +++ b/src/analysis/normed_space/affine_isometry.lean @@ -17,7 +17,7 @@ isometric equivalence between `P` and `P₂`. We also prove basic lemmas and provide convenience constructors. The choice of these lemmas and constructors is closely modelled on those for the `linear_isometry` and `affine_map` theories. -Since many elementary properties don't require `∥x∥ = 0 → x = 0` we initially set up the theory for +Since many elementary properties don't require `‖x‖ = 0 → x = 0` we initially set up the theory for `seminormed_add_comm_group` and specialize to `normed_add_comm_group` only when needed. ## Notation @@ -47,7 +47,7 @@ include V V₂ /-- An `𝕜`-affine isometric embedding of one normed add-torsor over a normed `𝕜`-space into another. -/ structure affine_isometry extends P →ᵃ[𝕜] P₂ := -(norm_map : ∀ x : V, ∥linear x∥ = ∥x∥) +(norm_map : ∀ x : V, ‖linear x‖ = ‖x‖) omit V V₂ variables {𝕜 P P₂} @@ -229,7 +229,7 @@ include V V₂ /-- A affine isometric equivalence between two normed vector spaces. -/ structure affine_isometry_equiv extends P ≃ᵃ[𝕜] P₂ := -(norm_map : ∀ x, ∥linear x∥ = ∥x∥) +(norm_map : ∀ x, ‖linear x‖ = ‖x‖) variables {𝕜 P P₂} omit V V₂ @@ -252,7 +252,7 @@ by { ext, refl } include V V₂ instance : has_coe_to_fun (P ≃ᵃⁱ[𝕜] P₂) (λ _, P → P₂) := ⟨λ f, f.to_fun⟩ -@[simp] lemma coe_mk (e : P ≃ᵃ[𝕜] P₂) (he : ∀ x, ∥e.linear x∥ = ∥x∥) : +@[simp] lemma coe_mk (e : P ≃ᵃ[𝕜] P₂) (he : ∀ x, ‖e.linear x‖ = ‖x‖) : ⇑(mk e he) = e := rfl @@ -531,11 +531,11 @@ to_affine_equiv_injective $ affine_equiv.point_reflection_symm 𝕜 x by rw [← (point_reflection 𝕜 x).dist_map y x, point_reflection_self] lemma dist_point_reflection_self' (x y : P) : - dist (point_reflection 𝕜 x y) y = ∥bit0 (x -ᵥ y)∥ := + dist (point_reflection 𝕜 x y) y = ‖bit0 (x -ᵥ y)‖ := by rw [point_reflection_apply, dist_eq_norm_vsub V, vadd_vsub_assoc, bit0] lemma dist_point_reflection_self (x y : P) : - dist (point_reflection 𝕜 x y) y = ∥(2:𝕜)∥ * dist x y := + dist (point_reflection 𝕜 x y) y = ‖(2:𝕜)‖ * dist x y := by rw [dist_point_reflection_self', ← two_smul' 𝕜 (x -ᵥ y), norm_smul, ← dist_eq_norm_vsub V] lemma point_reflection_fixed_iff [invertible (2:𝕜)] {x y : P} : diff --git a/src/analysis/normed_space/algebra.lean b/src/analysis/normed_space/algebra.lean index d253557f7d743..62336f754f7b8 100644 --- a/src/analysis/normed_space/algebra.lean +++ b/src/analysis/normed_space/algebra.lean @@ -36,14 +36,14 @@ variables [nontrivially_normed_field 𝕜] [normed_ring A] [normed_algebra 𝕜 A] [complete_space A] lemma norm_le_norm_one (φ : character_space 𝕜 A) : - ∥to_normed_dual (φ : weak_dual 𝕜 A)∥ ≤ ∥(1 : A)∥ := + ‖to_normed_dual (φ : weak_dual 𝕜 A)‖ ≤ ‖(1 : A)‖ := continuous_linear_map.op_norm_le_bound _ (norm_nonneg (1 : A)) $ - λ a, mul_comm (∥a∥) (∥(1 : A)∥) ▸ spectrum.norm_le_norm_mul_of_mem (apply_mem_spectrum φ a) + λ a, mul_comm (‖a‖) (‖(1 : A)‖) ▸ spectrum.norm_le_norm_mul_of_mem (apply_mem_spectrum φ a) instance [proper_space 𝕜] : compact_space (character_space 𝕜 A) := begin rw [←is_compact_iff_compact_space], - have h : character_space 𝕜 A ⊆ to_normed_dual ⁻¹' metric.closed_ball 0 (∥(1 : A)∥), + have h : character_space 𝕜 A ⊆ to_normed_dual ⁻¹' metric.closed_ball 0 (‖(1 : A)‖), { intros φ hφ, rw [set.mem_preimage, mem_closed_ball_zero_iff], exact (norm_le_norm_one ⟨φ, ⟨hφ.1, hφ.2⟩⟩ : _), }, diff --git a/src/analysis/normed_space/banach.lean b/src/analysis/normed_space/banach.lean index c806938b67e63..d086f173af75a 100644 --- a/src/analysis/normed_space/banach.lean +++ b/src/analysis/normed_space/banach.lean @@ -26,13 +26,13 @@ include 𝕜 namespace continuous_linear_map /-- A (possibly nonlinear) right inverse to a continuous linear map, which doesn't have to be -linear itself but which satisfies a bound `∥inverse x∥ ≤ C * ∥x∥`. A surjective continuous linear +linear itself but which satisfies a bound `‖inverse x‖ ≤ C * ‖x‖`. A surjective continuous linear map doesn't always have a continuous linear right inverse, but it always has a nonlinear inverse in this sense, by Banach's open mapping theorem. -/ structure nonlinear_right_inverse := (to_fun : F → E) (nnnorm : ℝ≥0) -(bound' : ∀ y, ∥to_fun y∥ ≤ nnnorm * ∥y∥) +(bound' : ∀ y, ‖to_fun y‖ ≤ nnnorm * ‖y‖) (right_inv' : ∀ y, f (to_fun y) = y) instance : has_coe_to_fun (nonlinear_right_inverse f) (λ _, F → E) := ⟨λ fsymm, fsymm.to_fun⟩ @@ -42,7 +42,7 @@ instance : has_coe_to_fun (nonlinear_right_inverse f) (λ _, F → E) := ⟨λ f fsymm.right_inv' y lemma nonlinear_right_inverse.bound {f : E →L[𝕜] F} (fsymm : nonlinear_right_inverse f) (y : F) : - ∥fsymm y∥ ≤ fsymm.nnnorm * ∥y∥ := + ‖fsymm y‖ ≤ fsymm.nnnorm * ‖y‖ := fsymm.bound' y end continuous_linear_map @@ -52,7 +52,7 @@ end continuous_linear_map noncomputable def continuous_linear_equiv.to_nonlinear_right_inverse (f : E ≃L[𝕜] F) : continuous_linear_map.nonlinear_right_inverse (f : E →L[𝕜] F) := { to_fun := f.inv_fun, - nnnorm := ∥(f.symm : F →L[𝕜] E)∥₊, + nnnorm := ‖(f.symm : F →L[𝕜] E)‖₊, bound' := λ y, continuous_linear_map.le_op_norm (f.symm : F →L[𝕜] E) _, right_inv' := f.apply_symm_apply } @@ -69,16 +69,16 @@ namespace continuous_linear_map First step of the proof of the Banach open mapping theorem (using completeness of `F`): by Baire's theorem, there exists a ball in `E` whose image closure has nonempty interior. Rescaling everything, it follows that any `y ∈ F` is arbitrarily well approached by -images of elements of norm at most `C * ∥y∥`. +images of elements of norm at most `C * ‖y‖`. For further use, we will only need such an element whose image -is within distance `∥y∥/2` of `y`, to apply an iterative process. -/ +is within distance `‖y‖/2` of `y`, to apply an iterative process. -/ lemma exists_approx_preimage_norm_le (surj : surjective f) : - ∃C ≥ 0, ∀y, ∃x, dist (f x) y ≤ 1/2 * ∥y∥ ∧ ∥x∥ ≤ C * ∥y∥ := + ∃C ≥ 0, ∀y, ∃x, dist (f x) y ≤ 1/2 * ‖y‖ ∧ ‖x‖ ≤ C * ‖y‖ := begin have A : (⋃n:ℕ, closure (f '' (ball 0 n))) = univ, { refine subset.antisymm (subset_univ _) (λy hy, _), rcases surj y with ⟨x, hx⟩, - rcases exists_nat_gt (∥x∥) with ⟨n, hn⟩, + rcases exists_nat_gt (‖x‖) with ⟨n, hn⟩, refine mem_Union.2 ⟨n, subset_closure _⟩, refine (mem_image _ _ _).2 ⟨x, ⟨_, hx⟩⟩, rwa [mem_ball, dist_eq_norm, sub_zero] }, @@ -87,13 +87,13 @@ begin simp only [mem_interior_iff_mem_nhds, metric.mem_nhds_iff] at this, rcases this with ⟨n, a, ε, ⟨εpos, H⟩⟩, rcases normed_field.exists_one_lt_norm 𝕜 with ⟨c, hc⟩, - refine ⟨(ε/2)⁻¹ * ∥c∥ * 2 * n, _, λy, _⟩, + refine ⟨(ε/2)⁻¹ * ‖c‖ * 2 * n, _, λy, _⟩, { refine mul_nonneg (mul_nonneg (mul_nonneg _ (norm_nonneg _)) (by norm_num)) _, exacts [inv_nonneg.2 (div_nonneg (le_of_lt εpos) (by norm_num)), n.cast_nonneg] }, { by_cases hy : y = 0, { use 0, simp [hy] }, { rcases rescale_to_shell hc (half_pos εpos) hy with ⟨d, hd, ydlt, leyd, dinv⟩, - let δ := ∥d∥ * ∥y∥/4, + let δ := ‖d‖ * ‖y‖/4, have δpos : 0 < δ := div_pos (mul_pos (norm_pos_iff.2 hd) (norm_pos_iff.2 hy)) (by norm_num), have : a + d • y ∈ ball a ε, @@ -108,10 +108,10 @@ begin rw ← xz₂ at h₂, rw [mem_ball, dist_eq_norm, sub_zero] at hx₂, let x := x₁ - x₂, - have I : ∥f x - d • y∥ ≤ 2 * δ := calc - ∥f x - d • y∥ = ∥f x₁ - (a + d • y) - (f x₂ - a)∥ : + have I : ‖f x - d • y‖ ≤ 2 * δ := calc + ‖f x - d • y‖ = ‖f x₁ - (a + d • y) - (f x₂ - a)‖ : by { congr' 1, simp only [x, f.map_sub], abel } - ... ≤ ∥f x₁ - (a + d • y)∥ + ∥f x₂ - a∥ : + ... ≤ ‖f x₁ - (a + d • y)‖ + ‖f x₂ - a‖ : norm_sub_le _ _ ... ≤ δ + δ : begin apply add_le_add, @@ -119,29 +119,29 @@ begin { rw [← dist_eq_norm, dist_comm], exact le_of_lt h₂ } end ... = 2 * δ : (two_mul _).symm, - have J : ∥f (d⁻¹ • x) - y∥ ≤ 1/2 * ∥y∥ := calc - ∥f (d⁻¹ • x) - y∥ = ∥d⁻¹ • f x - (d⁻¹ * d) • y∥ : + have J : ‖f (d⁻¹ • x) - y‖ ≤ 1/2 * ‖y‖ := calc + ‖f (d⁻¹ • x) - y‖ = ‖d⁻¹ • f x - (d⁻¹ * d) • y‖ : by rwa [f.map_smul _, inv_mul_cancel, one_smul] - ... = ∥d⁻¹ • (f x - d • y)∥ : by rw [mul_smul, smul_sub] - ... = ∥d∥⁻¹ * ∥f x - d • y∥ : by rw [norm_smul, norm_inv] - ... ≤ ∥d∥⁻¹ * (2 * δ) : begin + ... = ‖d⁻¹ • (f x - d • y)‖ : by rw [mul_smul, smul_sub] + ... = ‖d‖⁻¹ * ‖f x - d • y‖ : by rw [norm_smul, norm_inv] + ... ≤ ‖d‖⁻¹ * (2 * δ) : begin apply mul_le_mul_of_nonneg_left I, rw inv_nonneg, exact norm_nonneg _ end - ... = (∥d∥⁻¹ * ∥d∥) * ∥y∥ /2 : by { simp only [δ], ring } - ... = ∥y∥/2 : by { rw [inv_mul_cancel, one_mul], simp [norm_eq_zero, hd] } - ... = (1/2) * ∥y∥ : by ring, + ... = (‖d‖⁻¹ * ‖d‖) * ‖y‖ /2 : by { simp only [δ], ring } + ... = ‖y‖/2 : by { rw [inv_mul_cancel, one_mul], simp [norm_eq_zero, hd] } + ... = (1/2) * ‖y‖ : by ring, rw ← dist_eq_norm at J, - have K : ∥d⁻¹ • x∥ ≤ (ε / 2)⁻¹ * ∥c∥ * 2 * ↑n * ∥y∥ := calc - ∥d⁻¹ • x∥ = ∥d∥⁻¹ * ∥x₁ - x₂∥ : by rw [norm_smul, norm_inv] - ... ≤ ((ε / 2)⁻¹ * ∥c∥ * ∥y∥) * (n + n) : begin + have K : ‖d⁻¹ • x‖ ≤ (ε / 2)⁻¹ * ‖c‖ * 2 * ↑n * ‖y‖ := calc + ‖d⁻¹ • x‖ = ‖d‖⁻¹ * ‖x₁ - x₂‖ : by rw [norm_smul, norm_inv] + ... ≤ ((ε / 2)⁻¹ * ‖c‖ * ‖y‖) * (n + n) : begin refine mul_le_mul dinv _ (norm_nonneg _) _, { exact le_trans (norm_sub_le _ _) (add_le_add (le_of_lt hx₁) (le_of_lt hx₂)) }, { apply mul_nonneg (mul_nonneg _ (norm_nonneg _)) (norm_nonneg _), exact inv_nonneg.2 (le_of_lt (half_pos εpos)) } end - ... = (ε / 2)⁻¹ * ∥c∥ * 2 * ↑n * ∥y∥ : by ring, + ... = (ε / 2)⁻¹ * ‖c‖ * 2 * ↑n * ‖y‖ : by ring, exact ⟨d⁻¹ • x, J, K⟩ } }, end @@ -150,7 +150,7 @@ variable [complete_space E] /-- The Banach open mapping theorem: if a bounded linear map between Banach spaces is onto, then any point has a preimage with controlled norm. -/ theorem exists_preimage_norm_le (surj : surjective f) : - ∃C > 0, ∀y, ∃x, f x = y ∧ ∥x∥ ≤ C * ∥y∥ := + ∃C > 0, ∀y, ∃x, f x = y ∧ ‖x‖ ≤ C * ‖y‖ := begin obtain ⟨C, C0, hC⟩ := exists_approx_preimage_norm_le f surj, /- Second step of the proof: starting from `y`, we want an exact preimage of `y`. Let `g y` be @@ -161,12 +161,12 @@ begin preimage of `y`. This uses completeness of `E`. -/ choose g hg using hC, let h := λy, y - f (g y), - have hle : ∀y, ∥h y∥ ≤ (1/2) * ∥y∥, + have hle : ∀y, ‖h y‖ ≤ (1/2) * ‖y‖, { assume y, rw [← dist_eq_norm, dist_comm], exact (hg y).1 }, refine ⟨2 * C + 1, by linarith, λy, _⟩, - have hnle : ∀n:ℕ, ∥(h^[n]) y∥ ≤ (1/2)^n * ∥y∥, + have hnle : ∀n:ℕ, ‖(h^[n]) y‖ ≤ (1/2)^n * ‖y‖, { assume n, induction n with n IH, { simp only [one_div, nat.nat_zero_eq_zero, one_mul, iterate_zero_apply, @@ -177,24 +177,24 @@ begin apply mul_le_mul_of_nonneg_left IH, norm_num } }, let u := λn, g((h^[n]) y), - have ule : ∀n, ∥u n∥ ≤ (1/2)^n * (C * ∥y∥), + have ule : ∀n, ‖u n‖ ≤ (1/2)^n * (C * ‖y‖), { assume n, apply le_trans (hg _).2 _, - calc C * ∥(h^[n]) y∥ ≤ C * ((1/2)^n * ∥y∥) : mul_le_mul_of_nonneg_left (hnle n) C0 - ... = (1 / 2) ^ n * (C * ∥y∥) : by ring }, - have sNu : summable (λn, ∥u n∥), + calc C * ‖(h^[n]) y‖ ≤ C * ((1/2)^n * ‖y‖) : mul_le_mul_of_nonneg_left (hnle n) C0 + ... = (1 / 2) ^ n * (C * ‖y‖) : by ring }, + have sNu : summable (λn, ‖u n‖), { refine summable_of_nonneg_of_le (λn, norm_nonneg _) ule _, exact summable.mul_right _ (summable_geometric_of_lt_1 (by norm_num) (by norm_num)) }, have su : summable u := summable_of_summable_norm sNu, let x := tsum u, - have x_ineq : ∥x∥ ≤ (2 * C + 1) * ∥y∥ := calc - ∥x∥ ≤ ∑'n, ∥u n∥ : norm_tsum_le_tsum_norm sNu - ... ≤ ∑'n, (1/2)^n * (C * ∥y∥) : + have x_ineq : ‖x‖ ≤ (2 * C + 1) * ‖y‖ := calc + ‖x‖ ≤ ∑'n, ‖u n‖ : norm_tsum_le_tsum_norm sNu + ... ≤ ∑'n, (1/2)^n * (C * ‖y‖) : tsum_le_tsum ule sNu (summable.mul_right _ summable_geometric_two) - ... = (∑'n, (1/2)^n) * (C * ∥y∥) : tsum_mul_right - ... = 2 * C * ∥y∥ : by rw [tsum_geometric_two, mul_assoc] - ... ≤ 2 * C * ∥y∥ + ∥y∥ : le_add_of_nonneg_right (norm_nonneg y) - ... = (2 * C + 1) * ∥y∥ : by ring, + ... = (∑'n, (1/2)^n) * (C * ‖y‖) : tsum_mul_right + ... = 2 * C * ‖y‖ : by rw [tsum_geometric_two, mul_assoc] + ... ≤ 2 * C * ‖y‖ + ‖y‖ : le_add_of_nonneg_right (norm_nonneg y) + ... = (2 * C + 1) * ‖y‖ : by ring, have fsumeq : ∀n:ℕ, f (∑ i in finset.range n, u i) = y - (h^[n]) y, { assume n, induction n with n IH, @@ -210,7 +210,7 @@ begin rw tendsto_iff_norm_tendsto_zero, simp only [sub_zero], refine squeeze_zero (λ_, norm_nonneg _) hnle _, - rw [← zero_mul ∥y∥], + rw [← zero_mul ‖y‖], refine (tendsto_pow_at_top_nhds_0_of_lt_1 _ _).mul tendsto_const_nhds; norm_num }, have feq : f x = y - 0 := tendsto_nhds_unique L₁ L₂, rw sub_zero at feq, @@ -231,8 +231,8 @@ begin have : f (x + w) = z, by { rw [f.map_add, wim, fxy, add_sub_cancel'_right] }, rw ← this, have : x + w ∈ ball x ε := calc - dist (x+w) x = ∥w∥ : by { rw dist_eq_norm, simp } - ... ≤ C * ∥z - y∥ : wnorm + dist (x+w) x = ‖w‖ : by { rw dist_eq_norm, simp } + ... ≤ C * ‖z - y‖ : wnorm ... < C * (ε/C) : begin apply mul_lt_mul_of_pos_left _ Cpos, rwa [mem_ball, dist_eq_norm] at hz, diff --git a/src/analysis/normed_space/banach_steinhaus.lean b/src/analysis/normed_space/banach_steinhaus.lean index be992f565df5c..4ef0b2647f13e 100644 --- a/src/analysis/normed_space/banach_steinhaus.lean +++ b/src/analysis/normed_space/banach_steinhaus.lean @@ -33,11 +33,11 @@ variables If a family of continuous linear maps from a Banach space into a normed space is pointwise bounded, then the norms of these linear maps are uniformly bounded. -/ theorem banach_steinhaus {ι : Type*} [complete_space E] {g : ι → E →SL[σ₁₂] F} - (h : ∀ x, ∃ C, ∀ i, ∥g i x∥ ≤ C) : - ∃ C', ∀ i, ∥g i∥ ≤ C' := + (h : ∀ x, ∃ C, ∀ i, ‖g i x‖ ≤ C) : + ∃ C', ∀ i, ‖g i‖ ≤ C' := begin - /- sequence of subsets consisting of those `x : E` with norms `∥g i x∥` bounded by `n` -/ - let e : ℕ → set E := λ n, (⋂ i : ι, { x : E | ∥g i x∥ ≤ n }), + /- sequence of subsets consisting of those `x : E` with norms `‖g i x‖` bounded by `n` -/ + let e : ℕ → set E := λ n, (⋂ i : ι, { x : E | ‖g i x‖ ≤ n }), /- each of these sets is closed -/ have hc : ∀ n : ℕ, is_closed (e n), from λ i, is_closed_Inter (λ i, is_closed_le (continuous.norm (g i).cont) continuous_const), @@ -52,41 +52,41 @@ begin rcases metric.is_open_iff.mp is_open_interior x hx with ⟨ε, ε_pos, hε⟩, obtain ⟨k, hk⟩ := normed_field.exists_one_lt_norm 𝕜, /- show all elements in the ball have norm bounded by `m` after applying any `g i` -/ - have real_norm_le : ∀ z : E, z ∈ metric.ball x ε → ∀ i : ι, ∥g i z∥ ≤ m, + have real_norm_le : ∀ z : E, z ∈ metric.ball x ε → ∀ i : ι, ‖g i z‖ ≤ m, { intros z hz i, replace hz := mem_Inter.mp (interior_Inter_subset _ (hε hz)) i, apply interior_subset hz }, - have εk_pos : 0 < ε / ∥k∥ := div_pos ε_pos (zero_lt_one.trans hk), - refine ⟨(m + m : ℕ) / (ε / ∥k∥), λ i, continuous_linear_map.op_norm_le_of_shell ε_pos _ hk _⟩, + have εk_pos : 0 < ε / ‖k‖ := div_pos ε_pos (zero_lt_one.trans hk), + refine ⟨(m + m : ℕ) / (ε / ‖k‖), λ i, continuous_linear_map.op_norm_le_of_shell ε_pos _ hk _⟩, { exact div_nonneg (nat.cast_nonneg _) εk_pos.le }, intros y le_y y_lt, - calc ∥g i y∥ - = ∥g i (y + x) - g i x∥ : by rw [continuous_linear_map.map_add, add_sub_cancel] - ... ≤ ∥g i (y + x)∥ + ∥g i x∥ : norm_sub_le _ _ + calc ‖g i y‖ + = ‖g i (y + x) - g i x‖ : by rw [continuous_linear_map.map_add, add_sub_cancel] + ... ≤ ‖g i (y + x)‖ + ‖g i x‖ : norm_sub_le _ _ ... ≤ m + m : add_le_add (real_norm_le (y + x) (by rwa [add_comm, add_mem_ball_iff_norm]) i) (real_norm_le x (metric.mem_ball_self ε_pos) i) ... = (m + m : ℕ) : (m.cast_add m).symm - ... ≤ (m + m : ℕ) * (∥y∥ / (ε / ∥k∥)) + ... ≤ (m + m : ℕ) * (‖y‖ / (ε / ‖k‖)) : le_mul_of_one_le_right (nat.cast_nonneg _) ((one_le_div $ div_pos ε_pos (zero_lt_one.trans hk)).2 le_y) - ... = (m + m : ℕ) / (ε / ∥k∥) * ∥y∥ : (mul_comm_div _ _ _).symm, + ... = (m + m : ℕ) / (ε / ‖k‖) * ‖y‖ : (mul_comm_div _ _ _).symm, end open_locale ennreal open ennreal -/-- This version of Banach-Steinhaus is stated in terms of suprema of `↑∥⬝∥₊ : ℝ≥0∞` +/-- This version of Banach-Steinhaus is stated in terms of suprema of `↑‖⬝‖₊ : ℝ≥0∞` for convenience. -/ theorem banach_steinhaus_supr_nnnorm {ι : Type*} [complete_space E] {g : ι → E →SL[σ₁₂] F} - (h : ∀ x, (⨆ i, ↑∥g i x∥₊) < ∞) : - (⨆ i, ↑∥g i∥₊) < ∞ := + (h : ∀ x, (⨆ i, ↑‖g i x‖₊) < ∞) : + (⨆ i, ↑‖g i‖₊) < ∞ := begin - have h' : ∀ x : E, ∃ C : ℝ, ∀ i : ι, ∥g i x∥ ≤ C, + have h' : ∀ x : E, ∃ C : ℝ, ∀ i : ι, ‖g i x‖ ≤ C, { intro x, rcases lt_iff_exists_coe.mp (h x) with ⟨p, hp₁, _⟩, refine ⟨p, (λ i, _)⟩, exact_mod_cast - calc (∥g i x∥₊ : ℝ≥0∞) ≤ ⨆ j, ∥g j x∥₊ : le_supr _ i + calc (‖g i x‖₊ : ℝ≥0∞) ≤ ⨆ j, ‖g j x‖₊ : le_supr _ i ... = p : hp₁ }, cases banach_steinhaus h' with C' hC', refine (supr_le $ λ i, _).trans_lt (@coe_lt_top C'.to_nnreal), @@ -109,21 +109,21 @@ def continuous_linear_map_of_tendsto [complete_space E] [t2_space F] cont := begin /- show that the maps are pointwise bounded and apply `banach_steinhaus`-/ - have h_point_bdd : ∀ x : E, ∃ C : ℝ, ∀ n : ℕ, ∥g n x∥ ≤ C, + have h_point_bdd : ∀ x : E, ∃ C : ℝ, ∀ n : ℕ, ‖g n x‖ ≤ C, { intro x, rcases cauchy_seq_bdd (tendsto_pi_nhds.mp h x).cauchy_seq with ⟨C, C_pos, hC⟩, - refine ⟨C + ∥g 0 x∥, (λ n, _)⟩, + refine ⟨C + ‖g 0 x‖, (λ n, _)⟩, simp_rw dist_eq_norm at hC, - calc ∥g n x∥ ≤ ∥g 0 x∥ + ∥g n x - g 0 x∥ : norm_le_insert' _ _ - ... ≤ C + ∥g 0 x∥ : by linarith [hC n 0] }, + calc ‖g n x‖ ≤ ‖g 0 x‖ + ‖g n x - g 0 x‖ : norm_le_insert' _ _ + ... ≤ C + ‖g 0 x‖ : by linarith [hC n 0] }, cases banach_steinhaus h_point_bdd with C' hC', /- show the uniform bound from `banach_steinhaus` is a norm bound of the limit map by allowing "an `ε` of room." -/ refine add_monoid_hom_class.continuous_of_bound (linear_map_of_tendsto _ _ h) C' (λ x, le_of_forall_pos_lt_add (λ ε ε_pos, _)), cases metric.tendsto_at_top.mp (tendsto_pi_nhds.mp h x) ε ε_pos with n hn, - have lt_ε : ∥g n x - f x∥ < ε, by {rw ←dist_eq_norm, exact hn n (le_refl n)}, - calc ∥f x∥ ≤ ∥g n x∥ + ∥g n x - f x∥ : norm_le_insert _ _ - ... < ∥g n∥ * ∥x∥ + ε : by linarith [lt_ε, (g n).le_op_norm x] - ... ≤ C' * ∥x∥ + ε : by nlinarith [hC' n, norm_nonneg x], + have lt_ε : ‖g n x - f x‖ < ε, by {rw ←dist_eq_norm, exact hn n (le_refl n)}, + calc ‖f x‖ ≤ ‖g n x‖ + ‖g n x - f x‖ : norm_le_insert _ _ + ... < ‖g n‖ * ‖x‖ + ε : by linarith [lt_ε, (g n).le_op_norm x] + ... ≤ C' * ‖x‖ + ε : by nlinarith [hC' n, norm_nonneg x], end } diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index 7c98210c6745f..315c0d6d6a289 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -28,15 +28,15 @@ set_option extends_priority 920 -- to take precedence over `semiring.to_module` as this leads to instance paths with better -- unification properties. /-- A normed space over a normed field is a vector space endowed with a norm which satisfies the -equality `∥c • x∥ = ∥c∥ ∥x∥`. We require only `∥c • x∥ ≤ ∥c∥ ∥x∥` in the definition, then prove -`∥c • x∥ = ∥c∥ ∥x∥` in `norm_smul`. +equality `‖c • x‖ = ‖c‖ ‖x‖`. We require only `‖c • x‖ ≤ ‖c‖ ‖x‖` in the definition, then prove +`‖c • x‖ = ‖c‖ ‖x‖` in `norm_smul`. Note that since this requires `seminormed_add_comm_group` and not `normed_add_comm_group`, this typeclass can be used for "semi normed spaces" too, just as `module` can be used for "semi modules". -/ class normed_space (α : Type*) (β : Type*) [normed_field α] [seminormed_add_comm_group β] extends module α β := -(norm_smul_le : ∀ (a:α) (b:β), ∥a • b∥ ≤ ∥a∥ * ∥b∥) +(norm_smul_le : ∀ (a:α) (b:β), ‖a • b‖ ≤ ‖a‖ * ‖b‖) end prio variables [normed_field α] [seminormed_add_comm_group β] @@ -51,52 +51,52 @@ instance normed_space.has_bounded_smul [normed_space α β] : has_bounded_smul instance normed_field.to_normed_space : normed_space α α := { norm_smul_le := λ a b, le_of_eq (norm_mul a b) } -lemma norm_smul [normed_space α β] (s : α) (x : β) : ∥s • x∥ = ∥s∥ * ∥x∥ := +lemma norm_smul [normed_space α β] (s : α) (x : β) : ‖s • x‖ = ‖s‖ * ‖x‖ := begin by_cases h : s = 0, { simp [h] }, { refine le_antisymm (normed_space.norm_smul_le s x) _, - calc ∥s∥ * ∥x∥ = ∥s∥ * ∥s⁻¹ • s • x∥ : by rw [inv_smul_smul₀ h] - ... ≤ ∥s∥ * (∥s⁻¹∥ * ∥s • x∥) : + calc ‖s‖ * ‖x‖ = ‖s‖ * ‖s⁻¹ • s • x‖ : by rw [inv_smul_smul₀ h] + ... ≤ ‖s‖ * (‖s⁻¹‖ * ‖s • x‖) : mul_le_mul_of_nonneg_left (normed_space.norm_smul_le _ _) (norm_nonneg _) - ... = ∥s • x∥ : + ... = ‖s • x‖ : by rw [norm_inv, ← mul_assoc, mul_inv_cancel (mt norm_eq_zero.1 h), one_mul] } end lemma norm_zsmul (α) [normed_field α] [normed_space α β] (n : ℤ) (x : β) : - ∥n • x∥ = ∥(n : α)∥ * ∥x∥ := + ‖n • x‖ = ‖(n : α)‖ * ‖x‖ := by rw [← norm_smul, ← int.smul_one_eq_coe, smul_assoc, one_smul] -@[simp] lemma abs_norm_eq_norm (z : β) : |∥z∥| = ∥z∥ := +@[simp] lemma abs_norm_eq_norm (z : β) : |‖z‖| = ‖z‖ := (abs_eq (norm_nonneg z)).mpr (or.inl rfl) lemma inv_norm_smul_mem_closed_unit_ball [normed_space ℝ β] (x : β) : - ∥x∥⁻¹ • x ∈ closed_ball (0 : β) 1 := + ‖x‖⁻¹ • x ∈ closed_ball (0 : β) 1 := by simp only [mem_closed_ball_zero_iff, norm_smul, norm_inv, norm_norm, ← div_eq_inv_mul, div_self_le_one] -lemma dist_smul [normed_space α β] (s : α) (x y : β) : dist (s • x) (s • y) = ∥s∥ * dist x y := +lemma dist_smul [normed_space α β] (s : α) (x y : β) : dist (s • x) (s • y) = ‖s‖ * dist x y := by simp only [dist_eq_norm, (norm_smul _ _).symm, smul_sub] -lemma nnnorm_smul [normed_space α β] (s : α) (x : β) : ∥s • x∥₊ = ∥s∥₊ * ∥x∥₊ := +lemma nnnorm_smul [normed_space α β] (s : α) (x : β) : ‖s • x‖₊ = ‖s‖₊ * ‖x‖₊ := nnreal.eq $ norm_smul s x lemma nndist_smul [normed_space α β] (s : α) (x y : β) : - nndist (s • x) (s • y) = ∥s∥₊ * nndist x y := + nndist (s • x) (s • y) = ‖s‖₊ * nndist x y := nnreal.eq $ dist_smul s x y -lemma lipschitz_with_smul [normed_space α β] (s : α) : lipschitz_with ∥s∥₊ ((•) s : β → β) := +lemma lipschitz_with_smul [normed_space α β] (s : α) : lipschitz_with ‖s‖₊ ((•) s : β → β) := lipschitz_with_iff_dist_le_mul.2 $ λ x y, by rw [dist_smul, coe_nnnorm] lemma norm_smul_of_nonneg [normed_space ℝ β] {t : ℝ} (ht : 0 ≤ t) (x : β) : - ∥t • x∥ = t * ∥x∥ := by rw [norm_smul, real.norm_eq_abs, abs_of_nonneg ht] + ‖t • x‖ = t * ‖x‖ := by rw [norm_smul, real.norm_eq_abs, abs_of_nonneg ht] variables {E : Type*} [seminormed_add_comm_group E] [normed_space α E] variables {F : Type*} [seminormed_add_comm_group F] [normed_space α F] theorem eventually_nhds_norm_smul_sub_lt (c : α) (x : E) {ε : ℝ} (h : 0 < ε) : - ∀ᶠ y in 𝓝 x, ∥c • (y - x)∥ < ε := -have tendsto (λ y, ∥c • (y - x)∥) (𝓝 x) (𝓝 0), + ∀ᶠ y in 𝓝 x, ‖c • (y - x)‖ < ε := +have tendsto (λ y, ‖c • (y - x)‖) (𝓝 x) (𝓝 0), from ((continuous_id.sub continuous_const).const_smul _).norm.tendsto' _ _ (by simp), this.eventually (gt_mem_nhds h) @@ -166,7 +166,7 @@ begin rcases eq_or_ne e 0 with rfl | he, { rw [add_subgroup.zmultiples_zero_eq_bot], apply_instance, }, { rw [discrete_topology_iff_open_singleton_zero, is_open_induced_iff], - refine ⟨metric.ball 0 (∥e∥), metric.is_open_ball, _⟩, + refine ⟨metric.ball 0 (‖e‖), metric.is_open_ball, _⟩, ext ⟨x, hx⟩, obtain ⟨k, rfl⟩ := add_subgroup.mem_zmultiples_iff.mp hx, rw [mem_preimage, mem_ball_zero_iff, add_subgroup.coe_mk, mem_singleton_iff, @@ -177,7 +177,7 @@ begin end /-- A (semi) normed real vector space is homeomorphic to the unit ball in the same space. -This homeomorphism sends `x : E` to `(1 + ∥x∥²)^(- ½) • x`. +This homeomorphism sends `x : E` to `(1 + ‖x‖²)^(- ½) • x`. In many cases the actual implementation is not important, so we don't mark the projection lemmas `homeomorph_unit_ball_apply_coe` and `homeomorph_unit_ball_symm_apply` as `@[simp]`. @@ -187,35 +187,35 @@ smoothness properties that hold when `E` is an inner-product space. -/ @[simps { attrs := [] }] def homeomorph_unit_ball [normed_space ℝ E] : E ≃ₜ ball (0 : E) 1 := -{ to_fun := λ x, ⟨(1 + ∥x∥^2).sqrt⁻¹ • x, begin - have : 0 < 1 + ∥x∥ ^ 2, by positivity, +{ to_fun := λ x, ⟨(1 + ‖x‖^2).sqrt⁻¹ • x, begin + have : 0 < 1 + ‖x‖ ^ 2, by positivity, rw [mem_ball_zero_iff, norm_smul, real.norm_eq_abs, abs_inv, ← div_eq_inv_mul, div_lt_one (abs_pos.mpr $ real.sqrt_ne_zero'.mpr this), ← abs_norm_eq_norm x, ← sq_lt_sq, abs_norm_eq_norm, real.sq_sqrt this.le], exact lt_one_add _, end⟩, - inv_fun := λ y, (1 - ∥(y : E)∥^2).sqrt⁻¹ • (y : E), + inv_fun := λ y, (1 - ‖(y : E)‖^2).sqrt⁻¹ • (y : E), left_inv := λ x, by field_simp [norm_smul, smul_smul, (zero_lt_one_add_norm_sq x).ne', real.sq_sqrt (zero_lt_one_add_norm_sq x).le, ← real.sqrt_div (zero_lt_one_add_norm_sq x).le], right_inv := λ y, begin - have : 0 < 1 - ∥(y : E)∥ ^ 2 := - by nlinarith [norm_nonneg (y : E), (mem_ball_zero_iff.1 y.2 : ∥(y : E)∥ < 1)], + have : 0 < 1 - ‖(y : E)‖ ^ 2 := + by nlinarith [norm_nonneg (y : E), (mem_ball_zero_iff.1 y.2 : ‖(y : E)‖ < 1)], field_simp [norm_smul, smul_smul, this.ne', real.sq_sqrt this.le, ← real.sqrt_div this.le], end, continuous_to_fun := begin - suffices : continuous (λ x, (1 + ∥x∥^2).sqrt⁻¹), from (this.smul continuous_id).subtype_mk _, + suffices : continuous (λ x, (1 + ‖x‖^2).sqrt⁻¹), from (this.smul continuous_id).subtype_mk _, refine continuous.inv₀ _ (λ x, real.sqrt_ne_zero'.mpr (by positivity)), continuity, end, continuous_inv_fun := begin - suffices : ∀ (y : ball (0 : E) 1), (1 - ∥(y : E)∥ ^ 2).sqrt ≠ 0, { continuity, }, + suffices : ∀ (y : ball (0 : E) 1), (1 - ‖(y : E)‖ ^ 2).sqrt ≠ 0, { continuity, }, intros y, rw real.sqrt_ne_zero', - nlinarith [norm_nonneg (y : E), (mem_ball_zero_iff.1 y.2 : ∥(y : E)∥ < 1)], + nlinarith [norm_nonneg (y : E), (mem_ball_zero_iff.1 y.2 : ‖(y : E)‖ < 1)], end } @[simp] lemma coe_homeomorph_unit_ball_apply_zero [normed_space ℝ E] : @@ -239,8 +239,8 @@ instance prod.normed_space : normed_space α (E × F) := instance pi.normed_space {E : ι → Type*} [fintype ι] [∀i, seminormed_add_comm_group (E i)] [∀i, normed_space α (E i)] : normed_space α (Πi, E i) := { norm_smul_le := λ a f, le_of_eq $ - show (↑(finset.sup finset.univ (λ (b : ι), ∥a • f b∥₊)) : ℝ) = - ∥a∥₊ * ↑(finset.sup finset.univ (λ (b : ι), ∥f b∥₊)), + show (↑(finset.sup finset.univ (λ (b : ι), ‖a • f b‖₊)) : ℝ) = + ‖a‖₊ * ↑(finset.sup finset.univ (λ (b : ι), ‖f b‖₊)), by simp only [(nnreal.coe_mul _ _).symm, nnreal.mul_finset_sup, nnnorm_smul] } /-- A subspace of a normed space is also a normed space, with the restriction of the norm. -/ @@ -250,29 +250,29 @@ instance submodule.normed_space {𝕜 R : Type*} [has_smul 𝕜 R] [normed_field normed_space 𝕜 s := { norm_smul_le := λc x, le_of_eq $ norm_smul c (x : E) } -/-- If there is a scalar `c` with `∥c∥>1`, then any element with nonzero norm can be -moved by scalar multiplication to any shell of width `∥c∥`. Also recap information on the norm of +/-- If there is a scalar `c` with `‖c‖>1`, then any element with nonzero norm can be +moved by scalar multiplication to any shell of width `‖c‖`. Also recap information on the norm of the rescaling element that shows up in applications. -/ -lemma rescale_to_shell_semi_normed {c : α} (hc : 1 < ∥c∥) {ε : ℝ} (εpos : 0 < ε) {x : E} - (hx : ∥x∥ ≠ 0) : ∃d:α, d ≠ 0 ∧ ∥d • x∥ < ε ∧ (ε/∥c∥ ≤ ∥d • x∥) ∧ (∥d∥⁻¹ ≤ ε⁻¹ * ∥c∥ * ∥x∥) := +lemma rescale_to_shell_semi_normed {c : α} (hc : 1 < ‖c‖) {ε : ℝ} (εpos : 0 < ε) {x : E} + (hx : ‖x‖ ≠ 0) : ∃d:α, d ≠ 0 ∧ ‖d • x‖ < ε ∧ (ε/‖c‖ ≤ ‖d • x‖) ∧ (‖d‖⁻¹ ≤ ε⁻¹ * ‖c‖ * ‖x‖) := begin - have xεpos : 0 < ∥x∥/ε := div_pos ((ne.symm hx).le_iff_lt.1 (norm_nonneg x)) εpos, + have xεpos : 0 < ‖x‖/ε := div_pos ((ne.symm hx).le_iff_lt.1 (norm_nonneg x)) εpos, rcases exists_mem_Ico_zpow xεpos hc with ⟨n, hn⟩, - have cpos : 0 < ∥c∥ := lt_trans (zero_lt_one : (0 :ℝ) < 1) hc, - have cnpos : 0 < ∥c^(n+1)∥ := by { rw norm_zpow, exact lt_trans xεpos hn.2 }, + have cpos : 0 < ‖c‖ := lt_trans (zero_lt_one : (0 :ℝ) < 1) hc, + have cnpos : 0 < ‖c^(n+1)‖ := by { rw norm_zpow, exact lt_trans xεpos hn.2 }, refine ⟨(c^(n+1))⁻¹, _, _, _, _⟩, show (c ^ (n + 1))⁻¹ ≠ 0, by rwa [ne.def, inv_eq_zero, ← ne.def, ← norm_pos_iff], - show ∥(c ^ (n + 1))⁻¹ • x∥ < ε, + show ‖(c ^ (n + 1))⁻¹ • x‖ < ε, { rw [norm_smul, norm_inv, ← div_eq_inv_mul, div_lt_iff cnpos, mul_comm, norm_zpow], exact (div_lt_iff εpos).1 (hn.2) }, - show ε / ∥c∥ ≤ ∥(c ^ (n + 1))⁻¹ • x∥, + show ε / ‖c‖ ≤ ‖(c ^ (n + 1))⁻¹ • x‖, { rw [div_le_iff cpos, norm_smul, norm_inv, norm_zpow, zpow_add₀ (ne_of_gt cpos), zpow_one, mul_inv_rev, mul_comm, ← mul_assoc, ← mul_assoc, mul_inv_cancel (ne_of_gt cpos), one_mul, ← div_eq_inv_mul, le_div_iff (zpow_pos_of_pos cpos _), mul_comm], exact (le_div_iff εpos).1 hn.1 }, - show ∥(c ^ (n + 1))⁻¹∥⁻¹ ≤ ε⁻¹ * ∥c∥ * ∥x∥, - { have : ε⁻¹ * ∥c∥ * ∥x∥ = ε⁻¹ * ∥x∥ * ∥c∥, by ring, + show ‖(c ^ (n + 1))⁻¹‖⁻¹ ≤ ε⁻¹ * ‖c‖ * ‖x‖, + { have : ε⁻¹ * ‖c‖ * ‖x‖ = ε⁻¹ * ‖x‖ * ‖c‖, by ring, rw [norm_inv, inv_inv, norm_zpow, zpow_add₀ (ne_of_gt cpos), zpow_one, this, ← div_eq_inv_mul], exact mul_le_mul_of_nonneg_right hn.1 (norm_nonneg _) } end @@ -318,11 +318,11 @@ section surj variables (E) [normed_space ℝ E] [nontrivial E] -lemma exists_norm_eq {c : ℝ} (hc : 0 ≤ c) : ∃ x : E, ∥x∥ = c := +lemma exists_norm_eq {c : ℝ} (hc : 0 ≤ c) : ∃ x : E, ‖x‖ = c := begin rcases exists_ne (0 : E) with ⟨x, hx⟩, rw ← norm_ne_zero_iff at hx, - use c • ∥x∥⁻¹ • x, + use c • ‖x‖⁻¹ • x, simp [norm_smul, real.norm_of_nonneg hc, hx] end @@ -351,11 +351,11 @@ by rw [frontier, closure_closed_ball, interior_closed_ball' x r, closed_ball_dif variables {α} -/-- If there is a scalar `c` with `∥c∥>1`, then any element can be moved by scalar multiplication to -any shell of width `∥c∥`. Also recap information on the norm of the rescaling element that shows +/-- If there is a scalar `c` with `‖c‖>1`, then any element can be moved by scalar multiplication to +any shell of width `‖c‖`. Also recap information on the norm of the rescaling element that shows up in applications. -/ -lemma rescale_to_shell {c : α} (hc : 1 < ∥c∥) {ε : ℝ} (εpos : 0 < ε) {x : E} (hx : x ≠ 0) : - ∃d:α, d ≠ 0 ∧ ∥d • x∥ < ε ∧ (ε/∥c∥ ≤ ∥d • x∥) ∧ (∥d∥⁻¹ ≤ ε⁻¹ * ∥c∥ * ∥x∥) := +lemma rescale_to_shell {c : α} (hc : 1 < ‖c‖) {ε : ℝ} (εpos : 0 < ε) {x : E} (hx : x ≠ 0) : + ∃d:α, d ≠ 0 ∧ ‖d • x‖ < ε ∧ (ε/‖c‖ ≤ ‖d • x‖) ∧ (‖d‖⁻¹ ≤ ε⁻¹ * ‖c‖ * ‖x‖) := rescale_to_shell_semi_normed hc εpos (ne_of_lt (norm_pos_iff.2 hx)).symm end normed_add_comm_group @@ -369,10 +369,10 @@ include 𝕜 /-- If `E` is a nontrivial normed space over a nontrivially normed field `𝕜`, then `E` is unbounded: for any `c : ℝ`, there exists a vector `x : E` with norm strictly greater than `c`. -/ -lemma normed_space.exists_lt_norm (c : ℝ) : ∃ x : E, c < ∥x∥ := +lemma normed_space.exists_lt_norm (c : ℝ) : ∃ x : E, c < ‖x‖ := begin rcases exists_ne (0 : E) with ⟨x, hx⟩, - rcases normed_field.exists_lt_norm 𝕜 (c / ∥x∥) with ⟨r, hr⟩, + rcases normed_field.exists_lt_norm 𝕜 (c / ‖x‖) with ⟨r, hr⟩, use r • x, rwa [norm_smul, ← div_lt_iff], rwa norm_pos_iff @@ -413,7 +413,7 @@ variables [normed_module 𝕜 𝕜'] [smul_comm_class 𝕜 𝕜' 𝕜'] [is_scal -/ class normed_algebra (𝕜 : Type*) (𝕜' : Type*) [normed_field 𝕜] [semi_normed_ring 𝕜'] extends algebra 𝕜 𝕜' := -(norm_smul_le : ∀ (r : 𝕜) (x : 𝕜'), ∥r • x∥ ≤ ∥r∥ * ∥x∥) +(norm_smul_le : ∀ (r : 𝕜) (x : 𝕜'), ‖r • x‖ ≤ ‖r‖ * ‖x‖) variables {𝕜 : Type*} (𝕜' : Type*) [normed_field 𝕜] [semi_normed_ring 𝕜'] [normed_algebra 𝕜 𝕜'] @@ -437,29 +437,29 @@ See `normed_space.to_module'` for a similar situation. -/ instance normed_algebra.to_normed_space' {𝕜'} [normed_ring 𝕜'] [normed_algebra 𝕜 𝕜'] : normed_space 𝕜 𝕜' := by apply_instance -lemma norm_algebra_map (x : 𝕜) : ∥algebra_map 𝕜 𝕜' x∥ = ∥x∥ * ∥(1 : 𝕜')∥ := +lemma norm_algebra_map (x : 𝕜) : ‖algebra_map 𝕜 𝕜' x‖ = ‖x‖ * ‖(1 : 𝕜')‖ := begin rw algebra.algebra_map_eq_smul_one, exact norm_smul _ _, end -lemma nnnorm_algebra_map (x : 𝕜) : ∥algebra_map 𝕜 𝕜' x∥₊ = ∥x∥₊ * ∥(1 : 𝕜')∥₊ := +lemma nnnorm_algebra_map (x : 𝕜) : ‖algebra_map 𝕜 𝕜' x‖₊ = ‖x‖₊ * ‖(1 : 𝕜')‖₊ := subtype.ext $ norm_algebra_map 𝕜' x -@[simp] lemma norm_algebra_map' [norm_one_class 𝕜'] (x : 𝕜) : ∥algebra_map 𝕜 𝕜' x∥ = ∥x∥ := +@[simp] lemma norm_algebra_map' [norm_one_class 𝕜'] (x : 𝕜) : ‖algebra_map 𝕜 𝕜' x‖ = ‖x‖ := by rw [norm_algebra_map, norm_one, mul_one] -@[simp] lemma nnnorm_algebra_map' [norm_one_class 𝕜'] (x : 𝕜) : ∥algebra_map 𝕜 𝕜' x∥₊ = ∥x∥₊ := +@[simp] lemma nnnorm_algebra_map' [norm_one_class 𝕜'] (x : 𝕜) : ‖algebra_map 𝕜 𝕜' x‖₊ = ‖x‖₊ := subtype.ext $ norm_algebra_map' _ _ section nnreal variables [norm_one_class 𝕜'] [normed_algebra ℝ 𝕜'] -@[simp] lemma norm_algebra_map_nnreal (x : ℝ≥0) : ∥algebra_map ℝ≥0 𝕜' x∥ = x := +@[simp] lemma norm_algebra_map_nnreal (x : ℝ≥0) : ‖algebra_map ℝ≥0 𝕜' x‖ = x := (norm_algebra_map' 𝕜' (x : ℝ)).symm ▸ real.norm_of_nonneg x.prop -@[simp] lemma nnnorm_algebra_map_nnreal (x : ℝ≥0) : ∥algebra_map ℝ≥0 𝕜' x∥₊ = x := +@[simp] lemma nnnorm_algebra_map_nnreal (x : ℝ≥0) : ‖algebra_map ℝ≥0 𝕜' x‖₊ = x := subtype.ext $ norm_algebra_map_nnreal 𝕜' x end nnreal diff --git a/src/analysis/normed_space/bounded_linear_maps.lean b/src/analysis/normed_space/bounded_linear_maps.lean index 01be37cf60501..af8cbd4355ee5 100644 --- a/src/analysis/normed_space/bounded_linear_maps.lean +++ b/src/analysis/normed_space/bounded_linear_maps.lean @@ -13,16 +13,16 @@ import analysis.asymptotics.asymptotics This file defines a class stating that a map between normed vector spaces is (bi)linear and continuous. Instead of asking for continuity, the definition takes the equivalent condition (because the space -is normed) that `∥f x∥` is bounded by a multiple of `∥x∥`. Hence the "bounded" in the name refers to -`∥f x∥/∥x∥` rather than `∥f x∥` itself. +is normed) that `‖f x‖` is bounded by a multiple of `‖x‖`. Hence the "bounded" in the name refers to +`‖f x‖/‖x‖` rather than `‖f x‖` itself. ## Main definitions -* `is_bounded_linear_map`: Class stating that a map `f : E → F` is linear and has `∥f x∥` bounded - by a multiple of `∥x∥`. +* `is_bounded_linear_map`: Class stating that a map `f : E → F` is linear and has `‖f x‖` bounded + by a multiple of `‖x‖`. * `is_bounded_bilinear_map`: Class stating that a map `f : E × F → G` is bilinear and continuous, - but through the simpler to provide statement that `∥f (x, y)∥` is bounded by a multiple of - `∥x∥ * ∥y∥` + but through the simpler to provide statement that `‖f (x, y)‖` is bounded by a multiple of + `‖x‖ * ‖y‖` * `is_bounded_bilinear_map.linear_deriv`: Derivative of a continuous bilinear map as a linear map. * `is_bounded_bilinear_map.deriv`: Derivative of a continuous bilinear map as a continuous linear map. The proof that it is indeed the derivative is `is_bounded_bilinear_map.has_fderiv_at` in @@ -59,15 +59,15 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] /-- A function `f` satisfies `is_bounded_linear_map 𝕜 f` if it is linear and satisfies the -inequality `∥f x∥ ≤ M * ∥x∥` for some positive constant `M`. -/ +inequality `‖f x‖ ≤ M * ‖x‖` for some positive constant `M`. -/ structure is_bounded_linear_map (𝕜 : Type*) [normed_field 𝕜] {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] (f : E → F) extends is_linear_map 𝕜 f : Prop := -(bound : ∃ M, 0 < M ∧ ∀ x : E, ∥f x∥ ≤ M * ∥x∥) +(bound : ∃ M, 0 < M ∧ ∀ x : E, ‖f x‖ ≤ M * ‖x‖) lemma is_linear_map.with_bound - {f : E → F} (hf : is_linear_map 𝕜 f) (M : ℝ) (h : ∀ x : E, ∥f x∥ ≤ M * ∥x∥) : + {f : E → F} (hf : is_linear_map 𝕜 f) (M : ℝ) (h : ∀ x : E, ‖f x‖ ≤ M * ‖x‖) : is_bounded_linear_map 𝕜 f := ⟨ hf, classical.by_cases (assume : M ≤ 0, ⟨1, zero_lt_one, λ x, @@ -116,10 +116,10 @@ variables {f g : E → F} lemma smul (c : 𝕜) (hf : is_bounded_linear_map 𝕜 f) : is_bounded_linear_map 𝕜 (c • f) := let ⟨hlf, M, hMp, hM⟩ := hf in -(c • hlf.mk' f).is_linear.with_bound (∥c∥ * M) $ λ x, - calc ∥c • f x∥ = ∥c∥ * ∥f x∥ : norm_smul c (f x) - ... ≤ ∥c∥ * (M * ∥x∥) : mul_le_mul_of_nonneg_left (hM _) (norm_nonneg _) - ... = (∥c∥ * M) * ∥x∥ : (mul_assoc _ _ _).symm +(c • hlf.mk' f).is_linear.with_bound (‖c‖ * M) $ λ x, + calc ‖c • f x‖ = ‖c‖ * ‖f x‖ : norm_smul c (f x) + ... ≤ ‖c‖ * (M * ‖x‖) : mul_le_mul_of_nonneg_left (hM _) (norm_nonneg _) + ... = (‖c‖ * M) * ‖x‖ : (mul_assoc _ _ _).symm lemma neg (hf : is_bounded_linear_map 𝕜 f) : is_bounded_linear_map 𝕜 (λ e, -f e) := @@ -133,8 +133,8 @@ lemma add (hf : is_bounded_linear_map 𝕜 f) (hg : is_bounded_linear_map 𝕜 g let ⟨hlf, Mf, hMfp, hMf⟩ := hf in let ⟨hlg, Mg, hMgp, hMg⟩ := hg in (hlf.mk' _ + hlg.mk' _).is_linear.with_bound (Mf + Mg) $ λ x, - calc ∥f x + g x∥ ≤ Mf * ∥x∥ + Mg * ∥x∥ : norm_add_le_of_le (hMf x) (hMg x) - ... ≤ (Mf + Mg) * ∥x∥ : by rw add_mul + calc ‖f x + g x‖ ≤ Mf * ‖x‖ + Mg * ‖x‖ : norm_add_le_of_le (hMf x) (hMg x) + ... ≤ (Mf + Mg) * ‖x‖ : by rw add_mul lemma sub (hf : is_bounded_linear_map 𝕜 f) (hg : is_bounded_linear_map 𝕜 g) : is_bounded_linear_map 𝕜 (λ e, f e - g e) := @@ -151,9 +151,9 @@ let ⟨hf, M, hMp, hM⟩ := hf in tendsto_iff_norm_tendsto_zero.2 $ squeeze_zero (λ e, norm_nonneg _) (λ e, - calc ∥f e - f x∥ = ∥hf.mk' f (e - x)∥ : by rw (hf.mk' _).map_sub e x; refl - ... ≤ M * ∥e - x∥ : hM (e - x)) - (suffices tendsto (λ (e : E), M * ∥e - x∥) (𝓝 x) (𝓝 (M * 0)), by simpa, + calc ‖f e - f x‖ = ‖hf.mk' f (e - x)‖ : by rw (hf.mk' _).map_sub e x; refl + ... ≤ M * ‖e - x‖ : hM (e - x)) + (suffices tendsto (λ (e : E), M * ‖e - x‖) (𝓝 x) (𝓝 (M * 0)), by simpa, tendsto_const_nhds.mul (tendsto_norm_sub_self _)) lemma continuous (hf : is_bounded_linear_map 𝕜 f) : continuous f := @@ -211,16 +211,16 @@ lemma is_bounded_linear_map_continuous_multilinear_map_comp_linear (g : G →L[ f.comp_continuous_linear_map (λ _, g)) := begin refine is_linear_map.with_bound ⟨λ f₁ f₂, by { ext m, refl }, λ c f, by { ext m, refl }⟩ - (∥g∥ ^ (fintype.card ι)) (λ f, _), + (‖g‖ ^ (fintype.card ι)) (λ f, _), apply continuous_multilinear_map.op_norm_le_bound _ _ (λ m, _), { apply_rules [mul_nonneg, pow_nonneg, norm_nonneg] }, - calc ∥f (g ∘ m)∥ ≤ - ∥f∥ * ∏ i, ∥g (m i)∥ : f.le_op_norm _ - ... ≤ ∥f∥ * ∏ i, (∥g∥ * ∥m i∥) : begin + calc ‖f (g ∘ m)‖ ≤ + ‖f‖ * ∏ i, ‖g (m i)‖ : f.le_op_norm _ + ... ≤ ‖f‖ * ∏ i, (‖g‖ * ‖m i‖) : begin apply mul_le_mul_of_nonneg_left _ (norm_nonneg _), exact finset.prod_le_prod (λ i hi, norm_nonneg _) (λ i hi, g.le_op_norm _) end - ... = ∥g∥ ^ fintype.card ι * ∥f∥ * ∏ i, ∥m i∥ : + ... = ‖g‖ ^ fintype.card ι * ‖f‖ * ∏ i, ‖m i‖ : by { simp [finset.prod_mul_distrib, finset.card_univ], ring } end @@ -288,7 +288,7 @@ structure is_bounded_bilinear_map (f : E × F → G) : Prop := (smul_left : ∀ (c : 𝕜) (x : E) (y : F), f (c • x, y) = c • f (x, y)) (add_right : ∀ (x : E) (y₁ y₂ : F), f (x, y₁ + y₂) = f (x, y₁) + f (x, y₂)) (smul_right : ∀ (c : 𝕜) (x : E) (y : F), f (x, c • y) = c • f (x,y)) -(bound : ∃ C > 0, ∀ (x : E) (y : F), ∥f (x, y)∥ ≤ C * ∥x∥ * ∥y∥) +(bound : ∃ C > 0, ∀ (x : E) (y : F), ‖f (x, y)‖ ≤ C * ‖x‖ * ‖y‖) variable {𝕜} variable {f : E × F → G} @@ -299,22 +299,22 @@ lemma continuous_linear_map.is_bounded_bilinear_map (f : E →L[𝕜] F →L[ smul_left := f.map_smul₂, add_right := λ x, (f x).map_add, smul_right := λ c x, (f x).map_smul c, - bound := ⟨max ∥f∥ 1, zero_lt_one.trans_le (le_max_right _ _), + bound := ⟨max ‖f‖ 1, zero_lt_one.trans_le (le_max_right _ _), λ x y, (f.le_op_norm₂ x y).trans $ by apply_rules [mul_le_mul_of_nonneg_right, norm_nonneg, le_max_left]⟩ } protected lemma is_bounded_bilinear_map.is_O (h : is_bounded_bilinear_map 𝕜 f) : - f =O[⊤] (λ p : E × F, ∥p.1∥ * ∥p.2∥) := + f =O[⊤] (λ p : E × F, ‖p.1‖ * ‖p.2‖) := let ⟨C, Cpos, hC⟩ := h.bound in asymptotics.is_O.of_bound _ $ filter.eventually_of_forall $ λ ⟨x, y⟩, by simpa [mul_assoc] using hC x y lemma is_bounded_bilinear_map.is_O_comp {α : Type*} (H : is_bounded_bilinear_map 𝕜 f) {g : α → E} {h : α → F} {l : filter α} : - (λ x, f (g x, h x)) =O[l] (λ x, ∥g x∥ * ∥h x∥) := + (λ x, f (g x, h x)) =O[l] (λ x, ‖g x‖ * ‖h x‖) := H.is_O.comp_tendsto le_top protected lemma is_bounded_bilinear_map.is_O' (h : is_bounded_bilinear_map 𝕜 f) : - f =O[⊤] (λ p : E × F, ∥p∥ * ∥p∥) := + f =O[⊤] (λ p : E × F, ‖p‖ * ‖p‖) := h.is_O.trans $ (@asymptotics.is_O_fst_prod' _ E F _ _ _ _).norm_norm.mul (@asymptotics.is_O_snd_prod' _ E F _ _ _ _).norm_norm @@ -338,7 +338,7 @@ begin obtain ⟨C, (Cpos : 0 < C), hC⟩ := h.bound, rw continuous_iff_continuous_at, intros x, - have H : ∀ (a:E) (b:F), ∥f (a, b)∥ ≤ C * ∥∥a∥ * ∥b∥∥, + have H : ∀ (a:E) (b:F), ‖f (a, b)‖ ≤ C * ‖‖a‖ * ‖b‖‖, { intros a b, simpa [mul_assoc] using hC a b }, have h₁ : (λ e : E × F, f (e.1 - x.1, e.2)) =o[𝓝 x] (λ e, (1:ℝ)), @@ -381,12 +381,12 @@ lemma is_bounded_bilinear_map.is_bounded_linear_map_left (h : is_bounded_bilinea map_smul := λ c x, h.smul_left _ _ _, bound := begin rcases h.bound with ⟨C, C_pos, hC⟩, - refine ⟨C * (∥y∥ + 1), mul_pos C_pos (lt_of_lt_of_le (zero_lt_one) (by simp)), λ x, _⟩, - have : ∥y∥ ≤ ∥y∥ + 1, by simp [zero_le_one], - calc ∥f (x, y)∥ ≤ C * ∥x∥ * ∥y∥ : hC x y - ... ≤ C * ∥x∥ * (∥y∥ + 1) : + refine ⟨C * (‖y‖ + 1), mul_pos C_pos (lt_of_lt_of_le (zero_lt_one) (by simp)), λ x, _⟩, + have : ‖y‖ ≤ ‖y‖ + 1, by simp [zero_le_one], + calc ‖f (x, y)‖ ≤ C * ‖x‖ * ‖y‖ : hC x y + ... ≤ C * ‖x‖ * (‖y‖ + 1) : by apply_rules [norm_nonneg, mul_le_mul_of_nonneg_left, le_of_lt C_pos, mul_nonneg] - ... = C * (∥y∥ + 1) * ∥x∥ : by ring + ... = C * (‖y‖ + 1) * ‖x‖ : by ring end } lemma is_bounded_bilinear_map.is_bounded_linear_map_right @@ -396,10 +396,10 @@ lemma is_bounded_bilinear_map.is_bounded_linear_map_right map_smul := λ c y, h.smul_right _ _ _, bound := begin rcases h.bound with ⟨C, C_pos, hC⟩, - refine ⟨C * (∥x∥ + 1), mul_pos C_pos (lt_of_lt_of_le (zero_lt_one) (by simp)), λ y, _⟩, - have : ∥x∥ ≤ ∥x∥ + 1, by simp [zero_le_one], - calc ∥f (x, y)∥ ≤ C * ∥x∥ * ∥y∥ : hC x y - ... ≤ C * (∥x∥ + 1) * ∥y∥ : + refine ⟨C * (‖x‖ + 1), mul_pos C_pos (lt_of_lt_of_le (zero_lt_one) (by simp)), λ y, _⟩, + have : ‖x‖ ≤ ‖x‖ + 1, by simp [zero_le_one], + calc ‖f (x, y)‖ ≤ C * ‖x‖ * ‖y‖ : hC x y + ... ≤ C * (‖x‖ + 1) * ‖y‖ : by apply_rules [mul_le_mul_of_nonneg_right, norm_nonneg, mul_le_mul_of_nonneg_left, le_of_lt C_pos] end } @@ -470,17 +470,17 @@ from `E × F` to `G`. The statement that this is indeed the derivative of `f` is def is_bounded_bilinear_map.deriv (h : is_bounded_bilinear_map 𝕜 f) (p : E × F) : E × F →L[𝕜] G := (h.linear_deriv p).mk_continuous_of_exists_bound $ begin rcases h.bound with ⟨C, Cpos, hC⟩, - refine ⟨C * ∥p.1∥ + C * ∥p.2∥, λ q, _⟩, - calc ∥f (p.1, q.2) + f (q.1, p.2)∥ - ≤ C * ∥p.1∥ * ∥q.2∥ + C * ∥q.1∥ * ∥p.2∥ : norm_add_le_of_le (hC _ _) (hC _ _) - ... ≤ C * ∥p.1∥ * ∥q∥ + C * ∥q∥ * ∥p.2∥ : begin + refine ⟨C * ‖p.1‖ + C * ‖p.2‖, λ q, _⟩, + calc ‖f (p.1, q.2) + f (q.1, p.2)‖ + ≤ C * ‖p.1‖ * ‖q.2‖ + C * ‖q.1‖ * ‖p.2‖ : norm_add_le_of_le (hC _ _) (hC _ _) + ... ≤ C * ‖p.1‖ * ‖q‖ + C * ‖q‖ * ‖p.2‖ : begin apply add_le_add, exact mul_le_mul_of_nonneg_left (le_max_right _ _) (mul_nonneg (le_of_lt Cpos) (norm_nonneg _)), apply mul_le_mul_of_nonneg_right _ (norm_nonneg _), exact mul_le_mul_of_nonneg_left (le_max_left _ _) (le_of_lt Cpos), end - ... = (C * ∥p.1∥ + C * ∥p.2∥) * ∥q∥ : by ring + ... = (C * ‖p.1‖ + C * ‖p.2‖) * ‖q‖ : by ring end @[simp] lemma is_bounded_bilinear_map_deriv_coe (h : is_bounded_bilinear_map 𝕜 f) (p q : E × F) : @@ -510,11 +510,11 @@ begin is_bounded_bilinear_map_deriv_coe, prod.smul_fst, prod.smul_snd, coe_smul', pi.smul_apply] }, { refine continuous_linear_map.op_norm_le_bound _ (mul_nonneg (add_nonneg Cpos.le Cpos.le) (norm_nonneg _)) (λ q, _), - calc ∥f (p.1, q.2) + f (q.1, p.2)∥ - ≤ C * ∥p.1∥ * ∥q.2∥ + C * ∥q.1∥ * ∥p.2∥ : norm_add_le_of_le (hC _ _) (hC _ _) - ... ≤ C * ∥p∥ * ∥q∥ + C * ∥q∥ * ∥p∥ : by apply_rules [add_le_add, mul_le_mul, norm_nonneg, + calc ‖f (p.1, q.2) + f (q.1, p.2)‖ + ≤ C * ‖p.1‖ * ‖q.2‖ + C * ‖q.1‖ * ‖p.2‖ : norm_add_le_of_le (hC _ _) (hC _ _) + ... ≤ C * ‖p‖ * ‖q‖ + C * ‖q‖ * ‖p‖ : by apply_rules [add_le_add, mul_le_mul, norm_nonneg, Cpos.le, le_refl, le_max_left, le_max_right, mul_nonneg] - ... = (C + C) * ∥p∥ * ∥q∥ : by ring }, + ... = (C + C) * ‖p‖ * ‖q‖ : by ring }, end end bilinear_map diff --git a/src/analysis/normed_space/compact_operator.lean b/src/analysis/normed_space/compact_operator.lean index a643d74a688ce..a2e78130ad594 100644 --- a/src/analysis/normed_space/compact_operator.lean +++ b/src/analysis/normed_space/compact_operator.lean @@ -332,9 +332,9 @@ begin -- neighborhood of `0` in `M₁`. rcases hf with ⟨K, hK, hKf⟩, -- But any compact set is totally bounded, hence Von-Neumann bounded. Thus, `K` absorbs `U`. - -- This gives `r > 0` such that `∀ a : 𝕜₂, r ≤ ∥a∥ → K ⊆ a • U`. + -- This gives `r > 0` such that `∀ a : 𝕜₂, r ≤ ‖a‖ → K ⊆ a • U`. rcases hK.totally_bounded.is_vonN_bounded 𝕜₂ hU with ⟨r, hr, hrU⟩, - -- Choose `c : 𝕜₂` with `r < ∥c∥`. + -- Choose `c : 𝕜₂` with `r < ‖c‖`. rcases normed_field.exists_lt_norm 𝕜₁ r with ⟨c, hc⟩, have hcnz : c ≠ 0 := ne_zero_of_norm_ne_zero (hr.trans hc).ne.symm, -- We have `f ⁻¹' ((σ₁₂ c⁻¹) • K) = c⁻¹ • f ⁻¹' K ∈ 𝓝 0`. Thus, showing that @@ -346,8 +346,8 @@ begin apply_instance }, -- Since `σ₁₂ c⁻¹` = `(σ₁₂ c)⁻¹`, we have to prove that `K ⊆ σ₁₂ c • U`. rw [map_inv₀, ← subset_set_smul_iff₀ ((map_ne_zero σ₁₂).mpr hcnz)], - -- But `σ₁₂` is isometric, so `∥σ₁₂ c∥ = ∥c∥ > r`, which concludes the argument since - -- `∀ a : 𝕜₂, r ≤ ∥a∥ → K ⊆ a • U`. + -- But `σ₁₂` is isometric, so `‖σ₁₂ c‖ = ‖c‖ > r`, which concludes the argument since + -- `∀ a : 𝕜₂, r ≤ ‖a‖ → K ⊆ a • U`. refine hrU (σ₁₂ c) _, rw ring_hom_isometric.is_iso, exact hc.le @@ -407,9 +407,9 @@ begin linarith [dist_triangle (u x) (v x) t] }, rw mem_closed_ball_zero_iff at hx, calc dist (u x) (v x) - = ∥u x - v x∥ : dist_eq_norm _ _ - ... = ∥(u - v) x∥ : by rw continuous_linear_map.sub_apply; refl - ... ≤ ∥u - v∥ : (u - v).unit_le_op_norm x hx + = ‖u x - v x‖ : dist_eq_norm _ _ + ... = ‖(u - v) x‖ : by rw continuous_linear_map.sub_apply; refl + ... ≤ ‖u - v‖ : (u - v).unit_le_op_norm x hx ... = dist u v : (dist_eq_norm _ _).symm ... < ε/2 : huv end diff --git a/src/analysis/normed_space/completion.lean b/src/analysis/normed_space/completion.lean index 0715c852b22ce..11b159a3fcf99 100644 --- a/src/analysis/normed_space/completion.lean +++ b/src/analysis/normed_space/completion.lean @@ -56,7 +56,7 @@ to_complₗᵢ.to_continuous_linear_map @[simp] lemma coe_to_complL : ⇑(to_complL : E →L[𝕜] completion E) = coe := rfl @[simp] lemma norm_to_complL {𝕜 E : Type*} [nontrivially_normed_field 𝕜] [normed_add_comm_group E] - [normed_space 𝕜 E] [nontrivial E] : ∥(to_complL : E →L[𝕜] completion E)∥ = 1 := + [normed_space 𝕜 E] [nontrivial E] : ‖(to_complL : E →L[𝕜] completion E)‖ = 1 := (to_complₗᵢ : E →ₗᵢ[𝕜] completion E).norm_to_continuous_linear_map section algebra diff --git a/src/analysis/normed_space/continuous_affine_map.lean b/src/analysis/normed_space/continuous_affine_map.lean index d654bc2508802..61513ff3a0f20 100644 --- a/src/analysis/normed_space/continuous_affine_map.lean +++ b/src/analysis/normed_space/continuous_affine_map.lean @@ -15,7 +15,7 @@ spaces. In the particular case that the affine spaces are just normed vector spaces `V`, `W`, we define a norm on the space of continuous affine maps by defining the norm of `f : V →A[𝕜] W` to be -`∥f∥ = max ∥f 0∥ ∥f.cont_linear∥`. This is chosen so that we have a linear isometry: +`‖f‖ = max ‖f 0‖ ‖f.cont_linear‖`. This is chosen so that we have a linear isometry: `(V →A[𝕜] W) ≃ₗᵢ[𝕜] W × (V →L[𝕜] W)`. The abstract picture is that for an affine space `P` modelled on a vector space `V`, together with @@ -27,7 +27,7 @@ take `P = V`, using `0 : V` as the base point provides a splitting, and we prove isometric decomposition. On the other hand, choosing a base point breaks the affine invariance so the norm fails to be -submultiplicative: for a composition of maps, we have only `∥f.comp g∥ ≤ ∥f∥ * ∥g∥ + ∥f 0∥`. +submultiplicative: for a composition of maps, we have only `‖f.comp g‖ ≤ ‖f‖ * ‖g‖ + ‖f 0‖`. ## Main definitions: @@ -143,23 +143,23 @@ section normed_space_structure variables (f : V →A[𝕜] W) /-- Note that unlike the operator norm for linear maps, this norm is _not_ submultiplicative: -we do _not_ necessarily have `∥f.comp g∥ ≤ ∥f∥ * ∥g∥`. See `norm_comp_le` for what we can say. -/ -noncomputable instance has_norm : has_norm (V →A[𝕜] W) := ⟨λ f, max ∥f 0∥ ∥f.cont_linear∥⟩ +we do _not_ necessarily have `‖f.comp g‖ ≤ ‖f‖ * ‖g‖`. See `norm_comp_le` for what we can say. -/ +noncomputable instance has_norm : has_norm (V →A[𝕜] W) := ⟨λ f, max ‖f 0‖ ‖f.cont_linear‖⟩ -lemma norm_def : ∥f∥ = (max ∥f 0∥ ∥f.cont_linear∥) := rfl +lemma norm_def : ‖f‖ = (max ‖f 0‖ ‖f.cont_linear‖) := rfl -lemma norm_cont_linear_le : ∥f.cont_linear∥ ≤ ∥f∥ := le_max_right _ _ +lemma norm_cont_linear_le : ‖f.cont_linear‖ ≤ ‖f‖ := le_max_right _ _ -lemma norm_image_zero_le : ∥f 0∥ ≤ ∥f∥ := le_max_left _ _ +lemma norm_image_zero_le : ‖f 0‖ ≤ ‖f‖ := le_max_left _ _ -@[simp] lemma norm_eq (h : f 0 = 0) : ∥f∥ = ∥f.cont_linear∥ := -calc ∥f∥ = (max ∥f 0∥ ∥f.cont_linear∥) : by rw norm_def - ... = (max 0 ∥f.cont_linear∥) : by rw [h, norm_zero] - ... = ∥f.cont_linear∥ : max_eq_right (norm_nonneg _) +@[simp] lemma norm_eq (h : f 0 = 0) : ‖f‖ = ‖f.cont_linear‖ := +calc ‖f‖ = (max ‖f 0‖ ‖f.cont_linear‖) : by rw norm_def + ... = (max 0 ‖f.cont_linear‖) : by rw [h, norm_zero] + ... = ‖f.cont_linear‖ : max_eq_right (norm_nonneg _) noncomputable instance : normed_add_comm_group (V →A[𝕜] W) := add_group_norm.to_normed_add_comm_group -{ to_fun := λ f, max ∥f 0∥ ∥f.cont_linear∥, +{ to_fun := λ f, max ‖f 0‖ ‖f.cont_linear‖, map_zero' := by simp, neg' := λ f, by simp, add_le' := λ f g, begin @@ -187,23 +187,23 @@ instance : normed_space 𝕜 (V →A[𝕜] W) := norm_smul, ← mul_max_of_nonneg _ _ (norm_nonneg t)], } lemma norm_comp_le (g : W₂ →A[𝕜] V) : - ∥f.comp g∥ ≤ ∥f∥ * ∥g∥ + ∥f 0∥ := + ‖f.comp g‖ ≤ ‖f‖ * ‖g‖ + ‖f 0‖ := begin rw [norm_def, max_le_iff], split, - { calc ∥f.comp g 0∥ = ∥f (g 0)∥ : by simp - ... = ∥f.cont_linear (g 0) + f 0∥ : by { rw f.decomp, simp, } - ... ≤ ∥f.cont_linear∥ * ∥g 0∥ + ∥f 0∥ : + { calc ‖f.comp g 0‖ = ‖f (g 0)‖ : by simp + ... = ‖f.cont_linear (g 0) + f 0‖ : by { rw f.decomp, simp, } + ... ≤ ‖f.cont_linear‖ * ‖g 0‖ + ‖f 0‖ : (norm_add_le _ _).trans (add_le_add_right (f.cont_linear.le_op_norm _) _) - ... ≤ ∥f∥ * ∥g∥ + ∥f 0∥ : + ... ≤ ‖f‖ * ‖g‖ + ‖f 0‖ : add_le_add_right (mul_le_mul f.norm_cont_linear_le g.norm_image_zero_le (norm_nonneg _) (norm_nonneg _)) _, }, - { calc ∥(f.comp g).cont_linear∥ ≤ ∥f.cont_linear∥ * ∥g.cont_linear∥ : + { calc ‖(f.comp g).cont_linear‖ ≤ ‖f.cont_linear‖ * ‖g.cont_linear‖ : (g.comp_cont_linear f).symm ▸ f.cont_linear.op_norm_comp_le _ - ... ≤ ∥f∥ * ∥g∥ : + ... ≤ ‖f‖ * ‖g‖ : mul_le_mul f.norm_cont_linear_le g.norm_cont_linear_le (norm_nonneg _) (norm_nonneg _) - ... ≤ ∥f∥ * ∥g∥ + ∥f 0∥ : + ... ≤ ‖f‖ * ‖g‖ + ‖f 0‖ : by { rw le_add_iff_nonneg_right, apply norm_nonneg, }, }, end diff --git a/src/analysis/normed_space/dual.lean b/src/analysis/normed_space/dual.lean index bdd00bb90571d..230af5dd29c23 100644 --- a/src/analysis/normed_space/dual.lean +++ b/src/analysis/normed_space/dual.lean @@ -26,7 +26,7 @@ theory for `seminormed_add_comm_group` and we specialize to `normed_add_comm_gro * `inclusion_in_double_dual` and `inclusion_in_double_dual_li` are the inclusion of a normed space in its double dual, considered as a bounded linear map and as a linear isometry, respectively. * `polar 𝕜 s` is the subset of `dual 𝕜 E` consisting of those functionals `x'` for which - `∥x' z∥ ≤ 1` for every `z ∈ s`. + `‖x' z‖ ≤ 1` for every `z ∈ s`. ## Tags @@ -65,13 +65,13 @@ continuous_linear_map.apply 𝕜 𝕜 @[simp] lemma dual_def (x : E) (f : dual 𝕜 E) : inclusion_in_double_dual 𝕜 E x f = f x := rfl lemma inclusion_in_double_dual_norm_eq : - ∥inclusion_in_double_dual 𝕜 E∥ = ∥(continuous_linear_map.id 𝕜 (dual 𝕜 E))∥ := + ‖inclusion_in_double_dual 𝕜 E‖ = ‖(continuous_linear_map.id 𝕜 (dual 𝕜 E))‖ := continuous_linear_map.op_norm_flip _ -lemma inclusion_in_double_dual_norm_le : ∥inclusion_in_double_dual 𝕜 E∥ ≤ 1 := +lemma inclusion_in_double_dual_norm_le : ‖inclusion_in_double_dual 𝕜 E‖ ≤ 1 := by { rw inclusion_in_double_dual_norm_eq, exact continuous_linear_map.norm_id_le } -lemma double_dual_bound (x : E) : ∥(inclusion_in_double_dual 𝕜 E) x∥ ≤ ∥x∥ := +lemma double_dual_bound (x : E) : ‖(inclusion_in_double_dual 𝕜 E) x‖ ≤ ‖x‖ := by simpa using continuous_linear_map.le_of_op_norm_le _ (inclusion_in_double_dual_norm_le 𝕜 E) x /-- The dual pairing as a bilinear form. -/ @@ -94,16 +94,16 @@ variables (𝕜 : Type v) [is_R_or_C 𝕜] /-- If one controls the norm of every `f x`, then one controls the norm of `x`. Compare `continuous_linear_map.op_norm_le_bound`. -/ -lemma norm_le_dual_bound (x : E) {M : ℝ} (hMp: 0 ≤ M) (hM : ∀ (f : dual 𝕜 E), ∥f x∥ ≤ M * ∥f∥) : - ∥x∥ ≤ M := +lemma norm_le_dual_bound (x : E) {M : ℝ} (hMp: 0 ≤ M) (hM : ∀ (f : dual 𝕜 E), ‖f x‖ ≤ M * ‖f‖) : + ‖x‖ ≤ M := begin classical, by_cases h : x = 0, { simp only [h, hMp, norm_zero] }, - { obtain ⟨f, hf₁, hfx⟩ : ∃ f : E →L[𝕜] 𝕜, ∥f∥ = 1 ∧ f x = ∥x∥ := exists_dual_vector 𝕜 x h, - calc ∥x∥ = ∥(∥x∥ : 𝕜)∥ : is_R_or_C.norm_coe_norm.symm - ... = ∥f x∥ : by rw hfx - ... ≤ M * ∥f∥ : hM f + { obtain ⟨f, hf₁, hfx⟩ : ∃ f : E →L[𝕜] 𝕜, ‖f‖ = 1 ∧ f x = ‖x‖ := exists_dual_vector 𝕜 x h, + calc ‖x‖ = ‖(‖x‖ : 𝕜)‖ : is_R_or_C.norm_coe_norm.symm + ... = ‖f x‖ : by rw hfx + ... ≤ M * ‖f‖ : hM f ... = M : by rw [hf₁, mul_one] } end @@ -150,7 +150,7 @@ def polar (𝕜 : Type*) [nontrivially_normed_field 𝕜] variables (𝕜 : Type*) [nontrivially_normed_field 𝕜] variables {E : Type*} [seminormed_add_comm_group E] [normed_space 𝕜 E] -lemma mem_polar_iff {x' : dual 𝕜 E} (s : set E) : x' ∈ polar 𝕜 s ↔ ∀ z ∈ s, ∥x' z∥ ≤ 1 := iff.rfl +lemma mem_polar_iff {x' : dual 𝕜 E} (s : set E) : x' ∈ polar 𝕜 s ↔ ∀ z ∈ s, ‖x' z‖ ≤ 1 := iff.rfl @[simp] lemma polar_univ : polar 𝕜 (univ : set E) = {(0 : dual 𝕜 E)} := (dual_pairing 𝕜 E).flip.polar_univ @@ -173,34 +173,34 @@ end variables {𝕜} -/-- If `x'` is a dual element such that the norms `∥x' z∥` are bounded for `z ∈ s`, then a +/-- If `x'` is a dual element such that the norms `‖x' z‖` are bounded for `z ∈ s`, then a small scalar multiple of `x'` is in `polar 𝕜 s`. -/ lemma smul_mem_polar {s : set E} {x' : dual 𝕜 E} {c : 𝕜} - (hc : ∀ z, z ∈ s → ∥ x' z ∥ ≤ ∥c∥) : c⁻¹ • x' ∈ polar 𝕜 s := + (hc : ∀ z, z ∈ s → ‖ x' z ‖ ≤ ‖c‖) : c⁻¹ • x' ∈ polar 𝕜 s := begin by_cases c_zero : c = 0, { simp only [c_zero, inv_zero, zero_smul], exact (dual_pairing 𝕜 E).flip.zero_mem_polar _ }, - have eq : ∀ z, ∥ c⁻¹ • (x' z) ∥ = ∥ c⁻¹ ∥ * ∥ x' z ∥ := λ z, norm_smul c⁻¹ _, - have le : ∀ z, z ∈ s → ∥ c⁻¹ • (x' z) ∥ ≤ ∥ c⁻¹ ∥ * ∥ c ∥, + have eq : ∀ z, ‖ c⁻¹ • (x' z) ‖ = ‖ c⁻¹ ‖ * ‖ x' z ‖ := λ z, norm_smul c⁻¹ _, + have le : ∀ z, z ∈ s → ‖ c⁻¹ • (x' z) ‖ ≤ ‖ c⁻¹ ‖ * ‖ c ‖, { intros z hzs, rw eq z, apply mul_le_mul (le_of_eq rfl) (hc z hzs) (norm_nonneg _) (norm_nonneg _), }, - have cancel : ∥ c⁻¹ ∥ * ∥ c ∥ = 1, + have cancel : ‖ c⁻¹ ‖ * ‖ c ‖ = 1, by simp only [c_zero, norm_eq_zero, ne.def, not_false_iff, inv_mul_cancel, norm_inv], rwa cancel at le, end -lemma polar_ball_subset_closed_ball_div {c : 𝕜} (hc : 1 < ∥c∥) {r : ℝ} (hr : 0 < r) : - polar 𝕜 (ball (0 : E) r) ⊆ closed_ball (0 : dual 𝕜 E) (∥c∥ / r) := +lemma polar_ball_subset_closed_ball_div {c : 𝕜} (hc : 1 < ‖c‖) {r : ℝ} (hr : 0 < r) : + polar 𝕜 (ball (0 : E) r) ⊆ closed_ball (0 : dual 𝕜 E) (‖c‖ / r) := begin intros x' hx', rw mem_polar_iff at hx', simp only [polar, mem_set_of_eq, mem_closed_ball_zero_iff, mem_ball_zero_iff] at *, - have hcr : 0 < ∥c∥ / r, from div_pos (zero_lt_one.trans hc) hr, + have hcr : 0 < ‖c‖ / r, from div_pos (zero_lt_one.trans hc) hr, refine continuous_linear_map.op_norm_le_of_shell hr hcr.le hc (λ x h₁ h₂, _), - calc ∥x' x∥ ≤ 1 : hx' _ h₂ - ... ≤ (∥c∥ / r) * ∥x∥ : (inv_pos_le_iff_one_le_mul' hcr).1 (by rwa inv_div) + calc ‖x' x‖ ≤ 1 : hx' _ h₂ + ... ≤ (‖c‖ / r) * ‖x‖ : (inv_pos_le_iff_one_le_mul' hcr).1 (by rwa inv_div) end variables (𝕜) @@ -208,7 +208,7 @@ variables (𝕜) lemma closed_ball_inv_subset_polar_closed_ball {r : ℝ} : closed_ball (0 : dual 𝕜 E) r⁻¹ ⊆ polar 𝕜 (closed_ball (0 : E) r) := λ x' hx' x hx, -calc ∥x' x∥ ≤ ∥x'∥ * ∥x∥ : x'.le_op_norm x +calc ‖x' x‖ ≤ ‖x'‖ * ‖x‖ : x'.le_op_norm x ... ≤ r⁻¹ * r : mul_le_mul (mem_closed_ball_zero_iff.1 hx') (mem_closed_ball_zero_iff.1 hx) (norm_nonneg _) (dist_nonneg.trans hx') @@ -233,7 +233,7 @@ of all elements of the polar `polar 𝕜 s` are bounded by a constant. -/ lemma bounded_polar_of_mem_nhds_zero {s : set E} (s_nhd : s ∈ 𝓝 (0 : E)) : bounded (polar 𝕜 s) := begin - obtain ⟨a, ha⟩ : ∃ a : 𝕜, 1 < ∥a∥ := normed_field.exists_one_lt_norm 𝕜, + obtain ⟨a, ha⟩ : ∃ a : 𝕜, 1 < ‖a‖ := normed_field.exists_one_lt_norm 𝕜, obtain ⟨r, r_pos, r_ball⟩ : ∃ (r : ℝ) (hr : 0 < r), ball 0 r ⊆ s := metric.mem_nhds_iff.1 s_nhd, exact bounded_closed_ball.mono (((dual_pairing 𝕜 E).flip.polar_antitone r_ball).trans $ diff --git a/src/analysis/normed_space/enorm.lean b/src/analysis/normed_space/enorm.lean index 4af078e526972..5326b4c7d8729 100644 --- a/src/analysis/normed_space/enorm.lean +++ b/src/analysis/normed_space/enorm.lean @@ -35,12 +35,12 @@ local attribute [instance, priority 1001] classical.prop_decidable open_locale ennreal /-- Extended norm on a vector space. As in the case of normed spaces, we require only -`∥c • x∥ ≤ ∥c∥ * ∥x∥` in the definition, then prove an equality in `map_smul`. -/ +`‖c • x‖ ≤ ‖c‖ * ‖x‖` in the definition, then prove an equality in `map_smul`. -/ structure enorm (𝕜 : Type*) (V : Type*) [normed_field 𝕜] [add_comm_group V] [module 𝕜 V] := (to_fun : V → ℝ≥0∞) (eq_zero' : ∀ x, to_fun x = 0 → x = 0) (map_add_le' : ∀ x y : V, to_fun (x + y) ≤ to_fun x + to_fun y) -(map_smul_le' : ∀ (c : 𝕜) (x : V), to_fun (c • x) ≤ ∥c∥₊ * to_fun x) +(map_smul_le' : ∀ (c : 𝕜) (x : V), to_fun (c • x) ≤ ‖c‖₊ * to_fun x) namespace enorm @@ -61,12 +61,12 @@ lemma ext_iff {e₁ e₂ : enorm 𝕜 V} : e₁ = e₂ ↔ ∀ x, e₁ x = e₂ @[simp, norm_cast] lemma coe_inj {e₁ e₂ : enorm 𝕜 V} : (e₁ : V → ℝ≥0∞) = e₂ ↔ e₁ = e₂ := coe_fn_injective.eq_iff -@[simp] lemma map_smul (c : 𝕜) (x : V) : e (c • x) = ∥c∥₊ * e x := +@[simp] lemma map_smul (c : 𝕜) (x : V) : e (c • x) = ‖c‖₊ * e x := le_antisymm (e.map_smul_le' c x) $ begin by_cases hc : c = 0, { simp [hc] }, - calc (∥c∥₊ : ℝ≥0∞) * e x = ∥c∥₊ * e (c⁻¹ • c • x) : by rw [inv_smul_smul₀ hc] - ... ≤ ∥c∥₊ * (∥c⁻¹∥₊ * e (c • x)) : _ + calc (‖c‖₊ : ℝ≥0∞) * e x = ‖c‖₊ * e (c⁻¹ • c • x) : by rw [inv_smul_smul₀ hc] + ... ≤ ‖c‖₊ * (‖c⁻¹‖₊ * e (c • x)) : _ ... = e (c • x) : _, { exact ennreal.mul_le_mul le_rfl (e.map_smul_le' _ _) }, { rw [← mul_assoc, nnnorm_inv, ennreal.coe_inv, @@ -80,7 +80,7 @@ by { rw [← zero_smul 𝕜 (0:V), e.map_smul], norm_num } ⟨e.eq_zero' x, λ h, h.symm ▸ e.map_zero⟩ @[simp] lemma map_neg (x : V) : e (-x) = e x := -calc e (-x) = ∥(-1 : 𝕜)∥₊ * e x : by rw [← map_smul, neg_one_smul] +calc e (-x) = ‖(-1 : 𝕜)‖₊ * e x : by rw [← map_smul, neg_one_smul] ... = e x : by simp lemma map_sub_rev (x y : V) : e (x - y) = e (y - x) := @@ -162,7 +162,7 @@ def finite_subspace : subspace 𝕜 V := zero_mem' := by simp, add_mem' := λ x y hx hy, lt_of_le_of_lt (e.map_add_le x y) (ennreal.add_lt_top.2 ⟨hx, hy⟩), smul_mem' := λ c x (hx : _ < _), - calc e (c • x) = ∥c∥₊ * e x : e.map_smul c x + calc e (c • x) = ‖c‖₊ * e x : e.map_smul c x ... < ⊤ : ennreal.mul_lt_top ennreal.coe_ne_top hx.ne } /-- Metric space structure on `e.finite_subspace`. We use `emetric_space.to_metric_space` @@ -185,7 +185,7 @@ instance : normed_add_comm_group e.finite_subspace := dist_eq := λ x y, rfl, .. finite_subspace.metric_space e, .. submodule.add_comm_group _ } -lemma finite_norm_eq (x : e.finite_subspace) : ∥x∥ = (e x).to_real := rfl +lemma finite_norm_eq (x : e.finite_subspace) : ‖x‖ = (e x).to_real := rfl /-- Normed space instance on `e.finite_subspace`. -/ instance : normed_space 𝕜 e.finite_subspace := diff --git a/src/analysis/normed_space/exponential.lean b/src/analysis/normed_space/exponential.lean index c009801d5afc0..af316d21047c5 100644 --- a/src/analysis/normed_space/exponential.lean +++ b/src/analysis/normed_space/exponential.lean @@ -166,12 +166,12 @@ variables [normed_ring 𝔸] [normed_ring 𝔹] [normed_algebra 𝕂 𝔸] [norm lemma norm_exp_series_summable_of_mem_ball (x : 𝔸) (hx : x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : - summable (λ n, ∥exp_series 𝕂 𝔸 n (λ _, x)∥) := + summable (λ n, ‖exp_series 𝕂 𝔸 n (λ _, x)‖) := (exp_series 𝕂 𝔸).summable_norm_apply hx lemma norm_exp_series_summable_of_mem_ball' (x : 𝔸) (hx : x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : - summable (λ n, ∥(n!⁻¹ : 𝕂) • x^n∥) := + summable (λ n, ‖(n!⁻¹ : 𝕂) • x^n‖) := begin change summable (norm ∘ _), rw ← exp_series_apply_eq', @@ -301,7 +301,7 @@ variables (𝕂) lemma norm_exp_series_div_summable_of_mem_ball (x : 𝔸) (hx : x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : - summable (λ n, ∥x^n / n!∥) := + summable (λ n, ‖x^n / n!‖) := begin change summable (norm ∘ _), rw ← exp_series_apply_eq_div' x, @@ -363,7 +363,7 @@ begin filter_upwards [eventually_cofinite_ne 0] with n hn, rw [norm_mul, norm_norm (exp_series 𝕂 𝔸 n), exp_series, norm_smul, norm_inv, norm_pow, nnreal.norm_eq, norm_eq_abs, abs_cast_nat, mul_comm, ←mul_assoc, ←div_eq_mul_inv], - have : ∥continuous_multilinear_map.mk_pi_algebra_fin 𝕂 n 𝔸∥ ≤ 1 := + have : ‖continuous_multilinear_map.mk_pi_algebra_fin 𝕂 n 𝔸‖ ≤ 1 := norm_mk_pi_algebra_fin_le_of_pos (nat.pos_of_ne_zero hn), exact mul_le_of_le_one_right (div_nonneg (pow_nonneg r.coe_nonneg n) n!.cast_nonneg) this end @@ -376,10 +376,10 @@ end variables {𝕂 𝔸 𝔹} -lemma norm_exp_series_summable (x : 𝔸) : summable (λ n, ∥exp_series 𝕂 𝔸 n (λ _, x)∥) := +lemma norm_exp_series_summable (x : 𝔸) : summable (λ n, ‖exp_series 𝕂 𝔸 n (λ _, x)‖) := norm_exp_series_summable_of_mem_ball x ((exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _) -lemma norm_exp_series_summable' (x : 𝔸) : summable (λ n, ∥(n!⁻¹ : 𝕂) • x^n∥) := +lemma norm_exp_series_summable' (x : 𝔸) : summable (λ n, ‖(n!⁻¹ : 𝕂) • x^n‖) := norm_exp_series_summable_of_mem_ball' x ((exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _) section complete_algebra @@ -539,7 +539,7 @@ variables {𝕂 𝔸 : Type*} [is_R_or_C 𝕂] [normed_division_ring 𝔸] [norm variables (𝕂) -lemma norm_exp_series_div_summable (x : 𝔸) : summable (λ n, ∥x^n / n!∥) := +lemma norm_exp_series_div_summable (x : 𝔸) : summable (λ n, ‖x^n / n!‖) := norm_exp_series_div_summable_of_mem_ball 𝕂 x ((exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _) diff --git a/src/analysis/normed_space/extend.lean b/src/analysis/normed_space/extend.lean index 66bfa28a352c4..90ea75b806679 100644 --- a/src/analysis/normed_space/extend.lean +++ b/src/analysis/normed_space/extend.lean @@ -36,7 +36,7 @@ variables {𝕜 : Type*} [is_R_or_C 𝕜] {F : Type*} [seminormed_add_comm_group local notation `abs𝕜` := @is_R_or_C.abs 𝕜 _ /-- Extend `fr : F →ₗ[ℝ] ℝ` to `F →ₗ[𝕜] 𝕜` in a way that will also be continuous and have its norm -bounded by `∥fr∥` if `fr` is continuous. -/ +bounded by `‖fr‖` if `fr` is continuous. -/ noncomputable def linear_map.extend_to_𝕜' [module ℝ F] [is_scalar_tower ℝ 𝕜 F] (fr : F →ₗ[ℝ] ℝ) : F →ₗ[𝕜] 𝕜 := begin @@ -76,14 +76,14 @@ lemma linear_map.extend_to_𝕜'_apply [module ℝ F] [is_scalar_tower ℝ 𝕜 (fr : F →ₗ[ℝ] ℝ) (x : F) : fr.extend_to_𝕜' x = (fr x : 𝕜) - (I : 𝕜) * fr ((I : 𝕜) • x) := rfl -/-- The norm of the extension is bounded by `∥fr∥`. -/ +/-- The norm of the extension is bounded by `‖fr‖`. -/ lemma norm_bound [normed_space ℝ F] [is_scalar_tower ℝ 𝕜 F] (fr : F →L[ℝ] ℝ) (x : F) : - ∥(fr.to_linear_map.extend_to_𝕜' x : 𝕜)∥ ≤ ∥fr∥ * ∥x∥ := + ‖(fr.to_linear_map.extend_to_𝕜' x : 𝕜)‖ ≤ ‖fr‖ * ‖x‖ := begin let lm : F →ₗ[𝕜] 𝕜 := fr.to_linear_map.extend_to_𝕜', -- We aim to find a `t : 𝕜` such that -- * `lm (t • x) = fr (t • x)` (so `lm (t • x) = t * lm x ∈ ℝ`) - -- * `∥lm x∥ = ∥lm (t • x)∥` (so `t.abs` must be 1) + -- * `‖lm x‖ = ‖lm (t • x)‖` (so `t.abs` must be 1) -- If `lm x ≠ 0`, `(lm x)⁻¹` satisfies the first requirement, and after normalizing, it -- satisfies the second. -- (If `lm x = 0`, the goal is trivial.) @@ -107,21 +107,21 @@ begin ... = im (1 / (abs𝕜 (lm x)⁻¹ : 𝕜)) : by rw [div_mul_eq_mul_div, inv_mul_cancel h] ... = 0 : by rw [← of_real_one, ← of_real_div, of_real_im] ... = im (fr (t • x) : 𝕜) : by rw [of_real_im] } }, - calc ∥lm x∥ = abs𝕜 t * ∥lm x∥ : by rw [ht, one_mul] - ... = ∥t * lm x∥ : by rw [← norm_eq_abs, norm_mul] - ... = ∥lm (t • x)∥ : by rw [←smul_eq_mul, lm.map_smul] - ... = ∥(fr (t • x) : 𝕜)∥ : by rw h1 - ... = ∥fr (t • x)∥ : by rw [norm_eq_abs, abs_of_real, norm_eq_abs, abs_to_real] - ... ≤ ∥fr∥ * ∥t • x∥ : continuous_linear_map.le_op_norm _ _ - ... = ∥fr∥ * (∥t∥ * ∥x∥) : by rw norm_smul - ... ≤ ∥fr∥ * ∥x∥ : by rw [norm_eq_abs, ht, one_mul] + calc ‖lm x‖ = abs𝕜 t * ‖lm x‖ : by rw [ht, one_mul] + ... = ‖t * lm x‖ : by rw [← norm_eq_abs, norm_mul] + ... = ‖lm (t • x)‖ : by rw [←smul_eq_mul, lm.map_smul] + ... = ‖(fr (t • x) : 𝕜)‖ : by rw h1 + ... = ‖fr (t • x)‖ : by rw [norm_eq_abs, abs_of_real, norm_eq_abs, abs_to_real] + ... ≤ ‖fr‖ * ‖t • x‖ : continuous_linear_map.le_op_norm _ _ + ... = ‖fr‖ * (‖t‖ * ‖x‖) : by rw norm_smul + ... ≤ ‖fr‖ * ‖x‖ : by rw [norm_eq_abs, ht, one_mul] end /-- Extend `fr : F →L[ℝ] ℝ` to `F →L[𝕜] 𝕜`. -/ noncomputable def continuous_linear_map.extend_to_𝕜' [normed_space ℝ F] [is_scalar_tower ℝ 𝕜 F] (fr : F →L[ℝ] ℝ) : F →L[𝕜] 𝕜 := -linear_map.mk_continuous _ (∥fr∥) (norm_bound _) +linear_map.mk_continuous _ (‖fr‖) (norm_bound _) lemma continuous_linear_map.extend_to_𝕜'_apply [normed_space ℝ F] [is_scalar_tower ℝ 𝕜 F] (fr : F →L[ℝ] ℝ) (x : F) : diff --git a/src/analysis/normed_space/extr.lean b/src/analysis/normed_space/extr.lean index 603a1020ad2b3..b458f8d5dc8aa 100644 --- a/src/analysis/normed_space/extr.lean +++ b/src/analysis/normed_space/extr.lean @@ -11,7 +11,7 @@ import topology.local_extr In this file we prove the following lemma, see `is_max_filter.norm_add_same_ray`. If `f : α → E` is a function such that `norm ∘ f` has a maximum along a filter `l` at a point `c` and `y` is a vector -on the same ray as `f c`, then the function `λ x, ∥f x + y∥` has a maximul along `l` at `c`. +on the same ray as `f c`, then the function `λ x, ‖f x + y‖` has a maximul along `l` at `c`. Then we specialize it to the case `y = f c` and to different special cases of `is_max_filter`: `is_max_on`, `is_local_max_on`, and `is_local_max`. @@ -28,31 +28,31 @@ section variables {f : α → E} {l : filter α} {s : set α} {c : α} {y : E} /-- If `f : α → E` is a function such that `norm ∘ f` has a maximum along a filter `l` at a point -`c` and `y` is a vector on the same ray as `f c`, then the function `λ x, ∥f x + y∥` has a maximul +`c` and `y` is a vector on the same ray as `f c`, then the function `λ x, ‖f x + y‖` has a maximul along `l` at `c`. -/ lemma is_max_filter.norm_add_same_ray (h : is_max_filter (norm ∘ f) l c) (hy : same_ray ℝ (f c) y) : - is_max_filter (λ x, ∥f x + y∥) l c := + is_max_filter (λ x, ‖f x + y‖) l c := h.mono $ λ x hx, -calc ∥f x + y∥ ≤ ∥f x∥ + ∥y∥ : norm_add_le _ _ - ... ≤ ∥f c∥ + ∥y∥ : add_le_add_right hx _ - ... = ∥f c + y∥ : hy.norm_add.symm +calc ‖f x + y‖ ≤ ‖f x‖ + ‖y‖ : norm_add_le _ _ + ... ≤ ‖f c‖ + ‖y‖ : add_le_add_right hx _ + ... = ‖f c + y‖ : hy.norm_add.symm /-- If `f : α → E` is a function such that `norm ∘ f` has a maximum along a filter `l` at a point -`c`, then the function `λ x, ∥f x + f c∥` has a maximul along `l` at `c`. -/ +`c`, then the function `λ x, ‖f x + f c‖` has a maximul along `l` at `c`. -/ lemma is_max_filter.norm_add_self (h : is_max_filter (norm ∘ f) l c) : - is_max_filter (λ x, ∥f x + f c∥) l c := + is_max_filter (λ x, ‖f x + f c‖) l c := h.norm_add_same_ray same_ray.rfl /-- If `f : α → E` is a function such that `norm ∘ f` has a maximum on a set `s` at a point `c` and -`y` is a vector on the same ray as `f c`, then the function `λ x, ∥f x + y∥` has a maximul on `s` at +`y` is a vector on the same ray as `f c`, then the function `λ x, ‖f x + y‖` has a maximul on `s` at `c`. -/ lemma is_max_on.norm_add_same_ray (h : is_max_on (norm ∘ f) s c) (hy : same_ray ℝ (f c) y) : - is_max_on (λ x, ∥f x + y∥) s c := + is_max_on (λ x, ‖f x + y‖) s c := h.norm_add_same_ray hy /-- If `f : α → E` is a function such that `norm ∘ f` has a maximum on a set `s` at a point `c`, -then the function `λ x, ∥f x + f c∥` has a maximul on `s` at `c`. -/ -lemma is_max_on.norm_add_self (h : is_max_on (norm ∘ f) s c) : is_max_on (λ x, ∥f x + f c∥) s c := +then the function `λ x, ‖f x + f c‖` has a maximul on `s` at `c`. -/ +lemma is_max_on.norm_add_self (h : is_max_on (norm ∘ f) s c) : is_max_on (λ x, ‖f x + f c‖) s c := h.norm_add_self end @@ -60,26 +60,26 @@ end variables {f : X → E} {s : set X} {c : X} {y : E} /-- If `f : α → E` is a function such that `norm ∘ f` has a local maximum on a set `s` at a point -`c` and `y` is a vector on the same ray as `f c`, then the function `λ x, ∥f x + y∥` has a local +`c` and `y` is a vector on the same ray as `f c`, then the function `λ x, ‖f x + y‖` has a local maximul on `s` at `c`. -/ lemma is_local_max_on.norm_add_same_ray (h : is_local_max_on (norm ∘ f) s c) - (hy : same_ray ℝ (f c) y) : is_local_max_on (λ x, ∥f x + y∥) s c := + (hy : same_ray ℝ (f c) y) : is_local_max_on (λ x, ‖f x + y‖) s c := h.norm_add_same_ray hy /-- If `f : α → E` is a function such that `norm ∘ f` has a local maximum on a set `s` at a point -`c`, then the function `λ x, ∥f x + f c∥` has a local maximul on `s` at `c`. -/ +`c`, then the function `λ x, ‖f x + f c‖` has a local maximul on `s` at `c`. -/ lemma is_local_max_on.norm_add_self (h : is_local_max_on (norm ∘ f) s c) : - is_local_max_on (λ x, ∥f x + f c∥) s c := + is_local_max_on (λ x, ‖f x + f c‖) s c := h.norm_add_self /-- If `f : α → E` is a function such that `norm ∘ f` has a local maximum at a point `c` and `y` is -a vector on the same ray as `f c`, then the function `λ x, ∥f x + y∥` has a local maximul at `c`. -/ +a vector on the same ray as `f c`, then the function `λ x, ‖f x + y‖` has a local maximul at `c`. -/ lemma is_local_max.norm_add_same_ray (h : is_local_max (norm ∘ f) c) - (hy : same_ray ℝ (f c) y) : is_local_max (λ x, ∥f x + y∥) c := + (hy : same_ray ℝ (f c) y) : is_local_max (λ x, ‖f x + y‖) c := h.norm_add_same_ray hy /-- If `f : α → E` is a function such that `norm ∘ f` has a local maximum at a point `c`, then the -function `λ x, ∥f x + f c∥` has a local maximul at `c`. -/ +function `λ x, ‖f x + f c‖` has a local maximul at `c`. -/ lemma is_local_max.norm_add_self (h : is_local_max (norm ∘ f) c) : - is_local_max (λ x, ∥f x + f c∥) c := + is_local_max (λ x, ‖f x + f c‖) c := h.norm_add_self diff --git a/src/analysis/normed_space/finite_dimension.lean b/src/analysis/normed_space/finite_dimension.lean index 7281dbea2265e..87caaba6ba994 100644 --- a/src/analysis/normed_space/finite_dimension.lean +++ b/src/analysis/normed_space/finite_dimension.lean @@ -171,7 +171,7 @@ as `lipschitz_extension_constant E'`. -/ @[irreducible] def lipschitz_extension_constant (E' : Type*) [normed_add_comm_group E'] [normed_space ℝ E'] [finite_dimensional ℝ E'] : ℝ≥0 := let A := (basis.of_vector_space ℝ E').equiv_fun.to_continuous_linear_equiv in - max (∥A.symm.to_continuous_linear_map∥₊ * ∥A.to_continuous_linear_map∥₊) 1 + max (‖A.symm.to_continuous_linear_map‖₊ * ‖A.to_continuous_linear_map‖₊) 1 lemma lipschitz_extension_constant_pos (E' : Type*) [normed_add_comm_group E'] [normed_space ℝ E'] [finite_dimensional ℝ E'] : @@ -191,13 +191,13 @@ begin `E'` and such a space to transfer the result to `E'`. -/ let ι : Type* := basis.of_vector_space_index ℝ E', let A := (basis.of_vector_space ℝ E').equiv_fun.to_continuous_linear_equiv, - have LA : lipschitz_with (∥A.to_continuous_linear_map∥₊) A, by apply A.lipschitz, - have L : lipschitz_on_with (∥A.to_continuous_linear_map∥₊ * K) (A ∘ f) s := + have LA : lipschitz_with (‖A.to_continuous_linear_map‖₊) A, by apply A.lipschitz, + have L : lipschitz_on_with (‖A.to_continuous_linear_map‖₊ * K) (A ∘ f) s := LA.comp_lipschitz_on_with hf, - obtain ⟨g, hg, gs⟩ : ∃ g : α → (ι → ℝ), lipschitz_with (∥A.to_continuous_linear_map∥₊ * K) g ∧ + obtain ⟨g, hg, gs⟩ : ∃ g : α → (ι → ℝ), lipschitz_with (‖A.to_continuous_linear_map‖₊ * K) g ∧ eq_on (A ∘ f) g s := L.extend_pi, refine ⟨A.symm ∘ g, _, _⟩, - { have LAsymm : lipschitz_with (∥A.symm.to_continuous_linear_map∥₊) A.symm, + { have LAsymm : lipschitz_with (‖A.symm.to_continuous_linear_map‖₊) A.symm, by apply A.symm.lipschitz, apply (LAsymm.comp hg).weaken, rw [lipschitz_extension_constant, ← mul_assoc], @@ -223,12 +223,12 @@ begin casesI nonempty_fintype ι, simp only [fintype.linear_independent_iff'] at hf ⊢, rcases linear_map.exists_antilipschitz_with _ hf with ⟨K, K0, hK⟩, - have : tendsto (λ g : ι → E, ∑ i, ∥g i - f i∥) (𝓝 f) (𝓝 $ ∑ i, ∥f i - f i∥), + have : tendsto (λ g : ι → E, ∑ i, ‖g i - f i‖) (𝓝 f) (𝓝 $ ∑ i, ‖f i - f i‖), from tendsto_finset_sum _ (λ i hi, tendsto.norm $ ((continuous_apply i).tendsto _).sub tendsto_const_nhds), simp only [sub_self, norm_zero, finset.sum_const_zero] at this, refine (this.eventually (gt_mem_nhds $ inv_pos.2 K0)).mono (λ g hg, _), - replace hg : ∑ i, ∥g i - f i∥₊ < K⁻¹, by { rw ← nnreal.coe_lt_coe, push_cast, exact hg }, + replace hg : ∑ i, ‖g i - f i‖₊ < K⁻¹, by { rw ← nnreal.coe_lt_coe, push_cast, exact hg }, rw linear_map.ker_eq_bot, refine (hK.add_sub_lipschitz_with (lipschitz_with.of_dist_le_mul $ λ v u, _) hg).injective, simp only [dist_eq_norm, linear_map.lsum_apply, pi.sub_apply, linear_map.sum_apply, @@ -307,42 +307,42 @@ v.constr_apply_fintype 𝕜 _ _ v.constr_basis 𝕜 _ _ lemma basis.op_nnnorm_le {ι : Type*} [fintype ι] (v : basis ι 𝕜 E) {u : E →L[𝕜] F} (M : ℝ≥0) - (hu : ∀ i, ∥u (v i)∥₊ ≤ M) : - ∥u∥₊ ≤ fintype.card ι • ∥v.equiv_funL.to_continuous_linear_map∥₊ * M := + (hu : ∀ i, ‖u (v i)‖₊ ≤ M) : + ‖u‖₊ ≤ fintype.card ι • ‖v.equiv_funL.to_continuous_linear_map‖₊ * M := u.op_nnnorm_le_bound _ $ λ e, begin set φ := v.equiv_funL.to_continuous_linear_map, calc - ∥u e∥₊ = ∥u (∑ i, v.equiv_fun e i • v i)∥₊ : by rw [v.sum_equiv_fun] - ... = ∥∑ i, (v.equiv_fun e i) • (u $ v i)∥₊ : by simp [u.map_sum, linear_map.map_smul] - ... ≤ ∑ i, ∥(v.equiv_fun e i) • (u $ v i)∥₊ : nnnorm_sum_le _ _ - ... = ∑ i, ∥v.equiv_fun e i∥₊ * ∥u (v i)∥₊ : by simp only [nnnorm_smul] - ... ≤ ∑ i, ∥v.equiv_fun e i∥₊ * M : finset.sum_le_sum (λ i hi, + ‖u e‖₊ = ‖u (∑ i, v.equiv_fun e i • v i)‖₊ : by rw [v.sum_equiv_fun] + ... = ‖∑ i, (v.equiv_fun e i) • (u $ v i)‖₊ : by simp [u.map_sum, linear_map.map_smul] + ... ≤ ∑ i, ‖(v.equiv_fun e i) • (u $ v i)‖₊ : nnnorm_sum_le _ _ + ... = ∑ i, ‖v.equiv_fun e i‖₊ * ‖u (v i)‖₊ : by simp only [nnnorm_smul] + ... ≤ ∑ i, ‖v.equiv_fun e i‖₊ * M : finset.sum_le_sum (λ i hi, mul_le_mul_of_nonneg_left (hu i) (zero_le _)) - ... = (∑ i, ∥v.equiv_fun e i∥₊) * M : finset.sum_mul.symm - ... ≤ fintype.card ι • (∥φ∥₊ * ∥e∥₊) * M : + ... = (∑ i, ‖v.equiv_fun e i‖₊) * M : finset.sum_mul.symm + ... ≤ fintype.card ι • (‖φ‖₊ * ‖e‖₊) * M : (suffices _, from mul_le_mul_of_nonneg_right this (zero_le M), - calc ∑ i, ∥v.equiv_fun e i∥₊ - ≤ fintype.card ι • ∥φ e∥₊ : pi.sum_nnnorm_apply_le_nnnorm _ - ... ≤ fintype.card ι • (∥φ∥₊ * ∥e∥₊) : nsmul_le_nsmul_of_le_right (φ.le_op_nnnorm e) _) - ... = fintype.card ι • ∥φ∥₊ * M * ∥e∥₊ : by simp only [smul_mul_assoc, mul_right_comm], + calc ∑ i, ‖v.equiv_fun e i‖₊ + ≤ fintype.card ι • ‖φ e‖₊ : pi.sum_nnnorm_apply_le_nnnorm _ + ... ≤ fintype.card ι • (‖φ‖₊ * ‖e‖₊) : nsmul_le_nsmul_of_le_right (φ.le_op_nnnorm e) _) + ... = fintype.card ι • ‖φ‖₊ * M * ‖e‖₊ : by simp only [smul_mul_assoc, mul_right_comm], end lemma basis.op_norm_le {ι : Type*} [fintype ι] (v : basis ι 𝕜 E) {u : E →L[𝕜] F} {M : ℝ} - (hM : 0 ≤ M) (hu : ∀ i, ∥u (v i)∥ ≤ M) : - ∥u∥ ≤ fintype.card ι • ∥v.equiv_funL.to_continuous_linear_map∥ * M := + (hM : 0 ≤ M) (hu : ∀ i, ‖u (v i)‖ ≤ M) : + ‖u‖ ≤ fintype.card ι • ‖v.equiv_funL.to_continuous_linear_map‖ * M := by simpa using nnreal.coe_le_coe.mpr (v.op_nnnorm_le ⟨M, hM⟩ hu) /-- A weaker version of `basis.op_nnnorm_le` that abstracts away the value of `C`. -/ lemma basis.exists_op_nnnorm_le {ι : Type*} [finite ι] (v : basis ι 𝕜 E) : - ∃ C > (0 : ℝ≥0), ∀ {u : E →L[𝕜] F} (M : ℝ≥0), (∀ i, ∥u (v i)∥₊ ≤ M) → ∥u∥₊ ≤ C*M := + ∃ C > (0 : ℝ≥0), ∀ {u : E →L[𝕜] F} (M : ℝ≥0), (∀ i, ‖u (v i)‖₊ ≤ M) → ‖u‖₊ ≤ C*M := by casesI nonempty_fintype ι; exact - ⟨max (fintype.card ι • ∥v.equiv_funL.to_continuous_linear_map∥₊) 1, + ⟨max (fintype.card ι • ‖v.equiv_funL.to_continuous_linear_map‖₊) 1, zero_lt_one.trans_le (le_max_right _ _), λ u M hu, (v.op_nnnorm_le M hu).trans $ mul_le_mul_of_nonneg_right (le_max_left _ _) (zero_le M)⟩ /-- A weaker version of `basis.op_norm_le` that abstracts away the value of `C`. -/ lemma basis.exists_op_norm_le {ι : Type*} [finite ι] (v : basis ι 𝕜 E) : - ∃ C > (0 : ℝ), ∀ {u : E →L[𝕜] F} {M : ℝ}, 0 ≤ M → (∀ i, ∥u (v i)∥ ≤ M) → ∥u∥ ≤ C*M := + ∃ C > (0 : ℝ), ∀ {u : E →L[𝕜] F} {M : ℝ}, 0 ≤ M → (∀ i, ‖u (v i)‖ ≤ M) → ‖u‖ ≤ C*M := let ⟨C, hC, h⟩ := v.exists_op_nnnorm_le in ⟨C, hC, λ u, subtype.forall'.mpr h⟩ instance [finite_dimensional 𝕜 E] [second_countable_topology F] : @@ -357,24 +357,24 @@ begin obtain ⟨u : ℕ → F, hu : dense_range u⟩ := exists_dense_seq F, let v := finite_dimensional.fin_basis 𝕜 E, obtain ⟨C : ℝ, C_pos : 0 < C, - hC : ∀ {φ : E →L[𝕜] F} {M : ℝ}, 0 ≤ M → (∀ i, ∥φ (v i)∥ ≤ M) → ∥φ∥ ≤ C * M⟩ := + hC : ∀ {φ : E →L[𝕜] F} {M : ℝ}, 0 ≤ M → (∀ i, ‖φ (v i)‖ ≤ M) → ‖φ‖ ≤ C * M⟩ := v.exists_op_norm_le, have h_2C : 0 < 2*C := mul_pos zero_lt_two C_pos, have hε2C : 0 < ε/(2*C) := div_pos ε_pos h_2C, - have : ∀ φ : E →L[𝕜] F, ∃ n : fin d → ℕ, ∥φ - (v.constrL $ u ∘ n)∥ ≤ ε/2, + have : ∀ φ : E →L[𝕜] F, ∃ n : fin d → ℕ, ‖φ - (v.constrL $ u ∘ n)‖ ≤ ε/2, { intros φ, - have : ∀ i, ∃ n, ∥φ (v i) - u n∥ ≤ ε/(2*C), + have : ∀ i, ∃ n, ‖φ (v i) - u n‖ ≤ ε/(2*C), { simp only [norm_sub_rev], intro i, have : φ (v i) ∈ closure (range u) := hu _, - obtain ⟨n, hn⟩ : ∃ n, ∥u n - φ (v i)∥ < ε / (2 * C), + obtain ⟨n, hn⟩ : ∃ n, ‖u n - φ (v i)‖ < ε / (2 * C), { rw mem_closure_iff_nhds_basis metric.nhds_basis_ball at this, specialize this (ε/(2*C)) hε2C, simpa [dist_eq_norm] }, exact ⟨n, le_of_lt hn⟩ }, choose n hn using this, use n, - replace hn : ∀ i : fin d, ∥(φ - (v.constrL $ u ∘ n)) (v i)∥ ≤ ε / (2 * C), by simp [hn], + replace hn : ∀ i : fin d, ‖(φ - (v.constrL $ u ∘ n)) (v i)‖ ≤ ε / (2 * C), by simp [hn], have : C * (ε / (2 * C)) = ε/2, { rw [eq_div_iff (two_ne_zero : (2 : ℝ) ≠ 0), mul_comm, ← mul_assoc, mul_div_cancel' _ (ne_of_gt h_2C)] }, @@ -422,9 +422,9 @@ section riesz /-- In an infinite dimensional space, given a finite number of points, one may find a point with norm at most `R` which is at distance at least `1` of all these points. -/ -theorem exists_norm_le_le_norm_sub_of_finset {c : 𝕜} (hc : 1 < ∥c∥) {R : ℝ} (hR : ∥c∥ < R) +theorem exists_norm_le_le_norm_sub_of_finset {c : 𝕜} (hc : 1 < ‖c‖) {R : ℝ} (hR : ‖c‖ < R) (h : ¬ (finite_dimensional 𝕜 E)) (s : finset E) : - ∃ (x : E), ∥x∥ ≤ R ∧ ∀ y ∈ s, 1 ≤ ∥y - x∥ := + ∃ (x : E), ‖x‖ ≤ R ∧ ∀ y ∈ s, 1 ≤ ‖y - x‖ := begin let F := submodule.span 𝕜 (s : set E), haveI : finite_dimensional 𝕜 F := module.finite_def.2 @@ -435,9 +435,9 @@ begin have : (⊤ : submodule 𝕜 E) = F, by { ext x, simp [h] }, have : finite_dimensional 𝕜 (⊤ : submodule 𝕜 E), by rwa this, refine module.finite_def.2 ((submodule.fg_top _).1 (module.finite_def.1 this)) }, - obtain ⟨x, xR, hx⟩ : ∃ (x : E), ∥x∥ ≤ R ∧ ∀ (y : E), y ∈ F → 1 ≤ ∥x - y∥ := + obtain ⟨x, xR, hx⟩ : ∃ (x : E), ‖x‖ ≤ R ∧ ∀ (y : E), y ∈ F → 1 ≤ ‖x - y‖ := riesz_lemma_of_norm_lt hc hR Fclosed this, - have hx' : ∀ (y : E), y ∈ F → 1 ≤ ∥y - x∥, + have hx' : ∀ (y : E), y ∈ F → 1 ≤ ‖y - x‖, { assume y hy, rw ← norm_neg, simpa using hx y hy }, exact ⟨x, xR, λ y hy, hx' _ (submodule.subset_span hy)⟩, end @@ -445,27 +445,27 @@ end /-- In an infinite-dimensional normed space, there exists a sequence of points which are all bounded by `R` and at distance at least `1`. For a version not assuming `c` and `R`, see `exists_seq_norm_le_one_le_norm_sub`. -/ -theorem exists_seq_norm_le_one_le_norm_sub' {c : 𝕜} (hc : 1 < ∥c∥) {R : ℝ} (hR : ∥c∥ < R) +theorem exists_seq_norm_le_one_le_norm_sub' {c : 𝕜} (hc : 1 < ‖c‖) {R : ℝ} (hR : ‖c‖ < R) (h : ¬ (finite_dimensional 𝕜 E)) : - ∃ f : ℕ → E, (∀ n, ∥f n∥ ≤ R) ∧ (∀ m n, m ≠ n → 1 ≤ ∥f m - f n∥) := + ∃ f : ℕ → E, (∀ n, ‖f n‖ ≤ R) ∧ (∀ m n, m ≠ n → 1 ≤ ‖f m - f n‖) := begin - haveI : is_symm E (λ (x y : E), 1 ≤ ∥x - y∥), + haveI : is_symm E (λ (x y : E), 1 ≤ ‖x - y‖), { constructor, assume x y hxy, rw ← norm_neg, simpa }, - apply exists_seq_of_forall_finset_exists' (λ (x : E), ∥x∥ ≤ R) (λ (x : E) (y : E), 1 ≤ ∥x - y∥), + apply exists_seq_of_forall_finset_exists' (λ (x : E), ‖x‖ ≤ R) (λ (x : E) (y : E), 1 ≤ ‖x - y‖), assume s hs, exact exists_norm_le_le_norm_sub_of_finset hc hR h s, end theorem exists_seq_norm_le_one_le_norm_sub (h : ¬ (finite_dimensional 𝕜 E)) : - ∃ (R : ℝ) (f : ℕ → E), (1 < R) ∧ (∀ n, ∥f n∥ ≤ R) ∧ (∀ m n, m ≠ n → 1 ≤ ∥f m - f n∥) := + ∃ (R : ℝ) (f : ℕ → E), (1 < R) ∧ (∀ n, ‖f n‖ ≤ R) ∧ (∀ m n, m ≠ n → 1 ≤ ‖f m - f n‖) := begin - obtain ⟨c, hc⟩ : ∃ (c : 𝕜), 1 < ∥c∥ := normed_field.exists_one_lt_norm 𝕜, - have A : ∥c∥ < ∥c∥ + 1, by linarith, + obtain ⟨c, hc⟩ : ∃ (c : 𝕜), 1 < ‖c‖ := normed_field.exists_one_lt_norm 𝕜, + have A : ‖c‖ < ‖c‖ + 1, by linarith, rcases exists_seq_norm_le_one_le_norm_sub' hc A h with ⟨f, hf⟩, - exact ⟨∥c∥ + 1, f, hc.trans A, hf.1, hf.2⟩ + exact ⟨‖c‖ + 1, f, hc.trans A, hf.1, hf.2⟩ end variable (𝕜) @@ -477,29 +477,29 @@ theorem finite_dimensional_of_is_compact_closed_ball₀ {r : ℝ} (rpos : 0 < r) begin by_contra hfin, obtain ⟨R, f, Rgt, fle, lef⟩ : - ∃ (R : ℝ) (f : ℕ → E), (1 < R) ∧ (∀ n, ∥f n∥ ≤ R) ∧ (∀ m n, m ≠ n → 1 ≤ ∥f m - f n∥) := + ∃ (R : ℝ) (f : ℕ → E), (1 < R) ∧ (∀ n, ‖f n‖ ≤ R) ∧ (∀ m n, m ≠ n → 1 ≤ ‖f m - f n‖) := exists_seq_norm_le_one_le_norm_sub hfin, have rRpos : 0 < r / R := div_pos rpos (zero_lt_one.trans Rgt), - obtain ⟨c, hc⟩ : ∃ (c : 𝕜), 0 < ∥c∥ ∧ ∥c∥ < (r / R) := normed_field.exists_norm_lt _ rRpos, + obtain ⟨c, hc⟩ : ∃ (c : 𝕜), 0 < ‖c‖ ∧ ‖c‖ < (r / R) := normed_field.exists_norm_lt _ rRpos, let g := λ (n : ℕ), c • f n, have A : ∀ n, g n ∈ metric.closed_ball (0 : E) r, { assume n, simp only [norm_smul, dist_zero_right, metric.mem_closed_ball], - calc ∥c∥ * ∥f n∥ ≤ (r / R) * R : mul_le_mul hc.2.le (fle n) (norm_nonneg _) rRpos.le + calc ‖c‖ * ‖f n‖ ≤ (r / R) * R : mul_le_mul hc.2.le (fle n) (norm_nonneg _) rRpos.le ... = r : by field_simp [(zero_lt_one.trans Rgt).ne'] }, obtain ⟨x, hx, φ, φmono, φlim⟩ : ∃ (x : E) (H : x ∈ metric.closed_ball (0 : E) r) (φ : ℕ → ℕ), strict_mono φ ∧ tendsto (g ∘ φ) at_top (𝓝 x) := h.tendsto_subseq A, have B : cauchy_seq (g ∘ φ) := φlim.cauchy_seq, - obtain ⟨N, hN⟩ : ∃ (N : ℕ), ∀ (n : ℕ), N ≤ n → dist ((g ∘ φ) n) ((g ∘ φ) N) < ∥c∥ := - metric.cauchy_seq_iff'.1 B (∥c∥) hc.1, - apply lt_irrefl (∥c∥), - calc ∥c∥ ≤ dist (g (φ (N+1))) (g (φ N)) : begin - conv_lhs { rw [← mul_one (∥c∥)] }, + obtain ⟨N, hN⟩ : ∃ (N : ℕ), ∀ (n : ℕ), N ≤ n → dist ((g ∘ φ) n) ((g ∘ φ) N) < ‖c‖ := + metric.cauchy_seq_iff'.1 B (‖c‖) hc.1, + apply lt_irrefl (‖c‖), + calc ‖c‖ ≤ dist (g (φ (N+1))) (g (φ N)) : begin + conv_lhs { rw [← mul_one (‖c‖)] }, simp only [g, dist_eq_norm, ←smul_sub, norm_smul, -mul_one], apply mul_le_mul_of_nonneg_left (lef _ _ (ne_of_gt _)) (norm_nonneg _), exact φmono (nat.lt_succ_self N) end - ... < ∥c∥ : hN (N+1) (nat.le_succ N) + ... < ‖c‖ : hN (N+1) (nat.le_succ N) end /-- **Riesz's theorem**: if a closed ball of positive radius is compact in a vector space, then the @@ -673,25 +673,25 @@ begin rw [metric.inf_dist_zero_of_mem_closure hx'.2, dist_self] }, end -/-- In a finite dimensional vector space over `ℝ`, the series `∑ x, ∥f x∥` is unconditionally +/-- In a finite dimensional vector space over `ℝ`, the series `∑ x, ‖f x‖` is unconditionally summable if and only if the series `∑ x, f x` is unconditionally summable. One implication holds in any complete normed space, while the other holds only in finite dimensional spaces. -/ lemma summable_norm_iff {α E : Type*} [normed_add_comm_group E] [normed_space ℝ E] - [finite_dimensional ℝ E] {f : α → E} : summable (λ x, ∥f x∥) ↔ summable f := + [finite_dimensional ℝ E] {f : α → E} : summable (λ x, ‖f x‖) ↔ summable f := begin refine ⟨summable_of_summable_norm, λ hf, _⟩, -- First we use a finite basis to reduce the problem to the case `E = fin N → ℝ` - suffices : ∀ {N : ℕ} {g : α → fin N → ℝ}, summable g → summable (λ x, ∥g x∥), + suffices : ∀ {N : ℕ} {g : α → fin N → ℝ}, summable g → summable (λ x, ‖g x‖), { obtain v := fin_basis ℝ E, set e := v.equiv_funL, - have : summable (λ x, ∥e (f x)∥) := this (e.summable.2 hf), + have : summable (λ x, ‖e (f x)‖) := this (e.summable.2 hf), refine summable_of_norm_bounded _ (this.mul_left - ↑(∥(e.symm : (fin (finrank ℝ E) → ℝ) →L[ℝ] E)∥₊)) (λ i, _), + ↑(‖(e.symm : (fin (finrank ℝ E) → ℝ) →L[ℝ] E)‖₊)) (λ i, _), simpa using (e.symm : (fin (finrank ℝ E) → ℝ) →L[ℝ] E).le_op_norm (e $ f i) }, unfreezingI { clear_dependent E }, -- Now we deal with `g : α → fin N → ℝ` intros N g hg, - have : ∀ i, summable (λ x, ∥g x i∥) := λ i, (pi.summable.1 hg i).abs, + have : ∀ i, summable (λ x, ‖g x i‖) := λ i, (pi.summable.1 hg i).abs, refine summable_of_norm_bounded _ (summable_sum (λ i (hi : i ∈ finset.univ), this i)) (λ x, _), rw [norm_norm, pi_norm_le_iff_of_nonneg], { refine λ i, finset.single_le_sum (λ i hi, _) (finset.mem_univ i), diff --git a/src/analysis/normed_space/hahn_banach/extension.lean b/src/analysis/normed_space/hahn_banach/extension.lean index 1210370615d58..74898d0c531ca 100644 --- a/src/analysis/normed_space/hahn_banach/extension.lean +++ b/src/analysis/normed_space/hahn_banach/extension.lean @@ -23,7 +23,7 @@ In order to state and prove the corollaries uniformly, we prove the statements f satisfying `is_R_or_C 𝕜`. In this setting, `exists_dual_vector` states that, for any nonzero `x`, there exists a continuous -linear form `g` of norm `1` with `g x = ∥x∥` (where the norm has to be interpreted as an element +linear form `g` of norm `1` with `g x = ‖x‖` (where the norm has to be interpreted as an element of `𝕜`). -/ @@ -35,13 +35,13 @@ variables {E : Type*} [seminormed_add_comm_group E] [normed_space ℝ E] /-- Hahn-Banach theorem for continuous linear functions over `ℝ`. -/ theorem exists_extension_norm_eq (p : subspace ℝ E) (f : p →L[ℝ] ℝ) : - ∃ g : E →L[ℝ] ℝ, (∀ x : p, g x = f x) ∧ ∥g∥ = ∥f∥ := + ∃ g : E →L[ℝ] ℝ, (∀ x : p, g x = f x) ∧ ‖g‖ = ‖f‖ := begin - rcases exists_extension_of_le_sublinear ⟨p, f⟩ (λ x, ∥f∥ * ∥x∥) + rcases exists_extension_of_le_sublinear ⟨p, f⟩ (λ x, ‖f‖ * ‖x‖) (λ c hc x, by simp only [norm_smul c x, real.norm_eq_abs, abs_of_pos hc, mul_left_comm]) (λ x y, _) (λ x, le_trans (le_abs_self _) (f.le_op_norm _)) with ⟨g, g_eq, g_le⟩, - set g' := g.mk_continuous (∥f∥) + set g' := g.mk_continuous (‖f‖) (λ x, abs_le.2 ⟨neg_le.1 $ g.map_neg x ▸ norm_neg x ▸ g_le (-x), g_le x⟩), { refine ⟨g', g_eq, _⟩, { apply le_antisymm (g.mk_continuous_norm_le (norm_nonneg f) _), @@ -62,7 +62,7 @@ variables {𝕜 : Type*} [is_R_or_C 𝕜] {F : Type*} [seminormed_add_comm_group /-- Hahn-Banach theorem for continuous linear functions over `𝕜` satisyfing `is_R_or_C 𝕜`. -/ theorem exists_extension_norm_eq (p : subspace 𝕜 F) (f : p →L[𝕜] 𝕜) : - ∃ g : F →L[𝕜] 𝕜, (∀ x : p, g x = f x) ∧ ∥g∥ = ∥f∥ := + ∃ g : F →L[𝕜] 𝕜, (∀ x : p, g x = f x) ∧ ‖g‖ = ‖f‖ := begin letI : module ℝ F := restrict_scalars.module ℝ 𝕜 F, letI : is_scalar_tower ℝ 𝕜 F := restrict_scalars.is_scalar_tower _ _ _, @@ -91,11 +91,11 @@ begin sub_neg_eq_add, continuous_linear_map.map_smul] } }, -- And we derive the equality of the norms by bounding on both sides. refine ⟨h, le_antisymm _ _⟩, - { calc ∥g.extend_to_𝕜∥ - ≤ ∥g∥ : g.extend_to_𝕜.op_norm_le_bound g.op_norm_nonneg (norm_bound _) - ... = ∥fr∥ : hnormeq - ... ≤ ∥re_clm∥ * ∥f∥ : continuous_linear_map.op_norm_comp_le _ _ - ... = ∥f∥ : by rw [re_clm_norm, one_mul] }, + { calc ‖g.extend_to_𝕜‖ + ≤ ‖g‖ : g.extend_to_𝕜.op_norm_le_bound g.op_norm_nonneg (norm_bound _) + ... = ‖fr‖ : hnormeq + ... ≤ ‖re_clm‖ * ‖f‖ : continuous_linear_map.op_norm_comp_le _ _ + ... = ‖f‖ : by rw [re_clm_norm, one_mul] }, { exact f.op_norm_le_bound g.extend_to_𝕜.op_norm_nonneg (λ x, h x ▸ g.extend_to_𝕜.le_op_norm x) } end @@ -108,31 +108,31 @@ variables {E : Type u} [normed_add_comm_group E] [normed_space 𝕜 E] open continuous_linear_equiv submodule open_locale classical -lemma coord_norm' {x : E} (h : x ≠ 0) : ∥(∥x∥ : 𝕜) • coord 𝕜 x h∥ = 1 := +lemma coord_norm' {x : E} (h : x ≠ 0) : ‖(‖x‖ : 𝕜) • coord 𝕜 x h‖ = 1 := by rw [norm_smul, is_R_or_C.norm_coe_norm, coord_norm, mul_inv_cancel (mt norm_eq_zero.mp h)] /-- Corollary of Hahn-Banach. Given a nonzero element `x` of a normed space, there exists an - element of the dual space, of norm `1`, whose value on `x` is `∥x∥`. -/ -theorem exists_dual_vector (x : E) (h : x ≠ 0) : ∃ g : E →L[𝕜] 𝕜, ∥g∥ = 1 ∧ g x = ∥x∥ := + element of the dual space, of norm `1`, whose value on `x` is `‖x‖`. -/ +theorem exists_dual_vector (x : E) (h : x ≠ 0) : ∃ g : E →L[𝕜] 𝕜, ‖g‖ = 1 ∧ g x = ‖x‖ := begin let p : submodule 𝕜 E := 𝕜 ∙ x, - let f := (∥x∥ : 𝕜) • coord 𝕜 x h, + let f := (‖x‖ : 𝕜) • coord 𝕜 x h, obtain ⟨g, hg⟩ := exists_extension_norm_eq p f, refine ⟨g, _, _⟩, { rw [hg.2, coord_norm'] }, { calc g x = g (⟨x, mem_span_singleton_self x⟩ : 𝕜 ∙ x) : by rw coe_mk - ... = ((∥x∥ : 𝕜) • coord 𝕜 x h) (⟨x, mem_span_singleton_self x⟩ : 𝕜 ∙ x) : by rw ← hg.1 - ... = ∥x∥ : by simp } + ... = ((‖x‖ : 𝕜) • coord 𝕜 x h) (⟨x, mem_span_singleton_self x⟩ : 𝕜 ∙ x) : by rw ← hg.1 + ... = ‖x‖ : by simp } end /-- Variant of Hahn-Banach, eliminating the hypothesis that `x` be nonzero, and choosing the dual element arbitrarily when `x = 0`. -/ theorem exists_dual_vector' [nontrivial E] (x : E) : - ∃ g : E →L[𝕜] 𝕜, ∥g∥ = 1 ∧ g x = ∥x∥ := + ∃ g : E →L[𝕜] 𝕜, ‖g‖ = 1 ∧ g x = ‖x‖ := begin by_cases hx : x = 0, { obtain ⟨y, hy⟩ := exists_ne (0 : E), - obtain ⟨g, hg⟩ : ∃ g : E →L[𝕜] 𝕜, ∥g∥ = 1 ∧ g y = ∥y∥ := exists_dual_vector 𝕜 y hy, + obtain ⟨g, hg⟩ : ∃ g : E →L[𝕜] 𝕜, ‖g‖ = 1 ∧ g y = ‖y‖ := exists_dual_vector 𝕜 y hy, refine ⟨g, hg.left, _⟩, simp [hx] }, { exact exists_dual_vector 𝕜 x hx } @@ -142,7 +142,7 @@ end the dual element has norm at most `1` (this can not be improved for the trivial vector space). -/ theorem exists_dual_vector'' (x : E) : - ∃ g : E →L[𝕜] 𝕜, ∥g∥ ≤ 1 ∧ g x = ∥x∥ := + ∃ g : E →L[𝕜] 𝕜, ‖g‖ ≤ 1 ∧ g x = ‖x‖ := begin by_cases hx : x = 0, { refine ⟨0, by simp, _⟩, diff --git a/src/analysis/normed_space/indicator_function.lean b/src/analysis/normed_space/indicator_function.lean index a27af2a3975c7..0214d2e6582ca 100644 --- a/src/analysis/normed_space/indicator_function.lean +++ b/src/analysis/normed_space/indicator_function.lean @@ -20,22 +20,22 @@ variables {α E : Type*} [seminormed_add_comm_group E] {s t : set α} (f : α open set lemma norm_indicator_eq_indicator_norm : - ∥indicator s f a∥ = indicator s (λa, ∥f a∥) a := + ‖indicator s f a‖ = indicator s (λa, ‖f a‖) a := flip congr_fun a (indicator_comp_of_zero norm_zero).symm lemma nnnorm_indicator_eq_indicator_nnnorm : - ∥indicator s f a∥₊ = indicator s (λa, ∥f a∥₊) a := + ‖indicator s f a‖₊ = indicator s (λa, ‖f a‖₊) a := flip congr_fun a (indicator_comp_of_zero nnnorm_zero).symm lemma norm_indicator_le_of_subset (h : s ⊆ t) (f : α → E) (a : α) : - ∥indicator s f a∥ ≤ ∥indicator t f a∥ := + ‖indicator s f a‖ ≤ ‖indicator t f a‖ := begin simp only [norm_indicator_eq_indicator_norm], exact indicator_le_indicator_of_subset ‹_› (λ _, norm_nonneg _) _ end -lemma indicator_norm_le_norm_self : indicator s (λa, ∥f a∥) a ≤ ∥f a∥ := +lemma indicator_norm_le_norm_self : indicator s (λa, ‖f a‖) a ≤ ‖f a‖ := indicator_le_self' (λ _ _, norm_nonneg _) a -lemma norm_indicator_le_norm_self : ∥indicator s f a∥ ≤ ∥f a∥ := +lemma norm_indicator_le_norm_self : ‖indicator s f a‖ ≤ ‖f a‖ := by { rw norm_indicator_eq_indicator_norm, apply indicator_norm_le_norm_self } diff --git a/src/analysis/normed_space/int.lean b/src/analysis/normed_space/int.lean index 6db4eff292aff..786f3f7d5fe07 100644 --- a/src/analysis/normed_space/int.lean +++ b/src/analysis/normed_space/int.lean @@ -10,33 +10,33 @@ import analysis.normed.field.basic This file contains basic facts about the integers as normed ring. -Recall that `∥n∥` denotes the norm of `n` as real number. +Recall that `‖n‖` denotes the norm of `n` as real number. This norm is always nonnegative, so we can bundle the norm together with this fact, to obtain a term of type `nnreal` (the nonnegative real numbers). -The resulting nonnegative real number is denoted by `∥n∥₊`. +The resulting nonnegative real number is denoted by `‖n‖₊`. -/ open_locale big_operators namespace int -lemma nnnorm_coe_units (e : ℤˣ) : ∥(e : ℤ)∥₊ = 1 := +lemma nnnorm_coe_units (e : ℤˣ) : ‖(e : ℤ)‖₊ = 1 := begin obtain (rfl|rfl) := int.units_eq_one_or e; simp only [units.coe_neg_one, units.coe_one, nnnorm_neg, nnnorm_one], end -lemma norm_coe_units (e : ℤˣ) : ∥(e : ℤ)∥ = 1 := +lemma norm_coe_units (e : ℤˣ) : ‖(e : ℤ)‖ = 1 := by rw [← coe_nnnorm, int.nnnorm_coe_units, nnreal.coe_one] -@[simp] lemma nnnorm_coe_nat (n : ℕ) : ∥(n : ℤ)∥₊ = n := real.nnnorm_coe_nat _ +@[simp] lemma nnnorm_coe_nat (n : ℕ) : ‖(n : ℤ)‖₊ = n := real.nnnorm_coe_nat _ -@[simp] lemma norm_coe_nat (n : ℕ) : ∥(n : ℤ)∥ = n := real.norm_coe_nat _ +@[simp] lemma norm_coe_nat (n : ℕ) : ‖(n : ℤ)‖ = n := real.norm_coe_nat _ -@[simp] lemma to_nat_add_to_nat_neg_eq_nnnorm (n : ℤ) : ↑(n.to_nat) + ↑((-n).to_nat) = ∥n∥₊ := +@[simp] lemma to_nat_add_to_nat_neg_eq_nnnorm (n : ℤ) : ↑(n.to_nat) + ↑((-n).to_nat) = ‖n‖₊ := by rw [← nat.cast_add, to_nat_add_to_nat_neg_eq_nat_abs, nnreal.coe_nat_abs] -@[simp] lemma to_nat_add_to_nat_neg_eq_norm (n : ℤ) : ↑(n.to_nat) + ↑((-n).to_nat) = ∥n∥ := +@[simp] lemma to_nat_add_to_nat_neg_eq_norm (n : ℤ) : ↑(n.to_nat) + ↑((-n).to_nat) = ‖n‖ := by simpa only [nnreal.coe_nat_cast, nnreal.coe_add] using congr_arg (coe : _ → ℝ) (to_nat_add_to_nat_neg_eq_nnnorm n) diff --git a/src/analysis/normed_space/is_R_or_C.lean b/src/analysis/normed_space/is_R_or_C.lean index ee2efcf98847a..149808aefc381 100644 --- a/src/analysis/normed_space/is_R_or_C.lean +++ b/src/analysis/normed_space/is_R_or_C.lean @@ -30,38 +30,38 @@ open metric variables {𝕜 : Type*} [is_R_or_C 𝕜] {E : Type*} [normed_add_comm_group E] -lemma is_R_or_C.norm_coe_norm {z : E} : ∥(∥z∥ : 𝕜)∥ = ∥z∥ := by simp +lemma is_R_or_C.norm_coe_norm {z : E} : ‖(‖z‖ : 𝕜)‖ = ‖z‖ := by simp variables [normed_space 𝕜 E] /-- Lemma to normalize a vector in a normed space `E` over either `ℂ` or `ℝ` to unit length. -/ -@[simp] lemma norm_smul_inv_norm {x : E} (hx : x ≠ 0) : ∥(∥x∥⁻¹ : 𝕜) • x∥ = 1 := +@[simp] lemma norm_smul_inv_norm {x : E} (hx : x ≠ 0) : ‖(‖x‖⁻¹ : 𝕜) • x‖ = 1 := begin - have : ∥x∥ ≠ 0 := by simp [hx], + have : ‖x‖ ≠ 0 := by simp [hx], field_simp [norm_smul] end /-- Lemma to normalize a vector in a normed space `E` over either `ℂ` or `ℝ` to length `r`. -/ lemma norm_smul_inv_norm' {r : ℝ} (r_nonneg : 0 ≤ r) {x : E} (hx : x ≠ 0) : - ∥(r * ∥x∥⁻¹ : 𝕜) • x∥ = r := + ‖(r * ‖x‖⁻¹ : 𝕜) • x‖ = r := begin - have : ∥x∥ ≠ 0 := by simp [hx], + have : ‖x‖ ≠ 0 := by simp [hx], field_simp [norm_smul, is_R_or_C.norm_eq_abs, r_nonneg] with is_R_or_C_simps end lemma linear_map.bound_of_sphere_bound - {r : ℝ} (r_pos : 0 < r) (c : ℝ) (f : E →ₗ[𝕜] 𝕜) (h : ∀ z ∈ sphere (0 : E) r, ∥f z∥ ≤ c) (z : E) : - ∥f z∥ ≤ c / r * ∥z∥ := + {r : ℝ} (r_pos : 0 < r) (c : ℝ) (f : E →ₗ[𝕜] 𝕜) (h : ∀ z ∈ sphere (0 : E) r, ‖f z‖ ≤ c) (z : E) : + ‖f z‖ ≤ c / r * ‖z‖ := begin by_cases z_zero : z = 0, { rw z_zero, simp only [linear_map.map_zero, norm_zero, mul_zero], }, - set z₁ := (r * ∥z∥⁻¹ : 𝕜) • z with hz₁, - have norm_f_z₁ : ∥f z₁∥ ≤ c, + set z₁ := (r * ‖z‖⁻¹ : 𝕜) • z with hz₁, + have norm_f_z₁ : ‖f z₁‖ ≤ c, { apply h, rw mem_sphere_zero_iff_norm, exact norm_smul_inv_norm' r_pos.le z_zero }, have r_ne_zero : (r : 𝕜) ≠ 0 := is_R_or_C.of_real_ne_zero.mpr r_pos.ne', - have eq : f z = ∥z∥ / r * (f z₁), + have eq : f z = ‖z‖ / r * (f z₁), { rw [hz₁, linear_map.map_smul, smul_eq_mul], rw [← mul_assoc, ← mul_assoc, div_mul_cancel _ r_ne_zero, mul_inv_cancel, one_mul], simp only [z_zero, is_R_or_C.of_real_eq_zero, norm_eq_zero, ne.def, not_false_iff], }, @@ -76,13 +76,13 @@ end `linear_map.bound_of_ball_bound` is a version of this over arbitrary nontrivially normed fields. It produces a less precise bound so we keep both versions. -/ lemma linear_map.bound_of_ball_bound' {r : ℝ} (r_pos : 0 < r) (c : ℝ) (f : E →ₗ[𝕜] 𝕜) - (h : ∀ z ∈ closed_ball (0 : E) r, ∥f z∥ ≤ c) (z : E) : - ∥f z∥ ≤ c / r * ∥z∥ := + (h : ∀ z ∈ closed_ball (0 : E) r, ‖f z‖ ≤ c) (z : E) : + ‖f z‖ ≤ c / r * ‖z‖ := f.bound_of_sphere_bound r_pos c (λ z hz, h z hz.le) z lemma continuous_linear_map.op_norm_bound_of_ball_bound - {r : ℝ} (r_pos : 0 < r) (c : ℝ) (f : E →L[𝕜] 𝕜) (h : ∀ z ∈ closed_ball (0 : E) r, ∥f z∥ ≤ c) : - ∥f∥ ≤ c / r := + {r : ℝ} (r_pos : 0 < r) (c : ℝ) (f : E →L[𝕜] 𝕜) (h : ∀ z ∈ closed_ball (0 : E) r, ‖f z‖ ≤ c) : + ‖f‖ ≤ c / r := begin apply continuous_linear_map.op_norm_le_bound, { apply div_nonneg _ r_pos.le, diff --git a/src/analysis/normed_space/linear_isometry.lean b/src/analysis/normed_space/linear_isometry.lean index 799ca8e08e90a..197d40bd68a70 100644 --- a/src/analysis/normed_space/linear_isometry.lean +++ b/src/analysis/normed_space/linear_isometry.lean @@ -18,7 +18,7 @@ the star-linear versions. We also prove some trivial lemmas and provide convenience constructors. -Since a lot of elementary properties don't require `∥x∥ = 0 → x = 0` we start setting up the +Since a lot of elementary properties don't require `‖x‖ = 0 → x = 0` we start setting up the theory for `seminormed_add_comm_group` and we specialize to `normed_add_comm_group` when needed. -/ open function set @@ -44,7 +44,7 @@ variables {R R₂ R₃ R₄ E E₂ E₃ E₄ F 𝓕 : Type*} [semiring R] [semir /-- A `σ₁₂`-semilinear isometric embedding of a normed `R`-module into an `R₂`-module. -/ structure linear_isometry (σ₁₂ : R →+* R₂) (E E₂ : Type*) [seminormed_add_comm_group E] [seminormed_add_comm_group E₂] [module R E] [module R₂ E₂] extends E →ₛₗ[σ₁₂] E₂ := -(norm_map' : ∀ x, ∥to_linear_map x∥ = ∥x∥) +(norm_map' : ∀ x, ‖to_linear_map x‖ = ‖x‖) notation E ` →ₛₗᵢ[`:25 σ₁₂:25 `] `:0 E₂:0 := linear_isometry σ₁₂ E E₂ notation E ` →ₗᵢ[`:25 R:25 `] `:0 E₂:0 := linear_isometry (ring_hom.id R) E E₂ @@ -63,7 +63,7 @@ class semilinear_isometry_class (𝓕 : Type*) {R R₂ : out_param Type*} [semir (σ₁₂ : out_param $ R →+* R₂) (E E₂ : out_param Type*) [seminormed_add_comm_group E] [seminormed_add_comm_group E₂] [module R E] [module R₂ E₂] extends semilinear_map_class 𝓕 σ₁₂ E E₂ := -(norm_map : ∀ (f : 𝓕) (x : E), ∥f x∥ = ∥x∥) +(norm_map : ∀ (f : 𝓕) (x : E), ‖f x‖ = ‖x‖) /-- `linear_isometry_class F R E E₂` asserts `F` is a type of bundled `R`-linear isometries `M → M₂`. @@ -86,7 +86,7 @@ add_monoid_hom_class.isometry_of_norm _ (norm_map _) (semilinear_isometry_class.isometry f).continuous @[simp] lemma nnnorm_map [semilinear_isometry_class 𝓕 σ₁₂ E E₂] (f : 𝓕) (x : E) : - ∥f x∥₊ = ∥x∥₊ := + ‖f x‖₊ = ‖x‖₊ := nnreal.eq $ norm_map f x protected lemma lipschitz [semilinear_isometry_class 𝓕 σ₁₂ E E₂] (f : 𝓕) : @@ -182,9 +182,9 @@ f.to_linear_map.map_smulₛₗ c x f (c • x) = c • f x := f.to_linear_map.map_smul c x -@[simp] lemma norm_map (x : E) : ∥f x∥ = ∥x∥ := semilinear_isometry_class.norm_map f x +@[simp] lemma norm_map (x : E) : ‖f x‖ = ‖x‖ := semilinear_isometry_class.norm_map f x -@[simp] lemma nnnorm_map (x : E) : ∥f x∥₊ = ∥x∥₊ := nnreal.eq $ norm_map f x +@[simp] lemma nnnorm_map (x : E) : ‖f x‖₊ = ‖x‖₊ := nnreal.eq $ norm_map f x protected lemma isometry : isometry f := add_monoid_hom_class.isometry_of_norm _ (norm_map _) @@ -338,7 +338,7 @@ end submodule structure linear_isometry_equiv (σ₁₂ : R →+* R₂) {σ₂₁ : R₂ →+* R} [ring_hom_inv_pair σ₁₂ σ₂₁] [ring_hom_inv_pair σ₂₁ σ₁₂] (E E₂ : Type*) [seminormed_add_comm_group E] [seminormed_add_comm_group E₂] [module R E] [module R₂ E₂] extends E ≃ₛₗ[σ₁₂] E₂ := -(norm_map' : ∀ x, ∥to_linear_equiv x∥ = ∥x∥) +(norm_map' : ∀ x, ‖to_linear_equiv x‖ = ‖x‖) notation E ` ≃ₛₗᵢ[`:25 σ₁₂:25 `] `:0 E₂:0 := linear_isometry_equiv σ₁₂ E E₂ notation E ` ≃ₗᵢ[`:25 R:25 `] `:0 E₂:0 := linear_isometry_equiv (ring_hom.id R) E E₂ @@ -359,7 +359,7 @@ class semilinear_isometry_equiv_class (𝓕 : Type*) {R R₂ : out_param Type*} [ring_hom_inv_pair σ₁₂ σ₂₁] [ring_hom_inv_pair σ₂₁ σ₁₂] (E E₂ : out_param Type*) [seminormed_add_comm_group E] [seminormed_add_comm_group E₂] [module R E] [module R₂ E₂] extends semilinear_equiv_class 𝓕 σ₁₂ E E₂ := -(norm_map : ∀ (f : 𝓕) (x : E), ∥f x∥ = ∥x∥) +(norm_map : ∀ (f : 𝓕) (x : E), ‖f x‖ = ‖x‖) /-- `linear_isometry_equiv_class F R E E₂` asserts `F` is a type of bundled `R`-linear isometries `M → M₂`. @@ -419,7 +419,7 @@ instance : has_coe_to_fun (E ≃ₛₗᵢ[σ₁₂] E₂) (λ _, E → E₂) := lemma coe_injective : @function.injective (E ≃ₛₗᵢ[σ₁₂] E₂) (E → E₂) coe_fn := fun_like.coe_injective -@[simp] lemma coe_mk (e : E ≃ₛₗ[σ₁₂] E₂) (he : ∀ x, ∥e x∥ = ∥x∥) : +@[simp] lemma coe_mk (e : E ≃ₛₗ[σ₁₂] E₂) (he : ∀ x, ‖e x‖ = ‖x‖) : ⇑(mk e he) = e := rfl @@ -434,12 +434,12 @@ protected lemma congr_arg {f : E ≃ₛₗᵢ[σ₁₂] E₂} : Π {x x' : E}, x protected lemma congr_fun {f g : E ≃ₛₗᵢ[σ₁₂] E₂} (h : f = g) (x : E) : f x = g x := h ▸ rfl /-- Construct a `linear_isometry_equiv` from a `linear_equiv` and two inequalities: -`∀ x, ∥e x∥ ≤ ∥x∥` and `∀ y, ∥e.symm y∥ ≤ ∥y∥`. -/ -def of_bounds (e : E ≃ₛₗ[σ₁₂] E₂) (h₁ : ∀ x, ∥e x∥ ≤ ∥x∥) (h₂ : ∀ y, ∥e.symm y∥ ≤ ∥y∥) : +`∀ x, ‖e x‖ ≤ ‖x‖` and `∀ y, ‖e.symm y‖ ≤ ‖y‖`. -/ +def of_bounds (e : E ≃ₛₗ[σ₁₂] E₂) (h₁ : ∀ x, ‖e x‖ ≤ ‖x‖) (h₂ : ∀ y, ‖e.symm y‖ ≤ ‖y‖) : E ≃ₛₗᵢ[σ₁₂] E₂ := ⟨e, λ x, le_antisymm (h₁ x) $ by simpa only [e.symm_apply_apply] using h₂ (e x)⟩ -@[simp] lemma norm_map (x : E) : ∥e x∥ = ∥x∥ := e.norm_map' x +@[simp] lemma norm_map (x : E) : ‖e x‖ = ‖x‖ := e.norm_map' x /-- Reinterpret a `linear_isometry_equiv` as a `linear_isometry`. -/ def to_linear_isometry : E →ₛₗᵢ[σ₁₂] E₂ := ⟨e.1, e.2⟩ @@ -647,7 +647,7 @@ omit σ₂₁ @[simp] lemma map_smul [module R E₂] {e : E ≃ₗᵢ[R] E₂} (c : R) (x : E) : e (c • x) = c • e x := e.1.map_smul c x -@[simp] lemma nnnorm_map (x : E) : ∥e x∥₊ = ∥x∥₊ := semilinear_isometry_class.nnnorm_map e x +@[simp] lemma nnnorm_map (x : E) : ‖e x‖₊ = ‖x‖₊ := semilinear_isometry_class.nnnorm_map e x @[simp] lemma dist_map (x y : E) : dist (e x) (e y) = dist x y := e.to_linear_isometry.dist_map x y diff --git a/src/analysis/normed_space/lp_equiv.lean b/src/analysis/normed_space/lp_equiv.lean index 8a761c9788197..ce69795dbb044 100644 --- a/src/analysis/normed_space/lp_equiv.lean +++ b/src/analysis/normed_space/lp_equiv.lean @@ -46,7 +46,7 @@ lemma mem_ℓp.all [finite α] (f : Π i, E i) : mem_ℓp f p := begin rcases p.trichotomy with (rfl | rfl | h), { exact mem_ℓp_zero_iff.mpr {i : α | f i ≠ 0}.to_finite, }, - { exact mem_ℓp_infty_iff.mpr (set.finite.bdd_above (set.range (λ (i : α), ∥f i∥)).to_finite) }, + { exact mem_ℓp_infty_iff.mpr (set.finite.bdd_above (set.range (λ (i : α), ‖f i‖)).to_finite) }, { casesI nonempty_fintype α, exact mem_ℓp_gen ⟨finset.univ.sum _, has_sum_fintype _⟩ } end @@ -62,7 +62,7 @@ def equiv.lp_pi_Lp : lp E p ≃ pi_Lp p E := lemma coe_equiv_lp_pi_Lp (f : lp E p) : equiv.lp_pi_Lp f = f := rfl lemma coe_equiv_lp_pi_Lp_symm (f : pi_Lp p E) : (equiv.lp_pi_Lp.symm f : Π i, E i) = f := rfl -lemma equiv_lp_pi_Lp_norm (f : lp E p) : ∥equiv.lp_pi_Lp f∥ = ∥f∥ := +lemma equiv_lp_pi_Lp_norm (f : lp E p) : ‖equiv.lp_pi_Lp f‖ = ‖f‖ := begin unfreezingI { rcases p.trichotomy with (rfl | rfl | h) }, { rw [pi_Lp.norm_eq_card, lp.norm_eq_card_dsupport], refl }, @@ -118,7 +118,7 @@ section normed_add_comm_group /-- The canonical map between `lp (λ (_ : α), E) ∞` and `α →ᵇ E` as an `add_equiv`. -/ noncomputable def add_equiv.lp_bcf : lp (λ (_ : α), E) ∞ ≃+ (α →ᵇ E) := -{ to_fun := λ f, of_normed_add_comm_group_discrete f (∥f∥) $ le_csupr (mem_ℓp_infty_iff.mp f.prop), +{ to_fun := λ f, of_normed_add_comm_group_discrete f (‖f‖) $ le_csupr (mem_ℓp_infty_iff.mp f.prop), inv_fun := λ f, ⟨f, f.bdd_above_range_norm_comp⟩, left_inv := λ f, lp.ext rfl, right_inv := λ f, ext $ λ x, rfl, diff --git a/src/analysis/normed_space/lp_space.lean b/src/analysis/normed_space/lp_space.lean index e43273cd02880..52dae33088188 100644 --- a/src/analysis/normed_space/lp_space.lean +++ b/src/analysis/normed_space/lp_space.lean @@ -11,11 +11,11 @@ import topology.algebra.order.liminf_limsup # ℓp space This file describes properties of elements `f` of a pi-type `Π i, E i` with finite "norm", -defined for `p:ℝ≥0∞` as the size of the support of `f` if `p=0`, `(∑' a, ∥f a∥^p) ^ (1/p)` for -`0 < p < ∞` and `⨆ a, ∥f a∥` for `p=∞`. +defined for `p:ℝ≥0∞` as the size of the support of `f` if `p=0`, `(∑' a, ‖f a‖^p) ^ (1/p)` for +`0 < p < ∞` and `⨆ a, ‖f a‖` for `p=∞`. The Prop-valued `mem_ℓp f p` states that a function `f : Π i, E i` has finite norm according -to the above definition; that is, `f` has finite support if `p = 0`, `summable (λ a, ∥f a∥^p)` if +to the above definition; that is, `f` has finite support if `p = 0`, `summable (λ a, ‖f a‖^p)` if `0 < p < ∞`, and `bdd_above (norm '' (set.range f))` if `p = ∞`. The space `lp E p` is the subtype of elements of `Π i : α, E i` which satisfy `mem_ℓp f p`. For @@ -24,7 +24,7 @@ The space `lp E p` is the subtype of elements of `Π i : α, E i` which satisfy ## Main definitions * `mem_ℓp f p` : property that the function `f` satisfies, as appropriate, `f` finitely supported - if `p = 0`, `summable (λ a, ∥f a∥^p)` if `0 < p < ∞`, and `bdd_above (norm '' (set.range f))` if + if `p = 0`, `summable (λ a, ‖f a‖^p)` if `0 < p < ∞`, and `bdd_above (norm '' (set.range f))` if `p = ∞`. * `lp E p` : elements of `Π i : α, E i` such that `mem_ℓp f p`. Defined as an `add_subgroup` of a type synonym `pre_lp` for `Π i : α, E i`, and equipped with a `normed_add_comm_group` structure. @@ -43,12 +43,12 @@ The space `lp E p` is the subtype of elements of `Π i : α, E i` which satisfy ## Implementation Since `lp` is defined as an `add_subgroup`, dot notation does not work. Use `lp.norm_neg f` to -say that `∥-f∥ = ∥f∥`, instead of the non-working `f.norm_neg`. +say that `‖-f‖ = ‖f‖`, instead of the non-working `f.norm_neg`. ## TODO * More versions of Hölder's inequality (for example: the case `p = 1`, `q = ∞`; a version for normed - rings which has `∥∑' i, f i * g i∥` rather than `∑' i, ∥f i∥ * g i∥` on the RHS; a version for + rings which has `‖∑' i, f i * g i‖` rather than `∑' i, ‖f i‖ * g i‖` on the RHS; a version for three exponents satisfying `1 / r = 1 / p + 1 / q`) -/ @@ -65,11 +65,11 @@ variables {α : Type*} {E : α → Type*} {p q : ℝ≥0∞} [Π i, normed_add_c /-- The property that `f : Π i : α, E i` * is finitely supported, if `p = 0`, or -* admits an upper bound for `set.range (λ i, ∥f i∥)`, if `p = ∞`, or -* has the series `∑' i, ∥f i∥ ^ p` be summable, if `0 < p < ∞`. -/ +* admits an upper bound for `set.range (λ i, ‖f i‖)`, if `p = ∞`, or +* has the series `∑' i, ‖f i‖ ^ p` be summable, if `0 < p < ∞`. -/ def mem_ℓp (f : Π i, E i) (p : ℝ≥0∞) : Prop := if p = 0 then (set.finite {i | f i ≠ 0}) else - (if p = ∞ then bdd_above (set.range (λ i, ∥f i∥)) else summable (λ i, ∥f i∥ ^ p.to_real)) + (if p = ∞ then bdd_above (set.range (λ i, ‖f i‖)) else summable (λ i, ‖f i‖ ^ p.to_real)) lemma mem_ℓp_zero_iff {f : Π i, E i} : mem_ℓp f 0 ↔ set.finite {i | f i ≠ 0} := by dsimp [mem_ℓp]; rw [if_pos rfl] @@ -77,21 +77,21 @@ by dsimp [mem_ℓp]; rw [if_pos rfl] lemma mem_ℓp_zero {f : Π i, E i} (hf : set.finite {i | f i ≠ 0}) : mem_ℓp f 0 := mem_ℓp_zero_iff.2 hf -lemma mem_ℓp_infty_iff {f : Π i, E i} : mem_ℓp f ∞ ↔ bdd_above (set.range (λ i, ∥f i∥)) := +lemma mem_ℓp_infty_iff {f : Π i, E i} : mem_ℓp f ∞ ↔ bdd_above (set.range (λ i, ‖f i‖)) := by dsimp [mem_ℓp]; rw [if_neg ennreal.top_ne_zero, if_pos rfl] -lemma mem_ℓp_infty {f : Π i, E i} (hf : bdd_above (set.range (λ i, ∥f i∥))) : mem_ℓp f ∞ := +lemma mem_ℓp_infty {f : Π i, E i} (hf : bdd_above (set.range (λ i, ‖f i‖))) : mem_ℓp f ∞ := mem_ℓp_infty_iff.2 hf lemma mem_ℓp_gen_iff (hp : 0 < p.to_real) {f : Π i, E i} : - mem_ℓp f p ↔ summable (λ i, ∥f i∥ ^ p.to_real) := + mem_ℓp f p ↔ summable (λ i, ‖f i‖ ^ p.to_real) := begin rw ennreal.to_real_pos_iff at hp, dsimp [mem_ℓp], rw [if_neg hp.1.ne', if_neg hp.2.ne], end -lemma mem_ℓp_gen {f : Π i, E i} (hf : summable (λ i, ∥f i∥ ^ p.to_real)) : +lemma mem_ℓp_gen {f : Π i, E i} (hf : summable (λ i, ‖f i‖ ^ p.to_real)) : mem_ℓp f p := begin rcases p.trichotomy with rfl | rfl | hp, @@ -100,15 +100,15 @@ begin exact (finite_of_summable_const (by norm_num) H).subset (set.subset_univ _) }, { apply mem_ℓp_infty, have H : summable (λ i : α, (1:ℝ)) := by simpa using hf, - simpa using ((finite_of_summable_const (by norm_num) H).image (λ i, ∥f i∥)).bdd_above }, + simpa using ((finite_of_summable_const (by norm_num) H).image (λ i, ‖f i‖)).bdd_above }, exact (mem_ℓp_gen_iff hp).2 hf end -lemma mem_ℓp_gen' {C : ℝ} {f : Π i, E i} (hf : ∀ s : finset α, ∑ i in s, ∥f i∥ ^ p.to_real ≤ C) : +lemma mem_ℓp_gen' {C : ℝ} {f : Π i, E i} (hf : ∀ s : finset α, ∑ i in s, ‖f i‖ ^ p.to_real ≤ C) : mem_ℓp f p := begin apply mem_ℓp_gen, - use ⨆ s : finset α, ∑ i in s, ∥f i∥ ^ p.to_real, + use ⨆ s : finset α, ∑ i in s, ‖f i‖ ^ p.to_real, apply has_sum_of_is_lub_of_nonneg, { intros b, exact real.rpow_nonneg_of_nonneg (norm_nonneg _) _ }, @@ -137,11 +137,11 @@ namespace mem_ℓp lemma finite_dsupport {f : Π i, E i} (hf : mem_ℓp f 0) : set.finite {i | f i ≠ 0} := mem_ℓp_zero_iff.1 hf -lemma bdd_above {f : Π i, E i} (hf : mem_ℓp f ∞) : bdd_above (set.range (λ i, ∥f i∥)) := +lemma bdd_above {f : Π i, E i} (hf : mem_ℓp f ∞) : bdd_above (set.range (λ i, ‖f i‖)) := mem_ℓp_infty_iff.1 hf lemma summable (hp : 0 < p.to_real) {f : Π i, E i} (hf : mem_ℓp f p) : - summable (λ i, ∥f i∥ ^ p.to_real) := + summable (λ i, ‖f i‖ ^ p.to_real) := (mem_ℓp_gen_iff hp).1 hf lemma neg {f : Π i, E i} (hf : mem_ℓp f p) : mem_ℓp (-f) p := @@ -166,14 +166,14 @@ begin | ⟨hq, hp, hpq'⟩, { exact hfq }, { apply mem_ℓp_infty, - obtain ⟨C, hC⟩ := (hfq.finite_dsupport.image (λ i, ∥f i∥)).bdd_above, + obtain ⟨C, hC⟩ := (hfq.finite_dsupport.image (λ i, ‖f i‖)).bdd_above, use max 0 C, rintros x ⟨i, rfl⟩, by_cases hi : f i = 0, { simp [hi] }, { exact (hC ⟨i, hi, rfl⟩).trans (le_max_right _ _) } }, { apply mem_ℓp_gen, - have : ∀ i ∉ hfq.finite_dsupport.to_finset, ∥f i∥ ^ p.to_real = 0, + have : ∀ i ∉ hfq.finite_dsupport.to_finset, ‖f i‖ ^ p.to_real = 0, { intros i hi, have : f i = 0 := by simpa using hi, simp [this, real.zero_rpow hp.ne'] }, @@ -183,19 +183,19 @@ begin obtain ⟨A, hA⟩ := (hfq.summable hq).tendsto_cofinite_zero.bdd_above_range_of_cofinite, use A ^ (q.to_real⁻¹), rintros x ⟨i, rfl⟩, - have : 0 ≤ ∥f i∥ ^ q.to_real := real.rpow_nonneg_of_nonneg (norm_nonneg _) _, + have : 0 ≤ ‖f i‖ ^ q.to_real := real.rpow_nonneg_of_nonneg (norm_nonneg _) _, simpa [← real.rpow_mul, mul_inv_cancel hq.ne'] using real.rpow_le_rpow this (hA ⟨i, rfl⟩) (inv_nonneg.mpr hq.le) }, { apply mem_ℓp_gen, have hf' := hfq.summable hq, - refine summable_of_norm_bounded_eventually _ hf' (@set.finite.subset _ {i | 1 ≤ ∥f i∥} _ _ _), - { have H : {x : α | 1 ≤ ∥f x∥ ^ q.to_real}.finite, + refine summable_of_norm_bounded_eventually _ hf' (@set.finite.subset _ {i | 1 ≤ ‖f i‖} _ _ _), + { have H : {x : α | 1 ≤ ‖f x‖ ^ q.to_real}.finite, { simpa using eventually_lt_of_tendsto_lt (by norm_num : (0:ℝ) < 1) hf'.tendsto_cofinite_zero }, exact H.subset (λ i hi, real.one_le_rpow hi hq.le) }, - { show ∀ i, ¬ (|∥f i∥ ^ p.to_real| ≤ ∥f i∥ ^ q.to_real) → 1 ≤ ∥f i∥, + { show ∀ i, ¬ (|‖f i‖ ^ p.to_real| ≤ ‖f i‖ ^ q.to_real) → 1 ≤ ‖f i‖, intros i hi, - have : 0 ≤ ∥f i∥ ^ p.to_real := real.rpow_nonneg_of_nonneg (norm_nonneg _) p.to_real, + have : 0 ≤ ‖f i‖ ^ p.to_real := real.rpow_nonneg_of_nonneg (norm_nonneg _) p.to_real, simp only [abs_of_nonneg, this] at hi, contrapose! hi, exact real.rpow_le_rpow_of_exponent_ge' (norm_nonneg _) hi.le hq.le hpq' } } @@ -223,8 +223,8 @@ begin { refine (real.rpow_le_rpow (norm_nonneg _) (norm_add_le _ _) hp.le).trans _, dsimp [C], split_ifs with h h, - { simpa using nnreal.coe_le_coe.2 (nnreal.rpow_add_le_add_rpow (∥f i∥₊) (∥g i∥₊) hp.le h.le) }, - { let F : fin 2 → ℝ≥0 := ![∥f i∥₊, ∥g i∥₊], + { simpa using nnreal.coe_le_coe.2 (nnreal.rpow_add_le_add_rpow (‖f i‖₊) (‖g i‖₊) hp.le h.le) }, + { let F : fin 2 → ℝ≥0 := ![‖f i‖₊, ‖g i‖₊], have : ∀ i, (0:ℝ) ≤ F i := λ i, (F i).coe_nonneg, simp only [not_lt] at h, simpa [F, fin.sum_univ_succ] using @@ -258,11 +258,11 @@ begin refine hf.finite_dsupport.subset (λ i, (_ : ¬c • f i = 0 → ¬f i = 0)), exact not_imp_not.mpr (λ hf', hf'.symm ▸ (smul_zero c)) }, { obtain ⟨A, hA⟩ := hf.bdd_above, - refine mem_ℓp_infty ⟨∥c∥ * A, _⟩, + refine mem_ℓp_infty ⟨‖c‖ * A, _⟩, rintros a ⟨i, rfl⟩, simpa [norm_smul] using mul_le_mul_of_nonneg_left (hA ⟨i, rfl⟩) (norm_nonneg c) }, { apply mem_ℓp_gen, - convert (hf.summable hp).mul_left (∥c∥ ^ p.to_real), + convert (hf.summable hp).mul_left (‖c‖ ^ p.to_real), ext i, simp [norm_smul, real.mul_rpow (norm_nonneg c) (norm_nonneg (f i))] }, end @@ -340,25 +340,25 @@ end instance : has_norm (lp E p) := { norm := λ f, if hp : p = 0 then by subst hp; exact (lp.mem_ℓp f).finite_dsupport.to_finset.card - else (if p = ∞ then ⨆ i, ∥f i∥ else (∑' i, ∥f i∥ ^ p.to_real) ^ (1/p.to_real)) } + else (if p = ∞ then ⨆ i, ‖f i‖ else (∑' i, ‖f i‖ ^ p.to_real) ^ (1/p.to_real)) } -lemma norm_eq_card_dsupport (f : lp E 0) : ∥f∥ = (lp.mem_ℓp f).finite_dsupport.to_finset.card := +lemma norm_eq_card_dsupport (f : lp E 0) : ‖f‖ = (lp.mem_ℓp f).finite_dsupport.to_finset.card := dif_pos rfl -lemma norm_eq_csupr (f : lp E ∞) : ∥f∥ = ⨆ i, ∥f i∥ := +lemma norm_eq_csupr (f : lp E ∞) : ‖f‖ = ⨆ i, ‖f i‖ := begin dsimp [norm], rw [dif_neg ennreal.top_ne_zero, if_pos rfl] end -lemma is_lub_norm [nonempty α] (f : lp E ∞) : is_lub (set.range (λ i, ∥f i∥)) ∥f∥ := +lemma is_lub_norm [nonempty α] (f : lp E ∞) : is_lub (set.range (λ i, ‖f i‖)) ‖f‖ := begin rw lp.norm_eq_csupr, exact is_lub_csupr (lp.mem_ℓp f) end lemma norm_eq_tsum_rpow (hp : 0 < p.to_real) (f : lp E p) : - ∥f∥ = (∑' i, ∥f i∥ ^ p.to_real) ^ (1/p.to_real) := + ‖f‖ = (∑' i, ‖f i‖ ^ p.to_real) ^ (1/p.to_real) := begin dsimp [norm], rw ennreal.to_real_pos_iff at hp, @@ -366,7 +366,7 @@ begin end lemma norm_rpow_eq_tsum (hp : 0 < p.to_real) (f : lp E p) : - ∥f∥ ^ p.to_real = ∑' i, ∥f i∥ ^ p.to_real := + ‖f‖ ^ p.to_real = ∑' i, ‖f i‖ ^ p.to_real := begin rw [norm_eq_tsum_rpow hp, ← real.rpow_mul], { field_simp [hp.ne'] }, @@ -377,13 +377,13 @@ begin end lemma has_sum_norm (hp : 0 < p.to_real) (f : lp E p) : - has_sum (λ i, ∥f i∥ ^ p.to_real) (∥f∥ ^ p.to_real) := + has_sum (λ i, ‖f i‖ ^ p.to_real) (‖f‖ ^ p.to_real) := begin rw norm_rpow_eq_tsum hp, exact ((lp.mem_ℓp f).summable hp).has_sum end -lemma norm_nonneg' (f : lp E p) : 0 ≤ ∥f∥ := +lemma norm_nonneg' (f : lp E p) : 0 ≤ ‖f‖ := begin rcases p.trichotomy with rfl | rfl | hp, { simp [lp.norm_eq_card_dsupport f] }, @@ -397,7 +397,7 @@ begin exact λ i, real.rpow_nonneg_of_nonneg (norm_nonneg _) _ }, end -@[simp] lemma norm_zero : ∥(0 : lp E p)∥ = 0 := +@[simp] lemma norm_zero : ‖(0 : lp E p)‖ = 0 := begin rcases p.trichotomy with rfl | rfl | hp, { simp [lp.norm_eq_card_dsupport] }, @@ -407,7 +407,7 @@ begin simpa [real.zero_rpow hp.ne'] using real.zero_rpow hp' } end -lemma norm_eq_zero_iff {f : lp E p} : ∥f∥ = 0 ↔ f = 0 := +lemma norm_eq_zero_iff {f : lp E p} : ‖f‖ = 0 ↔ f = 0 := begin classical, refine ⟨λ h, _, by { rintros rfl, exact norm_zero }⟩, @@ -418,15 +418,15 @@ begin tauto }, { cases is_empty_or_nonempty α with _i _i; resetI, { simp }, - have H : is_lub (set.range (λ i, ∥f i∥)) 0, + have H : is_lub (set.range (λ i, ‖f i‖)) 0, { simpa [h] using lp.is_lub_norm f }, ext i, - have : ∥f i∥ = 0 := le_antisymm (H.1 ⟨i, rfl⟩) (norm_nonneg _), + have : ‖f i‖ = 0 := le_antisymm (H.1 ⟨i, rfl⟩) (norm_nonneg _), simpa using this }, - { have hf : has_sum (λ (i : α), ∥f i∥ ^ p.to_real) 0, + { have hf : has_sum (λ (i : α), ‖f i‖ ^ p.to_real) 0, { have := lp.has_sum_norm hp f, rwa [h, real.zero_rpow hp.ne'] at this }, - have : ∀ i, 0 ≤ ∥f i∥ ^ p.to_real := λ i, real.rpow_nonneg_of_nonneg (norm_nonneg _) _, + have : ∀ i, 0 ≤ ‖f i‖ ^ p.to_real := λ i, real.rpow_nonneg_of_nonneg (norm_nonneg _) _, rw has_sum_zero_iff_of_nonneg this at hf, ext i, have : f i = 0 ∧ p.to_real ≠ 0, @@ -437,7 +437,7 @@ end lemma eq_zero_iff_coe_fn_eq_zero {f : lp E p} : f = 0 ↔ ⇑f = 0 := by rw [lp.ext_iff, coe_fn_zero] -@[simp] lemma norm_neg ⦃f : lp E p⦄ : ∥-f∥ = ∥f∥ := +@[simp] lemma norm_neg ⦃f : lp E p⦄ : ‖-f‖ = ‖f‖ := begin rcases p.trichotomy with rfl | rfl | hp, { simp [lp.norm_eq_card_dsupport] }, @@ -445,7 +445,7 @@ begin { simp [lp.eq_zero' f], }, apply (lp.is_lub_norm (-f)).unique, simpa using lp.is_lub_norm f }, - { suffices : ∥-f∥ ^ p.to_real = ∥f∥ ^ p.to_real, + { suffices : ‖-f‖ ^ p.to_real = ‖f‖ ^ p.to_real, { exact real.rpow_left_inj_on hp.ne' (norm_nonneg' _) (norm_nonneg' _) this }, apply (lp.has_sum_norm hp (-f)).unique, simpa using lp.has_sum_norm hp f } @@ -466,8 +466,8 @@ add_group_norm.to_normed_add_comm_group ⟨_, _, ⟨i, rfl⟩, ⟨i, rfl⟩, rfl⟩), exact norm_add_le (f i) (g i) }, { have hp'' : 0 < p.to_real := zero_lt_one.trans_le hp', - have hf₁ : ∀ i, 0 ≤ ∥f i∥ := λ i, norm_nonneg _, - have hg₁ : ∀ i, 0 ≤ ∥g i∥ := λ i, norm_nonneg _, + have hf₁ : ∀ i, 0 ≤ ‖f i‖ := λ i, norm_nonneg _, + have hg₁ : ∀ i, 0 ≤ ‖g i‖ := λ i, norm_nonneg _, have hf₂ := lp.has_sum_norm hp'' f, have hg₂ := lp.has_sum_norm hp'' g, -- apply Minkowski's inequality @@ -486,10 +486,10 @@ add_group_norm.to_normed_add_comm_group /-- Hölder inequality -/ protected lemma tsum_mul_le_mul_norm {p q : ℝ≥0∞} (hpq : p.to_real.is_conjugate_exponent q.to_real) (f : lp E p) (g : lp E q) : - summable (λ i, ∥f i∥ * ∥g i∥) ∧ ∑' i, ∥f i∥ * ∥g i∥ ≤ ∥f∥ * ∥g∥ := + summable (λ i, ‖f i‖ * ‖g i‖) ∧ ∑' i, ‖f i‖ * ‖g i‖ ≤ ‖f‖ * ‖g‖ := begin - have hf₁ : ∀ i, 0 ≤ ∥f i∥ := λ i, norm_nonneg _, - have hg₁ : ∀ i, 0 ≤ ∥g i∥ := λ i, norm_nonneg _, + have hf₁ : ∀ i, 0 ≤ ‖f i‖ := λ i, norm_nonneg _, + have hg₁ : ∀ i, 0 ≤ ‖g i‖ := λ i, norm_nonneg _, have hf₂ := lp.has_sum_norm hpq.pos f, have hg₂ := lp.has_sum_norm hpq.symm.pos g, obtain ⟨C, -, hC', hC⟩ := @@ -500,46 +500,46 @@ end protected lemma summable_mul {p q : ℝ≥0∞} (hpq : p.to_real.is_conjugate_exponent q.to_real) (f : lp E p) (g : lp E q) : - summable (λ i, ∥f i∥ * ∥g i∥) := + summable (λ i, ‖f i‖ * ‖g i‖) := (lp.tsum_mul_le_mul_norm hpq f g).1 protected lemma tsum_mul_le_mul_norm' {p q : ℝ≥0∞} (hpq : p.to_real.is_conjugate_exponent q.to_real) (f : lp E p) (g : lp E q) : - ∑' i, ∥f i∥ * ∥g i∥ ≤ ∥f∥ * ∥g∥ := + ∑' i, ‖f i‖ * ‖g i‖ ≤ ‖f‖ * ‖g‖ := (lp.tsum_mul_le_mul_norm hpq f g).2 section compare_pointwise -lemma norm_apply_le_norm (hp : p ≠ 0) (f : lp E p) (i : α) : ∥f i∥ ≤ ∥f∥ := +lemma norm_apply_le_norm (hp : p ≠ 0) (f : lp E p) (i : α) : ‖f i‖ ≤ ‖f‖ := begin rcases eq_or_ne p ∞ with rfl | hp', { haveI : nonempty α := ⟨i⟩, exact (is_lub_norm f).1 ⟨i, rfl⟩ }, have hp'' : 0 < p.to_real := ennreal.to_real_pos hp hp', - have : ∀ i, 0 ≤ ∥f i∥ ^ p.to_real, + have : ∀ i, 0 ≤ ‖f i‖ ^ p.to_real, { exact λ i, real.rpow_nonneg_of_nonneg (norm_nonneg _) _ }, rw ← real.rpow_le_rpow_iff (norm_nonneg _) (norm_nonneg' _) hp'', convert le_has_sum (has_sum_norm hp'' f) i (λ i hi, this i), end lemma sum_rpow_le_norm_rpow (hp : 0 < p.to_real) (f : lp E p) (s : finset α) : - ∑ i in s, ∥f i∥ ^ p.to_real ≤ ∥f∥ ^ p.to_real := + ∑ i in s, ‖f i‖ ^ p.to_real ≤ ‖f‖ ^ p.to_real := begin rw lp.norm_rpow_eq_tsum hp f, - have : ∀ i, 0 ≤ ∥f i∥ ^ p.to_real, + have : ∀ i, 0 ≤ ‖f i‖ ^ p.to_real, { exact λ i, real.rpow_nonneg_of_nonneg (norm_nonneg _) _ }, refine sum_le_tsum _ (λ i hi, this i) _, exact (lp.mem_ℓp f).summable hp end -lemma norm_le_of_forall_le' [nonempty α] {f : lp E ∞} (C : ℝ) (hCf : ∀ i, ∥f i∥ ≤ C) : ∥f∥ ≤ C := +lemma norm_le_of_forall_le' [nonempty α] {f : lp E ∞} (C : ℝ) (hCf : ∀ i, ‖f i‖ ≤ C) : ‖f‖ ≤ C := begin refine (is_lub_norm f).2 _, rintros - ⟨i, rfl⟩, exact hCf i, end -lemma norm_le_of_forall_le {f : lp E ∞} {C : ℝ} (hC : 0 ≤ C) (hCf : ∀ i, ∥f i∥ ≤ C) : ∥f∥ ≤ C := +lemma norm_le_of_forall_le {f : lp E ∞} {C : ℝ} (hC : 0 ≤ C) (hCf : ∀ i, ‖f i‖ ≤ C) : ‖f‖ ≤ C := begin casesI is_empty_or_nonempty α, { simpa [eq_zero' f] using hC, }, @@ -547,16 +547,16 @@ begin end lemma norm_le_of_tsum_le (hp : 0 < p.to_real) {C : ℝ} (hC : 0 ≤ C) {f : lp E p} - (hf : ∑' i, ∥f i∥ ^ p.to_real ≤ C ^ p.to_real) : - ∥f∥ ≤ C := + (hf : ∑' i, ‖f i‖ ^ p.to_real ≤ C ^ p.to_real) : + ‖f‖ ≤ C := begin rw [← real.rpow_le_rpow_iff (norm_nonneg' _) hC hp, norm_rpow_eq_tsum hp], exact hf, end lemma norm_le_of_forall_sum_le (hp : 0 < p.to_real) {C : ℝ} (hC : 0 ≤ C) {f : lp E p} - (hf : ∀ s : finset α, ∑ i in s, ∥f i∥ ^ p.to_real ≤ C ^ p.to_real) : - ∥f∥ ≤ C := + (hf : ∀ s : finset α, ∑ i in s, ‖f i‖ ^ p.to_real ≤ C ^ p.to_real) : + ‖f‖ ≤ C := norm_le_of_tsum_le hp hC (tsum_le_of_sum_le ((lp.mem_ℓp f).summable hp) hf) end compare_pointwise @@ -587,7 +587,7 @@ instance : module 𝕜 (lp E p) := @[simp] lemma coe_fn_smul (c : 𝕜) (f : lp E p) : ⇑(c • f) = c • f := rfl -lemma norm_const_smul (hp : p ≠ 0) {c : 𝕜} (f : lp E p) : ∥c • f∥ = ∥c∥ * ∥f∥ := +lemma norm_const_smul (hp : p ≠ 0) {c : 𝕜} (f : lp E p) : ‖c • f‖ = ‖c‖ * ‖f‖ := begin rcases p.trichotomy with rfl | rfl | hp, { exact absurd rfl hp }, @@ -597,14 +597,14 @@ begin convert (lp.is_lub_norm f).mul_left (norm_nonneg c), ext a, simp [coe_fn_smul, norm_smul] }, - { suffices : ∥c • f∥ ^ p.to_real = (∥c∥ * ∥f∥) ^ p.to_real, + { suffices : ‖c • f‖ ^ p.to_real = (‖c‖ * ‖f‖) ^ p.to_real, { refine real.rpow_left_inj_on hp.ne' _ _ this, { exact norm_nonneg' _ }, { exact mul_nonneg (norm_nonneg _) (norm_nonneg' _) } }, apply (lp.has_sum_norm hp (c • f)).unique, - convert (lp.has_sum_norm hp f).mul_left (∥c∥ ^ p.to_real), + convert (lp.has_sum_norm hp f).mul_left (‖c‖ ^ p.to_real), { simp [coe_fn_smul, norm_smul, real.mul_rpow (norm_nonneg c) (norm_nonneg _)] }, - have hf : 0 ≤ ∥f∥ := lp.norm_nonneg' f, + have hf : 0 ≤ ‖f‖ := lp.norm_nonneg' f, simp [coe_fn_smul, norm_smul, real.mul_rpow (norm_nonneg c) hf] } end @@ -684,7 +684,7 @@ begin obtain ⟨⟨Cf, hCf⟩, ⟨Cg, hCg⟩⟩ := ⟨hf.bdd_above, hg.bdd_above⟩, refine ⟨Cf * Cg, _⟩, rintros _ ⟨i, rfl⟩, - calc ∥(f * g) i∥ ≤ ∥f i∥ * ∥g i∥ : norm_mul_le (f i) (g i) + calc ‖(f * g) i‖ ≤ ‖f i‖ * ‖g i‖ : norm_mul_le (f i) (g i) ... ≤ Cf * Cg : mul_le_mul (hCf ⟨i, rfl⟩) (hCg ⟨i, rfl⟩) (norm_nonneg _) ((norm_nonneg _).trans (hCf ⟨i, rfl⟩)) end @@ -701,8 +701,8 @@ function.injective.non_unital_ring lp.has_coe_to_fun.coe (subtype.coe_injective) instance : non_unital_normed_ring (lp B ∞) := { norm_mul := λ f g, lp.norm_le_of_forall_le (mul_nonneg (norm_nonneg f) (norm_nonneg g)) - (λ i, calc ∥(f * g) i∥ ≤ ∥f i∥ * ∥g i∥ : norm_mul_le _ _ - ... ≤ ∥f∥ * ∥g∥ + (λ i, calc ‖(f * g) i‖ ≤ ‖f i‖ * ‖g i‖ : norm_mul_le _ _ + ... ≤ ‖f‖ * ‖g‖ : mul_le_mul (lp.norm_apply_le_norm ennreal.top_ne_zero f i) (lp.norm_apply_le_norm ennreal.top_ne_zero g i) (norm_nonneg _) (norm_nonneg _)), .. lp.normed_add_comm_group } @@ -733,12 +733,12 @@ instance infty_cstar_ring [∀ i, cstar_ring (B i)] : cstar_ring (lp B ∞) := begin apply le_antisymm, { rw ←sq, - refine lp.norm_le_of_forall_le (sq_nonneg ∥ f ∥) (λ i, _), + refine lp.norm_le_of_forall_le (sq_nonneg ‖ f ‖) (λ i, _), simp only [lp.star_apply, cstar_ring.norm_star_mul_self, ←sq, infty_coe_fn_mul, pi.mul_apply], refine sq_le_sq' _ (lp.norm_apply_le_norm ennreal.top_ne_zero _ _), linarith [norm_nonneg (f i), norm_nonneg f] }, { rw [←sq, ←real.le_sqrt (norm_nonneg _) (norm_nonneg _)], - refine lp.norm_le_of_forall_le (∥star f * f∥.sqrt_nonneg) (λ i, _), + refine lp.norm_le_of_forall_le (‖star f * f‖.sqrt_nonneg) (λ i, _), rw [real.le_sqrt (norm_nonneg _) (norm_nonneg _), sq, ←cstar_ring.norm_star_mul_self], exact lp.norm_apply_le_norm ennreal.top_ne_zero (star f * f) i, } end } @@ -895,38 +895,38 @@ begin end protected lemma norm_sum_single (hp : 0 < p.to_real) (f : Π i, E i) (s : finset α) : - ∥∑ i in s, lp.single p i (f i)∥ ^ p.to_real = ∑ i in s, ∥f i∥ ^ p.to_real := + ‖∑ i in s, lp.single p i (f i)‖ ^ p.to_real = ∑ i in s, ‖f i‖ ^ p.to_real := begin refine (has_sum_norm hp (∑ i in s, lp.single p i (f i))).unique _, simp only [lp.single_apply, coe_fn_sum, finset.sum_apply, finset.sum_dite_eq], - have h : ∀ i ∉ s, ∥ite (i ∈ s) (f i) 0∥ ^ p.to_real = 0, + have h : ∀ i ∉ s, ‖ite (i ∈ s) (f i) 0‖ ^ p.to_real = 0, { intros i hi, simp [if_neg hi, real.zero_rpow hp.ne'], }, - have h' : ∀ i ∈ s, ∥f i∥ ^ p.to_real = ∥ite (i ∈ s) (f i) 0∥ ^ p.to_real, + have h' : ∀ i ∈ s, ‖f i‖ ^ p.to_real = ‖ite (i ∈ s) (f i) 0‖ ^ p.to_real, { intros i hi, rw if_pos hi }, simpa [finset.sum_congr rfl h'] using has_sum_sum_of_ne_finset_zero h, end protected lemma norm_single (hp : 0 < p.to_real) (f : Π i, E i) (i : α) : - ∥lp.single p i (f i)∥ = ∥f i∥ := + ‖lp.single p i (f i)‖ = ‖f i‖ := begin refine real.rpow_left_inj_on hp.ne' (norm_nonneg' _) (norm_nonneg _) _, simpa using lp.norm_sum_single hp f {i}, end protected lemma norm_sub_norm_compl_sub_single (hp : 0 < p.to_real) (f : lp E p) (s : finset α) : - ∥f∥ ^ p.to_real - ∥f - ∑ i in s, lp.single p i (f i)∥ ^ p.to_real = ∑ i in s, ∥f i∥ ^ p.to_real := + ‖f‖ ^ p.to_real - ‖f - ∑ i in s, lp.single p i (f i)‖ ^ p.to_real = ∑ i in s, ‖f i‖ ^ p.to_real := begin refine ((has_sum_norm hp f).sub (has_sum_norm hp (f - ∑ i in s, lp.single p i (f i)))).unique _, - let F : α → ℝ := λ i, ∥f i∥ ^ p.to_real - ∥(f - ∑ i in s, lp.single p i (f i)) i∥ ^ p.to_real, + let F : α → ℝ := λ i, ‖f i‖ ^ p.to_real - ‖(f - ∑ i in s, lp.single p i (f i)) i‖ ^ p.to_real, have hF : ∀ i ∉ s, F i = 0, { intros i hi, - suffices : ∥f i∥ ^ p.to_real - ∥f i - ite (i ∈ s) (f i) 0∥ ^ p.to_real = 0, + suffices : ‖f i‖ ^ p.to_real - ‖f i - ite (i ∈ s) (f i) 0‖ ^ p.to_real = 0, { simpa only [F, coe_fn_sum, lp.single_apply, coe_fn_sub, pi.sub_apply, finset.sum_apply, finset.sum_dite_eq] using this, }, simp only [if_neg hi, sub_zero, sub_self] }, - have hF' : ∀ i ∈ s, F i = ∥f i∥ ^ p.to_real, + have hF' : ∀ i ∈ s, F i = ‖f i‖ ^ p.to_real, { intros i hi, simp only [F, coe_fn_sum, lp.single_apply, if_pos hi, sub_self, eq_self_iff_true, coe_fn_sub, pi.sub_apply, finset.sum_apply, finset.sum_dite_eq, sub_eq_self], @@ -936,7 +936,7 @@ begin end protected lemma norm_compl_sum_single (hp : 0 < p.to_real) (f : lp E p) (s : finset α) : - ∥f - ∑ i in s, lp.single p i (f i)∥ ^ p.to_real = ∥f∥ ^ p.to_real - ∑ i in s, ∥f i∥ ^ p.to_real := + ‖f - ∑ i in s, lp.single p i (f i)‖ ^ p.to_real = ‖f‖ ^ p.to_real - ∑ i in s, ‖f i‖ ^ p.to_real := by linarith [lp.norm_sub_norm_compl_sub_single hp f s] /-- The canonical finitely-supported approximations to an element `f` of `lp` converge to it, in the @@ -954,13 +954,13 @@ begin rw ← real.rpow_lt_rpow_iff dist_nonneg (le_of_lt hε) hp', rw dist_comm at hs, simp only [dist_eq_norm, real.norm_eq_abs] at hs ⊢, - have H : ∥∑ i in s, lp.single p i (f i : E i) - f∥ ^ p.to_real - = ∥f∥ ^ p.to_real - ∑ i in s, ∥f i∥ ^ p.to_real, + have H : ‖∑ i in s, lp.single p i (f i : E i) - f‖ ^ p.to_real + = ‖f‖ ^ p.to_real - ∑ i in s, ‖f i‖ ^ p.to_real, { simpa only [coe_fn_neg, pi.neg_apply, lp.single_neg, finset.sum_neg_distrib, neg_sub_neg, norm_neg, _root_.norm_neg] using lp.norm_compl_sum_single hp' (-f) s }, rw ← H at hs, - have : |∥∑ i in s, lp.single p i (f i : E i) - f∥ ^ p.to_real| - = ∥∑ i in s, lp.single p i (f i : E i) - f∥ ^ p.to_real, + have : |‖∑ i in s, lp.single p i (f i : E i) - f‖ ^ p.to_real| + = ‖∑ i in s, lp.single p i (f i : E i) - f‖ ^ p.to_real, { simp only [real.abs_rpow_of_nonneg (norm_nonneg _), abs_norm_eq_norm] }, linarith end @@ -982,18 +982,18 @@ begin normed_add_comm_group.uniformity_basis_dist, intros ε hε, refine ⟨ε, hε, _⟩, - rintros f g (hfg : ∥f - g∥ < ε), - have : ∥f i - g i∥ ≤ ∥f - g∥ := norm_apply_le_norm hp (f - g) i, + rintros f g (hfg : ‖f - g‖ < ε), + have : ‖f i - g i‖ ≤ ‖f - g‖ := norm_apply_le_norm hp (f - g) i, exact this.trans_lt hfg, end variables {ι : Type*} {l : filter ι} [filter.ne_bot l] -lemma norm_apply_le_of_tendsto {C : ℝ} {F : ι → lp E ∞} (hCF : ∀ᶠ k in l, ∥F k∥ ≤ C) +lemma norm_apply_le_of_tendsto {C : ℝ} {F : ι → lp E ∞} (hCF : ∀ᶠ k in l, ‖F k‖ ≤ C) {f : Π a, E a} (hf : tendsto (id (λ i, F i) : ι → Π a, E a) l (𝓝 f)) (a : α) : - ∥f a∥ ≤ C := + ‖f a‖ ≤ C := begin - have : tendsto (λ k, ∥F k a∥) l (𝓝 ∥f a∥) := + have : tendsto (λ k, ‖F k a‖) l (𝓝 ‖f a‖) := (tendsto.comp (continuous_apply a).continuous_at hf).norm, refine le_of_tendsto this (hCF.mono _), intros k hCFk, @@ -1004,13 +1004,13 @@ variables [_i : fact (1 ≤ p)] include _i -lemma sum_rpow_le_of_tendsto (hp : p ≠ ∞) {C : ℝ} {F : ι → lp E p} (hCF : ∀ᶠ k in l, ∥F k∥ ≤ C) +lemma sum_rpow_le_of_tendsto (hp : p ≠ ∞) {C : ℝ} {F : ι → lp E p} (hCF : ∀ᶠ k in l, ‖F k‖ ≤ C) {f : Π a, E a} (hf : tendsto (id (λ i, F i) : ι → Π a, E a) l (𝓝 f)) (s : finset α) : - ∑ (i : α) in s, ∥f i∥ ^ p.to_real ≤ C ^ p.to_real := + ∑ (i : α) in s, ‖f i‖ ^ p.to_real ≤ C ^ p.to_real := begin have hp' : p ≠ 0 := (ennreal.zero_lt_one.trans_le _i.elim).ne', have hp'' : 0 < p.to_real := ennreal.to_real_pos hp' hp, - let G : (Π a, E a) → ℝ := λ f, ∑ a in s, ∥f a∥ ^ p.to_real, + let G : (Π a, E a) → ℝ := λ f, ∑ a in s, ‖f a‖ ^ p.to_real, have hG : continuous G, { refine continuous_finset_sum s _, intros a ha, @@ -1025,9 +1025,9 @@ end /-- "Semicontinuity of the `lp` norm": If all sufficiently large elements of a sequence in `lp E p` have `lp` norm `≤ C`, then the pointwise limit, if it exists, also has `lp` norm `≤ C`. -/ -lemma norm_le_of_tendsto {C : ℝ} {F : ι → lp E p} (hCF : ∀ᶠ k in l, ∥F k∥ ≤ C) {f : lp E p} +lemma norm_le_of_tendsto {C : ℝ} {F : ι → lp E p} (hCF : ∀ᶠ k in l, ‖F k‖ ≤ C) {f : lp E p} (hf : tendsto (id (λ i, F i) : ι → Π a, E a) l (𝓝 f)) : - ∥f∥ ≤ C := + ‖f‖ ≤ C := begin obtain ⟨i, hi⟩ := hCF.exists, have hC : 0 ≤ C := (norm_nonneg _).trans hi, @@ -1046,7 +1046,7 @@ lemma mem_ℓp_of_tendsto {F : ι → lp E p} (hF : metric.bounded (set.range F) mem_ℓp f p := begin obtain ⟨C, hC, hCF'⟩ := hF.exists_pos_norm_le, - have hCF : ∀ k, ∥F k∥ ≤ C := λ k, hCF' _ ⟨k, rfl⟩, + have hCF : ∀ k, ‖F k‖ ≤ C := λ k, hCF' _ ⟨k, rfl⟩, unfreezingI { rcases eq_top_or_lt_top p with rfl | hp }, { apply mem_ℓp_infty, use C, @@ -1064,10 +1064,10 @@ lemma tendsto_lp_of_tendsto_pi {F : ℕ → lp E p} (hF : cauchy_seq F) {f : lp begin rw metric.nhds_basis_closed_ball.tendsto_right_iff, intros ε hε, - have hε' : {p : (lp E p) × (lp E p) | ∥p.1 - p.2∥ < ε} ∈ 𝓤 (lp E p), + have hε' : {p : (lp E p) × (lp E p) | ‖p.1 - p.2‖ < ε} ∈ 𝓤 (lp E p), { exact normed_add_comm_group.uniformity_basis_dist.mem_of_mem hε }, refine (hF.eventually_eventually hε').mono _, - rintros n (hn : ∀ᶠ l in at_top, ∥(λ f, F n - f) (F l)∥ < ε), + rintros n (hn : ∀ᶠ l in at_top, ‖(λ f, F n - f) (F l)‖ < ε), refine norm_le_of_tendsto (hn.mono (λ k hk, hk.le)) _, rw tendsto_pi_nhds, intros a, diff --git a/src/analysis/normed_space/mazur_ulam.lean b/src/analysis/normed_space/mazur_ulam.lean index 2906940622db4..b693a3234c559 100644 --- a/src/analysis/normed_space/mazur_ulam.lean +++ b/src/analysis/normed_space/mazur_ulam.lean @@ -108,7 +108,7 @@ We define a conversion to a `continuous_linear_equiv` first, then a conversion t over `ℝ` and `f 0 = 0`, then `f` is a linear isometry equivalence. -/ def to_real_linear_isometry_equiv_of_map_zero (f : E ≃ᵢ F) (h0 : f 0 = 0) : E ≃ₗᵢ[ℝ] F := -{ norm_map' := λ x, show ∥f x∥ = ∥x∥, by simp only [← dist_zero_right, ← h0, f.dist_eq], +{ norm_map' := λ x, show ‖f x‖ = ‖x‖, by simp only [← dist_zero_right, ← h0, f.dist_eq], .. ((add_monoid_hom.of_map_midpoint ℝ ℝ f h0 f.map_midpoint).to_real_linear_map f.continuous), .. f } diff --git a/src/analysis/normed_space/multilinear.lean b/src/analysis/normed_space/multilinear.lean index 95163f6eeffea..aba5f7d1d9f1c 100644 --- a/src/analysis/normed_space/multilinear.lean +++ b/src/analysis/normed_space/multilinear.lean @@ -9,8 +9,8 @@ import topology.algebra.module.multilinear /-! # Operator norm on the space of continuous multilinear maps -When `f` is a continuous multilinear map in finitely many variables, we define its norm `∥f∥` as the -smallest number such that `∥f m∥ ≤ ∥f∥ * ∏ i, ∥m i∥` for all `m`. +When `f` is a continuous multilinear map in finitely many variables, we define its norm `‖f‖` as the +smallest number such that `‖f m‖ ≤ ‖f‖ * ∏ i, ‖m i‖` for all `m`. We show that it is indeed a norm, and prove its basic properties. @@ -18,16 +18,16 @@ We show that it is indeed a norm, and prove its basic properties. Let `f` be a multilinear map in finitely many variables. * `exists_bound_of_continuous` asserts that, if `f` is continuous, then there exists `C > 0` - with `∥f m∥ ≤ C * ∏ i, ∥m i∥` for all `m`. + with `‖f m‖ ≤ C * ∏ i, ‖m i‖` for all `m`. * `continuous_of_bound`, conversely, asserts that this bound implies continuity. * `mk_continuous` constructs the associated continuous multilinear map. Let `f` be a continuous multilinear map in finitely many variables. -* `∥f∥` is its norm, i.e., the smallest number such that `∥f m∥ ≤ ∥f∥ * ∏ i, ∥m i∥` for +* `‖f‖` is its norm, i.e., the smallest number such that `‖f m‖ ≤ ‖f‖ * ∏ i, ‖m i‖` for all `m`. -* `le_op_norm f m` asserts the fundamental inequality `∥f m∥ ≤ ∥f∥ * ∏ i, ∥m i∥`. +* `le_op_norm f m` asserts the fundamental inequality `‖f m‖ ≤ ‖f‖ * ∏ i, ‖m i‖`. * `norm_image_sub_le f m₁ m₂` gives a control of the difference `f m₁ - f m₂` in terms of - `∥f∥` and `∥m₁ - m₂∥`. + `‖f‖` and `‖m₁ - m₂‖`. We also register isomorphisms corresponding to currying or uncurrying variables, transforming a continuous multilinear function `f` in `n+1` variables into a continuous linear function taking @@ -87,43 +87,43 @@ variables {𝕜 : Type u} {ι : Type v} {ι' : Type v'} {n : ℕ} /-! ### Continuity properties of multilinear maps -We relate continuity of multilinear maps to the inequality `∥f m∥ ≤ C * ∏ i, ∥m i∥`, in -both directions. Along the way, we prove useful bounds on the difference `∥f m₁ - f m₂∥`. +We relate continuity of multilinear maps to the inequality `‖f m‖ ≤ C * ∏ i, ‖m i‖`, in +both directions. Along the way, we prove useful bounds on the difference `‖f m₁ - f m₂‖`. -/ namespace multilinear_map variable (f : multilinear_map 𝕜 E G) /-- If a multilinear map in finitely many variables on normed spaces satisfies the inequality -`∥f m∥ ≤ C * ∏ i, ∥m i∥` on a shell `ε i / ∥c i∥ < ∥m i∥ < ε i` for some positive numbers `ε i` -and elements `c i : 𝕜`, `1 < ∥c i∥`, then it satisfies this inequality for all `m`. -/ -lemma bound_of_shell {ε : ι → ℝ} {C : ℝ} (hε : ∀ i, 0 < ε i) {c : ι → 𝕜} (hc : ∀ i, 1 < ∥c i∥) - (hf : ∀ m : Π i, E i, (∀ i, ε i / ∥c i∥ ≤ ∥m i∥) → (∀ i, ∥m i∥ < ε i) → ∥f m∥ ≤ C * ∏ i, ∥m i∥) - (m : Π i, E i) : ∥f m∥ ≤ C * ∏ i, ∥m i∥ := +`‖f m‖ ≤ C * ∏ i, ‖m i‖` on a shell `ε i / ‖c i‖ < ‖m i‖ < ε i` for some positive numbers `ε i` +and elements `c i : 𝕜`, `1 < ‖c i‖`, then it satisfies this inequality for all `m`. -/ +lemma bound_of_shell {ε : ι → ℝ} {C : ℝ} (hε : ∀ i, 0 < ε i) {c : ι → 𝕜} (hc : ∀ i, 1 < ‖c i‖) + (hf : ∀ m : Π i, E i, (∀ i, ε i / ‖c i‖ ≤ ‖m i‖) → (∀ i, ‖m i‖ < ε i) → ‖f m‖ ≤ C * ∏ i, ‖m i‖) + (m : Π i, E i) : ‖f m‖ ≤ C * ∏ i, ‖m i‖ := begin rcases em (∃ i, m i = 0) with ⟨i, hi⟩|hm; [skip, push_neg at hm], { simp [f.map_coord_zero i hi, prod_eq_zero (mem_univ i), hi] }, choose δ hδ0 hδm_lt hle_δm hδinv using λ i, rescale_to_shell (hc i) (hε i) (hm i), - have hδ0 : 0 < ∏ i, ∥δ i∥, from prod_pos (λ i _, norm_pos_iff.2 (hδ0 i)), + have hδ0 : 0 < ∏ i, ‖δ i‖, from prod_pos (λ i _, norm_pos_iff.2 (hδ0 i)), simpa [map_smul_univ, norm_smul, prod_mul_distrib, mul_left_comm C, mul_le_mul_left hδ0] using hf (λ i, δ i • m i) hle_δm hδm_lt, end /-- If a multilinear map in finitely many variables on normed spaces is continuous, then it -satisfies the inequality `∥f m∥ ≤ C * ∏ i, ∥m i∥`, for some `C` which can be chosen to be +satisfies the inequality `‖f m‖ ≤ C * ∏ i, ‖m i‖`, for some `C` which can be chosen to be positive. -/ theorem exists_bound_of_continuous (hf : continuous f) : - ∃ (C : ℝ), 0 < C ∧ (∀ m, ∥f m∥ ≤ C * ∏ i, ∥m i∥) := + ∃ (C : ℝ), 0 < C ∧ (∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) := begin casesI is_empty_or_nonempty ι, - { refine ⟨∥f 0∥ + 1, add_pos_of_nonneg_of_pos (norm_nonneg _) zero_lt_one, λ m, _⟩, + { refine ⟨‖f 0‖ + 1, add_pos_of_nonneg_of_pos (norm_nonneg _) zero_lt_one, λ m, _⟩, obtain rfl : m = 0, from funext (is_empty.elim ‹_›), simp [univ_eq_empty, zero_le_one] }, - obtain ⟨ε : ℝ, ε0 : 0 < ε, hε : ∀ m : Π i, E i, ∥m - 0∥ < ε → ∥f m - f 0∥ < 1⟩ := + obtain ⟨ε : ℝ, ε0 : 0 < ε, hε : ∀ m : Π i, E i, ‖m - 0‖ < ε → ‖f m - f 0‖ < 1⟩ := normed_add_comm_group.tendsto_nhds_nhds.1 (hf.tendsto 0) 1 zero_lt_one, simp only [sub_zero, f.map_zero] at hε, rcases normed_field.exists_one_lt_norm 𝕜 with ⟨c, hc⟩, - have : 0 < (∥c∥ / ε) ^ fintype.card ι, from pow_pos (div_pos (zero_lt_one.trans hc) ε0) _, + have : 0 < (‖c‖ / ε) ^ fintype.card ι, from pow_pos (div_pos (zero_lt_one.trans hc) ε0) _, refine ⟨_, this, _⟩, refine f.bound_of_shell (λ _, ε0) (λ _, hc) (λ m hcm hm, _), refine (hε m ((pi_norm_lt_iff ε0).2 hm)).le.trans _, @@ -134,20 +134,20 @@ end /-- If `f` satisfies a boundedness property around `0`, one can deduce a bound on `f m₁ - f m₂` using the multilinearity. Here, we give a precise but hard to use version. See `norm_image_sub_le_of_bound` for a less precise but more usable version. The bound reads -`∥f m - f m'∥ ≤ - C * ∥m 1 - m' 1∥ * max ∥m 2∥ ∥m' 2∥ * max ∥m 3∥ ∥m' 3∥ * ... * max ∥m n∥ ∥m' n∥ + ...`, +`‖f m - f m'‖ ≤ + C * ‖m 1 - m' 1‖ * max ‖m 2‖ ‖m' 2‖ * max ‖m 3‖ ‖m' 3‖ * ... * max ‖m n‖ ‖m' n‖ + ...`, where the other terms in the sum are the same products where `1` is replaced by any `i`. -/ lemma norm_image_sub_le_of_bound' {C : ℝ} (hC : 0 ≤ C) - (H : ∀ m, ∥f m∥ ≤ C * ∏ i, ∥m i∥) (m₁ m₂ : Πi, E i) : - ∥f m₁ - f m₂∥ ≤ - C * ∑ i, ∏ j, if j = i then ∥m₁ i - m₂ i∥ else max ∥m₁ j∥ ∥m₂ j∥ := + (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) (m₁ m₂ : Πi, E i) : + ‖f m₁ - f m₂‖ ≤ + C * ∑ i, ∏ j, if j = i then ‖m₁ i - m₂ i‖ else max ‖m₁ j‖ ‖m₂ j‖ := begin - have A : ∀(s : finset ι), ∥f m₁ - f (s.piecewise m₂ m₁)∥ - ≤ C * ∑ i in s, ∏ j, if j = i then ∥m₁ i - m₂ i∥ else max ∥m₁ j∥ ∥m₂ j∥, + have A : ∀(s : finset ι), ‖f m₁ - f (s.piecewise m₂ m₁)‖ + ≤ C * ∑ i in s, ∏ j, if j = i then ‖m₁ i - m₂ i‖ else max ‖m₁ j‖ ‖m₂ j‖, { refine finset.induction (by simp) _, assume i s his Hrec, - have I : ∥f (s.piecewise m₂ m₁) - f ((insert i s).piecewise m₂ m₁)∥ - ≤ C * ∏ j, if j = i then ∥m₁ i - m₂ i∥ else max ∥m₁ j∥ ∥m₂ j∥, + have I : ‖f (s.piecewise m₂ m₁) - f ((insert i s).piecewise m₂ m₁)‖ + ≤ C * ∏ j, if j = i then ‖m₁ i - m₂ i‖ else max ‖m₁ j‖ ‖m₂ j‖, { have A : ((insert i s).piecewise m₂ m₁) = function.update (s.piecewise m₂ m₁) i (m₂ i) := s.piecewise_insert _ _ _, have B : s.piecewise m₂ m₁ = function.update (s.piecewise m₂ m₁) i (m₁ i), @@ -162,13 +162,13 @@ begin { rw h, simp }, { by_cases h' : j ∈ s; simp [h', h, le_refl] } }, - calc ∥f m₁ - f ((insert i s).piecewise m₂ m₁)∥ ≤ - ∥f m₁ - f (s.piecewise m₂ m₁)∥ + ∥f (s.piecewise m₂ m₁) - f ((insert i s).piecewise m₂ m₁)∥ : + calc ‖f m₁ - f ((insert i s).piecewise m₂ m₁)‖ ≤ + ‖f m₁ - f (s.piecewise m₂ m₁)‖ + ‖f (s.piecewise m₂ m₁) - f ((insert i s).piecewise m₂ m₁)‖ : by { rw [← dist_eq_norm, ← dist_eq_norm, ← dist_eq_norm], exact dist_triangle _ _ _ } - ... ≤ (C * ∑ i in s, ∏ j, if j = i then ∥m₁ i - m₂ i∥ else max ∥m₁ j∥ ∥m₂ j∥) - + C * ∏ j, if j = i then ∥m₁ i - m₂ i∥ else max ∥m₁ j∥ ∥m₂ j∥ : + ... ≤ (C * ∑ i in s, ∏ j, if j = i then ‖m₁ i - m₂ i‖ else max ‖m₁ j‖ ‖m₂ j‖) + + C * ∏ j, if j = i then ‖m₁ i - m₂ i‖ else max ‖m₁ j‖ ‖m₂ j‖ : add_le_add Hrec I - ... = C * ∑ i in insert i s, ∏ j, if j = i then ∥m₁ i - m₂ i∥ else max ∥m₁ j∥ ∥m₂ j∥ : + ... = C * ∑ i in insert i s, ∏ j, if j = i then ‖m₁ i - m₂ i‖ else max ‖m₁ j‖ ‖m₂ j‖ : by simp [his, add_comm, left_distrib] }, convert A univ, simp @@ -177,16 +177,16 @@ end /-- If `f` satisfies a boundedness property around `0`, one can deduce a bound on `f m₁ - f m₂` using the multilinearity. Here, we give a usable but not very precise version. See `norm_image_sub_le_of_bound'` for a more precise but less usable version. The bound is -`∥f m - f m'∥ ≤ C * card ι * ∥m - m'∥ * (max ∥m∥ ∥m'∥) ^ (card ι - 1)`. -/ +`‖f m - f m'‖ ≤ C * card ι * ‖m - m'‖ * (max ‖m‖ ‖m'‖) ^ (card ι - 1)`. -/ lemma norm_image_sub_le_of_bound {C : ℝ} (hC : 0 ≤ C) - (H : ∀ m, ∥f m∥ ≤ C * ∏ i, ∥m i∥) (m₁ m₂ : Πi, E i) : - ∥f m₁ - f m₂∥ ≤ C * (fintype.card ι) * (max ∥m₁∥ ∥m₂∥) ^ (fintype.card ι - 1) * ∥m₁ - m₂∥ := + (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) (m₁ m₂ : Πi, E i) : + ‖f m₁ - f m₂‖ ≤ C * (fintype.card ι) * (max ‖m₁‖ ‖m₂‖) ^ (fintype.card ι - 1) * ‖m₁ - m₂‖ := begin - have A : ∀ (i : ι), ∏ j, (if j = i then ∥m₁ i - m₂ i∥ else max ∥m₁ j∥ ∥m₂ j∥) - ≤ ∥m₁ - m₂∥ * (max ∥m₁∥ ∥m₂∥) ^ (fintype.card ι - 1), + have A : ∀ (i : ι), ∏ j, (if j = i then ‖m₁ i - m₂ i‖ else max ‖m₁ j‖ ‖m₂ j‖) + ≤ ‖m₁ - m₂‖ * (max ‖m₁‖ ‖m₂‖) ^ (fintype.card ι - 1), { assume i, - calc ∏ j, (if j = i then ∥m₁ i - m₂ i∥ else max ∥m₁ j∥ ∥m₂ j∥) - ≤ ∏ j : ι, function.update (λ j, max ∥m₁∥ ∥m₂∥) i (∥m₁ - m₂∥) j : + calc ∏ j, (if j = i then ‖m₁ i - m₂ i‖ else max ‖m₁ j‖ ‖m₂ j‖) + ≤ ∏ j : ι, function.update (λ j, max ‖m₁‖ ‖m₂‖) i (‖m₁ - m₂‖) j : begin apply prod_le_prod, { assume j hj, by_cases h : j = i; simp [h, norm_nonneg] }, @@ -195,68 +195,68 @@ begin { rw h, simp, exact norm_le_pi_norm (m₁ - m₂) i }, { simp [h, max_le_max, norm_le_pi_norm (_ : Π i, E i)] } } end - ... = ∥m₁ - m₂∥ * (max ∥m₁∥ ∥m₂∥) ^ (fintype.card ι - 1) : + ... = ‖m₁ - m₂‖ * (max ‖m₁‖ ‖m₂‖) ^ (fintype.card ι - 1) : by { rw prod_update_of_mem (finset.mem_univ _), simp [card_univ_diff] } }, calc - ∥f m₁ - f m₂∥ - ≤ C * ∑ i, ∏ j, if j = i then ∥m₁ i - m₂ i∥ else max ∥m₁ j∥ ∥m₂ j∥ : + ‖f m₁ - f m₂‖ + ≤ C * ∑ i, ∏ j, if j = i then ‖m₁ i - m₂ i‖ else max ‖m₁ j‖ ‖m₂ j‖ : f.norm_image_sub_le_of_bound' hC H m₁ m₂ - ... ≤ C * ∑ i, ∥m₁ - m₂∥ * (max ∥m₁∥ ∥m₂∥) ^ (fintype.card ι - 1) : + ... ≤ C * ∑ i, ‖m₁ - m₂‖ * (max ‖m₁‖ ‖m₂‖) ^ (fintype.card ι - 1) : mul_le_mul_of_nonneg_left (sum_le_sum (λi hi, A i)) hC - ... = C * (fintype.card ι) * (max ∥m₁∥ ∥m₂∥) ^ (fintype.card ι - 1) * ∥m₁ - m₂∥ : + ... = C * (fintype.card ι) * (max ‖m₁‖ ‖m₂‖) ^ (fintype.card ι - 1) * ‖m₁ - m₂‖ : by { rw [sum_const, card_univ, nsmul_eq_mul], ring } end -/-- If a multilinear map satisfies an inequality `∥f m∥ ≤ C * ∏ i, ∥m i∥`, then it is +/-- If a multilinear map satisfies an inequality `‖f m‖ ≤ C * ∏ i, ‖m i‖`, then it is continuous. -/ -theorem continuous_of_bound (C : ℝ) (H : ∀ m, ∥f m∥ ≤ C * ∏ i, ∥m i∥) : +theorem continuous_of_bound (C : ℝ) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) : continuous f := begin let D := max C 1, have D_pos : 0 ≤ D := le_trans zero_le_one (le_max_right _ _), - replace H : ∀ m, ∥f m∥ ≤ D * ∏ i, ∥m i∥, + replace H : ∀ m, ‖f m‖ ≤ D * ∏ i, ‖m i‖, { assume m, apply le_trans (H m) (mul_le_mul_of_nonneg_right (le_max_left _ _) _), exact prod_nonneg (λ(i : ι) hi, norm_nonneg (m i)) }, refine continuous_iff_continuous_at.2 (λm, _), refine continuous_at_of_locally_lipschitz zero_lt_one - (D * (fintype.card ι) * (∥m∥ + 1) ^ (fintype.card ι - 1)) (λm' h', _), + (D * (fintype.card ι) * (‖m‖ + 1) ^ (fintype.card ι - 1)) (λm' h', _), rw [dist_eq_norm, dist_eq_norm], - have : 0 ≤ (max ∥m'∥ ∥m∥), by simp, - have : (max ∥m'∥ ∥m∥) ≤ ∥m∥ + 1, + have : 0 ≤ (max ‖m'‖ ‖m‖), by simp, + have : (max ‖m'‖ ‖m‖) ≤ ‖m‖ + 1, by simp [zero_le_one, norm_le_of_mem_closed_ball (le_of_lt h'), -add_comm], calc - ∥f m' - f m∥ - ≤ D * (fintype.card ι) * (max ∥m'∥ ∥m∥) ^ (fintype.card ι - 1) * ∥m' - m∥ : + ‖f m' - f m‖ + ≤ D * (fintype.card ι) * (max ‖m'‖ ‖m‖) ^ (fintype.card ι - 1) * ‖m' - m‖ : f.norm_image_sub_le_of_bound D_pos H m' m - ... ≤ D * (fintype.card ι) * (∥m∥ + 1) ^ (fintype.card ι - 1) * ∥m' - m∥ : + ... ≤ D * (fintype.card ι) * (‖m‖ + 1) ^ (fintype.card ι - 1) * ‖m' - m‖ : by apply_rules [mul_le_mul_of_nonneg_right, mul_le_mul_of_nonneg_left, mul_nonneg, norm_nonneg, nat.cast_nonneg, pow_le_pow_of_le_left] end /-- Constructing a continuous multilinear map from a multilinear map satisfying a boundedness condition. -/ -def mk_continuous (C : ℝ) (H : ∀ m, ∥f m∥ ≤ C * ∏ i, ∥m i∥) : +def mk_continuous (C : ℝ) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) : continuous_multilinear_map 𝕜 E G := { cont := f.continuous_of_bound C H, ..f } -@[simp] lemma coe_mk_continuous (C : ℝ) (H : ∀ m, ∥f m∥ ≤ C * ∏ i, ∥m i∥) : +@[simp] lemma coe_mk_continuous (C : ℝ) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) : ⇑(f.mk_continuous C H) = f := rfl /-- Given a multilinear map in `n` variables, if one restricts it to `k` variables putting `z` on the other coordinates, then the resulting restricted function satisfies an inequality -`∥f.restr v∥ ≤ C * ∥z∥^(n-k) * Π ∥v i∥` if the original function satisfies `∥f v∥ ≤ C * Π ∥v i∥`. -/ +`‖f.restr v‖ ≤ C * ‖z‖^(n-k) * Π ‖v i‖` if the original function satisfies `‖f v‖ ≤ C * Π ‖v i‖`. -/ lemma restr_norm_le {k n : ℕ} (f : (multilinear_map 𝕜 (λ i : fin n, G) G' : _)) (s : finset (fin n)) (hk : s.card = k) (z : G) {C : ℝ} - (H : ∀ m, ∥f m∥ ≤ C * ∏ i, ∥m i∥) (v : fin k → G) : - ∥f.restr s hk z v∥ ≤ C * ∥z∥ ^ (n - k) * ∏ i, ∥v i∥ := + (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) (v : fin k → G) : + ‖f.restr s hk z v‖ ≤ C * ‖z‖ ^ (n - k) * ∏ i, ‖v i‖ := begin rw [mul_right_comm, mul_assoc], convert H _ using 2, - simp only [apply_dite norm, fintype.prod_dite, prod_const (∥z∥), finset.card_univ, + simp only [apply_dite norm, fintype.prod_dite, prod_const (‖z‖), finset.card_univ, fintype.card_of_subtype sᶜ (λ x, mem_compl), card_compl, fintype.card_fin, hk, mk_coe, - ← (s.order_iso_of_fin hk).symm.bijective.prod_comp (λ x, ∥v x∥)], + ← (s.order_iso_of_fin hk).symm.bijective.prod_comp (λ x, ‖v x‖)], refl end @@ -265,21 +265,21 @@ end multilinear_map /-! ### Continuous multilinear maps -We define the norm `∥f∥` of a continuous multilinear map `f` in finitely many variables as the -smallest number such that `∥f m∥ ≤ ∥f∥ * ∏ i, ∥m i∥` for all `m`. We show that this +We define the norm `‖f‖` of a continuous multilinear map `f` in finitely many variables as the +smallest number such that `‖f m‖ ≤ ‖f‖ * ∏ i, ‖m i‖` for all `m`. We show that this defines a normed space structure on `continuous_multilinear_map 𝕜 E G`. -/ namespace continuous_multilinear_map variables (c : 𝕜) (f g : continuous_multilinear_map 𝕜 E G) (m : Πi, E i) -theorem bound : ∃ (C : ℝ), 0 < C ∧ (∀ m, ∥f m∥ ≤ C * ∏ i, ∥m i∥) := +theorem bound : ∃ (C : ℝ), 0 < C ∧ (∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) := f.to_multilinear_map.exists_bound_of_continuous f.2 open real /-- The operator norm of a continuous multilinear map is the inf of all its bounds. -/ -def op_norm := Inf {c | 0 ≤ (c : ℝ) ∧ ∀ m, ∥f m∥ ≤ c * ∏ i, ∥m i∥} +def op_norm := Inf {c | 0 ≤ (c : ℝ) ∧ ∀ m, ‖f m‖ ≤ c * ∏ i, ‖m i‖} instance has_op_norm : has_norm (continuous_multilinear_map 𝕜 E G) := ⟨op_norm⟩ /-- An alias of `continuous_multilinear_map.has_op_norm` with non-dependent types to help typeclass @@ -287,26 +287,26 @@ search. -/ instance has_op_norm' : has_norm (continuous_multilinear_map 𝕜 (λ (i : ι), G) G') := continuous_multilinear_map.has_op_norm -lemma norm_def : ∥f∥ = Inf {c | 0 ≤ (c : ℝ) ∧ ∀ m, ∥f m∥ ≤ c * ∏ i, ∥m i∥} := rfl +lemma norm_def : ‖f‖ = Inf {c | 0 ≤ (c : ℝ) ∧ ∀ m, ‖f m‖ ≤ c * ∏ i, ‖m i‖} := rfl -- So that invocations of `le_cInf` make sense: we show that the set of -- bounds is nonempty and bounded below. lemma bounds_nonempty {f : continuous_multilinear_map 𝕜 E G} : - ∃ c, c ∈ {c | 0 ≤ c ∧ ∀ m, ∥f m∥ ≤ c * ∏ i, ∥m i∥} := + ∃ c, c ∈ {c | 0 ≤ c ∧ ∀ m, ‖f m‖ ≤ c * ∏ i, ‖m i‖} := let ⟨M, hMp, hMb⟩ := f.bound in ⟨M, le_of_lt hMp, hMb⟩ lemma bounds_bdd_below {f : continuous_multilinear_map 𝕜 E G} : - bdd_below {c | 0 ≤ c ∧ ∀ m, ∥f m∥ ≤ c * ∏ i, ∥m i∥} := + bdd_below {c | 0 ≤ c ∧ ∀ m, ‖f m‖ ≤ c * ∏ i, ‖m i‖} := ⟨0, λ _ ⟨hn, _⟩, hn⟩ -lemma op_norm_nonneg : 0 ≤ ∥f∥ := +lemma op_norm_nonneg : 0 ≤ ‖f‖ := le_cInf bounds_nonempty (λ _ ⟨hx, _⟩, hx) /-- The fundamental property of the operator norm of a continuous multilinear map: -`∥f m∥` is bounded by `∥f∥` times the product of the `∥m i∥`. -/ -theorem le_op_norm : ∥f m∥ ≤ ∥f∥ * ∏ i, ∥m i∥ := +`‖f m‖` is bounded by `‖f‖` times the product of the `‖m i‖`. -/ +theorem le_op_norm : ‖f m‖ ≤ ‖f‖ * ∏ i, ‖m i‖ := begin - have A : 0 ≤ ∏ i, ∥m i∥ := prod_nonneg (λj hj, norm_nonneg _), + have A : 0 ≤ ∏ i, ‖m i‖ := prod_nonneg (λj hj, norm_nonneg _), cases A.eq_or_lt with h hlt, { rcases prod_eq_zero_iff.1 h.symm with ⟨i, _, hi⟩, rw norm_eq_zero at hi, @@ -318,43 +318,43 @@ begin rintro c ⟨_, hc⟩, rw [div_le_iff hlt], apply hc } end -theorem le_of_op_norm_le {C : ℝ} (h : ∥f∥ ≤ C) : ∥f m∥ ≤ C * ∏ i, ∥m i∥ := +theorem le_of_op_norm_le {C : ℝ} (h : ‖f‖ ≤ C) : ‖f m‖ ≤ C * ∏ i, ‖m i‖ := (f.le_op_norm m).trans $ mul_le_mul_of_nonneg_right h (prod_nonneg $ λ i _, norm_nonneg (m i)) -lemma ratio_le_op_norm : ∥f m∥ / ∏ i, ∥m i∥ ≤ ∥f∥ := +lemma ratio_le_op_norm : ‖f m‖ / ∏ i, ‖m i‖ ≤ ‖f‖ := div_le_of_nonneg_of_le_mul (prod_nonneg $ λ i _, norm_nonneg _) (op_norm_nonneg _) (f.le_op_norm m) /-- The image of the unit ball under a continuous multilinear map is bounded. -/ -lemma unit_le_op_norm (h : ∥m∥ ≤ 1) : ∥f m∥ ≤ ∥f∥ := +lemma unit_le_op_norm (h : ‖m‖ ≤ 1) : ‖f m‖ ≤ ‖f‖ := calc - ∥f m∥ ≤ ∥f∥ * ∏ i, ∥m i∥ : f.le_op_norm m - ... ≤ ∥f∥ * ∏ i : ι, 1 : + ‖f m‖ ≤ ‖f‖ * ∏ i, ‖m i‖ : f.le_op_norm m + ... ≤ ‖f‖ * ∏ i : ι, 1 : mul_le_mul_of_nonneg_left (prod_le_prod (λi hi, norm_nonneg _) (λi hi, le_trans (norm_le_pi_norm (_ : Π i, E i) _) h)) (op_norm_nonneg f) - ... = ∥f∥ : by simp + ... = ‖f‖ : by simp /-- If one controls the norm of every `f x`, then one controls the norm of `f`. -/ -lemma op_norm_le_bound {M : ℝ} (hMp: 0 ≤ M) (hM : ∀ m, ∥f m∥ ≤ M * ∏ i, ∥m i∥) : - ∥f∥ ≤ M := +lemma op_norm_le_bound {M : ℝ} (hMp: 0 ≤ M) (hM : ∀ m, ‖f m‖ ≤ M * ∏ i, ‖m i‖) : + ‖f‖ ≤ M := cInf_le bounds_bdd_below ⟨hMp, hM⟩ /-- The operator norm satisfies the triangle inequality. -/ -theorem op_norm_add_le : ∥f + g∥ ≤ ∥f∥ + ∥g∥ := +theorem op_norm_add_le : ‖f + g‖ ≤ ‖f‖ + ‖g‖ := cInf_le bounds_bdd_below ⟨add_nonneg (op_norm_nonneg _) (op_norm_nonneg _), λ x, by { rw add_mul, exact norm_add_le_of_le (le_op_norm _ _) (le_op_norm _ _) }⟩ -lemma op_norm_zero : ∥(0 : continuous_multilinear_map 𝕜 E G)∥ = 0 := +lemma op_norm_zero : ‖(0 : continuous_multilinear_map 𝕜 E G)‖ = 0 := (op_norm_nonneg _).antisymm' $ op_norm_le_bound 0 le_rfl $ λ m, by simp /-- A continuous linear map is zero iff its norm vanishes. -/ -theorem op_norm_zero_iff : ∥f∥ = 0 ↔ f = 0 := +theorem op_norm_zero_iff : ‖f‖ = 0 ↔ f = 0 := ⟨λ h, by { ext m, simpa [h] using f.le_op_norm m }, by { rintro rfl, exact op_norm_zero }⟩ section variables {𝕜' : Type*} [normed_field 𝕜'] [normed_space 𝕜' G] [smul_comm_class 𝕜 𝕜' G] -lemma op_norm_smul_le (c : 𝕜') : ∥c • f∥ ≤ ∥c∥ * ∥f∥ := +lemma op_norm_smul_le (c : 𝕜') : ‖c • f‖ ≤ ‖c‖ * ‖f‖ := (c • f).op_norm_le_bound (mul_nonneg (norm_nonneg _) (op_norm_nonneg _)) begin @@ -363,7 +363,7 @@ lemma op_norm_smul_le (c : 𝕜') : ∥c • f∥ ≤ ∥c∥ * ∥f∥ := exact mul_le_mul_of_nonneg_left (le_op_norm _ _) (norm_nonneg _) end -lemma op_norm_neg : ∥-f∥ = ∥f∥ := by { rw norm_def, apply congr_arg, ext, simp } +lemma op_norm_neg : ‖-f‖ = ‖f‖ := by { rw norm_def, apply congr_arg, ext, simp } /-- Continuous multilinear maps themselves form a normed space with respect to the operator norm. -/ @@ -389,34 +389,34 @@ search. -/ instance normed_space' : normed_space 𝕜' (continuous_multilinear_map 𝕜 (λ i : ι, G') G) := continuous_multilinear_map.normed_space -theorem le_op_norm_mul_prod_of_le {b : ι → ℝ} (hm : ∀ i, ∥m i∥ ≤ b i) : ∥f m∥ ≤ ∥f∥ * ∏ i, b i := +theorem le_op_norm_mul_prod_of_le {b : ι → ℝ} (hm : ∀ i, ‖m i‖ ≤ b i) : ‖f m‖ ≤ ‖f‖ * ∏ i, b i := (f.le_op_norm m).trans $ mul_le_mul_of_nonneg_left (prod_le_prod (λ _ _, norm_nonneg _) (λ i _, hm i)) (norm_nonneg f) -theorem le_op_norm_mul_pow_card_of_le {b : ℝ} (hm : ∀ i, ∥m i∥ ≤ b) : - ∥f m∥ ≤ ∥f∥ * b ^ fintype.card ι := +theorem le_op_norm_mul_pow_card_of_le {b : ℝ} (hm : ∀ i, ‖m i‖ ≤ b) : + ‖f m‖ ≤ ‖f‖ * b ^ fintype.card ι := by simpa only [prod_const] using f.le_op_norm_mul_prod_of_le m hm theorem le_op_norm_mul_pow_of_le {Ei : fin n → Type*} [Π i, normed_add_comm_group (Ei i)] [Π i, normed_space 𝕜 (Ei i)] (f : continuous_multilinear_map 𝕜 Ei G) (m : Π i, Ei i) - {b : ℝ} (hm : ∥m∥ ≤ b) : - ∥f m∥ ≤ ∥f∥ * b ^ n := + {b : ℝ} (hm : ‖m‖ ≤ b) : + ‖f m‖ ≤ ‖f‖ * b ^ n := by simpa only [fintype.card_fin] using f.le_op_norm_mul_pow_card_of_le m (λ i, (norm_le_pi_norm m i).trans hm) /-- The fundamental property of the operator norm of a continuous multilinear map: -`∥f m∥` is bounded by `∥f∥` times the product of the `∥m i∥`, `nnnorm` version. -/ -theorem le_op_nnnorm : ∥f m∥₊ ≤ ∥f∥₊ * ∏ i, ∥m i∥₊ := +`‖f m‖` is bounded by `‖f‖` times the product of the `‖m i‖`, `nnnorm` version. -/ +theorem le_op_nnnorm : ‖f m‖₊ ≤ ‖f‖₊ * ∏ i, ‖m i‖₊ := nnreal.coe_le_coe.1 $ by { push_cast, exact f.le_op_norm m } -theorem le_of_op_nnnorm_le {C : ℝ≥0} (h : ∥f∥₊ ≤ C) : ∥f m∥₊ ≤ C * ∏ i, ∥m i∥₊ := +theorem le_of_op_nnnorm_le {C : ℝ≥0} (h : ‖f‖₊ ≤ C) : ‖f m‖₊ ≤ C * ∏ i, ‖m i‖₊ := (f.le_op_nnnorm m).trans $ mul_le_mul' h le_rfl lemma op_norm_prod (f : continuous_multilinear_map 𝕜 E G) (g : continuous_multilinear_map 𝕜 E G') : - ∥f.prod g∥ = max (∥f∥) (∥g∥) := + ‖f.prod g‖ = max (‖f‖) (‖g‖) := le_antisymm (op_norm_le_bound _ (norm_nonneg (f, g)) (λ m, - have H : 0 ≤ ∏ i, ∥m i∥, from prod_nonneg $ λ _ _, norm_nonneg _, + have H : 0 ≤ ∏ i, ‖m i‖, from prod_nonneg $ λ _ _, norm_nonneg _, by simpa only [prod_apply, prod.norm_def, max_mul_of_nonneg, H] using max_le_max (f.le_op_norm m) (g.le_op_norm m))) $ max_le @@ -425,7 +425,7 @@ le_antisymm lemma norm_pi {ι' : Type v'} [fintype ι'] {E' : ι' → Type wE'} [Π i', normed_add_comm_group (E' i')] [Π i', normed_space 𝕜 (E' i')] (f : Π i', continuous_multilinear_map 𝕜 E (E' i')) : - ∥pi f∥ = ∥f∥ := + ‖pi f‖ = ‖f‖ := begin apply le_antisymm, { refine (op_norm_le_bound _ (norm_nonneg f) (λ m, _)), @@ -480,7 +480,7 @@ variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 variables [normed_space 𝕜' G] [is_scalar_tower 𝕜' 𝕜 G] variables [Π i, normed_space 𝕜' (E i)] [∀ i, is_scalar_tower 𝕜' 𝕜 (E i)] -@[simp] lemma norm_restrict_scalars : ∥f.restrict_scalars 𝕜'∥ = ∥f∥ := +@[simp] lemma norm_restrict_scalars : ‖f.restrict_scalars 𝕜'‖ = ‖f‖ := by simp only [norm_def, coe_restrict_scalars] variable (𝕜') @@ -502,21 +502,21 @@ lemma continuous_restrict_scalars : end restrict_scalars -/-- The difference `f m₁ - f m₂` is controlled in terms of `∥f∥` and `∥m₁ - m₂∥`, precise version. +/-- The difference `f m₁ - f m₂` is controlled in terms of `‖f‖` and `‖m₁ - m₂‖`, precise version. For a less precise but more usable version, see `norm_image_sub_le`. The bound reads -`∥f m - f m'∥ ≤ - ∥f∥ * ∥m 1 - m' 1∥ * max ∥m 2∥ ∥m' 2∥ * max ∥m 3∥ ∥m' 3∥ * ... * max ∥m n∥ ∥m' n∥ + ...`, +`‖f m - f m'‖ ≤ + ‖f‖ * ‖m 1 - m' 1‖ * max ‖m 2‖ ‖m' 2‖ * max ‖m 3‖ ‖m' 3‖ * ... * max ‖m n‖ ‖m' n‖ + ...`, where the other terms in the sum are the same products where `1` is replaced by any `i`.-/ lemma norm_image_sub_le' (m₁ m₂ : Πi, E i) : - ∥f m₁ - f m₂∥ ≤ - ∥f∥ * ∑ i, ∏ j, if j = i then ∥m₁ i - m₂ i∥ else max ∥m₁ j∥ ∥m₂ j∥ := + ‖f m₁ - f m₂‖ ≤ + ‖f‖ * ∑ i, ∏ j, if j = i then ‖m₁ i - m₂ i‖ else max ‖m₁ j‖ ‖m₂ j‖ := f.to_multilinear_map.norm_image_sub_le_of_bound' (norm_nonneg _) f.le_op_norm _ _ -/-- The difference `f m₁ - f m₂` is controlled in terms of `∥f∥` and `∥m₁ - m₂∥`, less precise +/-- The difference `f m₁ - f m₂` is controlled in terms of `‖f‖` and `‖m₁ - m₂‖`, less precise version. For a more precise but less usable version, see `norm_image_sub_le'`. -The bound is `∥f m - f m'∥ ≤ ∥f∥ * card ι * ∥m - m'∥ * (max ∥m∥ ∥m'∥) ^ (card ι - 1)`.-/ +The bound is `‖f m - f m'‖ ≤ ‖f‖ * card ι * ‖m - m'‖ * (max ‖m‖ ‖m'‖) ^ (card ι - 1)`.-/ lemma norm_image_sub_le (m₁ m₂ : Πi, E i) : - ∥f m₁ - f m₂∥ ≤ ∥f∥ * (fintype.card ι) * (max ∥m₁∥ ∥m₂∥) ^ (fintype.card ι - 1) * ∥m₁ - m₂∥ := + ‖f m₁ - f m₂‖ ≤ ‖f‖ * (fintype.card ι) * (max ‖m₁‖ ‖m₂‖) ^ (fintype.card ι - 1) * ‖m₁ - m₂‖ := f.to_multilinear_map.norm_image_sub_le_of_bound (norm_nonneg _) f.le_op_norm _ _ /-- Applying a multilinear map to a vector is continuous in both coordinates. -/ @@ -525,27 +525,27 @@ lemma continuous_eval : begin apply continuous_iff_continuous_at.2 (λp, _), apply continuous_at_of_locally_lipschitz zero_lt_one - ((∥p∥ + 1) * (fintype.card ι) * (∥p∥ + 1) ^ (fintype.card ι - 1) + ∏ i, ∥p.2 i∥) + ((‖p‖ + 1) * (fintype.card ι) * (‖p‖ + 1) ^ (fintype.card ι - 1) + ∏ i, ‖p.2 i‖) (λq hq, _), - have : 0 ≤ (max ∥q.2∥ ∥p.2∥), by simp, - have : 0 ≤ ∥p∥ + 1 := zero_le_one.trans ((le_add_iff_nonneg_left 1).2 $ norm_nonneg p), - have A : ∥q∥ ≤ ∥p∥ + 1 := norm_le_of_mem_closed_ball hq.le, - have : (max ∥q.2∥ ∥p.2∥) ≤ ∥p∥ + 1 := + have : 0 ≤ (max ‖q.2‖ ‖p.2‖), by simp, + have : 0 ≤ ‖p‖ + 1 := zero_le_one.trans ((le_add_iff_nonneg_left 1).2 $ norm_nonneg p), + have A : ‖q‖ ≤ ‖p‖ + 1 := norm_le_of_mem_closed_ball hq.le, + have : (max ‖q.2‖ ‖p.2‖) ≤ ‖p‖ + 1 := (max_le_max (norm_snd_le q) (norm_snd_le p)).trans (by simp [A, -add_comm, zero_le_one]), - have : ∀ (i : ι), i ∈ univ → 0 ≤ ∥p.2 i∥ := λ i hi, norm_nonneg _, + have : ∀ (i : ι), i ∈ univ → 0 ≤ ‖p.2 i‖ := λ i hi, norm_nonneg _, calc dist (q.1 q.2) (p.1 p.2) ≤ dist (q.1 q.2) (q.1 p.2) + dist (q.1 p.2) (p.1 p.2) : dist_triangle _ _ _ - ... = ∥q.1 q.2 - q.1 p.2∥ + ∥q.1 p.2 - p.1 p.2∥ : by rw [dist_eq_norm, dist_eq_norm] - ... ≤ ∥q.1∥ * (fintype.card ι) * (max ∥q.2∥ ∥p.2∥) ^ (fintype.card ι - 1) * ∥q.2 - p.2∥ - + ∥q.1 - p.1∥ * ∏ i, ∥p.2 i∥ : + ... = ‖q.1 q.2 - q.1 p.2‖ + ‖q.1 p.2 - p.1 p.2‖ : by rw [dist_eq_norm, dist_eq_norm] + ... ≤ ‖q.1‖ * (fintype.card ι) * (max ‖q.2‖ ‖p.2‖) ^ (fintype.card ι - 1) * ‖q.2 - p.2‖ + + ‖q.1 - p.1‖ * ∏ i, ‖p.2 i‖ : add_le_add (norm_image_sub_le _ _ _) ((q.1 - p.1).le_op_norm p.2) - ... ≤ (∥p∥ + 1) * (fintype.card ι) * (∥p∥ + 1) ^ (fintype.card ι - 1) * ∥q - p∥ - + ∥q - p∥ * ∏ i, ∥p.2 i∥ : + ... ≤ (‖p‖ + 1) * (fintype.card ι) * (‖p‖ + 1) ^ (fintype.card ι - 1) * ‖q - p‖ + + ‖q - p‖ * ∏ i, ‖p.2 i‖ : by apply_rules [add_le_add, mul_le_mul, le_refl, le_trans (norm_fst_le q) A, nat.cast_nonneg, mul_nonneg, pow_le_pow_of_le_left, pow_nonneg, norm_snd_le (q - p), norm_nonneg, norm_fst_le (q - p), prod_nonneg] - ... = ((∥p∥ + 1) * (fintype.card ι) * (∥p∥ + 1) ^ (fintype.card ι - 1) - + (∏ i, ∥p.2 i∥)) * dist q p : by { rw dist_eq_norm, ring } + ... = ((‖p‖ + 1) * (fintype.card ι) * (‖p‖ + 1) ^ (fintype.card ι - 1) + + (∏ i, ‖p.2 i‖)) * dist q p : by { rw dist_eq_norm, ring } end lemma continuous_eval_left (m : Π i, E i) : @@ -576,7 +576,7 @@ case from the multilinear case via a currying isomorphism. However, this would m and it is more satisfactory to have the simplest case as a standalone proof. -/ instance [complete_space G] : complete_space (continuous_multilinear_map 𝕜 E G) := begin - have nonneg : ∀ (v : Π i, E i), 0 ≤ ∏ i, ∥v i∥ := + have nonneg : ∀ (v : Π i, E i), 0 ≤ ∏ i, ‖v i‖ := λ v, finset.prod_nonneg (λ i hi, norm_nonneg _), -- We show that every Cauchy sequence converges. refine metric.complete_of_cauchy_seq_tendsto (λ f hf, _), @@ -585,7 +585,7 @@ begin -- and establish that the evaluation at any point `v : Π i, E i` is Cauchy. have cau : ∀ v, cauchy_seq (λ n, f n v), { assume v, - apply cauchy_seq_iff_le_tendsto_0.2 ⟨λ n, b n * ∏ i, ∥v i∥, λ n, _, _, _⟩, + apply cauchy_seq_iff_le_tendsto_0.2 ⟨λ n, b n * ∏ i, ‖v i‖, λ n, _, _, _⟩, { exact mul_nonneg (b0 n) (nonneg v) }, { assume n m N hn hm, rw dist_eq_norm, @@ -611,16 +611,16 @@ begin simp at A B, exact tendsto_nhds_unique A B end }, - -- and that `F` has norm at most `(b 0 + ∥f 0∥)`. - have Fnorm : ∀ v, ∥F v∥ ≤ (b 0 + ∥f 0∥) * ∏ i, ∥v i∥, + -- and that `F` has norm at most `(b 0 + ‖f 0‖)`. + have Fnorm : ∀ v, ‖F v‖ ≤ (b 0 + ‖f 0‖) * ∏ i, ‖v i‖, { assume v, - have A : ∀ n, ∥f n v∥ ≤ (b 0 + ∥f 0∥) * ∏ i, ∥v i∥, + have A : ∀ n, ‖f n v‖ ≤ (b 0 + ‖f 0‖) * ∏ i, ‖v i‖, { assume n, apply le_trans ((f n).le_op_norm _) _, apply mul_le_mul_of_nonneg_right _ (nonneg v), - calc ∥f n∥ = ∥(f n - f 0) + f 0∥ : by { congr' 1, abel } - ... ≤ ∥f n - f 0∥ + ∥f 0∥ : norm_add_le _ _ - ... ≤ b 0 + ∥f 0∥ : begin + calc ‖f n‖ = ‖(f n - f 0) + f 0‖ : by { congr' 1, abel } + ... ≤ ‖f n - f 0‖ + ‖f 0‖ : norm_add_le _ _ + ... ≤ b 0 + ‖f 0‖ : begin apply add_le_add_right, simpa [dist_eq_norm] using b_bound n 0 0 (zero_le _) (zero_le _) end }, @@ -629,14 +629,14 @@ begin let Fcont := Fmult.mk_continuous _ Fnorm, use Fcont, -- Our last task is to establish convergence to `F` in norm. - have : ∀ n, ∥f n - Fcont∥ ≤ b n, + have : ∀ n, ‖f n - Fcont‖ ≤ b n, { assume n, apply op_norm_le_bound _ (b0 n) (λ v, _), - have A : ∀ᶠ m in at_top, ∥(f n - f m) v∥ ≤ b n * ∏ i, ∥v i∥, + have A : ∀ᶠ m in at_top, ‖(f n - f m) v‖ ≤ b n * ∏ i, ‖v i‖, { refine eventually_at_top.2 ⟨n, λ m hm, _⟩, apply le_trans ((f n - f m).le_op_norm _) _, exact mul_le_mul_of_nonneg_right (b_bound n m n le_rfl hm) (nonneg v) }, - have B : tendsto (λ m, ∥(f n - f m) v∥) at_top (𝓝 (∥(f n - Fcont) v∥)) := + have B : tendsto (λ m, ‖(f n - f m) v‖) at_top (𝓝 (‖(f n - Fcont) v‖)) := tendsto.norm (tendsto_const_nhds.sub (hF v)), exact le_of_tendsto B A }, erw tendsto_iff_norm_tendsto_zero, @@ -649,14 +649,14 @@ end continuous_multilinear_map `mk_continuous`, then its norm is bounded by the bound given to the constructor if it is nonnegative. -/ lemma multilinear_map.mk_continuous_norm_le (f : multilinear_map 𝕜 E G) {C : ℝ} (hC : 0 ≤ C) - (H : ∀ m, ∥f m∥ ≤ C * ∏ i, ∥m i∥) : ∥f.mk_continuous C H∥ ≤ C := + (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) : ‖f.mk_continuous C H‖ ≤ C := continuous_multilinear_map.op_norm_le_bound _ hC (λm, H m) /-- If a continuous multilinear map is constructed from a multilinear map via the constructor `mk_continuous`, then its norm is bounded by the bound given to the constructor if it is nonnegative. -/ lemma multilinear_map.mk_continuous_norm_le' (f : multilinear_map 𝕜 E G) {C : ℝ} - (H : ∀ m, ∥f m∥ ≤ C * ∏ i, ∥m i∥) : ∥f.mk_continuous C H∥ ≤ max C 0 := + (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) : ‖f.mk_continuous C H‖ ≤ max C 0 := continuous_multilinear_map.op_norm_le_bound _ (le_max_right _ _) $ λ m, (H m).trans $ mul_le_mul_of_nonneg_right (le_max_left _ _) (prod_nonneg $ λ _ _, norm_nonneg _) @@ -671,10 +671,10 @@ identification between `fin k` and `s` that we use is the canonical (increasing) def restr {k n : ℕ} (f : (G [×n]→L[𝕜] G' : _)) (s : finset (fin n)) (hk : s.card = k) (z : G) : G [×k]→L[𝕜] G' := (f.to_multilinear_map.restr s hk z).mk_continuous -(∥f∥ * ∥z∥^(n-k)) $ λ v, multilinear_map.restr_norm_le _ _ _ _ f.le_op_norm _ +(‖f‖ * ‖z‖^(n-k)) $ λ v, multilinear_map.restr_norm_le _ _ _ _ f.le_op_norm _ lemma norm_restr {k n : ℕ} (f : G [×n]→L[𝕜] G') (s : finset (fin n)) (hk : s.card = k) (z : G) : - ∥f.restr s hk z∥ ≤ ∥f∥ * ∥z∥ ^ (n - k) := + ‖f.restr s hk z‖ ≤ ‖f‖ * ‖z‖ ^ (n - k) := begin apply multilinear_map.mk_continuous_norm_le, exact mul_nonneg (norm_nonneg _) (pow_nonneg (norm_nonneg _) _) @@ -686,7 +686,7 @@ variables {𝕜 ι} {A : Type*} [normed_comm_ring A] [normed_algebra 𝕜 A] @[simp] lemma norm_mk_pi_algebra_le [nonempty ι] : - ∥continuous_multilinear_map.mk_pi_algebra 𝕜 ι A∥ ≤ 1 := + ‖continuous_multilinear_map.mk_pi_algebra 𝕜 ι A‖ ≤ 1 := begin have := λ f, @op_norm_le_bound 𝕜 ι (λ i, A) A _ _ _ _ _ _ _ f _ zero_le_one, refine this _ _, @@ -696,7 +696,7 @@ begin end lemma norm_mk_pi_algebra_of_empty [is_empty ι] : - ∥continuous_multilinear_map.mk_pi_algebra 𝕜 ι A∥ = ∥(1 : A)∥ := + ‖continuous_multilinear_map.mk_pi_algebra 𝕜 ι A‖ = ‖(1 : A)‖ := begin apply le_antisymm, { have := λ f, @op_norm_le_bound 𝕜 ι (λ i, A) A _ _ _ _ _ _ _ f _ (norm_nonneg (1 : A)), @@ -707,7 +707,7 @@ begin end @[simp] lemma norm_mk_pi_algebra [norm_one_class A] : - ∥continuous_multilinear_map.mk_pi_algebra 𝕜 ι A∥ = 1 := + ‖continuous_multilinear_map.mk_pi_algebra 𝕜 ι A‖ = 1 := begin casesI is_empty_or_nonempty ι, { simp [norm_mk_pi_algebra_of_empty] }, @@ -723,7 +723,7 @@ section variables {𝕜 n} {A : Type*} [normed_ring A] [normed_algebra 𝕜 A] lemma norm_mk_pi_algebra_fin_succ_le : - ∥continuous_multilinear_map.mk_pi_algebra_fin 𝕜 n.succ A∥ ≤ 1 := + ‖continuous_multilinear_map.mk_pi_algebra_fin 𝕜 n.succ A‖ ≤ 1 := begin have := λ f, @op_norm_le_bound 𝕜 (fin n.succ) (λ i, A) A _ _ _ _ _ _ _ f _ zero_le_one, refine this _ _, @@ -737,14 +737,14 @@ begin end lemma norm_mk_pi_algebra_fin_le_of_pos (hn : 0 < n) : - ∥continuous_multilinear_map.mk_pi_algebra_fin 𝕜 n A∥ ≤ 1 := + ‖continuous_multilinear_map.mk_pi_algebra_fin 𝕜 n A‖ ≤ 1 := begin obtain ⟨n, rfl⟩ := nat.exists_eq_succ_of_ne_zero hn.ne', exact norm_mk_pi_algebra_fin_succ_le end lemma norm_mk_pi_algebra_fin_zero : - ∥continuous_multilinear_map.mk_pi_algebra_fin 𝕜 0 A∥ = ∥(1 : A)∥ := + ‖continuous_multilinear_map.mk_pi_algebra_fin 𝕜 0 A‖ = ‖(1 : A)‖ := begin refine le_antisymm _ _, { have := λ f, @op_norm_le_bound 𝕜 (fin 0) (λ i, A) A _ _ _ _ _ _ _ f _ (norm_nonneg (1 : A)), @@ -755,7 +755,7 @@ begin end @[simp] lemma norm_mk_pi_algebra_fin [norm_one_class A] : - ∥continuous_multilinear_map.mk_pi_algebra_fin 𝕜 n A∥ = 1 := + ‖continuous_multilinear_map.mk_pi_algebra_fin 𝕜 n A‖ = 1 := begin cases n, { simp [norm_mk_pi_algebra_fin_zero] }, @@ -772,7 +772,7 @@ variables (𝕜 ι) `m i` (multiplied by a fixed reference element `z` in the target module) -/ protected def mk_pi_field (z : G) : continuous_multilinear_map 𝕜 (λ(i : ι), 𝕜) G := multilinear_map.mk_continuous - (multilinear_map.mk_pi_ring 𝕜 ι z) (∥z∥) + (multilinear_map.mk_pi_ring 𝕜 ι z) (‖z‖) (λ m, by simp only [multilinear_map.mk_pi_ring_apply, norm_smul, norm_prod, mul_comm]) @@ -785,7 +785,7 @@ lemma mk_pi_field_apply_one_eq_self (f : continuous_multilinear_map 𝕜 (λ(i : continuous_multilinear_map.mk_pi_field 𝕜 ι (f (λi, 1)) = f := to_multilinear_map_inj f.to_multilinear_map.mk_pi_ring_apply_one_eq_self -@[simp] lemma norm_mk_pi_field (z : G) : ∥continuous_multilinear_map.mk_pi_field 𝕜 ι z∥ = ∥z∥ := +@[simp] lemma norm_mk_pi_field (z : G) : ‖continuous_multilinear_map.mk_pi_field 𝕜 ι z‖ = ‖z‖ := (multilinear_map.mk_continuous_norm_le _ (norm_nonneg z) _).antisymm $ by simpa using (continuous_multilinear_map.mk_pi_field 𝕜 ι z).le_op_norm (λ _, 1) @@ -826,9 +826,9 @@ namespace continuous_linear_map lemma norm_comp_continuous_multilinear_map_le (g : G →L[𝕜] G') (f : continuous_multilinear_map 𝕜 E G) : - ∥g.comp_continuous_multilinear_map f∥ ≤ ∥g∥ * ∥f∥ := + ‖g.comp_continuous_multilinear_map f‖ ≤ ‖g‖ * ‖f‖ := continuous_multilinear_map.op_norm_le_bound _ (mul_nonneg (norm_nonneg _) (norm_nonneg _)) $ λ m, -calc ∥g (f m)∥ ≤ ∥g∥ * (∥f∥ * ∏ i, ∥m i∥) : g.le_op_norm_of_le $ f.le_op_norm _ +calc ‖g (f m)‖ ≤ ‖g‖ * (‖f‖ * ∏ i, ‖m i‖) : g.le_op_norm_of_le $ f.le_op_norm _ ... = _ : (mul_assoc _ _ _).symm variables (𝕜 E G G') @@ -852,7 +852,7 @@ multilinear_map.mk_continuous map_add' := λ x y, by simp only [map_add, continuous_multilinear_map.add_apply], map_smul' := λ c x, by simp only [continuous_multilinear_map.smul_apply, map_smul, ring_hom.id_apply] } - (∥f∥ * ∏ i, ∥m i∥) $ λ x, + (‖f‖ * ∏ i, ‖m i‖) $ λ x, by { rw mul_right_comm, exact (f x).le_of_op_norm_le _ (f.le_op_norm x) }, map_add' := λ m i x y, by { ext1, simp only [add_apply, continuous_multilinear_map.map_add, linear_map.coe_mk, @@ -860,7 +860,7 @@ multilinear_map.mk_continuous map_smul' := λ m i c x, by { ext1, simp only [coe_smul', continuous_multilinear_map.map_smul, linear_map.coe_mk, linear_map.mk_continuous_apply, pi.smul_apply]} } - ∥f∥ $ λ m, + ‖f‖ $ λ m, linear_map.mk_continuous_norm_le _ (mul_nonneg (norm_nonneg f) (prod_nonneg $ λ i hi, norm_nonneg (m i))) _ @@ -870,7 +870,7 @@ open continuous_multilinear_map namespace multilinear_map /-- Given a map `f : G →ₗ[𝕜] multilinear_map 𝕜 E G'` and an estimate -`H : ∀ x m, ∥f x m∥ ≤ C * ∥x∥ * ∏ i, ∥m i∥`, construct a continuous linear +`H : ∀ x m, ‖f x m‖ ≤ C * ‖x‖ * ∏ i, ‖m i‖`, construct a continuous linear map from `G` to `continuous_multilinear_map 𝕜 E G'`. In order to lift, e.g., a map `f : (multilinear_map 𝕜 E G) →ₗ[𝕜] multilinear_map 𝕜 E' G'` @@ -878,57 +878,57 @@ to a map `(continuous_multilinear_map 𝕜 E G) →L[𝕜] continuous_multilinea one can apply this construction to `f.comp continuous_multilinear_map.to_multilinear_map_linear` which is a linear map from `continuous_multilinear_map 𝕜 E G` to `multilinear_map 𝕜 E' G'`. -/ def mk_continuous_linear (f : G →ₗ[𝕜] multilinear_map 𝕜 E G') (C : ℝ) - (H : ∀ x m, ∥f x m∥ ≤ C * ∥x∥ * ∏ i, ∥m i∥) : + (H : ∀ x m, ‖f x m‖ ≤ C * ‖x‖ * ∏ i, ‖m i‖) : G →L[𝕜] continuous_multilinear_map 𝕜 E G' := linear_map.mk_continuous - { to_fun := λ x, (f x).mk_continuous (C * ∥x∥) $ H x, + { to_fun := λ x, (f x).mk_continuous (C * ‖x‖) $ H x, map_add' := λ x y, by { ext1, simp only [_root_.map_add], refl }, map_smul' := λ c x, by { ext1, simp only [smul_hom_class.map_smul], refl } } (max C 0) $ λ x, ((f x).mk_continuous_norm_le' _).trans_eq $ by rw [max_mul_of_nonneg _ _ (norm_nonneg x), zero_mul] lemma mk_continuous_linear_norm_le' (f : G →ₗ[𝕜] multilinear_map 𝕜 E G') (C : ℝ) - (H : ∀ x m, ∥f x m∥ ≤ C * ∥x∥ * ∏ i, ∥m i∥) : - ∥mk_continuous_linear f C H∥ ≤ max C 0 := + (H : ∀ x m, ‖f x m‖ ≤ C * ‖x‖ * ∏ i, ‖m i‖) : + ‖mk_continuous_linear f C H‖ ≤ max C 0 := begin dunfold mk_continuous_linear, exact linear_map.mk_continuous_norm_le _ (le_max_right _ _) _ end lemma mk_continuous_linear_norm_le (f : G →ₗ[𝕜] multilinear_map 𝕜 E G') {C : ℝ} (hC : 0 ≤ C) - (H : ∀ x m, ∥f x m∥ ≤ C * ∥x∥ * ∏ i, ∥m i∥) : - ∥mk_continuous_linear f C H∥ ≤ C := + (H : ∀ x m, ‖f x m‖ ≤ C * ‖x‖ * ∏ i, ‖m i‖) : + ‖mk_continuous_linear f C H‖ ≤ C := (mk_continuous_linear_norm_le' f C H).trans_eq (max_eq_left hC) /-- Given a map `f : multilinear_map 𝕜 E (multilinear_map 𝕜 E' G)` and an estimate -`H : ∀ m m', ∥f m m'∥ ≤ C * ∏ i, ∥m i∥ * ∏ i, ∥m' i∥`, upgrade all `multilinear_map`s in the type to +`H : ∀ m m', ‖f m m'‖ ≤ C * ∏ i, ‖m i‖ * ∏ i, ‖m' i‖`, upgrade all `multilinear_map`s in the type to `continuous_multilinear_map`s. -/ def mk_continuous_multilinear (f : multilinear_map 𝕜 E (multilinear_map 𝕜 E' G)) (C : ℝ) - (H : ∀ m₁ m₂, ∥f m₁ m₂∥ ≤ C * (∏ i, ∥m₁ i∥) * ∏ i, ∥m₂ i∥) : + (H : ∀ m₁ m₂, ‖f m₁ m₂‖ ≤ C * (∏ i, ‖m₁ i‖) * ∏ i, ‖m₂ i‖) : continuous_multilinear_map 𝕜 E (continuous_multilinear_map 𝕜 E' G) := mk_continuous - { to_fun := λ m, mk_continuous (f m) (C * ∏ i, ∥m i∥) $ H m, + { to_fun := λ m, mk_continuous (f m) (C * ∏ i, ‖m i‖) $ H m, map_add' := λ m i x y, by { ext1, simp }, map_smul' := λ m i c x, by { ext1, simp } } (max C 0) $ λ m, ((f m).mk_continuous_norm_le' _).trans_eq $ by { rw [max_mul_of_nonneg, zero_mul], exact prod_nonneg (λ _ _, norm_nonneg _) } @[simp] lemma mk_continuous_multilinear_apply (f : multilinear_map 𝕜 E (multilinear_map 𝕜 E' G)) - {C : ℝ} (H : ∀ m₁ m₂, ∥f m₁ m₂∥ ≤ C * (∏ i, ∥m₁ i∥) * ∏ i, ∥m₂ i∥) (m : Π i, E i) : + {C : ℝ} (H : ∀ m₁ m₂, ‖f m₁ m₂‖ ≤ C * (∏ i, ‖m₁ i‖) * ∏ i, ‖m₂ i‖) (m : Π i, E i) : ⇑(mk_continuous_multilinear f C H m) = f m := rfl lemma mk_continuous_multilinear_norm_le' (f : multilinear_map 𝕜 E (multilinear_map 𝕜 E' G)) (C : ℝ) - (H : ∀ m₁ m₂, ∥f m₁ m₂∥ ≤ C * (∏ i, ∥m₁ i∥) * ∏ i, ∥m₂ i∥) : - ∥mk_continuous_multilinear f C H∥ ≤ max C 0 := + (H : ∀ m₁ m₂, ‖f m₁ m₂‖ ≤ C * (∏ i, ‖m₁ i‖) * ∏ i, ‖m₂ i‖) : + ‖mk_continuous_multilinear f C H‖ ≤ max C 0 := begin dunfold mk_continuous_multilinear, exact mk_continuous_norm_le _ (le_max_right _ _) _ end lemma mk_continuous_multilinear_norm_le (f : multilinear_map 𝕜 E (multilinear_map 𝕜 E' G)) {C : ℝ} - (hC : 0 ≤ C) (H : ∀ m₁ m₂, ∥f m₁ m₂∥ ≤ C * (∏ i, ∥m₁ i∥) * ∏ i, ∥m₂ i∥) : - ∥mk_continuous_multilinear f C H∥ ≤ C := + (hC : 0 ≤ C) (H : ∀ m₁ m₂, ‖f m₁ m₂‖ ≤ C * (∏ i, ‖m₁ i‖) * ∏ i, ‖m₂ i‖) : + ‖mk_continuous_multilinear f C H‖ ≤ C := (mk_continuous_multilinear_norm_le' f C H).trans_eq (max_eq_left hC) end multilinear_map @@ -937,13 +937,13 @@ namespace continuous_multilinear_map lemma norm_comp_continuous_linear_le (g : continuous_multilinear_map 𝕜 E₁ G) (f : Π i, E i →L[𝕜] E₁ i) : - ∥g.comp_continuous_linear_map f∥ ≤ ∥g∥ * ∏ i, ∥f i∥ := + ‖g.comp_continuous_linear_map f‖ ≤ ‖g‖ * ∏ i, ‖f i‖ := op_norm_le_bound _ (mul_nonneg (norm_nonneg _) $ prod_nonneg $ λ i hi, norm_nonneg _) $ λ m, -calc ∥g (λ i, f i (m i))∥ ≤ ∥g∥ * ∏ i, ∥f i (m i)∥ : g.le_op_norm _ -... ≤ ∥g∥ * ∏ i, (∥f i∥ * ∥m i∥) : +calc ‖g (λ i, f i (m i))‖ ≤ ‖g‖ * ∏ i, ‖f i (m i)‖ : g.le_op_norm _ +... ≤ ‖g‖ * ∏ i, (‖f i‖ * ‖m i‖) : mul_le_mul_of_nonneg_left (prod_le_prod (λ _ _, norm_nonneg _) (λ i hi, (f i).le_op_norm (m i))) (norm_nonneg g) -... = (∥g∥ * ∏ i, ∥f i∥) * ∏ i, ∥m i∥ : by rw [prod_mul_distrib, mul_assoc] +... = (‖g‖ * ∏ i, ‖f i‖) * ∏ i, ‖m i‖ : by rw [prod_mul_distrib, mul_assoc] /-- `continuous_multilinear_map.comp_continuous_linear_map` as a bundled continuous linear map. This implementation fixes `f : Π i, E i →L[𝕜] E₁ i`. @@ -956,7 +956,7 @@ linear_map.mk_continuous { to_fun := λ g, g.comp_continuous_linear_map f, map_add' := λ g₁ g₂, rfl, map_smul' := λ c g, rfl } - (∏ i, ∥f i∥) $ λ g, (norm_comp_continuous_linear_le _ _).trans_eq (mul_comm _ _) + (∏ i, ‖f i‖) $ λ g, (norm_comp_continuous_linear_le _ _).trans_eq (mul_comm _ _) @[simp] lemma comp_continuous_linear_mapL_apply (g : continuous_multilinear_map 𝕜 E₁ G) (f : Π i, E i →L[𝕜] E₁ i) : @@ -964,7 +964,7 @@ linear_map.mk_continuous rfl lemma norm_comp_continuous_linear_mapL_le (f : Π i, E i →L[𝕜] E₁ i) : - ∥@comp_continuous_linear_mapL 𝕜 ι E E₁ G _ _ _ _ _ _ _ _ _ f∥ ≤ (∏ i, ∥f i∥) := + ‖@comp_continuous_linear_mapL 𝕜 ι E E₁ G _ _ _ _ _ _ _ _ _ f‖ ≤ (∏ i, ‖f i‖) := linear_map.mk_continuous_norm_le _ (prod_nonneg $ λ i _, norm_nonneg _) _ end continuous_multilinear_map @@ -997,38 +997,38 @@ open fin function lemma continuous_linear_map.norm_map_tail_le (f : Ei 0 →L[𝕜] (continuous_multilinear_map 𝕜 (λ(i : fin n), Ei i.succ) G)) (m : Πi, Ei i) : - ∥f (m 0) (tail m)∥ ≤ ∥f∥ * ∏ i, ∥m i∥ := + ‖f (m 0) (tail m)‖ ≤ ‖f‖ * ∏ i, ‖m i‖ := calc - ∥f (m 0) (tail m)∥ ≤ ∥f (m 0)∥ * ∏ i, ∥(tail m) i∥ : (f (m 0)).le_op_norm _ - ... ≤ (∥f∥ * ∥m 0∥) * ∏ i, ∥(tail m) i∥ : + ‖f (m 0) (tail m)‖ ≤ ‖f (m 0)‖ * ∏ i, ‖(tail m) i‖ : (f (m 0)).le_op_norm _ + ... ≤ (‖f‖ * ‖m 0‖) * ∏ i, ‖(tail m) i‖ : mul_le_mul_of_nonneg_right (f.le_op_norm _) (prod_nonneg (λi hi, norm_nonneg _)) - ... = ∥f∥ * (∥m 0∥ * ∏ i, ∥(tail m) i∥) : by ring - ... = ∥f∥ * ∏ i, ∥m i∥ : by { rw prod_univ_succ, refl } + ... = ‖f‖ * (‖m 0‖ * ∏ i, ‖(tail m) i‖) : by ring + ... = ‖f‖ * ∏ i, ‖m i‖ : by { rw prod_univ_succ, refl } lemma continuous_multilinear_map.norm_map_init_le (f : continuous_multilinear_map 𝕜 (λ(i : fin n), Ei i.cast_succ) (Ei (last n) →L[𝕜] G)) (m : Πi, Ei i) : - ∥f (init m) (m (last n))∥ ≤ ∥f∥ * ∏ i, ∥m i∥ := + ‖f (init m) (m (last n))‖ ≤ ‖f‖ * ∏ i, ‖m i‖ := calc - ∥f (init m) (m (last n))∥ ≤ ∥f (init m)∥ * ∥m (last n)∥ : (f (init m)).le_op_norm _ - ... ≤ (∥f∥ * (∏ i, ∥(init m) i∥)) * ∥m (last n)∥ : + ‖f (init m) (m (last n))‖ ≤ ‖f (init m)‖ * ‖m (last n)‖ : (f (init m)).le_op_norm _ + ... ≤ (‖f‖ * (∏ i, ‖(init m) i‖)) * ‖m (last n)‖ : mul_le_mul_of_nonneg_right (f.le_op_norm _) (norm_nonneg _) - ... = ∥f∥ * ((∏ i, ∥(init m) i∥) * ∥m (last n)∥) : mul_assoc _ _ _ - ... = ∥f∥ * ∏ i, ∥m i∥ : by { rw prod_univ_cast_succ, refl } + ... = ‖f‖ * ((∏ i, ‖(init m) i‖) * ‖m (last n)‖) : mul_assoc _ _ _ + ... = ‖f‖ * ∏ i, ‖m i‖ : by { rw prod_univ_cast_succ, refl } lemma continuous_multilinear_map.norm_map_cons_le (f : continuous_multilinear_map 𝕜 Ei G) (x : Ei 0) (m : Π(i : fin n), Ei i.succ) : - ∥f (cons x m)∥ ≤ ∥f∥ * ∥x∥ * ∏ i, ∥m i∥ := + ‖f (cons x m)‖ ≤ ‖f‖ * ‖x‖ * ∏ i, ‖m i‖ := calc - ∥f (cons x m)∥ ≤ ∥f∥ * ∏ i, ∥cons x m i∥ : f.le_op_norm _ - ... = (∥f∥ * ∥x∥) * ∏ i, ∥m i∥ : by { rw prod_univ_succ, simp [mul_assoc] } + ‖f (cons x m)‖ ≤ ‖f‖ * ∏ i, ‖cons x m i‖ : f.le_op_norm _ + ... = (‖f‖ * ‖x‖) * ∏ i, ‖m i‖ : by { rw prod_univ_succ, simp [mul_assoc] } lemma continuous_multilinear_map.norm_map_snoc_le (f : continuous_multilinear_map 𝕜 Ei G) (m : Π(i : fin n), Ei i.cast_succ) (x : Ei (last n)) : - ∥f (snoc m x)∥ ≤ ∥f∥ * (∏ i, ∥m i∥) * ∥x∥ := + ‖f (snoc m x)‖ ≤ ‖f‖ * (∏ i, ‖m i‖) * ‖x‖ := calc - ∥f (snoc m x)∥ ≤ ∥f∥ * ∏ i, ∥snoc m x i∥ : f.le_op_norm _ - ... = ∥f∥ * (∏ i, ∥m i∥) * ∥x∥ : by { rw prod_univ_cast_succ, simp [mul_assoc] } + ‖f (snoc m x)‖ ≤ ‖f‖ * ∏ i, ‖snoc m x i‖ : f.le_op_norm _ + ... = ‖f‖ * (∏ i, ‖m i‖) * ‖x‖ : by { rw prod_univ_cast_succ, simp [mul_assoc] } /-! #### Left currying -/ @@ -1040,7 +1040,7 @@ def continuous_linear_map.uncurry_left continuous_multilinear_map 𝕜 Ei G := (@linear_map.uncurry_left 𝕜 n Ei G _ _ _ _ _ (continuous_multilinear_map.to_multilinear_map_linear.comp f.to_linear_map)).mk_continuous - (∥f∥) (λm, continuous_linear_map.norm_map_tail_le f m) + (‖f‖) (λm, continuous_linear_map.norm_map_tail_le f m) @[simp] lemma continuous_linear_map.uncurry_left_apply (f : Ei 0 →L[𝕜] (continuous_multilinear_map 𝕜 (λ(i : fin n), Ei i.succ) G)) (m : Πi, Ei i) : @@ -1056,11 +1056,11 @@ linear_map.mk_continuous { -- define a linear map into `n` continuous multilinear maps from an `n+1` continuous multilinear -- map to_fun := λx, (f.to_multilinear_map.curry_left x).mk_continuous - (∥f∥ * ∥x∥) (f.norm_map_cons_le x), + (‖f‖ * ‖x‖) (f.norm_map_cons_le x), map_add' := λx y, by { ext m, exact f.cons_add m x y }, map_smul' := λc x, by { ext m, exact f.cons_smul m c x } } -- then register its continuity thanks to its boundedness properties. -(∥f∥) (λx, multilinear_map.mk_continuous_norm_le _ (mul_nonneg (norm_nonneg _) (norm_nonneg _)) _) +(‖f‖) (λx, multilinear_map.mk_continuous_norm_le _ (mul_nonneg (norm_nonneg _) (norm_nonneg _)) _) @[simp] lemma continuous_multilinear_map.curry_left_apply (f : continuous_multilinear_map 𝕜 Ei G) (x : Ei 0) (m : Π(i : fin n), Ei i.succ) : @@ -1114,12 +1114,12 @@ variables {𝕜 Ei G} (continuous_multilinear_curry_left_equiv 𝕜 Ei G).symm f x v = f (cons x v) := rfl @[simp] lemma continuous_multilinear_map.curry_left_norm - (f : continuous_multilinear_map 𝕜 Ei G) : ∥f.curry_left∥ = ∥f∥ := + (f : continuous_multilinear_map 𝕜 Ei G) : ‖f.curry_left‖ = ‖f‖ := (continuous_multilinear_curry_left_equiv 𝕜 Ei G).symm.norm_map f @[simp] lemma continuous_linear_map.uncurry_left_norm (f : Ei 0 →L[𝕜] (continuous_multilinear_map 𝕜 (λ(i : fin n), Ei i.succ) G)) : - ∥f.uncurry_left∥ = ∥f∥ := + ‖f.uncurry_left‖ = ‖f‖ := (continuous_multilinear_curry_left_equiv 𝕜 Ei G).norm_map f /-! #### Right currying -/ @@ -1135,7 +1135,7 @@ let f' : multilinear_map 𝕜 (λ(i : fin n), Ei i.cast_succ) (Ei (last n) → map_add' := λ m i x y, by simp, map_smul' := λ m i c x, by simp } in (@multilinear_map.uncurry_right 𝕜 n Ei G _ _ _ _ _ f').mk_continuous - (∥f∥) (λm, f.norm_map_init_le m) + (‖f‖) (λm, f.norm_map_init_le m) @[simp] lemma continuous_multilinear_map.uncurry_right_apply (f : continuous_multilinear_map 𝕜 (λ(i : fin n), Ei i.cast_succ) (Ei (last n) →L[𝕜] G)) @@ -1150,10 +1150,10 @@ def continuous_multilinear_map.curry_right continuous_multilinear_map 𝕜 (λ i : fin n, Ei i.cast_succ) (Ei (last n) →L[𝕜] G) := let f' : multilinear_map 𝕜 (λ(i : fin n), Ei i.cast_succ) (Ei (last n) →L[𝕜] G) := { to_fun := λm, (f.to_multilinear_map.curry_right m).mk_continuous - (∥f∥ * ∏ i, ∥m i∥) $ λx, f.norm_map_snoc_le m x, + (‖f‖ * ∏ i, ‖m i‖) $ λx, f.norm_map_snoc_le m x, map_add' := λ m i x y, by { simp, refl }, map_smul' := λ m i c x, by { simp, refl } } in -f'.mk_continuous (∥f∥) (λm, linear_map.mk_continuous_norm_le _ +f'.mk_continuous (‖f‖) (λm, linear_map.mk_continuous_norm_le _ (mul_nonneg (norm_nonneg _) (prod_nonneg (λj hj, norm_nonneg _))) _) @[simp] lemma continuous_multilinear_map.curry_right_apply @@ -1235,12 +1235,12 @@ variables {n 𝕜 G Ei G'} (continuous_multilinear_curry_right_equiv' 𝕜 n G G').symm f v x = f (snoc v x) := rfl @[simp] lemma continuous_multilinear_map.curry_right_norm - (f : continuous_multilinear_map 𝕜 Ei G) : ∥f.curry_right∥ = ∥f∥ := + (f : continuous_multilinear_map 𝕜 Ei G) : ‖f.curry_right‖ = ‖f‖ := (continuous_multilinear_curry_right_equiv 𝕜 Ei G).symm.norm_map f @[simp] lemma continuous_multilinear_map.uncurry_right_norm (f : continuous_multilinear_map 𝕜 (λ i : fin n, Ei i.cast_succ) (Ei (last n) →L[𝕜] G)) : - ∥f.uncurry_right∥ = ∥f∥ := + ‖f.uncurry_right‖ = ‖f‖ := (continuous_multilinear_curry_right_equiv 𝕜 Ei G).norm_map f /-! @@ -1290,7 +1290,7 @@ variables (𝕜 G) (continuous_multilinear_map.curry0 𝕜 G x).uncurry0 = x := rfl @[simp] lemma continuous_multilinear_map.curry0_norm (x : G') : - ∥continuous_multilinear_map.curry0 𝕜 G x∥ = ∥x∥ := + ‖continuous_multilinear_map.curry0 𝕜 G x‖ = ‖x‖ := begin apply le_antisymm, { exact continuous_multilinear_map.op_norm_le_bound _ (norm_nonneg _) (λm, by simp) }, @@ -1299,17 +1299,17 @@ end variables {𝕜 G} @[simp] lemma continuous_multilinear_map.fin0_apply_norm (f : G [×0]→L[𝕜] G') {x : fin 0 → G} : - ∥f x∥ = ∥f∥ := + ‖f x‖ = ‖f‖ := begin obtain rfl : x = 0 := subsingleton.elim _ _, refine le_antisymm (by simpa using f.le_op_norm 0) _, - have : ∥continuous_multilinear_map.curry0 𝕜 G (f.uncurry0)∥ ≤ ∥f.uncurry0∥ := + have : ‖continuous_multilinear_map.curry0 𝕜 G (f.uncurry0)‖ ≤ ‖f.uncurry0‖ := continuous_multilinear_map.op_norm_le_bound _ (norm_nonneg _) (λm, by simp [-continuous_multilinear_map.apply_zero_curry0]), simpa end -lemma continuous_multilinear_map.uncurry0_norm (f : G [×0]→L[𝕜] G') : ∥f.uncurry0∥ = ∥f∥ := +lemma continuous_multilinear_map.uncurry0_norm (f : G [×0]→L[𝕜] G') : ‖f.uncurry0‖ = ‖f‖ := by simp variables (𝕜 G G') @@ -1366,9 +1366,9 @@ def dom_dom_congr (σ : ι ≃ ι') : continuous_multilinear_map 𝕜 (λ _ : ι, G) G' ≃ₗᵢ[𝕜] continuous_multilinear_map 𝕜 (λ _ : ι', G) G' := linear_isometry_equiv.of_bounds - { to_fun := λ f, (multilinear_map.dom_dom_congr σ f.to_multilinear_map).mk_continuous ∥f∥ $ + { to_fun := λ f, (multilinear_map.dom_dom_congr σ f.to_multilinear_map).mk_continuous ‖f‖ $ λ m, (f.le_op_norm (λ i, m (σ i))).trans_eq $ by rw [← σ.prod_comp], - inv_fun := λ f, (multilinear_map.dom_dom_congr σ.symm f.to_multilinear_map).mk_continuous ∥f∥ $ + inv_fun := λ f, (multilinear_map.dom_dom_congr σ.symm f.to_multilinear_map).mk_continuous ‖f‖ $ λ m, (f.le_op_norm (λ i, m (σ.symm i))).trans_eq $ by rw [← σ.symm.prod_comp], left_inv := λ f, ext $ λ m, congr_arg f $ by simp only [σ.symm_apply_apply], right_inv := λ f, ext $ λ m, congr_arg f $ by simp only [σ.apply_symm_apply], @@ -1388,7 +1388,7 @@ map with variables indexed by `ι` taking values in the space of continuous mult variables indexed by `ι'`. -/ def curry_sum (f : continuous_multilinear_map 𝕜 (λ x : ι ⊕ ι', G) G') : continuous_multilinear_map 𝕜 (λ x : ι, G) (continuous_multilinear_map 𝕜 (λ x : ι', G) G') := -multilinear_map.mk_continuous_multilinear (multilinear_map.curry_sum f.to_multilinear_map) (∥f∥) $ +multilinear_map.mk_continuous_multilinear (multilinear_map.curry_sum f.to_multilinear_map) (‖f‖) $ λ m m', by simpa [fintype.prod_sum_type, mul_assoc] using f.le_op_norm (sum.elim m m') @[simp] lemma curry_sum_apply (f : continuous_multilinear_map 𝕜 (λ x : ι ⊕ ι', G) G') @@ -1403,7 +1403,7 @@ def uncurry_sum (f : continuous_multilinear_map 𝕜 (λ x : ι, G) (continuous_multilinear_map 𝕜 (λ x : ι', G) G')) : continuous_multilinear_map 𝕜 (λ x : ι ⊕ ι', G) G' := multilinear_map.mk_continuous - (to_multilinear_map_linear.comp_multilinear_map f.to_multilinear_map).uncurry_sum (∥f∥) $ λ m, + (to_multilinear_map_linear.comp_multilinear_map f.to_multilinear_map).uncurry_sum (‖f‖) $ λ m, by simpa [fintype.prod_sum_type, mul_assoc] using (f (m ∘ sum.inl)).le_of_op_norm_le (m ∘ sum.inr) (f.le_op_norm _) diff --git a/src/analysis/normed_space/operator_norm.lean b/src/analysis/normed_space/operator_norm.lean index 20c05fd89e9aa..cdc5fddc547ac 100644 --- a/src/analysis/normed_space/operator_norm.lean +++ b/src/analysis/normed_space/operator_norm.lean @@ -14,7 +14,7 @@ import topology.algebra.module.strong_topology Define the operator norm on the space of continuous (semi)linear maps between normed spaces, and prove its basic properties. In particular, show that this space is itself a normed space. -Since a lot of elementary properties don't require `∥x∥ = 0 → x = 0` we start setting up the +Since a lot of elementary properties don't require `‖x‖ = 0 → x = 0` we start setting up the theory for `seminormed_add_comm_group` and we specialize to `normed_add_comm_group` at the end. Note that most of statements that apply to semilinear maps only hold when the ring homomorphism @@ -37,7 +37,7 @@ open metric continuous_linear_map section normed_field /-! Most statements in this file require the field to be non-discrete, -as this is necessary to deduce an inequality `∥f x∥ ≤ C ∥x∥` from the continuity of f. +as this is necessary to deduce an inequality `‖f x‖ ≤ C ‖x‖` from the continuity of f. However, the other direction always holds. In this section, we just assume that `𝕜` is a normed field. In the remainder of the file, it will be non-discrete. -/ @@ -48,45 +48,45 @@ variables [normed_space 𝕜 G] {σ : 𝕜 →+* 𝕜₂} (f : E →ₛₗ[σ] F /-- Construct a continuous linear map from a linear map and a bound on this linear map. The fact that the norm of the continuous linear map is then controlled is given in `linear_map.mk_continuous_norm_le`. -/ -def linear_map.mk_continuous (C : ℝ) (h : ∀x, ∥f x∥ ≤ C * ∥x∥) : E →SL[σ] F := +def linear_map.mk_continuous (C : ℝ) (h : ∀x, ‖f x‖ ≤ C * ‖x‖) : E →SL[σ] F := ⟨f, add_monoid_hom_class.continuous_of_bound f C h⟩ /-- Reinterpret a linear map `𝕜 →ₗ[𝕜] E` as a continuous linear map. This construction is generalized to the case of any finite dimensional domain in `linear_map.to_continuous_linear_map`. -/ def linear_map.to_continuous_linear_map₁ (f : 𝕜 →ₗ[𝕜] E) : 𝕜 →L[𝕜] E := -f.mk_continuous (∥f 1∥) $ λ x, le_of_eq $ +f.mk_continuous (‖f 1‖) $ λ x, le_of_eq $ by { conv_lhs { rw ← mul_one x }, rw [← smul_eq_mul, f.map_smul, norm_smul, mul_comm] } /-- Construct a continuous linear map from a linear map and the existence of a bound on this linear map. If you have an explicit bound, use `linear_map.mk_continuous` instead, as a norm estimate will follow automatically in `linear_map.mk_continuous_norm_le`. -/ -def linear_map.mk_continuous_of_exists_bound (h : ∃C, ∀x, ∥f x∥ ≤ C * ∥x∥) : E →SL[σ] F := +def linear_map.mk_continuous_of_exists_bound (h : ∃C, ∀x, ‖f x‖ ≤ C * ‖x‖) : E →SL[σ] F := ⟨f, let ⟨C, hC⟩ := h in add_monoid_hom_class.continuous_of_bound f C hC⟩ lemma continuous_of_linear_of_boundₛₗ {f : E → F} (h_add : ∀ x y, f (x + y) = f x + f y) - (h_smul : ∀ (c : 𝕜) x, f (c • x) = (σ c) • f x) {C : ℝ} (h_bound : ∀ x, ∥f x∥ ≤ C*∥x∥) : + (h_smul : ∀ (c : 𝕜) x, f (c • x) = (σ c) • f x) {C : ℝ} (h_bound : ∀ x, ‖f x‖ ≤ C*‖x‖) : continuous f := let φ : E →ₛₗ[σ] F := { to_fun := f, map_add' := h_add, map_smul' := h_smul } in add_monoid_hom_class.continuous_of_bound φ C h_bound lemma continuous_of_linear_of_bound {f : E → G} (h_add : ∀ x y, f (x + y) = f x + f y) - (h_smul : ∀ (c : 𝕜) x, f (c • x) = c • f x) {C : ℝ} (h_bound : ∀ x, ∥f x∥ ≤ C*∥x∥) : + (h_smul : ∀ (c : 𝕜) x, f (c • x) = c • f x) {C : ℝ} (h_bound : ∀ x, ‖f x‖ ≤ C*‖x‖) : continuous f := let φ : E →ₗ[𝕜] G := { to_fun := f, map_add' := h_add, map_smul' := h_smul } in add_monoid_hom_class.continuous_of_bound φ C h_bound -@[simp, norm_cast] lemma linear_map.mk_continuous_coe (C : ℝ) (h : ∀x, ∥f x∥ ≤ C * ∥x∥) : +@[simp, norm_cast] lemma linear_map.mk_continuous_coe (C : ℝ) (h : ∀x, ‖f x‖ ≤ C * ‖x‖) : ((f.mk_continuous C h) : E →ₛₗ[σ] F) = f := rfl -@[simp] lemma linear_map.mk_continuous_apply (C : ℝ) (h : ∀x, ∥f x∥ ≤ C * ∥x∥) (x : E) : +@[simp] lemma linear_map.mk_continuous_apply (C : ℝ) (h : ∀x, ‖f x‖ ≤ C * ‖x‖) (x : E) : f.mk_continuous C h x = f x := rfl @[simp, norm_cast] lemma linear_map.mk_continuous_of_exists_bound_coe - (h : ∃C, ∀x, ∥f x∥ ≤ C * ∥x∥) : + (h : ∃C, ∀x, ‖f x‖ ≤ C * ‖x‖) : ((f.mk_continuous_of_exists_bound h) : E →ₛₗ[σ] F) = f := rfl -@[simp] lemma linear_map.mk_continuous_of_exists_bound_apply (h : ∃C, ∀x, ∥f x∥ ≤ C * ∥x∥) (x : E) : +@[simp] lemma linear_map.mk_continuous_of_exists_bound_apply (h : ∃C, ∀x, ‖f x‖ ≤ C * ‖x‖) (x : E) : f.mk_continuous_of_exists_bound h x = f x := rfl @[simp] lemma linear_map.to_continuous_linear_map₁_coe (f : 𝕜 →ₗ[𝕜] E) : @@ -105,9 +105,9 @@ variables [nontrivially_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] {σ₁₂ : 𝕜 →+* 𝕜₂} {σ₂₃ : 𝕜₂ →+* 𝕜₃} {σ₁₃ : 𝕜 →+* 𝕜₃} [ring_hom_comp_triple σ₁₂ σ₂₃ σ₁₃] -/-- If `∥x∥ = 0` and `f` is continuous then `∥f x∥ = 0`. -/ +/-- If `‖x‖ = 0` and `f` is continuous then `‖f x‖ = 0`. -/ lemma norm_image_of_norm_zero [semilinear_map_class 𝓕 σ₁₂ E F] (f : 𝓕) - (hf : continuous f) {x : E} (hx : ∥x∥ = 0) : ∥f x∥ = 0 := + (hf : continuous f) {x : E} (hx : ‖x‖ = 0) : ‖f x‖ = 0 := begin refine le_antisymm (le_of_forall_pos_le_add (λ ε hε, _)) (norm_nonneg (f x)), rcases normed_add_comm_group.tendsto_nhds_nhds.1 (hf.tendsto 0) ε hε with ⟨δ, δ_pos, hδ⟩, @@ -123,9 +123,9 @@ section variables [ring_hom_isometric σ₁₂] [ring_hom_isometric σ₂₃] lemma semilinear_map_class.bound_of_shell_semi_normed [semilinear_map_class 𝓕 σ₁₂ E F] - (f : 𝓕) {ε C : ℝ} (ε_pos : 0 < ε) {c : 𝕜} (hc : 1 < ∥c∥) - (hf : ∀ x, ε / ∥c∥ ≤ ∥x∥ → ∥x∥ < ε → ∥f x∥ ≤ C * ∥x∥) {x : E} (hx : ∥x∥ ≠ 0) : - ∥f x∥ ≤ C * ∥x∥ := + (f : 𝓕) {ε C : ℝ} (ε_pos : 0 < ε) {c : 𝕜} (hc : 1 < ‖c‖) + (hf : ∀ x, ε / ‖c‖ ≤ ‖x‖ → ‖x‖ < ε → ‖f x‖ ≤ C * ‖x‖) {x : E} (hx : ‖x‖ ≠ 0) : + ‖f x‖ ≤ C * ‖x‖ := begin rcases rescale_to_shell_semi_normed hc ε_pos hx with ⟨δ, hδ, δxle, leδx, δinv⟩, have := hf (δ • x) leδx δxle, @@ -138,14 +138,14 @@ normed. The continuity ensures boundedness on a ball of some radius `ε`. The no norm is then used to rescale any element into an element of norm in `[ε/C, ε]`, whose image has a controlled norm. The norm control for the original element follows by rescaling. -/ lemma semilinear_map_class.bound_of_continuous [semilinear_map_class 𝓕 σ₁₂ E F] (f : 𝓕) - (hf : continuous f) : ∃ C, 0 < C ∧ (∀ x : E, ∥f x∥ ≤ C * ∥x∥) := + (hf : continuous f) : ∃ C, 0 < C ∧ (∀ x : E, ‖f x‖ ≤ C * ‖x‖) := begin rcases normed_add_comm_group.tendsto_nhds_nhds.1 (hf.tendsto 0) 1 zero_lt_one with ⟨ε, ε_pos, hε⟩, simp only [sub_zero, map_zero] at hε, rcases normed_field.exists_one_lt_norm 𝕜 with ⟨c, hc⟩, - have : 0 < ∥c∥ / ε, from div_pos (zero_lt_one.trans hc) ε_pos, - refine ⟨∥c∥ / ε, this, λ x, _⟩, - by_cases hx : ∥x∥ = 0, + have : 0 < ‖c‖ / ε, from div_pos (zero_lt_one.trans hc) ε_pos, + refine ⟨‖c‖ / ε, this, λ x, _⟩, + by_cases hx : ‖x‖ = 0, { rw [hx, mul_zero], exact le_of_eq (norm_image_of_norm_zero f hf hx) }, refine semilinear_map_class.bound_of_shell_semi_normed f ε_pos hc (λ x hle hlt, _) hx, @@ -158,7 +158,7 @@ end namespace continuous_linear_map theorem bound [ring_hom_isometric σ₁₂] (f : E →SL[σ₁₂] F) : - ∃ C, 0 < C ∧ (∀ x : E, ∥f x∥ ≤ C * ∥x∥) := + ∃ C, 0 < C ∧ (∀ x : E, ‖f x‖ ≤ C * ‖x‖) := semilinear_map_class.bound_of_continuous f f.2 section @@ -169,19 +169,19 @@ open filter the corresponding theorem about isometries plus a theorem about scalar multiplication. Likewise for the other theorems about homotheties in this file. -/ -def of_homothety (f : E →ₛₗ[σ₁₂] F) (a : ℝ) (hf : ∀x, ∥f x∥ = a * ∥x∥) : E →SL[σ₁₂] F := +def of_homothety (f : E →ₛₗ[σ₁₂] F) (a : ℝ) (hf : ∀x, ‖f x‖ = a * ‖x‖) : E →SL[σ₁₂] F := f.mk_continuous a (λ x, le_of_eq (hf x)) variable (𝕜) lemma to_span_singleton_homothety (x : E) (c : 𝕜) : - ∥linear_map.to_span_singleton 𝕜 E x c∥ = ∥x∥ * ∥c∥ := + ‖linear_map.to_span_singleton 𝕜 E x c‖ = ‖x‖ * ‖c‖ := by {rw mul_comm, exact norm_smul _ _} /-- Given an element `x` of a normed space `E` over a field `𝕜`, the natural continuous linear map from `𝕜` to `E` by taking multiples of `x`.-/ def to_span_singleton (x : E) : 𝕜 →L[𝕜] E := -of_homothety (linear_map.to_span_singleton 𝕜 E x) ∥x∥ (to_span_singleton_homothety 𝕜 x) +of_homothety (linear_map.to_span_singleton 𝕜 E x) ‖x‖ (to_span_singleton_homothety 𝕜 x) lemma to_span_singleton_apply (x : E) (r : 𝕜) : to_span_singleton 𝕜 x r = r • x := by simp [to_span_singleton, of_homothety, linear_map.to_span_singleton] @@ -202,16 +202,16 @@ to_span_singleton_smul' 𝕜 𝕜 c x variables (𝕜 E) /-- Given a unit-length element `x` of a normed space `E` over a field `𝕜`, the natural linear isometry map from `𝕜` to `E` by taking multiples of `x`.-/ -def _root_.linear_isometry.to_span_singleton {v : E} (hv : ∥v∥ = 1) : 𝕜 →ₗᵢ[𝕜] E := +def _root_.linear_isometry.to_span_singleton {v : E} (hv : ‖v‖ = 1) : 𝕜 →ₗᵢ[𝕜] E := { norm_map' := λ x, by simp [norm_smul, hv], .. linear_map.to_span_singleton 𝕜 E v } variables {𝕜 E} -@[simp] lemma _root_.linear_isometry.to_span_singleton_apply {v : E} (hv : ∥v∥ = 1) (a : 𝕜) : +@[simp] lemma _root_.linear_isometry.to_span_singleton_apply {v : E} (hv : ‖v‖ = 1) (a : 𝕜) : linear_isometry.to_span_singleton 𝕜 E hv a = a • v := rfl -@[simp] lemma _root_.linear_isometry.coe_to_span_singleton {v : E} (hv : ∥v∥ = 1) : +@[simp] lemma _root_.linear_isometry.coe_to_span_singleton {v : E} (hv : ‖v‖ = 1) : (linear_isometry.to_span_singleton 𝕜 E hv).to_linear_map = linear_map.to_span_singleton 𝕜 E v := rfl @@ -221,52 +221,52 @@ section op_norm open set real /-- The operator norm of a continuous linear map is the inf of all its bounds. -/ -def op_norm (f : E →SL[σ₁₂] F) := Inf {c | 0 ≤ c ∧ ∀ x, ∥f x∥ ≤ c * ∥x∥} +def op_norm (f : E →SL[σ₁₂] F) := Inf {c | 0 ≤ c ∧ ∀ x, ‖f x‖ ≤ c * ‖x‖} instance has_op_norm : has_norm (E →SL[σ₁₂] F) := ⟨op_norm⟩ -lemma norm_def (f : E →SL[σ₁₂] F) : ∥f∥ = Inf {c | 0 ≤ c ∧ ∀ x, ∥f x∥ ≤ c * ∥x∥} := rfl +lemma norm_def (f : E →SL[σ₁₂] F) : ‖f‖ = Inf {c | 0 ≤ c ∧ ∀ x, ‖f x‖ ≤ c * ‖x‖} := rfl -- So that invocations of `le_cInf` make sense: we show that the set of -- bounds is nonempty and bounded below. lemma bounds_nonempty [ring_hom_isometric σ₁₂] {f : E →SL[σ₁₂] F} : - ∃ c, c ∈ { c | 0 ≤ c ∧ ∀ x, ∥f x∥ ≤ c * ∥x∥ } := + ∃ c, c ∈ { c | 0 ≤ c ∧ ∀ x, ‖f x‖ ≤ c * ‖x‖ } := let ⟨M, hMp, hMb⟩ := f.bound in ⟨M, le_of_lt hMp, hMb⟩ lemma bounds_bdd_below {f : E →SL[σ₁₂] F} : - bdd_below { c | 0 ≤ c ∧ ∀ x, ∥f x∥ ≤ c * ∥x∥ } := + bdd_below { c | 0 ≤ c ∧ ∀ x, ‖f x‖ ≤ c * ‖x‖ } := ⟨0, λ _ ⟨hn, _⟩, hn⟩ /-- If one controls the norm of every `A x`, then one controls the norm of `A`. -/ -lemma op_norm_le_bound (f : E →SL[σ₁₂] F) {M : ℝ} (hMp: 0 ≤ M) (hM : ∀ x, ∥f x∥ ≤ M * ∥x∥) : - ∥f∥ ≤ M := +lemma op_norm_le_bound (f : E →SL[σ₁₂] F) {M : ℝ} (hMp: 0 ≤ M) (hM : ∀ x, ‖f x‖ ≤ M * ‖x‖) : + ‖f‖ ≤ M := cInf_le bounds_bdd_below ⟨hMp, hM⟩ -/-- If one controls the norm of every `A x`, `∥x∥ ≠ 0`, then one controls the norm of `A`. -/ +/-- If one controls the norm of every `A x`, `‖x‖ ≠ 0`, then one controls the norm of `A`. -/ lemma op_norm_le_bound' (f : E →SL[σ₁₂] F) {M : ℝ} (hMp: 0 ≤ M) - (hM : ∀ x, ∥x∥ ≠ 0 → ∥f x∥ ≤ M * ∥x∥) : - ∥f∥ ≤ M := -op_norm_le_bound f hMp $ λ x, (ne_or_eq (∥x∥) 0).elim (hM x) $ + (hM : ∀ x, ‖x‖ ≠ 0 → ‖f x‖ ≤ M * ‖x‖) : + ‖f‖ ≤ M := +op_norm_le_bound f hMp $ λ x, (ne_or_eq (‖x‖) 0).elim (hM x) $ λ h, by simp only [h, mul_zero, norm_image_of_norm_zero f f.2 h] theorem op_norm_le_of_lipschitz {f : E →SL[σ₁₂] F} {K : ℝ≥0} (hf : lipschitz_with K f) : - ∥f∥ ≤ K := + ‖f‖ ≤ K := f.op_norm_le_bound K.2 $ λ x, by simpa only [dist_zero_right, f.map_zero] using hf.dist_le_mul x 0 lemma op_norm_eq_of_bounds {φ : E →SL[σ₁₂] F} {M : ℝ} (M_nonneg : 0 ≤ M) - (h_above : ∀ x, ∥φ x∥ ≤ M*∥x∥) (h_below : ∀ N ≥ 0, (∀ x, ∥φ x∥ ≤ N*∥x∥) → M ≤ N) : - ∥φ∥ = M := + (h_above : ∀ x, ‖φ x‖ ≤ M*‖x‖) (h_below : ∀ N ≥ 0, (∀ x, ‖φ x‖ ≤ N*‖x‖) → M ≤ N) : + ‖φ‖ = M := le_antisymm (φ.op_norm_le_bound M_nonneg h_above) ((le_cInf_iff continuous_linear_map.bounds_bdd_below ⟨M, M_nonneg, h_above⟩).mpr $ λ N ⟨N_nonneg, hN⟩, h_below N N_nonneg hN) -lemma op_norm_neg (f : E →SL[σ₁₂] F) : ∥-f∥ = ∥f∥ := by simp only [norm_def, neg_apply, norm_neg] +lemma op_norm_neg (f : E →SL[σ₁₂] F) : ‖-f‖ = ‖f‖ := by simp only [norm_def, neg_apply, norm_neg] -theorem antilipschitz_of_bound (f : E →SL[σ₁₂] F) {K : ℝ≥0} (h : ∀ x, ∥x∥ ≤ K * ∥f x∥) : +theorem antilipschitz_of_bound (f : E →SL[σ₁₂] F) {K : ℝ≥0} (h : ∀ x, ‖x‖ ≤ K * ‖f x‖) : antilipschitz_with K f := add_monoid_hom_class.antilipschitz_of_bound _ h lemma bound_of_antilipschitz (f : E →SL[σ₁₂] F) {K : ℝ≥0} (h : antilipschitz_with K f) (x) : - ∥x∥ ≤ K * ∥f x∥ := + ‖x‖ ≤ K * ‖f x‖ := add_monoid_hom_class.bound_of_antilipschitz _ h x section @@ -274,44 +274,44 @@ section variables [ring_hom_isometric σ₁₂] [ring_hom_isometric σ₂₃] (f g : E →SL[σ₁₂] F) (h : F →SL[σ₂₃] G) (x : E) -lemma op_norm_nonneg : 0 ≤ ∥f∥ := +lemma op_norm_nonneg : 0 ≤ ‖f‖ := le_cInf bounds_nonempty (λ _ ⟨hx, _⟩, hx) -/-- The fundamental property of the operator norm: `∥f x∥ ≤ ∥f∥ * ∥x∥`. -/ -theorem le_op_norm : ∥f x∥ ≤ ∥f∥ * ∥x∥ := +/-- The fundamental property of the operator norm: `‖f x‖ ≤ ‖f‖ * ‖x‖`. -/ +theorem le_op_norm : ‖f x‖ ≤ ‖f‖ * ‖x‖ := begin obtain ⟨C, Cpos, hC⟩ := f.bound, replace hC := hC x, - by_cases h : ∥x∥ = 0, + by_cases h : ‖x‖ = 0, { rwa [h, mul_zero] at ⊢ hC }, - have hlt : 0 < ∥x∥ := lt_of_le_of_ne (norm_nonneg x) (ne.symm h), + have hlt : 0 < ‖x‖ := lt_of_le_of_ne (norm_nonneg x) (ne.symm h), exact (div_le_iff hlt).mp (le_cInf bounds_nonempty (λ c ⟨_, hc⟩, (div_le_iff hlt).mpr $ by { apply hc })), end -theorem dist_le_op_norm (x y : E) : dist (f x) (f y) ≤ ∥f∥ * dist x y := +theorem dist_le_op_norm (x y : E) : dist (f x) (f y) ≤ ‖f‖ * dist x y := by simp_rw [dist_eq_norm, ← map_sub, f.le_op_norm] -theorem le_op_norm_of_le {c : ℝ} {x} (h : ∥x∥ ≤ c) : ∥f x∥ ≤ ∥f∥ * c := +theorem le_op_norm_of_le {c : ℝ} {x} (h : ‖x‖ ≤ c) : ‖f x‖ ≤ ‖f‖ * c := le_trans (f.le_op_norm x) (mul_le_mul_of_nonneg_left h f.op_norm_nonneg) -theorem le_of_op_norm_le {c : ℝ} (h : ∥f∥ ≤ c) (x : E) : ∥f x∥ ≤ c * ∥x∥ := +theorem le_of_op_norm_le {c : ℝ} (h : ‖f‖ ≤ c) (x : E) : ‖f x‖ ≤ c * ‖x‖ := (f.le_op_norm x).trans (mul_le_mul_of_nonneg_right h (norm_nonneg x)) -lemma ratio_le_op_norm : ∥f x∥ / ∥x∥ ≤ ∥f∥ := +lemma ratio_le_op_norm : ‖f x‖ / ‖x‖ ≤ ‖f‖ := div_le_of_nonneg_of_le_mul (norm_nonneg _) f.op_norm_nonneg (le_op_norm _ _) /-- The image of the unit ball under a continuous linear map is bounded. -/ -lemma unit_le_op_norm : ∥x∥ ≤ 1 → ∥f x∥ ≤ ∥f∥ := -mul_one ∥f∥ ▸ f.le_op_norm_of_le +lemma unit_le_op_norm : ‖x‖ ≤ 1 → ‖f x‖ ≤ ‖f‖ := +mul_one ‖f‖ ▸ f.le_op_norm_of_le lemma op_norm_le_of_shell {f : E →SL[σ₁₂] F} {ε C : ℝ} (ε_pos : 0 < ε) (hC : 0 ≤ C) - {c : 𝕜} (hc : 1 < ∥c∥) (hf : ∀ x, ε / ∥c∥ ≤ ∥x∥ → ∥x∥ < ε → ∥f x∥ ≤ C * ∥x∥) : - ∥f∥ ≤ C := + {c : 𝕜} (hc : 1 < ‖c‖) (hf : ∀ x, ε / ‖c‖ ≤ ‖x‖ → ‖x‖ < ε → ‖f x‖ ≤ C * ‖x‖) : + ‖f‖ ≤ C := f.op_norm_le_bound' hC $ λ x hx, semilinear_map_class.bound_of_shell_semi_normed f ε_pos hc hf hx lemma op_norm_le_of_ball {f : E →SL[σ₁₂] F} {ε : ℝ} {C : ℝ} (ε_pos : 0 < ε) (hC : 0 ≤ C) - (hf : ∀ x ∈ ball (0 : E) ε, ∥f x∥ ≤ C * ∥x∥) : ∥f∥ ≤ C := + (hf : ∀ x ∈ ball (0 : E) ε, ‖f x‖ ≤ C * ‖x‖) : ‖f‖ ≤ C := begin rcases normed_field.exists_one_lt_norm 𝕜 with ⟨c, hc⟩, refine op_norm_le_of_shell ε_pos hC hc (λ x _ hx, hf x _), @@ -319,12 +319,12 @@ begin end lemma op_norm_le_of_nhds_zero {f : E →SL[σ₁₂] F} {C : ℝ} (hC : 0 ≤ C) - (hf : ∀ᶠ x in 𝓝 (0 : E), ∥f x∥ ≤ C * ∥x∥) : ∥f∥ ≤ C := + (hf : ∀ᶠ x in 𝓝 (0 : E), ‖f x‖ ≤ C * ‖x‖) : ‖f‖ ≤ C := let ⟨ε, ε0, hε⟩ := metric.eventually_nhds_iff_ball.1 hf in op_norm_le_of_ball ε0 hC hε lemma op_norm_le_of_shell' {f : E →SL[σ₁₂] F} {ε C : ℝ} (ε_pos : 0 < ε) (hC : 0 ≤ C) - {c : 𝕜} (hc : ∥c∥ < 1) (hf : ∀ x, ε * ∥c∥ ≤ ∥x∥ → ∥x∥ < ε → ∥f x∥ ≤ C * ∥x∥) : - ∥f∥ ≤ C := + {c : 𝕜} (hc : ‖c‖ < 1) (hf : ∀ x, ε * ‖c‖ ≤ ‖x‖ → ‖x‖ < ε → ‖f x‖ ≤ C * ‖x‖) : + ‖f‖ ≤ C := begin by_cases h0 : c = 0, { refine op_norm_le_of_ball ε_pos hC (λ x hx, hf x _ _), @@ -336,43 +336,43 @@ begin rwa [norm_inv, div_eq_mul_inv, inv_inv] } end -/-- For a continuous real linear map `f`, if one controls the norm of every `f x`, `∥x∥ = 1`, then +/-- For a continuous real linear map `f`, if one controls the norm of every `f x`, `‖x‖ = 1`, then one controls the norm of `f`. -/ lemma op_norm_le_of_unit_norm [normed_space ℝ E] [normed_space ℝ F] {f : E →L[ℝ] F} {C : ℝ} - (hC : 0 ≤ C) (hf : ∀ x, ∥x∥ = 1 → ∥f x∥ ≤ C) : ∥f∥ ≤ C := + (hC : 0 ≤ C) (hf : ∀ x, ‖x‖ = 1 → ‖f x‖ ≤ C) : ‖f‖ ≤ C := begin refine op_norm_le_bound' f hC (λ x hx, _), - have H₁ : ∥(∥x∥⁻¹ • x)∥ = 1, by rw [norm_smul, norm_inv, norm_norm, inv_mul_cancel hx], + have H₁ : ‖(‖x‖⁻¹ • x)‖ = 1, by rw [norm_smul, norm_inv, norm_norm, inv_mul_cancel hx], have H₂ := hf _ H₁, rwa [map_smul, norm_smul, norm_inv, norm_norm, ← div_eq_inv_mul, div_le_iff] at H₂, exact (norm_nonneg x).lt_of_ne' hx end /-- The operator norm satisfies the triangle inequality. -/ -theorem op_norm_add_le : ∥f + g∥ ≤ ∥f∥ + ∥g∥ := +theorem op_norm_add_le : ‖f + g‖ ≤ ‖f‖ + ‖g‖ := (f + g).op_norm_le_bound (add_nonneg f.op_norm_nonneg g.op_norm_nonneg) $ λ x, (norm_add_le_of_le (f.le_op_norm x) (g.le_op_norm x)).trans_eq (add_mul _ _ _).symm /-- The norm of the `0` operator is `0`. -/ -theorem op_norm_zero : ∥(0 : E →SL[σ₁₂] F)∥ = 0 := +theorem op_norm_zero : ‖(0 : E →SL[σ₁₂] F)‖ = 0 := le_antisymm (cInf_le bounds_bdd_below ⟨le_rfl, λ _, le_of_eq (by { rw [zero_mul], exact norm_zero })⟩) (op_norm_nonneg _) /-- 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 := +lemma norm_id_le : ‖id 𝕜 E‖ ≤ 1 := op_norm_le_bound _ zero_le_one (λx, by simp) /-- If there is an element with norm different from `0`, then the norm of the identity equals `1`. (Since we are working with seminorms supposing that the space is non-trivial is not enough.) -/ -lemma norm_id_of_nontrivial_seminorm (h : ∃ (x : E), ∥x∥ ≠ 0) : ∥id 𝕜 E∥ = 1 := +lemma norm_id_of_nontrivial_seminorm (h : ∃ (x : E), ‖x‖ ≠ 0) : ‖id 𝕜 E‖ = 1 := le_antisymm norm_id_le $ let ⟨x, hx⟩ := h in have _ := (id 𝕜 E).ratio_le_op_norm x, by rwa [id_apply, div_self hx] at this lemma op_norm_smul_le {𝕜' : Type*} [normed_field 𝕜'] [normed_space 𝕜' F] - [smul_comm_class 𝕜₂ 𝕜' F] (c : 𝕜') (f : E →SL[σ₁₂] F) : ∥c • f∥ ≤ ∥c∥ * ∥f∥ := + [smul_comm_class 𝕜₂ 𝕜' F] (c : 𝕜') (f : E →SL[σ₁₂] F) : ‖c • f‖ ≤ ‖c‖ * ‖f‖ := ((c • f).op_norm_le_bound (mul_nonneg (norm_nonneg _) (op_norm_nonneg _)) (λ _, begin @@ -424,8 +424,8 @@ protected lemma tmp_closed_ball_div_subset {a b : ℝ} (ha : 0 < a) (hb : 0 < b) begin intros f hf x hx, rw mem_closed_ball_zero_iff at ⊢ hf hx, - calc ∥f x∥ - ≤ ∥f∥ * ∥x∥ : le_op_norm _ _ + calc ‖f x‖ + ≤ ‖f‖ * ‖x‖ : le_op_norm _ _ ... ≤ (a/b) * b : mul_le_mul hf hx (norm_nonneg _) (div_pos ha hb).le ... = a : div_mul_cancel a hb.ne.symm end @@ -440,7 +440,7 @@ begin ((@metric.nhds_basis_closed_ball _ continuous_linear_map.tmp_pseudo_metric_space 0).ext (continuous_linear_map.has_basis_nhds_zero_of_basis metric.nhds_basis_closed_ball) _ _), { rcases normed_field.exists_norm_lt_one 𝕜 with ⟨c, hc₀, hc₁⟩, - refine λ ε hε, ⟨⟨closed_ball 0 (1 / ∥c∥), ε⟩, + refine λ ε hε, ⟨⟨closed_ball 0 (1 / ‖c‖), ε⟩, ⟨normed_space.is_vonN_bounded_closed_ball _ _ _, hε⟩, λ f hf, _⟩, change ∀ x, _ at hf, simp_rw mem_closed_ball_zero_iff at hf, @@ -474,7 +474,7 @@ continuous_linear_map.tmp_pseudo_metric_space.replace_uniformity instance to_seminormed_add_comm_group : seminormed_add_comm_group (E →SL[σ₁₂] F) := { dist_eq := continuous_linear_map.tmp_seminormed_add_comm_group.dist_eq } -lemma nnnorm_def (f : E →SL[σ₁₂] F) : ∥f∥₊ = Inf {c | ∀ x, ∥f x∥₊ ≤ c * ∥x∥₊} := +lemma nnnorm_def (f : E →SL[σ₁₂] F) : ‖f‖₊ = Inf {c | ∀ x, ‖f x‖₊ ≤ c * ‖x‖₊} := begin ext, rw [nnreal.coe_Inf, coe_nnnorm, norm_def, nnreal.coe_image], @@ -483,28 +483,28 @@ begin end /-- If one controls the norm of every `A x`, then one controls the norm of `A`. -/ -lemma op_nnnorm_le_bound (f : E →SL[σ₁₂] F) (M : ℝ≥0) (hM : ∀ x, ∥f x∥₊ ≤ M * ∥x∥₊) : - ∥f∥₊ ≤ M := +lemma op_nnnorm_le_bound (f : E →SL[σ₁₂] F) (M : ℝ≥0) (hM : ∀ x, ‖f x‖₊ ≤ M * ‖x‖₊) : + ‖f‖₊ ≤ M := op_norm_le_bound f (zero_le M) hM -/-- If one controls the norm of every `A x`, `∥x∥₊ ≠ 0`, then one controls the norm of `A`. -/ -lemma op_nnnorm_le_bound' (f : E →SL[σ₁₂] F) (M : ℝ≥0) (hM : ∀ x, ∥x∥₊ ≠ 0 → ∥f x∥₊ ≤ M * ∥x∥₊) : - ∥f∥₊ ≤ M := +/-- If one controls the norm of every `A x`, `‖x‖₊ ≠ 0`, then one controls the norm of `A`. -/ +lemma op_nnnorm_le_bound' (f : E →SL[σ₁₂] F) (M : ℝ≥0) (hM : ∀ x, ‖x‖₊ ≠ 0 → ‖f x‖₊ ≤ M * ‖x‖₊) : + ‖f‖₊ ≤ M := op_norm_le_bound' f (zero_le M) $ λ x hx, hM x $ by rwa [← nnreal.coe_ne_zero] -/-- For a continuous real linear map `f`, if one controls the norm of every `f x`, `∥x∥₊ = 1`, then +/-- For a continuous real linear map `f`, if one controls the norm of every `f x`, `‖x‖₊ = 1`, then one controls the norm of `f`. -/ lemma op_nnnorm_le_of_unit_nnnorm [normed_space ℝ E] [normed_space ℝ F] {f : E →L[ℝ] F} {C : ℝ≥0} - (hf : ∀ x, ∥x∥₊ = 1 → ∥f x∥₊ ≤ C) : ∥f∥₊ ≤ C := + (hf : ∀ x, ‖x‖₊ = 1 → ‖f x‖₊ ≤ C) : ‖f‖₊ ≤ C := op_norm_le_of_unit_norm C.coe_nonneg $ λ x hx, hf x $ by rwa [← nnreal.coe_eq_one] theorem op_nnnorm_le_of_lipschitz {f : E →SL[σ₁₂] F} {K : ℝ≥0} (hf : lipschitz_with K f) : - ∥f∥₊ ≤ K := + ‖f‖₊ ≤ K := op_norm_le_of_lipschitz hf lemma op_nnnorm_eq_of_bounds {φ : E →SL[σ₁₂] F} (M : ℝ≥0) - (h_above : ∀ x, ∥φ x∥ ≤ M*∥x∥) (h_below : ∀ N, (∀ x, ∥φ x∥₊ ≤ N*∥x∥₊) → M ≤ N) : - ∥φ∥₊ = M := + (h_above : ∀ x, ‖φ x‖ ≤ M*‖x‖) (h_below : ∀ N, (∀ x, ‖φ x‖₊ ≤ N*‖x‖₊) → M ≤ N) : + ‖φ‖₊ = M := subtype.ext $ op_norm_eq_of_bounds (zero_le M) h_above $ subtype.forall'.mpr h_below instance to_normed_space {𝕜' : Type*} [normed_field 𝕜'] [normed_space 𝕜' F] @@ -513,12 +513,12 @@ instance to_normed_space {𝕜' : Type*} [normed_field 𝕜'] [normed_space 𝕜 include σ₁₃ /-- The operator norm is submultiplicative. -/ -lemma op_norm_comp_le (f : E →SL[σ₁₂] F) : ∥h.comp f∥ ≤ ∥h∥ * ∥f∥ := +lemma op_norm_comp_le (f : E →SL[σ₁₂] F) : ‖h.comp f‖ ≤ ‖h‖ * ‖f‖ := (cInf_le bounds_bdd_below ⟨mul_nonneg (op_norm_nonneg _) (op_norm_nonneg _), λ x, by { rw mul_assoc, exact h.le_op_norm_of_le (f.le_op_norm x) } ⟩) -lemma op_nnnorm_comp_le [ring_hom_isometric σ₁₃] (f : E →SL[σ₁₂] F) : ∥h.comp f∥₊ ≤ ∥h∥₊ * ∥f∥₊ := +lemma op_nnnorm_comp_le [ring_hom_isometric σ₁₃] (f : E →SL[σ₁₂] F) : ‖h.comp f‖₊ ≤ ‖h‖₊ * ‖f‖₊ := op_norm_comp_le h f omit σ₁₃ @@ -533,17 +533,17 @@ instance to_normed_algebra : normed_algebra 𝕜 (E →L[𝕜] E) := { .. continuous_linear_map.to_normed_space, .. continuous_linear_map.algebra } -theorem le_op_nnnorm : ∥f x∥₊ ≤ ∥f∥₊ * ∥x∥₊ := f.le_op_norm x +theorem le_op_nnnorm : ‖f x‖₊ ≤ ‖f‖₊ * ‖x‖₊ := f.le_op_norm x -theorem nndist_le_op_nnnorm (x y : E) : nndist (f x) (f y) ≤ ∥f∥₊ * nndist x y := +theorem nndist_le_op_nnnorm (x y : E) : nndist (f x) (f y) ≤ ‖f‖₊ * nndist x y := dist_le_op_norm f x y /-- continuous linear maps are Lipschitz continuous. -/ -theorem lipschitz : lipschitz_with ∥f∥₊ f := +theorem lipschitz : lipschitz_with ‖f‖₊ f := add_monoid_hom_class.lipschitz_of_bound_nnnorm f _ f.le_op_nnnorm /-- Evaluation of a continuous linear map `f` at a point is Lipschitz continuous in `f`. -/ -theorem lipschitz_apply (x : E) : lipschitz_with ∥x∥₊ (λ f : E →SL[σ₁₂] F, f x) := +theorem lipschitz_apply (x : E) : lipschitz_with ‖x‖₊ (λ f : E →SL[σ₁₂] F, f x) := lipschitz_with_iff_norm_sub_le.2 $ λ f g, ((f - g).le_op_norm x).trans_eq (mul_comm _ _) end @@ -552,36 +552,36 @@ section Sup variables [ring_hom_isometric σ₁₂] -lemma exists_mul_lt_apply_of_lt_op_nnnorm (f : E →SL[σ₁₂] F) {r : ℝ≥0} (hr : r < ∥f∥₊) : - ∃ x, r * ∥x∥₊ < ∥f x∥₊ := +lemma exists_mul_lt_apply_of_lt_op_nnnorm (f : E →SL[σ₁₂] F) {r : ℝ≥0} (hr : r < ‖f‖₊) : + ∃ x, r * ‖x‖₊ < ‖f x‖₊ := by simpa only [not_forall, not_le, set.mem_set_of] using not_mem_of_lt_cInf - (nnnorm_def f ▸ hr : r < Inf {c : ℝ≥0 | ∀ x, ∥f x∥₊ ≤ c * ∥x∥₊}) (order_bot.bdd_below _) + (nnnorm_def f ▸ hr : r < Inf {c : ℝ≥0 | ∀ x, ‖f x‖₊ ≤ c * ‖x‖₊}) (order_bot.bdd_below _) -lemma exists_mul_lt_of_lt_op_norm (f : E →SL[σ₁₂] F) {r : ℝ} (hr₀ : 0 ≤ r) (hr : r < ∥f∥) : - ∃ x, r * ∥x∥ < ∥f x∥ := +lemma exists_mul_lt_of_lt_op_norm (f : E →SL[σ₁₂] F) {r : ℝ} (hr₀ : 0 ≤ r) (hr : r < ‖f‖) : + ∃ x, r * ‖x‖ < ‖f x‖ := by { lift r to ℝ≥0 using hr₀, exact f.exists_mul_lt_apply_of_lt_op_nnnorm hr } lemma exists_lt_apply_of_lt_op_nnnorm {𝕜 𝕜₂ E F : Type*} [normed_add_comm_group E] [seminormed_add_comm_group F] [densely_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] {σ₁₂ : 𝕜 →+* 𝕜₂} [normed_space 𝕜 E] [normed_space 𝕜₂ F] [ring_hom_isometric σ₁₂] - (f : E →SL[σ₁₂] F) {r : ℝ≥0} (hr : r < ∥f∥₊) : ∃ x : E, ∥x∥₊ < 1 ∧ r < ∥f x∥₊ := + (f : E →SL[σ₁₂] F) {r : ℝ≥0} (hr : r < ‖f‖₊) : ∃ x : E, ‖x‖₊ < 1 ∧ r < ‖f x‖₊ := begin obtain ⟨y, hy⟩ := f.exists_mul_lt_apply_of_lt_op_nnnorm hr, - have hy' : ∥y∥₊ ≠ 0 := nnnorm_ne_zero_iff.2 + have hy' : ‖y‖₊ ≠ 0 := nnnorm_ne_zero_iff.2 (λ heq, by simpa only [heq, nnnorm_zero, map_zero, not_lt_zero'] using hy), - have hfy : ∥f y∥₊ ≠ 0 := (zero_le'.trans_lt hy).ne', - rw [←inv_inv (∥f y∥₊), nnreal.lt_inv_iff_mul_lt (inv_ne_zero hfy), mul_assoc, mul_comm (∥y∥₊), + have hfy : ‖f y‖₊ ≠ 0 := (zero_le'.trans_lt hy).ne', + rw [←inv_inv (‖f y‖₊), nnreal.lt_inv_iff_mul_lt (inv_ne_zero hfy), mul_assoc, mul_comm (‖y‖₊), ←mul_assoc, ←nnreal.lt_inv_iff_mul_lt hy'] at hy, obtain ⟨k, hk₁, hk₂⟩ := normed_field.exists_lt_nnnorm_lt 𝕜 hy, refine ⟨k • y, (nnnorm_smul k y).symm ▸ (nnreal.lt_inv_iff_mul_lt hy').1 hk₂, _⟩, - have : ∥σ₁₂ k∥₊ = ∥k∥₊ := subtype.ext ring_hom_isometric.is_iso, + have : ‖σ₁₂ k‖₊ = ‖k‖₊ := subtype.ext ring_hom_isometric.is_iso, rwa [map_smulₛₗ f, nnnorm_smul, ←nnreal.div_lt_iff hfy, div_eq_mul_inv, this], end lemma exists_lt_apply_of_lt_op_norm {𝕜 𝕜₂ E F : Type*} [normed_add_comm_group E] [seminormed_add_comm_group F] [densely_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] {σ₁₂ : 𝕜 →+* 𝕜₂} [normed_space 𝕜 E] [normed_space 𝕜₂ F] [ring_hom_isometric σ₁₂] - (f : E →SL[σ₁₂] F) {r : ℝ} (hr : r < ∥f∥) : ∃ x : E, ∥x∥ < 1 ∧ r < ∥f x∥ := + (f : E →SL[σ₁₂] F) {r : ℝ} (hr : r < ‖f‖) : ∃ x : E, ‖x‖ < 1 ∧ r < ‖f x‖ := begin by_cases hr₀ : r < 0, { exact ⟨0, by simpa using hr₀⟩, }, @@ -592,7 +592,7 @@ end lemma Sup_unit_ball_eq_nnnorm {𝕜 𝕜₂ E F : Type*} [normed_add_comm_group E] [seminormed_add_comm_group F] [densely_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] {σ₁₂ : 𝕜 →+* 𝕜₂} [normed_space 𝕜 E] [normed_space 𝕜₂ F] [ring_hom_isometric σ₁₂] - (f : E →SL[σ₁₂] F) : Sup ((λ x, ∥f x∥₊) '' ball 0 1) = ∥f∥₊ := + (f : E →SL[σ₁₂] F) : Sup ((λ x, ‖f x‖₊) '' ball 0 1) = ‖f‖₊ := begin refine cSup_eq_of_forall_le_of_forall_lt_exists_gt ((nonempty_ball.mpr zero_lt_one).image _) _ (λ ub hub, _), @@ -605,26 +605,26 @@ end lemma Sup_unit_ball_eq_norm {𝕜 𝕜₂ E F : Type*} [normed_add_comm_group E] [seminormed_add_comm_group F] [densely_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] {σ₁₂ : 𝕜 →+* 𝕜₂} [normed_space 𝕜 E] [normed_space 𝕜₂ F] [ring_hom_isometric σ₁₂] - (f : E →SL[σ₁₂] F) : Sup ((λ x, ∥f x∥) '' ball 0 1) = ∥f∥ := + (f : E →SL[σ₁₂] F) : Sup ((λ x, ‖f x‖) '' ball 0 1) = ‖f‖ := by simpa only [nnreal.coe_Sup, set.image_image] using nnreal.coe_eq.2 f.Sup_unit_ball_eq_nnnorm lemma Sup_closed_unit_ball_eq_nnnorm {𝕜 𝕜₂ E F : Type*} [normed_add_comm_group E] [seminormed_add_comm_group F] [densely_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] {σ₁₂ : 𝕜 →+* 𝕜₂} [normed_space 𝕜 E] [normed_space 𝕜₂ F] [ring_hom_isometric σ₁₂] - (f : E →SL[σ₁₂] F) : Sup ((λ x, ∥f x∥₊) '' closed_ball 0 1) = ∥f∥₊ := + (f : E →SL[σ₁₂] F) : Sup ((λ x, ‖f x‖₊) '' closed_ball 0 1) = ‖f‖₊ := begin - have hbdd : ∀ y ∈ (λ x, ∥f x∥₊) '' closed_ball 0 1, y ≤ ∥f∥₊, + have hbdd : ∀ y ∈ (λ x, ‖f x‖₊) '' closed_ball 0 1, y ≤ ‖f‖₊, { rintro - ⟨x, hx, rfl⟩, exact f.unit_le_op_norm x (mem_closed_ball_zero_iff.1 hx) }, refine le_antisymm (cSup_le ((nonempty_closed_ball.mpr zero_le_one).image _) hbdd) _, rw ←Sup_unit_ball_eq_nnnorm, - exact cSup_le_cSup ⟨∥f∥₊, hbdd⟩ ((nonempty_ball.2 zero_lt_one).image _) + exact cSup_le_cSup ⟨‖f‖₊, hbdd⟩ ((nonempty_ball.2 zero_lt_one).image _) (set.image_subset _ ball_subset_closed_ball), end lemma Sup_closed_unit_ball_eq_norm {𝕜 𝕜₂ E F : Type*} [normed_add_comm_group E] [seminormed_add_comm_group F] [densely_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] {σ₁₂ : 𝕜 →+* 𝕜₂} [normed_space 𝕜 E] [normed_space 𝕜₂ F] [ring_hom_isometric σ₁₂] - (f : E →SL[σ₁₂] F) : Sup ((λ x, ∥f x∥) '' closed_ball 0 1) = ∥f∥ := + (f : E →SL[σ₁₂] F) : Sup ((λ x, ‖f x‖) '' closed_ball 0 1) = ‖f‖ := by simpa only [nnreal.coe_Sup, set.image_image] using nnreal.coe_eq.2 f.Sup_closed_unit_ball_eq_nnnorm @@ -633,25 +633,25 @@ end Sup section lemma op_norm_ext [ring_hom_isometric σ₁₃] (f : E →SL[σ₁₂] F) (g : E →SL[σ₁₃] G) - (h : ∀ x, ∥f x∥ = ∥g x∥) : ∥f∥ = ∥g∥ := + (h : ∀ x, ‖f x‖ = ‖g x‖) : ‖f‖ = ‖g‖ := op_norm_eq_of_bounds (norm_nonneg _) (λ x, by { rw h x, exact le_op_norm _ _ }) (λ c hc h₂, op_norm_le_bound _ hc (λ z, by { rw ←h z, exact h₂ z })) variables [ring_hom_isometric σ₂₃] theorem op_norm_le_bound₂ (f : E →SL[σ₁₃] F →SL[σ₂₃] G) {C : ℝ} (h0 : 0 ≤ C) - (hC : ∀ x y, ∥f x y∥ ≤ C * ∥x∥ * ∥y∥) : - ∥f∥ ≤ C := + (hC : ∀ x y, ‖f x y‖ ≤ C * ‖x‖ * ‖y‖) : + ‖f‖ ≤ C := f.op_norm_le_bound h0 $ λ x, (f x).op_norm_le_bound (mul_nonneg h0 (norm_nonneg _)) $ hC x theorem le_op_norm₂ [ring_hom_isometric σ₁₃] (f : E →SL[σ₁₃] F →SL[σ₂₃] G) (x : E) (y : F) : - ∥f x y∥ ≤ ∥f∥ * ∥x∥ * ∥y∥ := + ‖f x y‖ ≤ ‖f‖ * ‖x‖ * ‖y‖ := (f x).le_of_op_norm_le (f.le_op_norm x) y end -@[simp] lemma op_norm_prod (f : E →L[𝕜] Fₗ) (g : E →L[𝕜] Gₗ) : ∥f.prod g∥ = ∥(f, g)∥ := +@[simp] lemma op_norm_prod (f : E →L[𝕜] Fₗ) (g : E →L[𝕜] Gₗ) : ‖f.prod g‖ = ‖(f, g)‖ := le_antisymm (op_norm_le_bound _ (norm_nonneg _) $ λ x, by simpa only [prod_apply, prod.norm_def, max_mul_of_nonneg, norm_nonneg] @@ -660,7 +660,7 @@ le_antisymm (op_norm_le_bound _ (norm_nonneg _) $ λ x, (le_max_left _ _).trans ((f.prod g).le_op_norm x)) (op_norm_le_bound _ (norm_nonneg _) $ λ x, (le_max_right _ _).trans ((f.prod g).le_op_norm x)) -@[simp] lemma op_nnnorm_prod (f : E →L[𝕜] Fₗ) (g : E →L[𝕜] Gₗ) : ∥f.prod g∥₊ = ∥(f, g)∥₊ := +@[simp] lemma op_nnnorm_prod (f : E →L[𝕜] Fₗ) (g : E →L[𝕜] Gₗ) : ‖f.prod g‖₊ = ‖(f, g)‖₊ := subtype.ext $ op_norm_prod f g /-- `continuous_linear_map.prod` as a `linear_isometry_equiv`. -/ @@ -672,7 +672,7 @@ def prodₗᵢ (R : Type*) [semiring R] [module R Fₗ] [module R Gₗ] variables [ring_hom_isometric σ₁₂] (f : E →SL[σ₁₂] F) -@[simp, nontriviality] lemma op_norm_subsingleton [subsingleton E] : ∥f∥ = 0 := +@[simp, nontriviality] lemma op_norm_subsingleton [subsingleton E] : ‖f‖ = 0 := begin refine le_antisymm _ (norm_nonneg _), apply op_norm_le_bound _ rfl.ge, @@ -689,7 +689,7 @@ variables [ring_hom_isometric σ₁₂] open asymptotics -theorem is_O_with_id (l : filter E) : is_O_with ∥f∥ l f (λ x, x) := +theorem is_O_with_id (l : filter E) : is_O_with ‖f‖ l f (λ x, x) := is_O_with_of_le' _ f.le_op_norm theorem is_O_id (l : filter E) : f =O[l] (λ x, x) := @@ -697,7 +697,7 @@ theorem is_O_id (l : filter E) : f =O[l] (λ x, x) := theorem is_O_with_comp [ring_hom_isometric σ₂₃] {α : Type*} (g : F →SL[σ₂₃] G) (f : α → F) (l : filter α) : - is_O_with ∥g∥ l (λ x', g (f x')) f := + is_O_with ‖g‖ l (λ x', g (f x')) f := (g.is_O_with_id ⊤).comp_tendsto le_top theorem is_O_comp [ring_hom_isometric σ₂₃] {α : Type*} (g : F →SL[σ₂₃] G) (f : α → F) @@ -706,7 +706,7 @@ theorem is_O_comp [ring_hom_isometric σ₂₃] {α : Type*} (g : F →SL[σ₂ (g.is_O_with_comp f l).is_O theorem is_O_with_sub (f : E →SL[σ₁₂] F) (l : filter E) (x : E) : - is_O_with ∥f∥ l (λ x', f (x' - x)) (λ x', x' - x) := + is_O_with ‖f‖ l (λ x', f (x' - x)) (λ x', x' - x) := f.is_O_with_comp _ l theorem is_O_sub (f : E →SL[σ₁₂] F) (l : filter E) (x : E) : @@ -720,7 +720,7 @@ end continuous_linear_map namespace linear_isometry lemma norm_to_continuous_linear_map_le (f : E →ₛₗᵢ[σ₁₂] F) : - ∥f.to_continuous_linear_map∥ ≤ 1 := + ‖f.to_continuous_linear_map‖ ≤ 1 := f.to_continuous_linear_map.op_norm_le_bound zero_le_one $ λ x, by simp end linear_isometry @@ -729,14 +729,14 @@ namespace linear_map /-- If a continuous linear map is constructed from a linear map via the constructor `mk_continuous`, then its norm is bounded by the bound given to the constructor if it is nonnegative. -/ -lemma mk_continuous_norm_le (f : E →ₛₗ[σ₁₂] F) {C : ℝ} (hC : 0 ≤ C) (h : ∀x, ∥f x∥ ≤ C * ∥x∥) : - ∥f.mk_continuous C h∥ ≤ C := +lemma mk_continuous_norm_le (f : E →ₛₗ[σ₁₂] F) {C : ℝ} (hC : 0 ≤ C) (h : ∀x, ‖f x‖ ≤ C * ‖x‖) : + ‖f.mk_continuous C h‖ ≤ C := continuous_linear_map.op_norm_le_bound _ hC h /-- If a continuous linear map is constructed from a linear map via the constructor `mk_continuous`, then its norm is bounded by the bound or zero if bound is negative. -/ -lemma mk_continuous_norm_le' (f : E →ₛₗ[σ₁₂] F) {C : ℝ} (h : ∀x, ∥f x∥ ≤ C * ∥x∥) : - ∥f.mk_continuous C h∥ ≤ max C 0 := +lemma mk_continuous_norm_le' (f : E →ₛₗ[σ₁₂] F) {C : ℝ} (h : ∀x, ‖f x‖ ≤ C * ‖x‖) : + ‖f.mk_continuous C h‖ ≤ max C 0 := continuous_linear_map.op_norm_le_bound _ (le_max_right _ _) $ λ x, (h x).trans $ mul_le_mul_of_nonneg_right (le_max_left _ _) (norm_nonneg x) @@ -746,28 +746,28 @@ variables [ring_hom_isometric σ₂₃] map and a bound on the norm of the image. The linear map can be constructed using `linear_map.mk₂`. -/ def mk_continuous₂ (f : E →ₛₗ[σ₁₃] F →ₛₗ[σ₂₃] G) (C : ℝ) - (hC : ∀ x y, ∥f x y∥ ≤ C * ∥x∥ * ∥y∥) : + (hC : ∀ x y, ‖f x y‖ ≤ C * ‖x‖ * ‖y‖) : E →SL[σ₁₃] F →SL[σ₂₃] G := linear_map.mk_continuous - { to_fun := λ x, (f x).mk_continuous (C * ∥x∥) (hC x), + { to_fun := λ x, (f x).mk_continuous (C * ‖x‖) (hC x), map_add' := λ x y, by { ext z, simp }, map_smul' := λ c x, by { ext z, simp } } (max C 0) $ λ x, (mk_continuous_norm_le' _ _).trans_eq $ by rw [max_mul_of_nonneg _ _ (norm_nonneg x), zero_mul] @[simp] lemma mk_continuous₂_apply (f : E →ₛₗ[σ₁₃] F →ₛₗ[σ₂₃] G) {C : ℝ} - (hC : ∀ x y, ∥f x y∥ ≤ C * ∥x∥ * ∥y∥) (x : E) (y : F) : + (hC : ∀ x y, ‖f x y‖ ≤ C * ‖x‖ * ‖y‖) (x : E) (y : F) : f.mk_continuous₂ C hC x y = f x y := rfl lemma mk_continuous₂_norm_le' (f : E →ₛₗ[σ₁₃] F →ₛₗ[σ₂₃] G) {C : ℝ} - (hC : ∀ x y, ∥f x y∥ ≤ C * ∥x∥ * ∥y∥) : - ∥f.mk_continuous₂ C hC∥ ≤ max C 0 := + (hC : ∀ x y, ‖f x y‖ ≤ C * ‖x‖ * ‖y‖) : + ‖f.mk_continuous₂ C hC‖ ≤ max C 0 := mk_continuous_norm_le _ (le_max_iff.2 $ or.inr le_rfl) _ lemma mk_continuous₂_norm_le (f : E →ₛₗ[σ₁₃] F →ₛₗ[σ₂₃] G) {C : ℝ} (h0 : 0 ≤ C) - (hC : ∀ x y, ∥f x y∥ ≤ C * ∥x∥ * ∥y∥) : - ∥f.mk_continuous₂ C hC∥ ≤ C := + (hC : ∀ x y, ‖f x y‖ ≤ C * ‖x‖ * ‖y‖) : + ‖f.mk_continuous₂ C hC‖ ≤ C := (f.mk_continuous₂_norm_le' hC).trans_eq $ max_eq_left h0 end linear_map @@ -786,10 +786,10 @@ linear_map.mk_continuous₂ (λ c y x, (f x).map_smulₛₗ c y) (λ z x y, by rw [f.map_add, add_apply]) (λ c y x, by rw [f.map_smulₛₗ, smul_apply])) - ∥f∥ + ‖f‖ (λ y x, (f.le_op_norm₂ x y).trans_eq $ by rw mul_right_comm) -private lemma le_norm_flip (f : E →SL[σ₁₃] F →SL[σ₂₃] G) : ∥f∥ ≤ ∥flip f∥ := +private lemma le_norm_flip (f : E →SL[σ₁₃] F →SL[σ₂₃] G) : ‖f‖ ≤ ‖flip f‖ := f.op_norm_le_bound₂ (norm_nonneg _) $ λ x y, by { rw mul_right_comm, exact (flip f).le_op_norm₂ y x } @@ -800,7 +800,7 @@ f.op_norm_le_bound₂ (norm_nonneg _) $ λ x y, by { ext, refl } @[simp] lemma op_norm_flip (f : E →SL[σ₁₃] F →SL[σ₂₃] G) : - ∥f.flip∥ = ∥f∥ := + ‖f.flip‖ = ‖f‖ := le_antisymm (by simpa only [flip_flip] using le_norm_flip f.flip) (le_norm_flip f) @[simp] lemma flip_add (f g : E →SL[σ₁₃] F →SL[σ₂₃] G) : @@ -1001,7 +1001,7 @@ def mul : 𝕜' →L[𝕜] 𝕜' →L[𝕜] 𝕜' := (linear_map.mul 𝕜 𝕜') @[simp] lemma mul_apply' (x y : 𝕜') : mul 𝕜 𝕜' x y = x * y := rfl -@[simp] lemma op_norm_mul_apply_le (x : 𝕜') : ∥mul 𝕜 𝕜' x∥ ≤ ∥x∥ := +@[simp] lemma op_norm_mul_apply_le (x : 𝕜') : ‖mul 𝕜 𝕜' x‖ ≤ ‖x‖ := (op_norm_le_bound _ (norm_nonneg x) (norm_mul_le x)) /-- Simultaneous left- and right-multiplication in a non-unital normed algebra, considered as a @@ -1014,19 +1014,19 @@ def mul_left_right : 𝕜' →L[𝕜] 𝕜' →L[𝕜] 𝕜' →L[𝕜] 𝕜' := mul_left_right 𝕜 𝕜' x y z = x * z * y := rfl lemma op_norm_mul_left_right_apply_apply_le (x y : 𝕜') : - ∥mul_left_right 𝕜 𝕜' x y∥ ≤ ∥x∥ * ∥y∥ := + ‖mul_left_right 𝕜 𝕜' x y‖ ≤ ‖x‖ * ‖y‖ := (op_norm_comp_le _ _).trans $ (mul_comm _ _).trans_le $ mul_le_mul (op_norm_mul_apply_le _ _ _) (op_norm_le_bound _ (norm_nonneg _) (λ _, (norm_mul_le _ _).trans_eq (mul_comm _ _))) (norm_nonneg _) (norm_nonneg _) lemma op_norm_mul_left_right_apply_le (x : 𝕜') : - ∥mul_left_right 𝕜 𝕜' x∥ ≤ ∥x∥ := + ‖mul_left_right 𝕜 𝕜' x‖ ≤ ‖x‖ := op_norm_le_bound _ (norm_nonneg x) (op_norm_mul_left_right_apply_apply_le 𝕜 𝕜' x) lemma op_norm_mul_left_right_le : - ∥mul_left_right 𝕜 𝕜'∥ ≤ 1 := -op_norm_le_bound _ zero_le_one (λ x, (one_mul ∥x∥).symm ▸ op_norm_mul_left_right_apply_le 𝕜 𝕜' x) + ‖mul_left_right 𝕜 𝕜'‖ ≤ 1 := +op_norm_le_bound _ zero_le_one (λ x, (one_mul ‖x‖).symm ▸ op_norm_mul_left_right_apply_le 𝕜 𝕜' x) end non_unital @@ -1043,7 +1043,7 @@ def mulₗᵢ : 𝕜' →ₗᵢ[𝕜] 𝕜' →L[𝕜] 𝕜' := @[simp] lemma coe_mulₗᵢ : ⇑(mulₗᵢ 𝕜 𝕜') = mul 𝕜 𝕜' := rfl -@[simp] lemma op_norm_mul_apply (x : 𝕜') : ∥mul 𝕜 𝕜' x∥ = ∥x∥ := +@[simp] lemma op_norm_mul_apply (x : 𝕜') : ‖mul 𝕜 𝕜' x‖ = ‖x‖ := (mulₗᵢ 𝕜 𝕜').norm_map x end unital @@ -1064,7 +1064,7 @@ def lsmul : 𝕜' →L[𝕜] E →L[𝕜] E := variables {𝕜'} -lemma norm_to_span_singleton (x : E) : ∥to_span_singleton 𝕜 x∥ = ∥x∥ := +lemma norm_to_span_singleton (x : E) : ‖to_span_singleton 𝕜 x‖ = ‖x‖ := begin refine op_norm_eq_of_bounds (norm_nonneg _) (λ x, _) (λ N hN_nonneg h, _), { rw [to_span_singleton_apply, norm_smul, mul_comm], }, @@ -1075,11 +1075,11 @@ end variables {𝕜} -lemma op_norm_lsmul_apply_le (x : 𝕜') : ∥(lsmul 𝕜 𝕜' x : E →L[𝕜] E)∥ ≤ ∥x∥ := +lemma op_norm_lsmul_apply_le (x : 𝕜') : ‖(lsmul 𝕜 𝕜' x : E →L[𝕜] E)‖ ≤ ‖x‖ := continuous_linear_map.op_norm_le_bound _ (norm_nonneg x) $ λ y, (norm_smul x y).le /-- The norm of `lsmul` is at most 1 in any semi-normed group. -/ -lemma op_norm_lsmul_le : ∥(lsmul 𝕜 𝕜' : 𝕜' →L[𝕜] E →L[𝕜] E)∥ ≤ 1 := +lemma op_norm_lsmul_le : ‖(lsmul 𝕜 𝕜' : 𝕜' →L[𝕜] E →L[𝕜] E)‖ ≤ 1 := begin refine continuous_linear_map.op_norm_le_bound _ zero_le_one (λ x, _), simp_rw [one_mul], @@ -1094,7 +1094,7 @@ variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 variables [normed_space 𝕜' E] [is_scalar_tower 𝕜' 𝕜 E] variables [normed_space 𝕜' Fₗ] [is_scalar_tower 𝕜' 𝕜 Fₗ] -@[simp] lemma norm_restrict_scalars (f : E →L[𝕜] Fₗ) : ∥f.restrict_scalars 𝕜'∥ = ∥f∥ := +@[simp] lemma norm_restrict_scalars (f : E →L[𝕜] Fₗ) : ‖f.restrict_scalars 𝕜'‖ = ‖f‖ := le_antisymm (op_norm_le_bound _ (norm_nonneg _) $ λ x, f.le_op_norm x) (op_norm_le_bound _ (norm_nonneg _) $ λ x, f.le_op_norm x) @@ -1138,7 +1138,7 @@ end continuous_linear_map namespace submodule -lemma norm_subtypeL_le (K : submodule 𝕜 E) : ∥K.subtypeL∥ ≤ 1 := +lemma norm_subtypeL_le (K : submodule 𝕜 E) : ‖K.subtypeL‖ ≤ 1 := K.subtypeₗᵢ.norm_to_continuous_linear_map_le end submodule @@ -1213,7 +1213,7 @@ variables {σ₂₁ : 𝕜₂ →+* 𝕜} [ring_hom_inv_pair σ₁₂ σ₂₁] variables (e : E ≃SL[σ₁₂] F) include σ₂₁ -protected lemma lipschitz : lipschitz_with (∥(e : E →SL[σ₁₂] F)∥₊) e := +protected lemma lipschitz : lipschitz_with (‖(e : E →SL[σ₁₂] F)‖₊) e := (e : E →SL[σ₁₂] F).lipschitz theorem is_O_comp {α : Type*} (f : α → E) (l : filter α) : (λ x', e (f x')) =O[l] f := @@ -1228,17 +1228,17 @@ variables {σ₂₁ : 𝕜₂ →+* 𝕜} [ring_hom_inv_pair σ₁₂ σ₂₁] include σ₂₁ lemma homothety_inverse (a : ℝ) (ha : 0 < a) (f : E ≃ₛₗ[σ₁₂] F) : - (∀ (x : E), ∥f x∥ = a * ∥x∥) → (∀ (y : F), ∥f.symm y∥ = a⁻¹ * ∥y∥) := + (∀ (x : E), ‖f x‖ = a * ‖x‖) → (∀ (y : F), ‖f.symm y‖ = a⁻¹ * ‖y‖) := begin intros hf y, - calc ∥(f.symm) y∥ = a⁻¹ * (a * ∥ (f.symm) y∥) : _ - ... = a⁻¹ * ∥f ((f.symm) y)∥ : by rw hf - ... = a⁻¹ * ∥y∥ : by simp, + calc ‖(f.symm) y‖ = a⁻¹ * (a * ‖ (f.symm) y‖) : _ + ... = a⁻¹ * ‖f ((f.symm) y)‖ : by rw hf + ... = a⁻¹ * ‖y‖ : by simp, rw [← mul_assoc, inv_mul_cancel (ne_of_lt ha).symm, one_mul], end /-- A linear equivalence which is a homothety is a continuous linear equivalence. -/ -def of_homothety (f : E ≃ₛₗ[σ₁₂] F) (a : ℝ) (ha : 0 < a) (hf : ∀x, ∥f x∥ = a * ∥x∥) : +def of_homothety (f : E ≃ₛₗ[σ₁₂] F) (a : ℝ) (ha : 0 < a) (hf : ∀x, ‖f x‖ = a * ‖x‖) : E ≃SL[σ₁₂] F := { to_linear_equiv := f, continuous_to_fun := add_monoid_hom_class.continuous_of_bound f a (λ x, le_of_eq (hf x)), @@ -1258,7 +1258,7 @@ omit σ₂₁ variable (𝕜) lemma to_span_nonzero_singleton_homothety (x : E) (h : x ≠ 0) (c : 𝕜) : - ∥linear_equiv.to_span_nonzero_singleton 𝕜 E x h c∥ = ∥x∥ * ∥c∥ := + ‖linear_equiv.to_span_nonzero_singleton 𝕜 E x h c‖ = ‖x‖ * ‖c‖ := continuous_linear_map.to_span_singleton_homothety _ _ _ end continuous_linear_equiv @@ -1269,7 +1269,7 @@ include σ₂₁ /-- Construct a continuous linear equivalence from a linear equivalence together with bounds in both directions. -/ def linear_equiv.to_continuous_linear_equiv_of_bounds (e : E ≃ₛₗ[σ₁₂] F) (C_to C_inv : ℝ) - (h_to : ∀ x, ∥e x∥ ≤ C_to * ∥x∥) (h_inv : ∀ x : F, ∥e.symm x∥ ≤ C_inv * ∥x∥) : E ≃SL[σ₁₂] F := + (h_to : ∀ x, ‖e x‖ ≤ C_to * ‖x‖) (h_inv : ∀ x : F, ‖e.symm x‖ ≤ C_inv * ‖x‖) : E ≃SL[σ₁₂] F := { to_linear_equiv := e, continuous_to_fun := add_monoid_hom_class.continuous_of_bound e C_to h_to, continuous_inv_fun := add_monoid_hom_class.continuous_of_bound e.symm C_inv h_inv } @@ -1332,9 +1332,9 @@ variables [nontrivially_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] (f g : E →SL[σ₁₂] F) (x y z : E) lemma linear_map.bound_of_shell [ring_hom_isometric σ₁₂] (f : E →ₛₗ[σ₁₂] F) {ε C : ℝ} - (ε_pos : 0 < ε) {c : 𝕜} (hc : 1 < ∥c∥) - (hf : ∀ x, ε / ∥c∥ ≤ ∥x∥ → ∥x∥ < ε → ∥f x∥ ≤ C * ∥x∥) (x : E) : - ∥f x∥ ≤ C * ∥x∥ := + (ε_pos : 0 < ε) {c : 𝕜} (hc : 1 < ‖c‖) + (hf : ∀ x, ε / ‖c‖ ≤ ‖x‖ → ‖x‖ < ε → ‖f x‖ ≤ C * ‖x‖) (x : E) : + ‖f x‖ ≤ C * ‖x‖ := begin by_cases hx : x = 0, { simp [hx] }, exact semilinear_map_class.bound_of_shell_semi_normed f ε_pos hc hf @@ -1346,15 +1346,15 @@ end that produces a concrete bound. -/ lemma linear_map.bound_of_ball_bound {r : ℝ} (r_pos : 0 < r) (c : ℝ) (f : E →ₗ[𝕜] Fₗ) - (h : ∀ z ∈ metric.ball (0 : E) r, ∥f z∥ ≤ c) : - ∃ C, ∀ (z : E), ∥f z∥ ≤ C * ∥z∥ := + (h : ∀ z ∈ metric.ball (0 : E) r, ‖f z‖ ≤ c) : + ∃ C, ∀ (z : E), ‖f z‖ ≤ C * ‖z‖ := begin cases @nontrivially_normed_field.non_trivial 𝕜 _ with k hk, - use c * (∥k∥ / r), + use c * (‖k‖ / r), intro z, refine linear_map.bound_of_shell _ r_pos hk (λ x hko hxo, _) _, - calc ∥f x∥ ≤ c : h _ (mem_ball_zero_iff.mpr hxo) - ... ≤ c * ((∥x∥ * ∥k∥) / r) : le_mul_of_one_le_right _ _ + calc ‖f x‖ ≤ c : h _ (mem_ball_zero_iff.mpr hxo) + ... ≤ c * ((‖x‖ * ‖k‖) / r) : le_mul_of_one_le_right _ _ ... = _ : by ring, { exact le_trans (norm_nonneg _) (h 0 (by simp [r_pos])) }, { rw [div_le_iff (zero_lt_one.trans hk)] at hko, @@ -1367,15 +1367,15 @@ section op_norm open set real /-- An operator is zero iff its norm vanishes. -/ -theorem op_norm_zero_iff [ring_hom_isometric σ₁₂] : ∥f∥ = 0 ↔ f = 0 := +theorem op_norm_zero_iff [ring_hom_isometric σ₁₂] : ‖f‖ = 0 ↔ f = 0 := iff.intro (λ hn, continuous_linear_map.ext (λ x, norm_le_zero_iff.1 - (calc _ ≤ ∥f∥ * ∥x∥ : le_op_norm _ _ + (calc _ ≤ ‖f‖ * ‖x‖ : le_op_norm _ _ ... = _ : by rw [hn, zero_mul]))) (by { rintro rfl, exact op_norm_zero }) /-- If a normed space is non-trivial, then the norm of the identity equals `1`. -/ -@[simp] lemma norm_id [nontrivial E] : ∥id 𝕜 E∥ = 1 := +@[simp] lemma norm_id [nontrivial E] : ‖id 𝕜 E‖ = 1 := begin refine norm_id_of_nontrivial_seminorm _, obtain ⟨x, hx⟩ := exists_ne (0 : E), @@ -1396,8 +1396,8 @@ instance to_normed_ring : normed_ring (E →L[𝕜] E) := variable {f} lemma homothety_norm [ring_hom_isometric σ₁₂] [nontrivial E] (f : E →SL[σ₁₂] F) {a : ℝ} - (hf : ∀x, ∥f x∥ = a * ∥x∥) : - ∥f∥ = a := + (hf : ∀x, ‖f x‖ = a * ‖x‖) : + ‖f‖ = a := begin obtain ⟨x, hx⟩ : ∃ (x : E), x ≠ 0 := exists_ne 0, rw ← norm_pos_iff at hx, @@ -1408,7 +1408,7 @@ end variable (f) -theorem uniform_embedding_of_bound {K : ℝ≥0} (hf : ∀ x, ∥x∥ ≤ K * ∥f x∥) : +theorem uniform_embedding_of_bound {K : ℝ≥0} (hf : ∀ x, ‖x‖ ≤ K * ‖f x‖) : uniform_embedding f := (add_monoid_hom_class.antilipschitz_of_bound f hf).uniform_embedding f.uniform_continuous @@ -1421,7 +1421,7 @@ begin from (uniform_embedding_iff.1 hf).2.2 1 zero_lt_one, let δ := ε/2, have δ_pos : δ > 0 := half_pos εpos, - have H : ∀{x}, ∥f x∥ ≤ δ → ∥x∥ ≤ 1, + have H : ∀{x}, ‖f x‖ ≤ δ → ‖x‖ ≤ 1, { assume x hx, have : dist x 0 ≤ 1, { refine (hε _).le, @@ -1429,7 +1429,7 @@ begin exact hx.trans_lt (half_lt_self εpos) }, simpa using this }, rcases normed_field.exists_one_lt_norm 𝕜 with ⟨c, hc⟩, - refine ⟨⟨δ⁻¹, _⟩ * ∥c∥₊, add_monoid_hom_class.antilipschitz_of_bound f $ λx, _⟩, + refine ⟨⟨δ⁻¹, _⟩ * ‖c‖₊, add_monoid_hom_class.antilipschitz_of_bound f $ λx, _⟩, exact inv_nonneg.2 (le_of_lt δ_pos), by_cases hx : f x = 0, { have : f x = f 0, by { simp [hx] }, @@ -1437,12 +1437,12 @@ begin simp [this] }, { rcases rescale_to_shell hc δ_pos hx with ⟨d, hd, dxlt, ledx, dinv⟩, rw [← f.map_smul d] at dxlt, - have : ∥d • x∥ ≤ 1 := H dxlt.le, - calc ∥x∥ = ∥d∥⁻¹ * ∥d • x∥ : + have : ‖d • x‖ ≤ 1 := H dxlt.le, + calc ‖x‖ = ‖d‖⁻¹ * ‖d • x‖ : by rwa [← norm_inv, ← norm_smul, ← mul_smul, inv_mul_cancel, one_smul] - ... ≤ ∥d∥⁻¹ * 1 : + ... ≤ ‖d‖⁻¹ * 1 : mul_le_mul_of_nonneg_left this (inv_nonneg.2 (norm_nonneg _)) - ... ≤ δ⁻¹ * ∥c∥ * ∥f x∥ : + ... ≤ δ⁻¹ * ‖c‖ * ‖f x‖ : by rwa [mul_one] } end @@ -1464,10 +1464,10 @@ begin -- `f` is a linear map due to `linear_map_of_mem_closure_range_coe` refine (linear_map_of_mem_closure_range_coe f _).mk_continuous_of_exists_bound _, { refine closure_mono (image_subset_iff.2 $ λ g hg, _) hf, exact ⟨g, rfl⟩ }, - { -- We need to show that `f` has bounded norm. Choose `C` such that `∥g∥ ≤ C` for all `g ∈ s`. + { -- We need to show that `f` has bounded norm. Choose `C` such that `‖g‖ ≤ C` for all `g ∈ s`. rcases bounded_iff_forall_norm_le.1 hs with ⟨C, hC⟩, - -- Then `∥g x∥ ≤ C * ∥x∥` for all `g ∈ s`, `x : E`, hence `∥f x∥ ≤ C * ∥x∥` for all `x`. - have : ∀ x, is_closed {g : E' → F | ∥g x∥ ≤ C * ∥x∥}, + -- Then `‖g x‖ ≤ C * ‖x‖` for all `g ∈ s`, `x : E`, hence `‖f x‖ ≤ C * ‖x‖` for all `x`. + have : ∀ x, is_closed {g : E' → F | ‖g x‖ ≤ C * ‖x‖}, from λ x, is_closed_Iic.preimage (@continuous_apply E' (λ _, F) _ x).norm, refine ⟨C, λ x, (this x).closure_subset_iff.2 (image_subset_iff.2 $ λ g hg, _) hf⟩, exact g.le_of_op_norm_le (hC _ hg) _ } @@ -1490,20 +1490,20 @@ lemma tendsto_of_tendsto_pointwise_of_cauchy_seq {f : ℕ → E' →SL[σ₁₂] (hg : tendsto (λ n x, f n x) at_top (𝓝 g)) (hf : cauchy_seq f) : tendsto f at_top (𝓝 g) := begin - /- Since `f` is a Cauchy sequence, there exists `b → 0` such that `∥f n - f m∥ ≤ b N` for any + /- Since `f` is a Cauchy sequence, there exists `b → 0` such that `‖f n - f m‖ ≤ b N` for any `m, n ≥ N`. -/ rcases cauchy_seq_iff_le_tendsto_0.1 hf with ⟨b, hb₀, hfb, hb_lim⟩, - -- Since `b → 0`, it suffices to show that `∥f n x - g x∥ ≤ b n * ∥x∥` for all `n` and `x`. - suffices : ∀ n x, ∥f n x - g x∥ ≤ b n * ∥x∥, + -- Since `b → 0`, it suffices to show that `‖f n x - g x‖ ≤ b n * ‖x‖` for all `n` and `x`. + suffices : ∀ n x, ‖f n x - g x‖ ≤ b n * ‖x‖, from tendsto_iff_norm_tendsto_zero.2 (squeeze_zero (λ n, norm_nonneg _) (λ n, op_norm_le_bound _ (hb₀ n) (this n)) hb_lim), intros n x, - -- Note that `f m x → g x`, hence `∥f n x - f m x∥ → ∥f n x - g x∥` as `m → ∞` - have : tendsto (λ m, ∥f n x - f m x∥) at_top (𝓝 (∥f n x - g x∥)), + -- Note that `f m x → g x`, hence `‖f n x - f m x‖ → ‖f n x - g x‖` as `m → ∞` + have : tendsto (λ m, ‖f n x - f m x‖) at_top (𝓝 (‖f n x - g x‖)), from (tendsto_const_nhds.sub $ tendsto_pi_nhds.1 hg _).norm, - -- Thus it suffices to verify `∥f n x - f m x∥ ≤ b n * ∥x∥` for `m ≥ n`. + -- Thus it suffices to verify `‖f n x - f m x‖ ≤ b n * ‖x‖` for `m ≥ n`. refine le_of_tendsto this (eventually_at_top.2 ⟨n, λ m hm, _⟩), - -- This inequality follows from `∥f n - f m∥ ≤ b n`. + -- This inequality follows from `‖f n - f m‖ ≤ b n`. exact (f n - f m).le_of_op_norm_le (hfb _ _ _ le_rfl hm) _ end @@ -1581,7 +1581,7 @@ begin have hr : 0 ≤ r, from nonempty_closed_ball.1 (nonempty_image_iff.1 (closure_nonempty_iff.1 ⟨_, hf⟩)), refine mem_closed_ball_iff_norm.2 (op_norm_le_bound _ hr $ λ x, _), - have : is_closed {g : E' → F | ∥g x - f₀ x∥ ≤ r * ∥x∥}, + have : is_closed {g : E' → F | ‖g x - f₀ x‖ ≤ r * ‖x‖}, from is_closed_Iic.preimage ((@continuous_apply E' (λ _, F) _ x).sub continuous_const).norm, refine this.closure_subset_iff.2 (image_subset_iff.2 $ λ g hg, _) hf, exact (g - f₀).le_of_op_norm_le (mem_closed_ball_iff_norm.1 hg) _ @@ -1649,13 +1649,13 @@ extend_unique _ _ _ _ _ (zero_comp _) end section -variables {N : ℝ≥0} (h_e : ∀x, ∥x∥ ≤ N * ∥e x∥) [ring_hom_isometric σ₁₂] +variables {N : ℝ≥0} (h_e : ∀x, ‖x‖ ≤ N * ‖e x‖) [ring_hom_isometric σ₁₂] local notation `ψ` := f.extend e h_dense (uniform_embedding_of_bound _ h_e).to_uniform_inducing /-- If a dense embedding `e : E →L[𝕜] G` expands the norm by a constant factor `N⁻¹`, then the -norm of the extension of `f` along `e` is bounded by `N * ∥f∥`. -/ -lemma op_norm_extend_le : ∥ψ∥ ≤ N * ∥f∥ := +norm of the extension of `f` along `e` is bounded by `N * ‖f‖`. -/ +lemma op_norm_extend_le : ‖ψ‖ ≤ N * ‖f‖ := begin have uni : uniform_inducing e := (uniform_embedding_of_bound _ h_e).to_uniform_inducing, have eq : ∀x, ψ (e x) = f x := uniformly_extend_of_ind uni h_dense f.uniform_continuous, @@ -1666,9 +1666,9 @@ begin { exact continuous_const.mul continuous_norm }, { assume x, rw eq, - calc ∥f x∥ ≤ ∥f∥ * ∥x∥ : le_op_norm _ _ - ... ≤ ∥f∥ * (N * ∥e x∥) : mul_le_mul_of_nonneg_left (h_e x) (norm_nonneg _) - ... ≤ N * ∥f∥ * ∥e x∥ : by rw [mul_comm ↑N ∥f∥, mul_assoc] } }, + calc ‖f x‖ ≤ ‖f‖ * ‖x‖ : le_op_norm _ _ + ... ≤ ‖f‖ * (N * ‖e x‖) : mul_le_mul_of_nonneg_left (h_e x) (norm_nonneg _) + ... ≤ N * ‖f‖ * ‖e x‖ : by rw [mul_comm ↑N ‖f‖, mul_assoc] } }, { have he : ∀ x : E, x = 0, { assume x, have N0 : N ≤ 0 := le_of_lt (lt_of_not_ge N0), @@ -1691,7 +1691,7 @@ namespace linear_isometry @[simp] lemma norm_to_continuous_linear_map [nontrivial E] [ring_hom_isometric σ₁₂] (f : E →ₛₗᵢ[σ₁₂] F) : - ∥f.to_continuous_linear_map∥ = 1 := + ‖f.to_continuous_linear_map‖ = 1 := f.to_continuous_linear_map.homothety_norm $ by simp variables {σ₁₃ : 𝕜 →+* 𝕜₃} [ring_hom_comp_triple σ₁₂ σ₂₃ σ₁₃] @@ -1701,7 +1701,7 @@ include σ₁₃ the operator norm. -/ lemma norm_to_continuous_linear_map_comp [ring_hom_isometric σ₁₂] (f : F →ₛₗᵢ[σ₂₃] G) {g : E →SL[σ₁₂] F} : - ∥f.to_continuous_linear_map.comp g∥ = ∥g∥ := + ‖f.to_continuous_linear_map.comp g‖ = ‖g‖ := op_norm_ext (f.to_continuous_linear_map.comp g) g (λ x, by simp only [norm_map, coe_to_continuous_linear_map, coe_comp']) omit σ₁₃ @@ -1728,7 +1728,7 @@ variables {𝕜₂' : Type*} [nontrivially_normed_field 𝕜₂'] {F' : Type*} [ include σ₂'' σ₂₃' /-- Precomposition with a linear isometry preserves the operator norm. -/ lemma op_norm_comp_linear_isometry_equiv (f : F →SL[σ₂₃] G) (g : F' ≃ₛₗᵢ[σ₂'] F) : - ∥f.comp g.to_linear_isometry.to_continuous_linear_map∥ = ∥f∥ := + ‖f.comp g.to_linear_isometry.to_continuous_linear_map‖ = ‖f‖ := begin casesI subsingleton_or_nontrivial F', { haveI := g.symm.to_linear_equiv.to_equiv.subsingleton, @@ -1748,30 +1748,30 @@ omit σ₂'' σ₂₃' /-- The norm of the tensor product of a scalar linear map and of an element of a normed space is the product of the norms. -/ @[simp] lemma norm_smul_right_apply (c : E →L[𝕜] 𝕜) (f : Fₗ) : - ∥smul_right c f∥ = ∥c∥ * ∥f∥ := + ‖smul_right c f‖ = ‖c‖ * ‖f‖ := begin refine le_antisymm _ _, { apply op_norm_le_bound _ (mul_nonneg (norm_nonneg _) (norm_nonneg _)) (λx, _), calc - ∥(c x) • f∥ = ∥c x∥ * ∥f∥ : norm_smul _ _ - ... ≤ (∥c∥ * ∥x∥) * ∥f∥ : + ‖(c x) • f‖ = ‖c x‖ * ‖f‖ : norm_smul _ _ + ... ≤ (‖c‖ * ‖x‖) * ‖f‖ : mul_le_mul_of_nonneg_right (le_op_norm _ _) (norm_nonneg _) - ... = ∥c∥ * ∥f∥ * ∥x∥ : by ring }, + ... = ‖c‖ * ‖f‖ * ‖x‖ : by ring }, { by_cases h : f = 0, { simp [h] }, - { have : 0 < ∥f∥ := norm_pos_iff.2 h, + { have : 0 < ‖f‖ := norm_pos_iff.2 h, rw ← le_div_iff this, apply op_norm_le_bound _ (div_nonneg (norm_nonneg _) (norm_nonneg f)) (λx, _), rw [div_mul_eq_mul_div, le_div_iff this], - calc ∥c x∥ * ∥f∥ = ∥c x • f∥ : (norm_smul _ _).symm - ... = ∥smul_right c f x∥ : rfl - ... ≤ ∥smul_right c f∥ * ∥x∥ : le_op_norm _ _ } }, + calc ‖c x‖ * ‖f‖ = ‖c x • f‖ : (norm_smul _ _).symm + ... = ‖smul_right c f x‖ : rfl + ... ≤ ‖smul_right c f‖ * ‖x‖ : le_op_norm _ _ } }, end /-- The non-negative norm of the tensor product of a scalar linear map and of an element of a normed space is the product of the non-negative norms. -/ @[simp] lemma nnnorm_smul_right_apply (c : E →L[𝕜] 𝕜) (f : Fₗ) : - ∥smul_right c f∥₊ = ∥c∥₊ * ∥f∥₊ := + ‖smul_right c f‖₊ = ‖c‖₊ * ‖f‖₊ := nnreal.eq $ c.norm_smul_right_apply f variables (𝕜 E Fₗ) @@ -1791,11 +1791,11 @@ linear_map.mk_continuous₂ variables {𝕜 E Fₗ} @[simp] lemma norm_smul_rightL_apply (c : E →L[𝕜] 𝕜) (f : Fₗ) : - ∥smul_rightL 𝕜 E Fₗ c f∥ = ∥c∥ * ∥f∥ := + ‖smul_rightL 𝕜 E Fₗ c f‖ = ‖c‖ * ‖f‖ := norm_smul_right_apply c f @[simp] lemma norm_smul_rightL (c : E →L[𝕜] 𝕜) [nontrivial Fₗ] : - ∥smul_rightL 𝕜 E Fₗ c∥ = ∥c∥ := + ‖smul_rightL 𝕜 E Fₗ c‖ = ‖c‖ := continuous_linear_map.homothety_norm _ c.norm_smul_right_apply variables (𝕜) (𝕜' : Type*) @@ -1803,7 +1803,7 @@ variables (𝕜) (𝕜' : Type*) section variables [normed_ring 𝕜'] [normed_algebra 𝕜 𝕜'] -@[simp] lemma op_norm_mul [norm_one_class 𝕜'] : ∥mul 𝕜 𝕜'∥ = 1 := +@[simp] lemma op_norm_mul [norm_one_class 𝕜'] : ‖mul 𝕜 𝕜'‖ = 1 := by haveI := norm_one_class.nontrivial 𝕜'; exact (mulₗᵢ 𝕜 𝕜').norm_to_continuous_linear_map end @@ -1813,7 +1813,7 @@ end This is `continuous_linear_map.op_norm_lsmul_le` as an equality. -/ @[simp] lemma op_norm_lsmul [normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] [normed_space 𝕜' E] [is_scalar_tower 𝕜 𝕜' E] [nontrivial E] : - ∥(lsmul 𝕜 𝕜' : 𝕜' →L[𝕜] E →L[𝕜] E)∥ = 1 := + ‖(lsmul 𝕜 𝕜' : 𝕜' →L[𝕜] E →L[𝕜] E)‖ = 1 := begin refine continuous_linear_map.op_norm_eq_of_bounds zero_le_one (λ x, _) (λ N hN h, _), { rw one_mul, @@ -1831,7 +1831,7 @@ namespace submodule variables [nontrivially_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] [nontrivially_normed_field 𝕜₃] [normed_space 𝕜 E] [normed_space 𝕜₂ F] {σ₁₂ : 𝕜 →+* 𝕜₂} -lemma norm_subtypeL (K : submodule 𝕜 E) [nontrivial K] : ∥K.subtypeL∥ = 1 := +lemma norm_subtypeL (K : submodule 𝕜 E) [nontrivial K] : ‖K.subtypeL‖ = 1 := K.subtypeₗᵢ.norm_to_continuous_linear_map end submodule @@ -1846,11 +1846,11 @@ section variables [ring_hom_isometric σ₂₁] protected lemma antilipschitz (e : E ≃SL[σ₁₂] F) : - antilipschitz_with ∥(e.symm : F →SL[σ₂₁] E)∥₊ e := + antilipschitz_with ‖(e.symm : F →SL[σ₂₁] E)‖₊ e := e.symm.lipschitz.to_right_inverse e.left_inv lemma one_le_norm_mul_norm_symm [ring_hom_isometric σ₁₂] [nontrivial E] (e : E ≃SL[σ₁₂] F) : - 1 ≤ ∥(e : E →SL[σ₁₂] F)∥ * ∥(e.symm : F →SL[σ₂₁] E)∥ := + 1 ≤ ‖(e : E →SL[σ₁₂] F)‖ * ‖(e.symm : F →SL[σ₂₁] E)‖ := begin rw [mul_comm], convert (e.symm : F →SL[σ₂₁] E).op_norm_comp_le (e : E →SL[σ₁₂] F), @@ -1859,20 +1859,20 @@ end include σ₂₁ lemma norm_pos [ring_hom_isometric σ₁₂] [nontrivial E] (e : E ≃SL[σ₁₂] F) : - 0 < ∥(e : E →SL[σ₁₂] F)∥ := + 0 < ‖(e : E →SL[σ₁₂] F)‖ := pos_of_mul_pos_left (lt_of_lt_of_le zero_lt_one e.one_le_norm_mul_norm_symm) (norm_nonneg _) omit σ₂₁ lemma norm_symm_pos [ring_hom_isometric σ₁₂] [nontrivial E] (e : E ≃SL[σ₁₂] F) : - 0 < ∥(e.symm : F →SL[σ₂₁] E)∥ := + 0 < ‖(e.symm : F →SL[σ₂₁] E)‖ := pos_of_mul_pos_right (zero_lt_one.trans_le e.one_le_norm_mul_norm_symm) (norm_nonneg _) lemma nnnorm_symm_pos [ring_hom_isometric σ₁₂] [nontrivial E] (e : E ≃SL[σ₁₂] F) : - 0 < ∥(e.symm : F →SL[σ₂₁] E)∥₊ := + 0 < ‖(e.symm : F →SL[σ₂₁] E)‖₊ := e.norm_symm_pos lemma subsingleton_or_norm_symm_pos [ring_hom_isometric σ₁₂] (e : E ≃SL[σ₁₂] F) : - subsingleton E ∨ 0 < ∥(e.symm : F →SL[σ₂₁] E)∥ := + subsingleton E ∨ 0 < ‖(e.symm : F →SL[σ₂₁] E)‖ := begin rcases subsingleton_or_nontrivial E with _i|_i; resetI, { left, apply_instance }, @@ -1880,7 +1880,7 @@ begin end lemma subsingleton_or_nnnorm_symm_pos [ring_hom_isometric σ₁₂] (e : E ≃SL[σ₁₂] F) : - subsingleton E ∨ 0 < ∥(e.symm : F →SL[σ₂₁] E)∥₊ := + subsingleton E ∨ 0 < ‖(e.symm : F →SL[σ₂₁] E)‖₊ := subsingleton_or_norm_symm_pos e variable (𝕜) @@ -1890,7 +1890,7 @@ variable (𝕜) def to_span_nonzero_singleton (x : E) (h : x ≠ 0) : 𝕜 ≃L[𝕜] (𝕜 ∙ x) := of_homothety (linear_equiv.to_span_nonzero_singleton 𝕜 E x h) - ∥x∥ + ‖x‖ (norm_pos_iff.mpr h) (to_span_nonzero_singleton_homothety 𝕜 x h) @@ -1909,9 +1909,9 @@ def coord (x : E) (h : x ≠ 0) : (𝕜 ∙ x) →L[𝕜] 𝕜 := (to_span_nonze to_span_nonzero_singleton 𝕜 x h (coord 𝕜 x h y) = y := (to_span_nonzero_singleton 𝕜 x h).apply_symm_apply y -@[simp] lemma coord_norm (x : E) (h : x ≠ 0) : ∥coord 𝕜 x h∥ = ∥x∥⁻¹ := +@[simp] lemma coord_norm (x : E) (h : x ≠ 0) : ‖coord 𝕜 x h‖ = ‖x‖⁻¹ := begin - have hx : 0 < ∥x∥ := (norm_pos_iff.mpr h), + have hx : 0 < ‖x‖ := (norm_pos_iff.mpr h), haveI : nontrivial (𝕜 ∙ x) := submodule.nontrivial_span_singleton h, exact continuous_linear_map.homothety_norm _ (λ y, homothety_inverse _ hx _ (to_span_nonzero_singleton_homothety 𝕜 x h) _) @@ -1969,9 +1969,9 @@ end normed /-- A bounded bilinear form `B` in a real normed space is *coercive* -if there is some positive constant C such that `C * ∥u∥ * ∥u∥ ≤ B u u`. +if there is some positive constant C such that `C * ‖u‖ * ‖u‖ ≤ B u u`. -/ def is_coercive [normed_add_comm_group E] [normed_space ℝ E] (B : E →L[ℝ] E →L[ℝ] ℝ) : Prop := -∃ C, (0 < C) ∧ ∀ u, C * ∥u∥ * ∥u∥ ≤ B u u +∃ C, (0 < C) ∧ ∀ u, C * ‖u‖ * ‖u‖ ≤ B u u diff --git a/src/analysis/normed_space/pi_Lp.lean b/src/analysis/normed_space/pi_Lp.lean index f45a505f3eaa0..12b33f1a0c6db 100644 --- a/src/analysis/normed_space/pi_Lp.lean +++ b/src/analysis/normed_space/pi_Lp.lean @@ -34,7 +34,7 @@ We only deal with the `L^p` distance on a product of finitely many metric spaces distinct. A closely related construction is `lp`, the `L^p` norm on a product of (possibly infinitely many) normed spaces, where the norm is $$ -\left(\sum ∥f (x)∥^p \right)^{1/p}. +\left(\sum ‖f (x)‖^p \right)^{1/p}. $$ However, the topology induced by this construction is not the product topology, and some functions have infinite `L^p` norm. These subtleties are not present in the case of finitely many metric @@ -43,7 +43,7 @@ spaces, hence it is worth devoting a file to this specific case which is particu Another related construction is `measure_theory.Lp`, the `L^p` norm on the space of functions from a measure space to a normed space, where the norm is $$ -\left(\int ∥f (x)∥^p dμ\right)^{1/p}. +\left(\int ‖f (x)‖^p dμ\right)^{1/p}. $$ This has all the same subtleties as `lp`, and the further subtlety that this only defines a seminorm (as almost everywhere zero functions have zero `L^p` norm). @@ -193,17 +193,17 @@ Registering this separately allows for a future norm-like structure on `pi_Lp p satisfying a relaxed triangle inequality. These are called *quasi-norms*. -/ instance has_norm : has_norm (pi_Lp p β) := { norm := λ f, if hp : p = 0 then {i | f i ≠ 0}.to_finite.to_finset.card - else (if p = ∞ then ⨆ i, ∥f i∥ else (∑ i, ∥f i∥ ^ p.to_real) ^ (1 / p.to_real)) } + else (if p = ∞ then ⨆ i, ‖f i‖ else (∑ i, ‖f i‖ ^ p.to_real) ^ (1 / p.to_real)) } variables {p β} -lemma norm_eq_card (f : pi_Lp 0 β) : ∥f∥ = {i | f i ≠ 0}.to_finite.to_finset.card := +lemma norm_eq_card (f : pi_Lp 0 β) : ‖f‖ = {i | f i ≠ 0}.to_finite.to_finset.card := if_pos rfl -lemma norm_eq_csupr (f : pi_Lp ∞ β) : ∥f∥ = ⨆ i, ∥f i∥ := +lemma norm_eq_csupr (f : pi_Lp ∞ β) : ‖f‖ = ⨆ i, ‖f i‖ := by { dsimp [norm], exact if_neg ennreal.top_ne_zero } lemma norm_eq_sum (hp : 0 < p.to_real) (f : pi_Lp p β) : - ∥f∥ = (∑ i, ∥f i∥ ^ p.to_real) ^ (1 / p.to_real) := + ‖f‖ = (∑ i, ‖f i‖ ^ p.to_real) ^ (1 / p.to_real) := let hp' := ennreal.to_real_pos_iff.mp hp in (if_neg hp'.1.ne').trans (if_neg hp'.2.ne) end norm @@ -466,16 +466,16 @@ instance normed_add_comm_group [Π i, normed_add_comm_group (α i)] : lemma nnnorm_eq_sum {p : ℝ≥0∞} [fact (1 ≤ p)] {β : ι → Type*} (hp : p ≠ ∞) [Π i, seminormed_add_comm_group (β i)] (f : pi_Lp p β) : - ∥f∥₊ = (∑ i, ∥f i∥₊ ^ p.to_real) ^ (1 / p.to_real) := + ‖f‖₊ = (∑ i, ‖f i‖₊ ^ p.to_real) ^ (1 / p.to_real) := by { ext, simp [nnreal.coe_sum, norm_eq_sum (p.to_real_pos_iff_ne_top.mpr hp)] } lemma nnnorm_eq_csupr {β : ι → Type*} [Π i, seminormed_add_comm_group (β i)] (f : pi_Lp ∞ β) : - ∥f∥₊ = ⨆ i, ∥f i∥₊ := + ‖f‖₊ = ⨆ i, ‖f i‖₊ := by { ext, simp [nnreal.coe_supr, norm_eq_csupr] } lemma norm_eq_of_nat {p : ℝ≥0∞} [fact (1 ≤ p)] {β : ι → Type*} [Π i, seminormed_add_comm_group (β i)] (n : ℕ) (h : p = n) (f : pi_Lp p β) : - ∥f∥ = (∑ i, ∥f i∥ ^ n) ^ (1/(n : ℝ)) := + ‖f‖ = (∑ i, ‖f i‖ ^ n) ^ (1/(n : ℝ)) := begin have := p.to_real_pos_iff_ne_top.mpr (ne_of_eq_of_ne h $ ennreal.nat_ne_top n), simp only [one_div, h, real.rpow_nat_cast, ennreal.to_real_nat, eq_self_iff_true, @@ -483,17 +483,17 @@ begin end lemma norm_eq_of_L2 {β : ι → Type*} [Π i, seminormed_add_comm_group (β i)] (x : pi_Lp 2 β) : - ∥x∥ = sqrt (∑ (i : ι), ∥x i∥ ^ 2) := + ‖x‖ = sqrt (∑ (i : ι), ‖x i‖ ^ 2) := by { convert norm_eq_of_nat 2 (by norm_cast) _, rw sqrt_eq_rpow, norm_cast } lemma nnnorm_eq_of_L2 {β : ι → Type*} [Π i, seminormed_add_comm_group (β i)] (x : pi_Lp 2 β) : - ∥x∥₊ = nnreal.sqrt (∑ (i : ι), ∥x i∥₊ ^ 2) := + ‖x‖₊ = nnreal.sqrt (∑ (i : ι), ‖x i‖₊ ^ 2) := subtype.ext $ by { push_cast, exact norm_eq_of_L2 x } lemma norm_sq_eq_of_L2 (β : ι → Type*) [Π i, seminormed_add_comm_group (β i)] (x : pi_Lp 2 β) : - ∥x∥ ^ 2 = ∑ (i : ι), ∥x i∥ ^ 2 := + ‖x‖ ^ 2 = ∑ (i : ι), ‖x i‖ ^ 2 := begin - suffices : ∥x∥₊ ^ 2 = ∑ (i : ι), ∥x i∥₊ ^ 2, + suffices : ‖x‖₊ ^ 2 = ∑ (i : ι), ‖x i‖₊ ^ 2, { simpa only [nnreal.coe_sum] using congr_arg (coe : ℝ≥0 → ℝ) this }, rw [nnnorm_eq_of_L2, nnreal.sq_sqrt], end @@ -519,7 +519,7 @@ instance normed_space [Π i, seminormed_add_comm_group (β i)] begin unfreezingI { rcases p.dichotomy with (rfl | hp) }, { letI : module 𝕜 (pi_Lp ∞ β) := pi.module ι β 𝕜, - suffices : ∥c • f∥₊ = ∥c∥₊ * ∥f∥₊, { exact_mod_cast nnreal.coe_mono this.le }, + suffices : ‖c • f‖₊ = ‖c‖₊ * ‖f‖₊, { exact_mod_cast nnreal.coe_mono this.le }, simpa only [nnnorm_eq_csupr, nnreal.mul_supr, ←nnnorm_smul] }, { have : p.to_real * (1 / p.to_real) = 1 := mul_div_cancel' 1 (zero_lt_one.trans_le hp).ne', simp only [norm_eq_sum (zero_lt_one.trans_le hp), norm_smul, mul_rpow, norm_nonneg, @@ -553,9 +553,9 @@ def equivₗᵢ : pi_Lp ∞ β ≃ₗᵢ[𝕜] Π i, β i := map_smul' := λ c f, rfl, norm_map' := λ f, begin - suffices : finset.univ.sup (λ i, ∥f i∥₊) = ⨆ i, ∥f i∥₊, + suffices : finset.univ.sup (λ i, ‖f i‖₊) = ⨆ i, ‖f i‖₊, { simpa only [nnreal.coe_supr] using congr_arg (coe : ℝ≥0 → ℝ) this }, - refine antisymm (finset.sup_le (λ i _, le_csupr (fintype.bdd_above_range (λ i, ∥f i∥₊)) _)) _, + refine antisymm (finset.sup_le (λ i _, le_csupr (fintype.bdd_above_range (λ i, ‖f i‖₊)) _)) _, casesI is_empty_or_nonempty ι, { simp only [csupr_of_empty, finset.univ_eq_empty, finset.sup_empty], }, { exact csupr_le (λ i, finset.le_sup (finset.mem_univ i)) }, @@ -624,12 +624,12 @@ end (pi_Lp.equiv p β).symm (c • x') = c • (pi_Lp.equiv p β).symm x' := rfl /-- When `p = ∞`, this lemma does not hold without the additional assumption `nonempty ι` because -the left-hand side simplifies to `0`, while the right-hand side simplifies to `∥b∥₊`. See +the left-hand side simplifies to `0`, while the right-hand side simplifies to `‖b‖₊`. See `pi_Lp.nnnorm_equiv_symm_const'` for a version which exchanges the hypothesis `p ≠ ∞` for `nonempty ι`. -/ lemma nnnorm_equiv_symm_const {β} [seminormed_add_comm_group β] (hp : p ≠ ∞) (b : β) : - ∥(pi_Lp.equiv p (λ _ : ι, β)).symm (function.const _ b)∥₊= - fintype.card ι ^ (1 / p).to_real * ∥b∥₊ := + ‖(pi_Lp.equiv p (λ _ : ι, β)).symm (function.const _ b)‖₊= + fintype.card ι ^ (1 / p).to_real * ‖b‖₊ := begin rcases p.dichotomy with (h | h), { exact false.elim (hp h) }, @@ -640,12 +640,12 @@ begin end /-- When `is_empty ι`, this lemma does not hold without the additional assumption `p ≠ ∞` because -the left-hand side simplifies to `0`, while the right-hand side simplifies to `∥b∥₊`. See +the left-hand side simplifies to `0`, while the right-hand side simplifies to `‖b‖₊`. See `pi_Lp.nnnorm_equiv_symm_const` for a version which exchanges the hypothesis `nonempty ι`. for `p ≠ ∞`. -/ lemma nnnorm_equiv_symm_const' {β} [seminormed_add_comm_group β] [nonempty ι] (b : β) : - ∥(pi_Lp.equiv p (λ _ : ι, β)).symm (function.const _ b)∥₊= - fintype.card ι ^ (1 / p).to_real * ∥b∥₊ := + ‖(pi_Lp.equiv p (λ _ : ι, β)).symm (function.const _ b)‖₊= + fintype.card ι ^ (1 / p).to_real * ‖b‖₊ := begin unfreezingI { rcases (em $ p = ∞) with (rfl | hp) }, { simp only [equiv_symm_apply, ennreal.div_top, ennreal.zero_to_real, nnreal.rpow_zero, one_mul, @@ -654,29 +654,29 @@ begin end /-- When `p = ∞`, this lemma does not hold without the additional assumption `nonempty ι` because -the left-hand side simplifies to `0`, while the right-hand side simplifies to `∥b∥₊`. See +the left-hand side simplifies to `0`, while the right-hand side simplifies to `‖b‖₊`. See `pi_Lp.norm_equiv_symm_const'` for a version which exchanges the hypothesis `p ≠ ∞` for `nonempty ι`. -/ lemma norm_equiv_symm_const {β} [seminormed_add_comm_group β] (hp : p ≠ ∞) (b : β) : - ∥(pi_Lp.equiv p (λ _ : ι, β)).symm (function.const _ b)∥ = - fintype.card ι ^ (1 / p).to_real * ∥b∥ := + ‖(pi_Lp.equiv p (λ _ : ι, β)).symm (function.const _ b)‖ = + fintype.card ι ^ (1 / p).to_real * ‖b‖ := (congr_arg coe $ nnnorm_equiv_symm_const hp b).trans $ by simp /-- When `is_empty ι`, this lemma does not hold without the additional assumption `p ≠ ∞` because -the left-hand side simplifies to `0`, while the right-hand side simplifies to `∥b∥₊`. See +the left-hand side simplifies to `0`, while the right-hand side simplifies to `‖b‖₊`. See `pi_Lp.norm_equiv_symm_const` for a version which exchanges the hypothesis `nonempty ι`. for `p ≠ ∞`. -/ lemma norm_equiv_symm_const' {β} [seminormed_add_comm_group β] [nonempty ι] (b : β) : - ∥(pi_Lp.equiv p (λ _ : ι, β)).symm (function.const _ b)∥ = - fintype.card ι ^ (1 / p).to_real * ∥b∥ := + ‖(pi_Lp.equiv p (λ _ : ι, β)).symm (function.const _ b)‖ = + fintype.card ι ^ (1 / p).to_real * ‖b‖ := (congr_arg coe $ nnnorm_equiv_symm_const' b).trans $ by simp lemma nnnorm_equiv_symm_one {β} [seminormed_add_comm_group β] (hp : p ≠ ∞) [has_one β] : - ∥(pi_Lp.equiv p (λ _ : ι, β)).symm 1∥₊ = fintype.card ι ^ (1 / p).to_real * ∥(1 : β)∥₊ := + ‖(pi_Lp.equiv p (λ _ : ι, β)).symm 1‖₊ = fintype.card ι ^ (1 / p).to_real * ‖(1 : β)‖₊ := (nnnorm_equiv_symm_const hp (1 : β)).trans rfl lemma norm_equiv_symm_one {β} [seminormed_add_comm_group β] (hp : p ≠ ∞) [has_one β] : - ∥(pi_Lp.equiv p (λ _ : ι, β)).symm 1∥ = fintype.card ι ^ (1 / p).to_real * ∥(1 : β)∥ := + ‖(pi_Lp.equiv p (λ _ : ι, β)).symm 1‖ = fintype.card ι ^ (1 / p).to_real * ‖(1 : β)‖ := (norm_equiv_symm_const hp (1 : β)).trans rfl variables (𝕜 p) diff --git a/src/analysis/normed_space/pointwise.lean b/src/analysis/normed_space/pointwise.lean index 0b2f1bde4a4fc..c28ea4ad6cb4b 100644 --- a/src/analysis/normed_space/pointwise.lean +++ b/src/analysis/normed_space/pointwise.lean @@ -24,7 +24,7 @@ section seminormed_add_comm_group variables [seminormed_add_comm_group E] [normed_space 𝕜 E] theorem smul_ball {c : 𝕜} (hc : c ≠ 0) (x : E) (r : ℝ) : - c • ball x r = ball (c • x) (∥c∥ * r) := + c • ball x r = ball (c • x) (‖c‖ * r) := begin ext y, rw mem_smul_set_iff_inv_smul_mem₀ hc, @@ -32,11 +32,11 @@ begin simp [← div_eq_inv_mul, div_lt_iff (norm_pos_iff.2 hc), mul_comm _ r, dist_smul], end -lemma smul_unit_ball {c : 𝕜} (hc : c ≠ 0) : c • ball (0 : E) (1 : ℝ) = ball (0 : E) (∥c∥) := +lemma smul_unit_ball {c : 𝕜} (hc : c ≠ 0) : c • ball (0 : E) (1 : ℝ) = ball (0 : E) (‖c‖) := by rw [smul_ball hc, smul_zero, mul_one] theorem smul_sphere' {c : 𝕜} (hc : c ≠ 0) (x : E) (r : ℝ) : - c • sphere x r = sphere (c • x) (∥c∥ * r) := + c • sphere x r = sphere (c • x) (‖c‖ * r) := begin ext y, rw mem_smul_set_iff_inv_smul_mem₀ hc, @@ -46,17 +46,17 @@ begin end theorem smul_closed_ball' {c : 𝕜} (hc : c ≠ 0) (x : E) (r : ℝ) : - c • closed_ball x r = closed_ball (c • x) (∥c∥ * r) := + c • closed_ball x r = closed_ball (c • x) (‖c‖ * r) := by simp only [← ball_union_sphere, set.smul_set_union, smul_ball hc, smul_sphere' hc] lemma metric.bounded.smul {s : set E} (hs : bounded s) (c : 𝕜) : bounded (c • s) := begin - obtain ⟨R, hR⟩ : ∃ (R : ℝ), ∀ x ∈ s, ∥x∥ ≤ R := hs.exists_norm_le, - refine bounded_iff_forall_norm_le.2 ⟨∥c∥ * R, λ z hz, _⟩, + obtain ⟨R, hR⟩ : ∃ (R : ℝ), ∀ x ∈ s, ‖x‖ ≤ R := hs.exists_norm_le, + refine bounded_iff_forall_norm_le.2 ⟨‖c‖ * R, λ z hz, _⟩, obtain ⟨y, ys, rfl⟩ : ∃ (y : E), y ∈ s ∧ c • y = z := mem_smul_set.1 hz, - calc ∥c • y∥ = ∥c∥ * ∥y∥ : norm_smul _ _ - ... ≤ ∥c∥ * R : mul_le_mul_of_nonneg_left (hR y ys) (norm_nonneg _) + calc ‖c • y‖ = ‖c‖ * ‖y‖ : norm_smul _ _ + ... ≤ ‖c‖ * R : mul_le_mul_of_nonneg_left (hR y ys) (norm_nonneg _) end /-- If `s` is a bounded set, then for small enough `r`, the set `{x} + r • s` is contained in any @@ -74,8 +74,8 @@ begin simp only [image_add_left, singleton_add], assume y hy, obtain ⟨z, zs, hz⟩ : ∃ (z : E), z ∈ s ∧ r • z = -x + y, by simpa [mem_smul_set] using hy, - have I : ∥r • z∥ ≤ ε := calc - ∥r • z∥ = ∥r∥ * ∥z∥ : norm_smul _ _ + have I : ‖r • z‖ ≤ ε := calc + ‖r • z‖ = ‖r‖ * ‖z‖ : norm_smul _ _ ... ≤ (ε / R) * R : mul_le_mul (mem_closed_ball_zero_iff.1 hr) (mem_closed_ball_zero_iff.1 (hR zs)) (norm_nonneg _) (div_pos εpos Rpos).le @@ -312,14 +312,14 @@ section normed_add_comm_group variables [normed_add_comm_group E] [normed_space 𝕜 E] theorem smul_closed_ball (c : 𝕜) (x : E) {r : ℝ} (hr : 0 ≤ r) : - c • closed_ball x r = closed_ball (c • x) (∥c∥ * r) := + c • closed_ball x r = closed_ball (c • x) (‖c‖ * r) := begin rcases eq_or_ne c 0 with rfl|hc, { simp [hr, zero_smul_set, set.singleton_zero, ← nonempty_closed_ball] }, { exact smul_closed_ball' hc x r } end -lemma smul_closed_unit_ball (c : 𝕜) : c • closed_ball (0 : E) (1 : ℝ) = closed_ball (0 : E) (∥c∥) := +lemma smul_closed_unit_ball (c : 𝕜) : c • closed_ball (0 : E) (1 : ℝ) = closed_ball (0 : E) (‖c‖) := by rw [smul_closed_ball _ _ zero_le_one, smul_zero, mul_one] variables [normed_space ℝ E] @@ -337,13 +337,13 @@ nonnegative. -/ begin obtain ⟨y, hy⟩ := exists_ne x, refine ⟨λ h, nonempty_closed_ball.1 (h.mono sphere_subset_closed_ball), λ hr, - ⟨r • ∥y - x∥⁻¹ • (y - x) + x, _⟩⟩, - have : ∥y - x∥ ≠ 0, by simpa [sub_eq_zero], + ⟨r • ‖y - x‖⁻¹ • (y - x) + x, _⟩⟩, + have : ‖y - x‖ ≠ 0, by simpa [sub_eq_zero], simp [norm_smul, this, real.norm_of_nonneg hr], end lemma smul_sphere [nontrivial E] (c : 𝕜) (x : E) {r : ℝ} (hr : 0 ≤ r) : - c • sphere x r = sphere (c • x) (∥c∥ * r) := + c • sphere x r = sphere (c • x) (‖c‖ * r) := begin rcases eq_or_ne c 0 with rfl|hc, { simp [zero_smul_set, set.singleton_zero, hr] }, diff --git a/src/analysis/normed_space/ray.lean b/src/analysis/normed_space/ray.lean index 338d77836a30d..eb01ab23957e8 100644 --- a/src/analysis/normed_space/ray.lean +++ b/src/analysis/normed_space/ray.lean @@ -11,7 +11,7 @@ import analysis.normed_space.basic In this file we prove some lemmas about the `same_ray` predicate in case of a real normed space. In this case, for two vectors `x y` in the same ray, the norm of their sum is equal to the sum of their -norms and `∥y∥ • x = ∥x∥ • y`. +norms and `‖y‖ • x = ‖x‖ • y`. -/ open real @@ -26,14 +26,14 @@ variables {x y : E} /-- If `x` and `y` are on the same ray, then the triangle inequality becomes the equality: the norm of `x + y` is the sum of the norms of `x` and `y`. The converse is true for a strictly convex space. -/ -lemma norm_add (h : same_ray ℝ x y) : ∥x + y∥ = ∥x∥ + ∥y∥ := +lemma norm_add (h : same_ray ℝ x y) : ‖x + y‖ = ‖x‖ + ‖y‖ := begin rcases h.exists_eq_smul with ⟨u, a, b, ha, hb, -, rfl, rfl⟩, rw [← add_smul, norm_smul_of_nonneg (add_nonneg ha hb), norm_smul_of_nonneg ha, norm_smul_of_nonneg hb, add_mul] end -lemma norm_sub (h : same_ray ℝ x y) : ∥x - y∥ = |∥x∥ - ∥y∥| := +lemma norm_sub (h : same_ray ℝ x y) : ‖x - y‖ = |‖x‖ - ‖y‖| := begin rcases h.exists_eq_smul with ⟨u, a, b, ha, hb, -, rfl, rfl⟩, wlog hab : b ≤ a := le_total b a using [a b, b a] tactic.skip, @@ -44,10 +44,10 @@ begin rw [norm_sub_rev, this hb ha hab.symm, abs_sub_comm] } end -lemma norm_smul_eq (h : same_ray ℝ x y) : ∥x∥ • y = ∥y∥ • x := +lemma norm_smul_eq (h : same_ray ℝ x y) : ‖x‖ • y = ‖y‖ • x := begin rcases h.exists_eq_smul with ⟨u, a, b, ha, hb, -, rfl, rfl⟩, - simp only [norm_smul_of_nonneg, *, mul_smul, smul_comm (∥u∥)], + simp only [norm_smul_of_nonneg, *, mul_smul, smul_comm (‖u‖)], apply smul_comm end @@ -68,22 +68,22 @@ end lemma norm_inj_on_ray_right (hy : y ≠ 0) : {x | same_ray ℝ x y}.inj_on norm := by simpa only [same_ray_comm] using norm_inj_on_ray_left hy -lemma same_ray_iff_norm_smul_eq : same_ray ℝ x y ↔ ∥x∥ • y = ∥y∥ • x := +lemma same_ray_iff_norm_smul_eq : same_ray ℝ x y ↔ ‖x‖ • y = ‖y‖ • x := ⟨same_ray.norm_smul_eq, λ h, or_iff_not_imp_left.2 $ λ hx, or_iff_not_imp_left.2 $ λ hy, - ⟨∥y∥, ∥x∥, norm_pos_iff.2 hy, norm_pos_iff.2 hx, h.symm⟩⟩ + ⟨‖y‖, ‖x‖, norm_pos_iff.2 hy, norm_pos_iff.2 hx, h.symm⟩⟩ /-- Two nonzero vectors `x y` in a real normed space are on the same ray if and only if the unit -vectors `∥x∥⁻¹ • x` and `∥y∥⁻¹ • y` are equal. -/ +vectors `‖x‖⁻¹ • x` and `‖y‖⁻¹ • y` are equal. -/ lemma same_ray_iff_inv_norm_smul_eq_of_ne (hx : x ≠ 0) (hy : y ≠ 0) : - same_ray ℝ x y ↔ ∥x∥⁻¹ • x = ∥y∥⁻¹ • y := + same_ray ℝ x y ↔ ‖x‖⁻¹ • x = ‖y‖⁻¹ • y := by rw [inv_smul_eq_iff₀, smul_comm, eq_comm, inv_smul_eq_iff₀, same_ray_iff_norm_smul_eq]; rwa norm_ne_zero_iff alias same_ray_iff_inv_norm_smul_eq_of_ne ↔ same_ray.inv_norm_smul_eq _ /-- Two vectors `x y` in a real normed space are on the ray if and only if one of them is zero or -the unit vectors `∥x∥⁻¹ • x` and `∥y∥⁻¹ • y` are equal. -/ -lemma same_ray_iff_inv_norm_smul_eq : same_ray ℝ x y ↔ x = 0 ∨ y = 0 ∨ ∥x∥⁻¹ • x = ∥y∥⁻¹ • y := +the unit vectors `‖x‖⁻¹ • x` and `‖y‖⁻¹ • y` are equal. -/ +lemma same_ray_iff_inv_norm_smul_eq : same_ray ℝ x y ↔ x = 0 ∨ y = 0 ∨ ‖x‖⁻¹ • x = ‖y‖⁻¹ • y := begin rcases eq_or_ne x 0 with rfl|hx, { simp [same_ray.zero_left] }, rcases eq_or_ne y 0 with rfl|hy, { simp [same_ray.zero_right] }, @@ -91,7 +91,7 @@ begin end /-- Two vectors of the same norm are on the same ray if and only if they are equal. -/ -lemma same_ray_iff_of_norm_eq (h : ∥x∥ = ∥y∥) : same_ray ℝ x y ↔ x = y := +lemma same_ray_iff_of_norm_eq (h : ‖x‖ = ‖y‖) : same_ray ℝ x y ↔ x = y := begin obtain rfl | hy := eq_or_ne y 0, { rw [norm_zero, norm_eq_zero] at h, @@ -99,13 +99,13 @@ begin { exact ⟨λ hxy, norm_inj_on_ray_right hy hxy same_ray.rfl h, λ hxy, hxy ▸ same_ray.rfl⟩ } end -lemma not_same_ray_iff_of_norm_eq (h : ∥x∥ = ∥y∥) : ¬ same_ray ℝ x y ↔ x ≠ y := +lemma not_same_ray_iff_of_norm_eq (h : ‖x‖ = ‖y‖) : ¬ same_ray ℝ x y ↔ x ≠ y := (same_ray_iff_of_norm_eq h).not /-- If two points on the same ray have the same norm, then they are equal. -/ -lemma same_ray.eq_of_norm_eq (h : same_ray ℝ x y) (hn : ∥x∥ = ∥y∥) : x = y := +lemma same_ray.eq_of_norm_eq (h : same_ray ℝ x y) (hn : ‖x‖ = ‖y‖) : x = y := (same_ray_iff_of_norm_eq hn).mp h /-- The norms of two vectors on the same ray are equal if and only if they are equal. -/ -lemma same_ray.norm_eq_iff (h : same_ray ℝ x y) : ∥x∥ = ∥y∥ ↔ x = y := +lemma same_ray.norm_eq_iff (h : same_ray ℝ x y) : ‖x‖ = ‖y‖ ↔ x = y := ⟨h.eq_of_norm_eq, λ h, h ▸ rfl⟩ diff --git a/src/analysis/normed_space/riesz_lemma.lean b/src/analysis/normed_space/riesz_lemma.lean index ac4c101c7b26d..7625ee42665ad 100644 --- a/src/analysis/normed_space/riesz_lemma.lean +++ b/src/analysis/normed_space/riesz_lemma.lean @@ -10,11 +10,11 @@ import topology.metric_space.hausdorff_distance # Applications of the Hausdorff distance in normed spaces Riesz's lemma, stated for a normed space over a normed field: for any -closed proper subspace `F` of `E`, there is a nonzero `x` such that `∥x - F∥` -is at least `r * ∥x∥` for any `r < 1`. This is `riesz_lemma`. +closed proper subspace `F` of `E`, there is a nonzero `x` such that `‖x - F‖` +is at least `r * ‖x‖` for any `r < 1`. This is `riesz_lemma`. -In a nontrivially normed field (with an element `c` of norm `> 1`) and any `R > ∥c∥`, one can -guarantee `∥x∥ ≤ R` and `∥x - y∥ ≥ 1` for any `y` in `F`. This is `riesz_lemma_of_norm_lt`. +In a nontrivially normed field (with an element `c` of norm `> 1`) and any `R > ‖c‖`, one can +guarantee `‖x‖ ≤ R` and `‖x - y‖ ≥ 1` for any `y` in `F`. This is `riesz_lemma_of_norm_lt`. A further lemma, `metric.closed_ball_inf_dist_compl_subset_closure`, finds a *closed* ball within the closure of a set `s` of optimal distance from a point in `x` to the frontier of `s`. @@ -35,7 +35,7 @@ is not guaranteed. For a variant giving an element with norm in `[1, R]`, see `riesz_lemma_of_norm_lt`. -/ lemma riesz_lemma {F : subspace 𝕜 E} (hFc : is_closed (F : set E)) (hF : ∃ x : E, x ∉ F) {r : ℝ} (hr : r < 1) : - ∃ x₀ : E, x₀ ∉ F ∧ ∀ y ∈ F, r * ∥x₀∥ ≤ ∥x₀ - y∥ := + ∃ x₀ : E, x₀ ∉ F ∧ ∀ y ∈ F, r * ‖x₀‖ ≤ ‖x₀ - y‖ := begin classical, obtain ⟨x, hx⟩ : ∃ x : E, x ∉ F := hF, @@ -57,10 +57,10 @@ begin refine ⟨x - y₀, x_ne_y₀, λy hy, le_of_lt _⟩, have hy₀y : y₀ + y ∈ F, from F.add_mem hy₀F hy, calc - r * ∥x - y₀∥ ≤ r' * ∥x - y₀∥ : mul_le_mul_of_nonneg_right (le_max_left _ _) (norm_nonneg _) + r * ‖x - y₀‖ ≤ r' * ‖x - y₀‖ : mul_le_mul_of_nonneg_right (le_max_left _ _) (norm_nonneg _) ... < d : by { rw ←dist_eq_norm, exact (lt_div_iff' hlt).1 hxy₀ } ... ≤ dist x (y₀ + y) : metric.inf_dist_le_dist_of_mem hy₀y - ... = ∥x - y₀ - y∥ : by { rw [sub_sub, dist_eq_norm] } + ... = ‖x - y₀ - y‖ : by { rw [sub_sub, dist_eq_norm] } end /-- @@ -71,31 +71,31 @@ strictly larger than the norm of an element of norm `> 1`. For a version without Since we are considering a general nontrivially normed field, there may be a gap in possible norms (for instance no element of norm in `(1,2)`). Hence, we can not allow `R` arbitrarily close to `1`, -and require `R > ∥c∥` for some `c : 𝕜` with norm `> 1`. +and require `R > ‖c‖` for some `c : 𝕜` with norm `> 1`. -/ lemma riesz_lemma_of_norm_lt - {c : 𝕜} (hc : 1 < ∥c∥) {R : ℝ} (hR : ∥c∥ < R) + {c : 𝕜} (hc : 1 < ‖c‖) {R : ℝ} (hR : ‖c‖ < R) {F : subspace 𝕜 E} (hFc : is_closed (F : set E)) (hF : ∃ x : E, x ∉ F) : - ∃ x₀ : E, ∥x₀∥ ≤ R ∧ ∀ y ∈ F, 1 ≤ ∥x₀ - y∥ := + ∃ x₀ : E, ‖x₀‖ ≤ R ∧ ∀ y ∈ F, 1 ≤ ‖x₀ - y‖ := begin have Rpos : 0 < R := (norm_nonneg _).trans_lt hR, - have : ∥c∥ / R < 1, by { rw div_lt_iff Rpos, simpa using hR }, + have : ‖c‖ / R < 1, by { rw div_lt_iff Rpos, simpa using hR }, rcases riesz_lemma hFc hF this with ⟨x, xF, hx⟩, have x0 : x ≠ 0 := λ H, by simpa [H] using xF, obtain ⟨d, d0, dxlt, ledx, -⟩ : - ∃ (d : 𝕜), d ≠ 0 ∧ ∥d • x∥ < R ∧ R / ∥c∥ ≤ ∥d • x∥ ∧ ∥d∥⁻¹ ≤ R⁻¹ * ∥c∥ * ∥x∥ := + ∃ (d : 𝕜), d ≠ 0 ∧ ‖d • x‖ < R ∧ R / ‖c‖ ≤ ‖d • x‖ ∧ ‖d‖⁻¹ ≤ R⁻¹ * ‖c‖ * ‖x‖ := rescale_to_shell hc Rpos x0, refine ⟨d • x, dxlt.le, λ y hy, _⟩, set y' := d⁻¹ • y with hy', have y'F : y' ∈ F, by simp [hy', submodule.smul_mem _ _ hy], have yy' : y = d • y', by simp [hy', smul_smul, mul_inv_cancel d0], - calc 1 = (∥c∥/R) * (R/∥c∥) : by field_simp [Rpos.ne', (zero_lt_one.trans hc).ne'] - ... ≤ (∥c∥/R) * (∥d • x∥) : + calc 1 = (‖c‖/R) * (R/‖c‖) : by field_simp [Rpos.ne', (zero_lt_one.trans hc).ne'] + ... ≤ (‖c‖/R) * (‖d • x‖) : mul_le_mul_of_nonneg_left ledx (div_nonneg (norm_nonneg _) Rpos.le) - ... = ∥d∥ * (∥c∥/R * ∥x∥) : by { simp [norm_smul], ring } - ... ≤ ∥d∥ * ∥x - y'∥ : + ... = ‖d‖ * (‖c‖/R * ‖x‖) : by { simp [norm_smul], ring } + ... ≤ ‖d‖ * ‖x - y'‖ : mul_le_mul_of_nonneg_left (hx y' (by simp [hy', submodule.smul_mem _ _ hy])) (norm_nonneg _) - ... = ∥d • x - y∥ : by simp [yy', ← smul_sub, norm_smul], + ... = ‖d • x - y‖ : by simp [yy', ← smul_sub, norm_smul], end lemma metric.closed_ball_inf_dist_compl_subset_closure {x : F} {s : set F} (hx : x ∈ s) : diff --git a/src/analysis/normed_space/spectrum.lean b/src/analysis/normed_space/spectrum.lean index b95cba3adcd7b..a7b03dc2b6df0 100644 --- a/src/analysis/normed_space/spectrum.lean +++ b/src/analysis/normed_space/spectrum.lean @@ -17,7 +17,7 @@ This file contains the basic theory for the resolvent and spectrum of a Banach a ## Main definitions -* `spectral_radius : ℝ≥0∞`: supremum of `∥k∥₊` for all `k ∈ spectrum 𝕜 a` +* `spectral_radius : ℝ≥0∞`: supremum of `‖k‖₊` for all `k ∈ spectrum 𝕜 a` * `normed_ring.alg_equiv_complex_of_complete`: **Gelfand-Mazur theorem** For a complex Banach division algebra, the natural `algebra_map ℂ A` is an algebra isomorphism whose inverse is given by selecting the (unique) element of `spectrum ℂ a` @@ -44,14 +44,14 @@ This file contains the basic theory for the resolvent and spectrum of a Banach a open_locale ennreal nnreal -/-- The *spectral radius* is the supremum of the `nnnorm` (`∥⬝∥₊`) of elements in the spectrum, +/-- The *spectral radius* is the supremum of the `nnnorm` (`‖⬝‖₊`) of elements in the spectrum, coerced into an element of `ℝ≥0∞`. Note that it is possible for `spectrum 𝕜 a = ∅`. In this case, `spectral_radius a = 0`. It is also possible that `spectrum 𝕜 a` be unbounded (though not for Banach algebras, see `spectrum.is_bounded`, below). In this case, `spectral_radius a = ∞`. -/ noncomputable def spectral_radius (𝕜 : Type*) {A : Type*} [normed_field 𝕜] [ring A] [algebra 𝕜 A] (a : A) : ℝ≥0∞ := -⨆ k ∈ spectrum 𝕜 a, ∥k∥₊ +⨆ k ∈ spectrum 𝕜 a, ‖k‖₊ variables {𝕜 : Type*} {A : Type*} @@ -74,7 +74,7 @@ by simp [spectral_radius] @[simp] lemma spectral_radius_zero : spectral_radius 𝕜 (0 : A) = 0 := by { nontriviality A, simp [spectral_radius] } -lemma mem_resolvent_set_of_spectral_radius_lt {a : A} {k : 𝕜} (h : spectral_radius 𝕜 a < ∥k∥₊) : +lemma mem_resolvent_set_of_spectral_radius_lt {a : A} {k : 𝕜} (h : spectral_radius 𝕜 a < ‖k‖₊) : k ∈ ρ a := not_not.mp $ λ hn, h.not_le $ le_supr₂ k hn @@ -86,7 +86,7 @@ units.is_open.preimage ((continuous_algebra_map 𝕜 A).sub continuous_const) protected lemma is_closed (a : A) : is_closed (σ a) := (is_open_resolvent_set a).is_closed_compl -lemma mem_resolvent_set_of_norm_lt_mul {a : A} {k : 𝕜} (h : ∥a∥ * ∥(1 : A)∥ < ∥k∥) : +lemma mem_resolvent_set_of_norm_lt_mul {a : A} {k : 𝕜} (h : ‖a‖ * ‖(1 : A)‖ < ‖k‖) : k ∈ ρ a := begin rw [resolvent_set, set.mem_set_of_eq, algebra.algebra_map_eq_smul_one], @@ -94,51 +94,51 @@ begin have hk : k ≠ 0, from ne_zero_of_norm_ne_zero ((mul_nonneg (norm_nonneg _) (norm_nonneg _)).trans_lt h).ne', let ku := units.map (↑ₐ).to_monoid_hom (units.mk0 k hk), - rw [←inv_inv (∥(1 : A)∥), mul_inv_lt_iff (inv_pos.2 $ norm_pos_iff.2 (one_ne_zero : (1 : A) ≠ 0))] + rw [←inv_inv (‖(1 : A)‖), mul_inv_lt_iff (inv_pos.2 $ norm_pos_iff.2 (one_ne_zero : (1 : A) ≠ 0))] at h, - have hku : ∥-a∥ < ∥(↑ku⁻¹:A)∥⁻¹ := by simpa [ku, norm_algebra_map] using h, + have hku : ‖-a‖ < ‖(↑ku⁻¹:A)‖⁻¹ := by simpa [ku, norm_algebra_map] using h, simpa [ku, sub_eq_add_neg, algebra.algebra_map_eq_smul_one] using (ku.add (-a) hku).is_unit, end -lemma mem_resolvent_set_of_norm_lt [norm_one_class A] {a : A} {k : 𝕜} (h : ∥a∥ < ∥k∥) : +lemma mem_resolvent_set_of_norm_lt [norm_one_class A] {a : A} {k : 𝕜} (h : ‖a‖ < ‖k‖) : k ∈ ρ a := mem_resolvent_set_of_norm_lt_mul (by rwa [norm_one, mul_one]) lemma norm_le_norm_mul_of_mem {a : A} {k : 𝕜} (hk : k ∈ σ a) : - ∥k∥ ≤ ∥a∥ * ∥(1 : A)∥ := + ‖k‖ ≤ ‖a‖ * ‖(1 : A)‖ := le_of_not_lt $ mt mem_resolvent_set_of_norm_lt_mul hk lemma norm_le_norm_of_mem [norm_one_class A] {a : A} {k : 𝕜} (hk : k ∈ σ a) : - ∥k∥ ≤ ∥a∥ := + ‖k‖ ≤ ‖a‖ := le_of_not_lt $ mt mem_resolvent_set_of_norm_lt hk lemma subset_closed_ball_norm_mul (a : A) : - σ a ⊆ metric.closed_ball (0 : 𝕜) (∥a∥ * ∥(1 : A)∥) := + σ a ⊆ metric.closed_ball (0 : 𝕜) (‖a‖ * ‖(1 : A)‖) := λ k hk, by simp [norm_le_norm_mul_of_mem hk] lemma subset_closed_ball_norm [norm_one_class A] (a : A) : - σ a ⊆ metric.closed_ball (0 : 𝕜) (∥a∥) := + σ a ⊆ metric.closed_ball (0 : 𝕜) (‖a‖) := λ k hk, by simp [norm_le_norm_of_mem hk] lemma is_bounded (a : A) : metric.bounded (σ a) := -(metric.bounded_iff_subset_ball 0).mpr ⟨∥a∥ * ∥(1 : A)∥, subset_closed_ball_norm_mul a⟩ +(metric.bounded_iff_subset_ball 0).mpr ⟨‖a‖ * ‖(1 : A)‖, subset_closed_ball_norm_mul a⟩ protected theorem is_compact [proper_space 𝕜] (a : A) : is_compact (σ a) := metric.is_compact_of_is_closed_bounded (spectrum.is_closed a) (is_bounded a) theorem spectral_radius_le_nnnorm [norm_one_class A] (a : A) : - spectral_radius 𝕜 a ≤ ∥a∥₊ := + spectral_radius 𝕜 a ≤ ‖a‖₊ := by { refine supr₂_le (λ k hk, _), exact_mod_cast norm_le_norm_of_mem hk } lemma exists_nnnorm_eq_spectral_radius_of_nonempty [proper_space 𝕜] {a : A} (ha : (σ a).nonempty) : - ∃ k ∈ σ a, (∥k∥₊ : ℝ≥0∞) = spectral_radius 𝕜 a := + ∃ k ∈ σ a, (‖k‖₊ : ℝ≥0∞) = spectral_radius 𝕜 a := begin obtain ⟨k, hk, h⟩ := (spectrum.is_compact a).exists_forall_ge ha continuous_nnnorm.continuous_on, exact ⟨k, hk, le_antisymm (le_supr₂ k hk) (supr₂_le $ by exact_mod_cast h)⟩, end lemma spectral_radius_lt_of_forall_lt_of_nonempty [proper_space 𝕜] {a : A} - (ha : (σ a).nonempty) {r : ℝ≥0} (hr : ∀ k ∈ σ a, ∥k∥₊ < r) : + (ha : (σ a).nonempty) {r : ℝ≥0} (hr : ∀ k ∈ σ a, ‖k‖₊ < r) : spectral_radius 𝕜 a < r := Sup_image.symm.trans_lt $ ((spectrum.is_compact a).Sup_lt_iff_of_continuous ha (ennreal.continuous_coe.comp continuous_nnnorm).continuous_on (r : ℝ≥0∞)).mpr @@ -148,7 +148,7 @@ open ennreal polynomial variable (𝕜) theorem spectral_radius_le_pow_nnnorm_pow_one_div (a : A) (n : ℕ) : - spectral_radius 𝕜 a ≤ (∥a ^ (n + 1)∥₊) ^ (1 / (n + 1) : ℝ) * (∥(1 : A)∥₊) ^ (1 / (n + 1) : ℝ) := + spectral_radius 𝕜 a ≤ (‖a ^ (n + 1)‖₊) ^ (1 / (n + 1) : ℝ) * (‖(1 : A)‖₊) ^ (1 / (n + 1) : ℝ) := begin refine supr₂_le (λ k hk, _), /- apply easy direction of the spectral mapping theorem for polynomials -/ @@ -156,7 +156,7 @@ begin by simpa only [one_mul, algebra.algebra_map_eq_smul_one, one_smul, aeval_monomial, one_mul, eval_monomial] using subset_polynomial_aeval a (monomial (n + 1) (1 : 𝕜)) ⟨k, hk, rfl⟩, /- power of the norm is bounded by norm of the power -/ - have nnnorm_pow_le : (↑(∥k∥₊ ^ (n + 1)) : ℝ≥0∞) ≤ ∥a ^ (n + 1)∥₊ * ∥(1 : A)∥₊, + have nnnorm_pow_le : (↑(‖k‖₊ ^ (n + 1)) : ℝ≥0∞) ≤ ‖a ^ (n + 1)‖₊ * ‖(1 : A)‖₊, { simpa only [real.to_nnreal_mul (norm_nonneg _), norm_to_nnreal, nnnorm_pow k (n + 1), ennreal.coe_mul] using coe_mono (real.to_nnreal_mono (norm_le_norm_mul_of_mem pow_mem)) }, /- take (n + 1)ᵗʰ roots and clean up the left-hand side -/ @@ -167,7 +167,7 @@ begin end theorem spectral_radius_le_liminf_pow_nnnorm_pow_one_div (a : A) : - spectral_radius 𝕜 a ≤ at_top.liminf (λ n : ℕ, (∥a ^ n∥₊ : ℝ≥0∞) ^ (1 / n : ℝ)) := + spectral_radius 𝕜 a ≤ at_top.liminf (λ n : ℕ, (‖a ^ n‖₊ : ℝ≥0∞) ^ (1 / n : ℝ)) := begin refine ennreal.le_of_forall_lt_one_mul_le (λ ε hε, _), by_cases ε = 0, @@ -178,7 +178,7 @@ begin liminf_eq_supr_infi_of_nat', ennreal.supr_mul, ennreal.infi_mul hε'], rw [←ennreal.inv_lt_inv, inv_one] at hε, obtain ⟨N, hN⟩ := eventually_at_top.mp - (ennreal.eventually_pow_one_div_le (ennreal.coe_ne_top : ↑∥(1 : A)∥₊ ≠ ∞) hε), + (ennreal.eventually_pow_one_div_le (ennreal.coe_ne_top : ↑‖(1 : A)‖₊ ≠ ∞) hε), refine (le_trans _ (le_supr _ (N + 1))), refine le_infi (λ n, _), simp only [←add_assoc], @@ -211,28 +211,28 @@ end version of this, for example: `tendsto (resolvent a) (cobounded 𝕜) (𝓝 0)` or more specifically `(resolvent a) =O[cobounded 𝕜] (λ z, z⁻¹)`. -/ lemma norm_resolvent_le_forall (a : A) : - ∀ ε > 0, ∃ R > 0, ∀ z : 𝕜, R ≤ ∥z∥ → ∥resolvent a z∥ ≤ ε := + ∀ ε > 0, ∃ R > 0, ∀ z : 𝕜, R ≤ ‖z‖ → ‖resolvent a z‖ ≤ ε := begin obtain ⟨c, c_pos, hc⟩ := (@normed_ring.inverse_one_sub_norm A _ _).exists_pos, rw [is_O_with_iff, eventually_iff, metric.mem_nhds_iff] at hc, rcases hc with ⟨δ, δ_pos, hδ⟩, simp only [cstar_ring.norm_one, mul_one] at hδ, intros ε hε, - have ha₁ : 0 < ∥a∥ + 1 := lt_of_le_of_lt (norm_nonneg a) (lt_add_one _), - have min_pos : 0 < min (δ * (∥a∥ + 1)⁻¹) (ε * c⁻¹), + have ha₁ : 0 < ‖a‖ + 1 := lt_of_le_of_lt (norm_nonneg a) (lt_add_one _), + have min_pos : 0 < min (δ * (‖a‖ + 1)⁻¹) (ε * c⁻¹), from lt_min (mul_pos δ_pos (inv_pos.mpr ha₁)) (mul_pos hε (inv_pos.mpr c_pos)), - refine ⟨(min (δ * (∥a∥ + 1)⁻¹) (ε * c⁻¹))⁻¹, inv_pos.mpr min_pos, (λ z hz, _)⟩, + refine ⟨(min (δ * (‖a‖ + 1)⁻¹) (ε * c⁻¹))⁻¹, inv_pos.mpr min_pos, (λ z hz, _)⟩, have hnz : z ≠ 0 := norm_pos_iff.mp (lt_of_lt_of_le (inv_pos.mpr min_pos) hz), replace hz := inv_le_of_inv_le min_pos hz, rcases (⟨units.mk0 z hnz, units.coe_mk0 hnz⟩ : is_unit z) with ⟨z, rfl⟩, - have lt_δ : ∥z⁻¹ • a∥ < δ, + have lt_δ : ‖z⁻¹ • a‖ < δ, { rw [units.smul_def, norm_smul, units.coe_inv, norm_inv], - calc ∥(z : 𝕜)∥⁻¹ * ∥a∥ ≤ δ * (∥a∥ + 1)⁻¹ * ∥a∥ + calc ‖(z : 𝕜)‖⁻¹ * ‖a‖ ≤ δ * (‖a‖ + 1)⁻¹ * ‖a‖ : mul_le_mul_of_nonneg_right (hz.trans (min_le_left _ _)) (norm_nonneg _) ... < δ : by { conv { rw mul_assoc, to_rhs, rw (mul_one δ).symm }, exact mul_lt_mul_of_pos_left - ((inv_mul_lt_iff ha₁).mpr ((mul_one (∥a∥ + 1)).symm ▸ (lt_add_one _))) δ_pos } }, + ((inv_mul_lt_iff ha₁).mpr ((mul_one (‖a‖ + 1)).symm ▸ (lt_add_one _))) δ_pos } }, rw [←inv_smul_smul z (resolvent a (z : 𝕜)), units_smul_resolvent_self, resolvent, algebra.algebra_map_eq_smul_one, one_smul, units.smul_def, norm_smul, units.coe_inv, norm_inv], calc _ ≤ ε * c⁻¹ * c : mul_le_mul (hz.trans (min_le_right _ _)) (hδ (mem_ball_zero_iff.mpr lt_δ)) @@ -253,19 +253,19 @@ variables variable (𝕜) /-- In a Banach algebra `A` over a nontrivially normed field `𝕜`, for any `a : A` the power series with coefficients `a ^ n` represents the function `(1 - z • a)⁻¹` in a disk of -radius `∥a∥₊⁻¹`. -/ +radius `‖a‖₊⁻¹`. -/ lemma has_fpower_series_on_ball_inverse_one_sub_smul [complete_space A] (a : A) : has_fpower_series_on_ball (λ z : 𝕜, ring.inverse (1 - z • a)) - (λ n, continuous_multilinear_map.mk_pi_field 𝕜 (fin n) (a ^ n)) 0 (∥a∥₊)⁻¹ := + (λ n, continuous_multilinear_map.mk_pi_field 𝕜 (fin n) (a ^ n)) 0 (‖a‖₊)⁻¹ := { r_le := begin - refine le_of_forall_nnreal_lt (λ r hr, le_radius_of_bound_nnreal _ (max 1 ∥(1 : A)∥₊) (λ n, _)), + refine le_of_forall_nnreal_lt (λ r hr, le_radius_of_bound_nnreal _ (max 1 ‖(1 : A)‖₊) (λ n, _)), rw [←norm_to_nnreal, norm_mk_pi_field, norm_to_nnreal], cases n, { simp only [le_refl, mul_one, or_true, le_max_iff, pow_zero] }, { refine le_trans (le_trans (mul_le_mul_right' (nnnorm_pow_le' a n.succ_pos) (r ^ n.succ)) _) (le_max_left _ _), - { by_cases ∥a∥₊ = 0, + { by_cases ‖a‖₊ = 0, { simp only [h, zero_mul, zero_le', pow_succ], }, { rw [←ennreal.coe_inv h, coe_lt_coe, nnreal.lt_inv_iff_mul_lt h] at hr, simpa only [←mul_pow, mul_comm] using pow_le_one' hr.le n.succ } } } @@ -273,10 +273,10 @@ lemma has_fpower_series_on_ball_inverse_one_sub_smul [complete_space A] (a : A) r_pos := ennreal.inv_pos.mpr coe_ne_top, has_sum := λ y hy, begin - have norm_lt : ∥y • a∥ < 1, - { by_cases h : ∥a∥₊ = 0, + have norm_lt : ‖y • a‖ < 1, + { by_cases h : ‖a‖₊ = 0, { simp only [nnnorm_eq_zero.mp h, norm_zero, zero_lt_one, smul_zero] }, - { have nnnorm_lt : ∥y∥₊ < ∥a∥₊⁻¹, + { have nnnorm_lt : ‖y‖₊ < ‖a‖₊⁻¹, { simpa only [←ennreal.coe_inv h, mem_ball_zero_iff, metric.emetric_ball_nnreal] using hy }, rwa [←coe_nnnorm, ←real.lt_to_nnreal_iff_coe_lt, real.to_nnreal_one, nnnorm_smul, ←nnreal.lt_inv_iff_mul_lt h] } }, @@ -285,7 +285,7 @@ lemma has_fpower_series_on_ball_inverse_one_sub_smul [complete_space A] (a : A) end } variable {𝕜} -lemma is_unit_one_sub_smul_of_lt_inv_radius {a : A} {z : 𝕜} (h : ↑∥z∥₊ < (spectral_radius 𝕜 a)⁻¹) : +lemma is_unit_one_sub_smul_of_lt_inv_radius {a : A} {z : 𝕜} (h : ↑‖z‖₊ < (spectral_radius 𝕜 a)⁻¹) : is_unit (1 - z • a) := begin by_cases hz : z = 0, @@ -327,7 +327,7 @@ variables /-- The `limsup` relationship for the spectral radius used to prove `spectrum.gelfand_formula`. -/ lemma limsup_pow_nnnorm_pow_one_div_le_spectral_radius (a : A) : - limsup (λ n : ℕ, ↑∥a ^ n∥₊ ^ (1 / n : ℝ)) at_top ≤ spectral_radius ℂ a := + limsup (λ n : ℕ, ↑‖a ^ n‖₊ ^ (1 / n : ℝ)) at_top ≤ spectral_radius ℂ a := begin refine ennreal.inv_le_inv.mp (le_of_forall_pos_nnreal_lt (λ r r_pos r_lt, _)), simp_rw [inv_limsup, ←one_div], @@ -338,25 +338,25 @@ begin simp only [p.radius_eq_liminf, ←norm_to_nnreal, norm_mk_pi_field], congr, ext n, - rw [norm_to_nnreal, ennreal.coe_rpow_def (∥a ^ n∥₊) (1 / n : ℝ), if_neg], + rw [norm_to_nnreal, ennreal.coe_rpow_def (‖a ^ n‖₊) (1 / n : ℝ), if_neg], exact λ ha, by linarith [ha.2, (one_div_nonneg.mpr n.cast_nonneg : 0 ≤ (1 / n : ℝ))], }, { have H₁ := (differentiable_on_inverse_one_sub_smul r_lt).has_fpower_series_on_ball r_pos, exact ((has_fpower_series_on_ball_inverse_one_sub_smul ℂ a).exchange_radius H₁).r_le, } end /-- **Gelfand's formula**: Given an element `a : A` of a complex Banach algebra, the -`spectral_radius` of `a` is the limit of the sequence `∥a ^ n∥₊ ^ (1 / n)` -/ +`spectral_radius` of `a` is the limit of the sequence `‖a ^ n‖₊ ^ (1 / n)` -/ theorem pow_nnnorm_pow_one_div_tendsto_nhds_spectral_radius (a : A) : - tendsto (λ n : ℕ, ((∥a ^ n∥₊ ^ (1 / n : ℝ)) : ℝ≥0∞)) at_top (𝓝 (spectral_radius ℂ a)) := + tendsto (λ n : ℕ, ((‖a ^ n‖₊ ^ (1 / n : ℝ)) : ℝ≥0∞)) at_top (𝓝 (spectral_radius ℂ a)) := tendsto_of_le_liminf_of_limsup_le (spectral_radius_le_liminf_pow_nnnorm_pow_one_div ℂ a) (limsup_pow_nnnorm_pow_one_div_le_spectral_radius a) /- This is the same as `pow_nnnorm_pow_one_div_tendsto_nhds_spectral_radius` but for `norm` instead of `nnnorm`. -/ /-- **Gelfand's formula**: Given an element `a : A` of a complex Banach algebra, the -`spectral_radius` of `a` is the limit of the sequence `∥a ^ n∥₊ ^ (1 / n)` -/ +`spectral_radius` of `a` is the limit of the sequence `‖a ^ n‖₊ ^ (1 / n)` -/ theorem pow_norm_pow_one_div_tendsto_nhds_spectral_radius (a : A) : - tendsto (λ n : ℕ, ennreal.of_real (∥a ^ n∥ ^ (1 / n : ℝ))) at_top (𝓝 (spectral_radius ℂ a)) := + tendsto (λ n : ℕ, ennreal.of_real (‖a ^ n‖ ^ (1 / n : ℝ))) at_top (𝓝 (spectral_radius ℂ a)) := begin convert pow_nnnorm_pow_one_div_tendsto_nhds_spectral_radius a, ext1, @@ -391,7 +391,7 @@ begin H₁.continuous.continuous_on with ⟨C, hC⟩, use max C 1, rintros _ ⟨w, rfl⟩, - refine or.elim (em (∥w∥ ≤ R)) (λ hw, _) (λ hw, _), + refine or.elim (em (‖w‖ ≤ R)) (λ hw, _) (λ hw, _), { exact (hC w (mem_closed_ball_zero_iff.mpr hw)).trans (le_max_left _ _) }, { exact (hR w (not_le.mp hw).le).trans (le_max_right _ _), }, }, /- `resolvent a 0 = 0`, which is a contradition because it isn't a unit. -/ @@ -407,12 +407,12 @@ end /-- In a complex Banach algebra, the spectral radius is always attained by some element of the spectrum. -/ -lemma exists_nnnorm_eq_spectral_radius : ∃ z ∈ spectrum ℂ a, (∥z∥₊ : ℝ≥0∞) = spectral_radius ℂ a := +lemma exists_nnnorm_eq_spectral_radius : ∃ z ∈ spectrum ℂ a, (‖z‖₊ : ℝ≥0∞) = spectral_radius ℂ a := exists_nnnorm_eq_spectral_radius_of_nonempty (spectrum.nonempty a) /-- In a complex Banach algebra, if every element of the spectrum has norm strictly less than `r : ℝ≥0`, then the spectral radius is also strictly less than `r`. -/ -lemma spectral_radius_lt_of_forall_lt {r : ℝ≥0} (hr : ∀ z ∈ spectrum ℂ a, ∥z∥₊ < r) : +lemma spectral_radius_lt_of_forall_lt {r : ℝ≥0} (hr : ∀ z ∈ spectrum ℂ a, ‖z‖₊ < r) : spectral_radius ℂ a < r := spectral_radius_lt_of_forall_lt_of_nonempty (spectrum.nonempty a) hr @@ -445,7 +445,7 @@ by rwa [mem_iff, hA, not_not, sub_eq_zero] at h is an algebra isomorphism whose inverse is given by selecting the (unique) element of `spectrum ℂ a`. In addition, `algebra_map_isometry` guarantees this map is an isometry. -Note: because `normed_division_ring` requires the field `norm_mul' : ∀ a b, ∥a * b∥ = ∥a∥ * ∥b∥`, we +Note: because `normed_division_ring` requires the field `norm_mul' : ∀ a b, ‖a * b‖ = ‖a‖ * ‖b‖`, we don't use this type class and instead opt for a `normed_ring` in which the nonzero elements are precisely the units. This allows for the application of this isomorphism in broader contexts, e.g., to the quotient of a complex Banach algebra by a maximal ideal. In the case when `A` is actually a @@ -476,7 +476,7 @@ begin sub_add_cancel] }, let b := ∑' n : ℕ, ((n + 1).factorial⁻¹ : 𝕜) • (a - ↑ₐz) ^ n, have hb : summable (λ n : ℕ, ((n + 1).factorial⁻¹ : 𝕜) • (a - ↑ₐz) ^ n), - { refine summable_of_norm_bounded_eventually _ (real.summable_pow_div_factorial ∥a - ↑ₐz∥) _, + { refine summable_of_norm_bounded_eventually _ (real.summable_pow_div_factorial ‖a - ↑ₐz‖) _, filter_upwards [filter.eventually_cofinite_ne 0] with n hn, rw [norm_smul, mul_comm, norm_inv, is_R_or_C.norm_eq_abs, is_R_or_C.abs_cast_nat, ←div_eq_mul_inv], @@ -512,8 +512,8 @@ local notation `↑ₐ` := algebra_map 𝕜 A automatically bounded). See note [lower instance priority] -/ @[priority 100] instance [alg_hom_class F 𝕜 A 𝕜] : continuous_linear_map_class F 𝕜 A 𝕜 := -{ map_continuous := λ φ, add_monoid_hom_class.continuous_of_bound φ ∥(1 : A)∥ $ - λ a, (mul_comm ∥a∥ ∥(1 : A)∥) ▸ spectrum.norm_le_norm_mul_of_mem (apply_mem_spectrum φ _), +{ map_continuous := λ φ, add_monoid_hom_class.continuous_of_bound φ ‖(1 : A)‖ $ + λ a, (mul_comm ‖a‖ ‖(1 : A)‖) ▸ spectrum.norm_le_norm_mul_of_mem (apply_mem_spectrum φ _), .. alg_hom_class.linear_map_class } /-- An algebra homomorphism into the base field, as a continuous linear map (since it is @@ -525,10 +525,10 @@ def to_continuous_linear_map (φ : A →ₐ[𝕜] 𝕜) : A →L[𝕜] 𝕜 := ⇑φ.to_continuous_linear_map = φ := rfl lemma norm_apply_le_self_mul_norm_one [alg_hom_class F 𝕜 A 𝕜] (f : F) (a : A) : - ∥f a∥ ≤ ∥a∥ * ∥(1 : A)∥ := + ‖f a‖ ≤ ‖a‖ * ‖(1 : A)‖ := spectrum.norm_le_norm_mul_of_mem (apply_mem_spectrum f _) -lemma norm_apply_le_self [norm_one_class A] [alg_hom_class F 𝕜 A 𝕜] (f : F) (a : A) : ∥f a∥ ≤ ∥a∥ := +lemma norm_apply_le_self [norm_one_class A] [alg_hom_class F 𝕜 A 𝕜] (f : F) (a : A) : ‖f a‖ ≤ ‖a‖ := spectrum.norm_le_norm_of_mem (apply_mem_spectrum f _) end normed_field @@ -538,9 +538,9 @@ variables [nontrivially_normed_field 𝕜] [normed_ring A] [normed_algebra 𝕜 local notation `↑ₐ` := algebra_map 𝕜 A @[simp] lemma to_continuous_linear_map_norm [norm_one_class A] (φ : A →ₐ[𝕜] 𝕜) : - ∥φ.to_continuous_linear_map∥ = 1 := + ‖φ.to_continuous_linear_map‖ = 1 := continuous_linear_map.op_norm_eq_of_bounds zero_le_one - (λ a, (one_mul ∥a∥).symm ▸ spectrum.norm_le_norm_of_mem (apply_mem_spectrum φ _)) + (λ a, (one_mul ‖a‖).symm ▸ spectrum.norm_le_norm_of_mem (apply_mem_spectrum φ _)) (λ _ _ h, by simpa only [coe_to_continuous_linear_map, map_one, norm_one, mul_one] using h 1) end nontrivially_normed_field diff --git a/src/analysis/normed_space/star/basic.lean b/src/analysis/normed_space/star/basic.lean index bbebcd299b61a..f504ce482ba0e 100644 --- a/src/analysis/normed_space/star/basic.lean +++ b/src/analysis/normed_space/star/basic.lean @@ -17,7 +17,7 @@ import algebra.star.unitary A normed star group is a normed group with a compatible `star` which is isometric. A C⋆-ring is a normed star group that is also a ring and that verifies the stronger -condition `∥x⋆ * x∥ = ∥x∥^2` for all `x`. If a C⋆-ring is also a star algebra, then it is a +condition `‖x⋆ * x‖ = ‖x‖^2` for all `x`. If a C⋆-ring is also a star algebra, then it is a C⋆-algebra. To get a C⋆-algebra `E` over field `𝕜`, use @@ -26,7 +26,7 @@ To get a C⋆-algebra `E` over field `𝕜`, use ## TODO -- Show that `∥x⋆ * x∥ = ∥x∥^2` is equivalent to `∥x⋆ * x∥ = ∥x⋆∥ * ∥x∥`, which is used as the +- Show that `‖x⋆ * x‖ = ‖x‖^2` is equivalent to `‖x⋆ * x‖ = ‖x⋆‖ * ‖x‖`, which is used as the definition of C*-algebras in some sources (e.g. Wikipedia). -/ @@ -37,7 +37,7 @@ local postfix `⋆`:std.prec.max_plus := star /-- A normed star group is a normed group with a compatible `star` which is isometric. -/ class normed_star_group (E : Type*) [seminormed_add_comm_group E] [star_add_monoid E] : Prop := -(norm_star : ∀ x : E, ∥x⋆∥ = ∥x∥) +(norm_star : ∀ x : E, ‖x⋆‖ = ‖x‖) export normed_star_group (norm_star) attribute [simp] norm_star @@ -47,7 +47,7 @@ variables {𝕜 E α : Type*} section normed_star_group variables [seminormed_add_comm_group E] [star_add_monoid E] [normed_star_group E] -@[simp] lemma nnnorm_star (x : E) : ∥star x∥₊ = ∥x∥₊ := subtype.ext $ norm_star _ +@[simp] lemma nnnorm_star (x : E) : ‖star x‖₊ = ‖x‖₊ := subtype.ext $ norm_star _ /-- The `star` map in a normed star group is a normed group homomorphism. -/ def star_normed_add_group_hom : normed_add_group_hom E E := @@ -58,7 +58,7 @@ def star_normed_add_group_hom : normed_add_group_hom E E := lemma star_isometry : isometry (star : E → E) := show isometry star_add_equiv, by exact add_monoid_hom_class.isometry_of_norm star_add_equiv - (show ∀ x, ∥x⋆∥ = ∥x∥, from norm_star) + (show ∀ x, ‖x⋆‖ = ‖x‖, from norm_star) @[priority 100] instance normed_star_group.to_has_continuous_star : has_continuous_star E := @@ -70,10 +70,10 @@ instance ring_hom_isometric.star_ring_end [normed_comm_ring E] [star_ring E] [normed_star_group E] : ring_hom_isometric (star_ring_end E) := ⟨norm_star⟩ -/-- A C*-ring is a normed star ring that satifies the stronger condition `∥x⋆ * x∥ = ∥x∥^2` +/-- A C*-ring is a normed star ring that satifies the stronger condition `‖x⋆ * x‖ = ‖x‖^2` for every `x`. -/ class cstar_ring (E : Type*) [non_unital_normed_ring E] [star_ring E] : Prop := -(norm_star_mul_self : ∀ {x : E}, ∥x⋆ * x∥ = ∥x∥ * ∥x∥) +(norm_star_mul_self : ∀ {x : E}, ‖x⋆ * x‖ = ‖x‖ * ‖x‖) instance : cstar_ring ℝ := { norm_star_mul_self := λ x, by simp only [star, id.def, norm_mul] } @@ -90,28 +90,28 @@ instance to_normed_star_group : normed_star_group E := intro x, by_cases htriv : x = 0, { simp only [htriv, star_zero] }, - { have hnt : 0 < ∥x∥ := norm_pos_iff.mpr htriv, - have hnt_star : 0 < ∥x⋆∥ := + { have hnt : 0 < ‖x‖ := norm_pos_iff.mpr htriv, + have hnt_star : 0 < ‖x⋆‖ := norm_pos_iff.mpr ((add_equiv.map_ne_zero_iff star_add_equiv).mpr htriv), have h₁ := calc - ∥x∥ * ∥x∥ = ∥x⋆ * x∥ : norm_star_mul_self.symm - ... ≤ ∥x⋆∥ * ∥x∥ : norm_mul_le _ _, + ‖x‖ * ‖x‖ = ‖x⋆ * x‖ : norm_star_mul_self.symm + ... ≤ ‖x⋆‖ * ‖x‖ : norm_mul_le _ _, have h₂ := calc - ∥x⋆∥ * ∥x⋆∥ = ∥x * x⋆∥ : by rw [←norm_star_mul_self, star_star] - ... ≤ ∥x∥ * ∥x⋆∥ : norm_mul_le _ _, + ‖x⋆‖ * ‖x⋆‖ = ‖x * x⋆‖ : by rw [←norm_star_mul_self, star_star] + ... ≤ ‖x‖ * ‖x⋆‖ : norm_mul_le _ _, exact le_antisymm (le_of_mul_le_mul_right h₂ hnt_star) (le_of_mul_le_mul_right h₁ hnt) }, end⟩ -lemma norm_self_mul_star {x : E} : ∥x * x⋆∥ = ∥x∥ * ∥x∥ := +lemma norm_self_mul_star {x : E} : ‖x * x⋆‖ = ‖x‖ * ‖x‖ := by { nth_rewrite 0 [←star_star x], simp only [norm_star_mul_self, norm_star] } -lemma norm_star_mul_self' {x : E} : ∥x⋆ * x∥ = ∥x⋆∥ * ∥x∥ := +lemma norm_star_mul_self' {x : E} : ‖x⋆ * x‖ = ‖x⋆‖ * ‖x‖ := by rw [norm_star_mul_self, norm_star] -lemma nnnorm_self_mul_star {x : E} : ∥x * star x∥₊ = ∥x∥₊ * ∥x∥₊ := +lemma nnnorm_self_mul_star {x : E} : ‖x * star x‖₊ = ‖x‖₊ * ‖x‖₊ := subtype.ext norm_self_mul_star -lemma nnnorm_star_mul_self {x : E} : ∥x⋆ * x∥₊ = ∥x∥₊ * ∥x∥₊ := +lemma nnnorm_star_mul_self {x : E} : ‖x⋆ * x‖₊ = ‖x‖₊ * ‖x‖₊ := subtype.ext norm_star_mul_self @[simp] @@ -154,7 +154,7 @@ instance _root_.prod.cstar_ring : cstar_ring (R₁ × R₂) := exact (le_max_left _ _).trans (le_abs_self _), exact (le_max_right _ _).trans (le_abs_self _) }, { rw le_sup_iff, - rcases le_total (∥x.fst∥) (∥x.snd∥) with (h | h); + rcases le_total (‖x.fst‖) (‖x.snd‖) with (h | h); simp [h] } end } @@ -174,55 +174,55 @@ end prod_pi section unital variables [normed_ring E] [star_ring E] [cstar_ring E] -@[simp] lemma norm_one [nontrivial E] : ∥(1 : E)∥ = 1 := +@[simp] lemma norm_one [nontrivial E] : ‖(1 : E)‖ = 1 := begin - have : 0 < ∥(1 : E)∥ := norm_pos_iff.mpr one_ne_zero, + have : 0 < ‖(1 : E)‖ := norm_pos_iff.mpr one_ne_zero, rw [←mul_left_inj' this.ne', ←norm_star_mul_self, mul_one, star_one, one_mul], end @[priority 100] -- see Note [lower instance priority] instance [nontrivial E] : norm_one_class E := ⟨norm_one⟩ -lemma norm_coe_unitary [nontrivial E] (U : unitary E) : ∥(U : E)∥ = 1 := +lemma norm_coe_unitary [nontrivial E] (U : unitary E) : ‖(U : E)‖ = 1 := begin rw [←sq_eq_sq (norm_nonneg _) zero_le_one, one_pow 2, sq, ←cstar_ring.norm_star_mul_self, unitary.coe_star_mul_self, cstar_ring.norm_one], end -@[simp] lemma norm_of_mem_unitary [nontrivial E] {U : E} (hU : U ∈ unitary E) : ∥U∥ = 1 := +@[simp] lemma norm_of_mem_unitary [nontrivial E] {U : E} (hU : U ∈ unitary E) : ‖U‖ = 1 := norm_coe_unitary ⟨U, hU⟩ -@[simp] lemma norm_coe_unitary_mul (U : unitary E) (A : E) : ∥(U : E) * A∥ = ∥A∥ := +@[simp] lemma norm_coe_unitary_mul (U : unitary E) (A : E) : ‖(U : E) * A‖ = ‖A‖ := begin nontriviality E, refine le_antisymm _ _, - { calc _ ≤ ∥(U : E)∥ * ∥A∥ : norm_mul_le _ _ - ... = ∥A∥ : by rw [norm_coe_unitary, one_mul] }, - { calc _ = ∥(U : E)⋆ * U * A∥ : by rw [unitary.coe_star_mul_self U, one_mul] - ... ≤ ∥(U : E)⋆∥ * ∥(U : E) * A∥ : by { rw [mul_assoc], exact norm_mul_le _ _ } - ... = ∥(U : E) * A∥ : by rw [norm_star, norm_coe_unitary, one_mul] }, + { calc _ ≤ ‖(U : E)‖ * ‖A‖ : norm_mul_le _ _ + ... = ‖A‖ : by rw [norm_coe_unitary, one_mul] }, + { calc _ = ‖(U : E)⋆ * U * A‖ : by rw [unitary.coe_star_mul_self U, one_mul] + ... ≤ ‖(U : E)⋆‖ * ‖(U : E) * A‖ : by { rw [mul_assoc], exact norm_mul_le _ _ } + ... = ‖(U : E) * A‖ : by rw [norm_star, norm_coe_unitary, one_mul] }, end -@[simp] lemma norm_unitary_smul (U : unitary E) (A : E) : ∥U • A∥ = ∥A∥ := +@[simp] lemma norm_unitary_smul (U : unitary E) (A : E) : ‖U • A‖ = ‖A‖ := norm_coe_unitary_mul U A -lemma norm_mem_unitary_mul {U : E} (A : E) (hU : U ∈ unitary E) : ∥U * A∥ = ∥A∥ := +lemma norm_mem_unitary_mul {U : E} (A : E) (hU : U ∈ unitary E) : ‖U * A‖ = ‖A‖ := norm_coe_unitary_mul ⟨U, hU⟩ A -@[simp] lemma norm_mul_coe_unitary (A : E) (U : unitary E) : ∥A * U∥ = ∥A∥ := -calc _ = ∥((U : E)⋆ * A⋆)⋆∥ : by simp only [star_star, star_mul] - ... = ∥(U : E)⋆ * A⋆∥ : by rw [norm_star] - ... = ∥A⋆∥ : norm_mem_unitary_mul (star A) (unitary.star_mem U.prop) - ... = ∥A∥ : norm_star _ +@[simp] lemma norm_mul_coe_unitary (A : E) (U : unitary E) : ‖A * U‖ = ‖A‖ := +calc _ = ‖((U : E)⋆ * A⋆)⋆‖ : by simp only [star_star, star_mul] + ... = ‖(U : E)⋆ * A⋆‖ : by rw [norm_star] + ... = ‖A⋆‖ : norm_mem_unitary_mul (star A) (unitary.star_mem U.prop) + ... = ‖A‖ : norm_star _ -lemma norm_mul_mem_unitary (A : E) {U : E} (hU : U ∈ unitary E) : ∥A * U∥ = ∥A∥ := +lemma norm_mul_mem_unitary (A : E) {U : E} (hU : U ∈ unitary E) : ‖A * U‖ = ‖A‖ := norm_mul_coe_unitary A ⟨U, hU⟩ end unital end cstar_ring lemma is_self_adjoint.nnnorm_pow_two_pow [normed_ring E] [star_ring E] - [cstar_ring E] {x : E} (hx : is_self_adjoint x) (n : ℕ) : ∥x ^ 2 ^ n∥₊ = ∥x∥₊ ^ (2 ^ n) := + [cstar_ring E] {x : E} (hx : is_self_adjoint x) (n : ℕ) : ‖x ^ 2 ^ n‖₊ = ‖x‖₊ ^ (2 ^ n) := begin induction n with k hk, { simp only [pow_zero, pow_one] }, @@ -232,7 +232,7 @@ begin end lemma self_adjoint.nnnorm_pow_two_pow [normed_ring E] [star_ring E] [cstar_ring E] - (x : self_adjoint E) (n : ℕ) : ∥x ^ 2 ^ n∥₊ = ∥x∥₊ ^ (2 ^ n) := + (x : self_adjoint E) (n : ℕ) : ‖x ^ 2 ^ n‖₊ = ‖x‖₊ ^ (2 ^ n) := x.prop.nnnorm_pow_two_pow _ section starₗᵢ @@ -265,7 +265,7 @@ variables [normed_space 𝕜 E] [is_scalar_tower 𝕜 E E] [smul_comm_class 𝕜 /-- In a C⋆-algebra `E`, either unital or non-unital, multiplication on the left by `a : E` has norm equal to the norm of `a`. -/ -@[simp] lemma op_nnnorm_mul : ∥mul 𝕜 E a∥₊ = ∥a∥₊ := +@[simp] lemma op_nnnorm_mul : ‖mul 𝕜 E a‖₊ = ‖a‖₊ := begin rw ←Sup_closed_unit_ball_eq_nnnorm, refine cSup_eq_of_forall_le_of_forall_lt_exists_gt _ _ (λ r hr, _), @@ -273,20 +273,20 @@ begin { rintro - ⟨x, hx, rfl⟩, exact ((mul 𝕜 E a).unit_le_op_norm x $ mem_closed_ball_zero_iff.mp hx).trans (op_norm_mul_apply_le 𝕜 E a) }, - { have ha : 0 < ∥a∥₊ := zero_le'.trans_lt hr, - rw [←inv_inv (∥a∥₊), nnreal.lt_inv_iff_mul_lt (inv_ne_zero ha.ne')] at hr, + { have ha : 0 < ‖a‖₊ := zero_le'.trans_lt hr, + rw [←inv_inv (‖a‖₊), nnreal.lt_inv_iff_mul_lt (inv_ne_zero ha.ne')] at hr, obtain ⟨k, hk₁, hk₂⟩ := normed_field.exists_lt_nnnorm_lt 𝕜 (mul_lt_mul_of_pos_right hr $ nnreal.inv_pos.2 ha), refine ⟨_, ⟨k • star a, _, rfl⟩, _⟩, { simpa only [mem_closed_ball_zero_iff, norm_smul, one_mul, norm_star] using - (nnreal.le_inv_iff_mul_le ha.ne').1 (one_mul ∥a∥₊⁻¹ ▸ hk₂.le : ∥k∥₊ ≤ ∥a∥₊⁻¹) }, + (nnreal.le_inv_iff_mul_le ha.ne').1 (one_mul ‖a‖₊⁻¹ ▸ hk₂.le : ‖k‖₊ ≤ ‖a‖₊⁻¹) }, { simp only [map_smul, nnnorm_smul, mul_apply', mul_smul_comm, cstar_ring.nnnorm_self_mul_star], rwa [←nnreal.div_lt_iff (mul_pos ha ha).ne', div_eq_mul_inv, mul_inv, ←mul_assoc] } }, end /-- In a C⋆-algebra `E`, either unital or non-unital, multiplication on the right by `a : E` has norm eqaul to the norm of `a`. -/ -@[simp] lemma op_nnnorm_mul_flip : ∥(mul 𝕜 E).flip a∥₊ = ∥a∥₊ := +@[simp] lemma op_nnnorm_mul_flip : ‖(mul 𝕜 E).flip a‖₊ = ‖a‖₊ := begin rw [←Sup_unit_ball_eq_nnnorm, ←nnnorm_star, ←@op_nnnorm_mul 𝕜 E, ←Sup_unit_ball_eq_nnnorm], congr' 1, diff --git a/src/analysis/normed_space/star/matrix.lean b/src/analysis/normed_space/star/matrix.lean index 4f9f24d3b3585..241239a3b3976 100644 --- a/src/analysis/normed_space/star/matrix.lean +++ b/src/analysis/normed_space/star/matrix.lean @@ -22,10 +22,10 @@ section entrywise_sup_norm variables [is_R_or_C 𝕜] [fintype n] [decidable_eq n] lemma entry_norm_bound_of_unitary {U : matrix n n 𝕜} (hU : U ∈ matrix.unitary_group n 𝕜) (i j : n): - ∥U i j∥ ≤ 1 := + ‖U i j‖ ≤ 1 := begin -- The norm squared of an entry is at most the L2 norm of its row. - have norm_sum : ∥ U i j ∥^2 ≤ (∑ x, ∥ U i x ∥^2), + have norm_sum : ‖ U i j ‖^2 ≤ (∑ x, ‖ U i x ‖^2), { apply multiset.single_le_sum, { intros x h_x, rw multiset.mem_map at h_x, @@ -36,11 +36,11 @@ begin use j, simp only [eq_self_iff_true, finset.mem_univ_val, and_self, sq_eq_sq] } }, -- The L2 norm of a row is a diagonal entry of U ⬝ Uᴴ - have diag_eq_norm_sum : (U ⬝ Uᴴ) i i = ∑ (x : n), ∥ U i x ∥^2, + have diag_eq_norm_sum : (U ⬝ Uᴴ) i i = ∑ (x : n), ‖ U i x ‖^2, { simp only [matrix.mul_apply, matrix.conj_transpose_apply, ←star_ring_end_apply, is_R_or_C.mul_conj, is_R_or_C.norm_sq_eq_def', is_R_or_C.of_real_pow] }, -- The L2 norm of a row is a diagonal entry of U ⬝ Uᴴ, real part - have re_diag_eq_norm_sum : is_R_or_C.re ((U ⬝ Uᴴ) i i) = ∑ (x : n), ∥ U i x ∥^2, + have re_diag_eq_norm_sum : is_R_or_C.re ((U ⬝ Uᴴ) i i) = ∑ (x : n), ‖ U i x ‖^2, { rw is_R_or_C.ext_iff at diag_eq_norm_sum, rw diag_eq_norm_sum.1, norm_cast }, @@ -57,7 +57,7 @@ local attribute [instance] matrix.normed_add_comm_group /-- The entrywise sup norm of a unitary matrix is at most 1. -/ lemma entrywise_sup_norm_bound_of_unitary {U : matrix n n 𝕜} (hU : U ∈ matrix.unitary_group n 𝕜) : - ∥ U ∥ ≤ 1 := + ‖ U ‖ ≤ 1 := begin simp_rw pi_norm_le_iff_of_nonneg zero_le_one, intros i j, diff --git a/src/analysis/normed_space/star/spectrum.lean b/src/analysis/normed_space/star/spectrum.lean index 2e431d57795b5..55bd525acdbf2 100644 --- a/src/analysis/normed_space/star/spectrum.lean +++ b/src/analysis/normed_space/star/spectrum.lean @@ -36,7 +36,7 @@ begin { rw ←unitary.coe_to_units_apply u at hk, have hnk := ne_zero_of_mem_of_unit hk, rw [←inv_inv (unitary.to_units u), ←spectrum.map_inv, set.mem_inv] at hk, - have : ∥k∥⁻¹ ≤ ∥↑((unitary.to_units u)⁻¹)∥, simpa only [norm_inv] using norm_le_norm_of_mem hk, + have : ‖k‖⁻¹ ≤ ‖↑((unitary.to_units u)⁻¹)‖, simpa only [norm_inv] using norm_le_norm_of_mem hk, simpa using inv_le_of_inv_le (norm_pos_iff.mpr hnk) this } end @@ -57,9 +57,9 @@ local notation `↑ₐ` := algebra_map ℂ A lemma is_self_adjoint.spectral_radius_eq_nnnorm {a : A} (ha : is_self_adjoint a) : - spectral_radius ℂ a = ∥a∥₊ := + spectral_radius ℂ a = ‖a‖₊ := begin - have hconst : tendsto (λ n : ℕ, (∥a∥₊ : ℝ≥0∞)) at_top _ := tendsto_const_nhds, + have hconst : tendsto (λ n : ℕ, (‖a‖₊ : ℝ≥0∞)) at_top _ := tendsto_const_nhds, refine tendsto_nhds_unique _ hconst, convert (spectrum.pow_nnnorm_pow_one_div_tendsto_nhds_spectral_radius (a : A)).comp (nat.tendsto_pow_at_top_at_top_of_one_lt one_lt_two), @@ -70,11 +70,11 @@ begin end lemma is_star_normal.spectral_radius_eq_nnnorm (a : A) [is_star_normal a] : - spectral_radius ℂ a = ∥a∥₊ := + spectral_radius ℂ a = ‖a‖₊ := begin refine (ennreal.pow_strict_mono two_ne_zero).injective _, - have heq : (λ n : ℕ, ((∥(a⋆ * a) ^ n∥₊ ^ (1 / n : ℝ)) : ℝ≥0∞)) - = (λ x, x ^ 2) ∘ (λ n : ℕ, ((∥a ^ n∥₊ ^ (1 / n : ℝ)) : ℝ≥0∞)), + have heq : (λ n : ℕ, ((‖(a⋆ * a) ^ n‖₊ ^ (1 / n : ℝ)) : ℝ≥0∞)) + = (λ x, x ^ 2) ∘ (λ n : ℕ, ((‖a ^ n‖₊ ^ (1 / n : ℝ)) : ℝ≥0∞)), { funext, rw [function.comp_apply, ←rpow_nat_cast, ←rpow_mul, mul_comm, rpow_mul, rpow_nat_cast, ←coe_pow, sq, ←nnnorm_star_mul_self, commute.mul_pow (star_comm_self' a), star_pow], }, @@ -128,9 +128,9 @@ variables {F A B : Type*} include hF /-- A star algebra homomorphism of complex C⋆-algebras is norm contractive. -/ -lemma nnnorm_apply_le (a : A) : ∥(φ a : B)∥₊ ≤ ∥a∥₊ := +lemma nnnorm_apply_le (a : A) : ‖(φ a : B)‖₊ ≤ ‖a‖₊ := begin - suffices : ∀ s : A, is_self_adjoint s → ∥φ s∥₊ ≤ ∥s∥₊, + suffices : ∀ s : A, is_self_adjoint s → ‖φ s‖₊ ≤ ‖s‖₊, { exact nonneg_le_nonneg_of_sq_le_sq zero_le' (by simpa only [nnnorm_star_mul_self, map_star, map_mul] using this _ (is_self_adjoint.star_mul_self a)) }, @@ -141,7 +141,7 @@ begin end /-- A star algebra homomorphism of complex C⋆-algebras is norm contractive. -/ -lemma norm_apply_le (a : A) : ∥(φ a : B)∥ ≤ ∥a∥ := nnnorm_apply_le φ a +lemma norm_apply_le (a : A) : ‖(φ a : B)‖ ≤ ‖a‖ := nnnorm_apply_le φ a /-- Star algebra homomorphisms between C⋆-algebras are continuous linear maps. See note [lower instance priority] -/ diff --git a/src/analysis/normed_space/units.lean b/src/analysis/normed_space/units.lean index 1dd2104580a84..653c13667e416 100644 --- a/src/analysis/normed_space/units.lean +++ b/src/analysis/normed_space/units.lean @@ -36,32 +36,32 @@ namespace units /-- In a complete normed ring, a perturbation of `1` by an element `t` of distance less than `1` from `1` is a unit. Here we construct its `units` structure. -/ @[simps coe] -def one_sub (t : R) (h : ∥t∥ < 1) : Rˣ := +def one_sub (t : R) (h : ‖t‖ < 1) : Rˣ := { val := 1 - t, inv := ∑' n : ℕ, t ^ n, val_inv := mul_neg_geom_series t h, inv_val := geom_series_mul_neg t h } /-- In a complete normed ring, a perturbation of a unit `x` by an element `t` of distance less than -`∥x⁻¹∥⁻¹` from `x` is a unit. Here we construct its `units` structure. -/ +`‖x⁻¹‖⁻¹` from `x` is a unit. Here we construct its `units` structure. -/ @[simps coe] -def add (x : Rˣ) (t : R) (h : ∥t∥ < ∥(↑x⁻¹ : R)∥⁻¹) : Rˣ := +def add (x : Rˣ) (t : R) (h : ‖t‖ < ‖(↑x⁻¹ : R)‖⁻¹) : Rˣ := units.copy -- to make `coe_add` true definitionally, for convenience (x * (units.one_sub (-(↑x⁻¹ * t)) begin nontriviality R using [zero_lt_one], - have hpos : 0 < ∥(↑x⁻¹ : R)∥ := units.norm_pos x⁻¹, - calc ∥-(↑x⁻¹ * t)∥ - = ∥↑x⁻¹ * t∥ : by { rw norm_neg } - ... ≤ ∥(↑x⁻¹ : R)∥ * ∥t∥ : norm_mul_le ↑x⁻¹ _ - ... < ∥(↑x⁻¹ : R)∥ * ∥(↑x⁻¹ : R)∥⁻¹ : by nlinarith only [h, hpos] + have hpos : 0 < ‖(↑x⁻¹ : R)‖ := units.norm_pos x⁻¹, + calc ‖-(↑x⁻¹ * t)‖ + = ‖↑x⁻¹ * t‖ : by { rw norm_neg } + ... ≤ ‖(↑x⁻¹ : R)‖ * ‖t‖ : norm_mul_le ↑x⁻¹ _ + ... < ‖(↑x⁻¹ : R)‖ * ‖(↑x⁻¹ : R)‖⁻¹ : by nlinarith only [h, hpos] ... = 1 : mul_inv_cancel (ne_of_gt hpos) end)) (x + t) (by simp [mul_add]) _ rfl -/-- In a complete normed ring, an element `y` of distance less than `∥x⁻¹∥⁻¹` from `x` is a unit. +/-- In a complete normed ring, an element `y` of distance less than `‖x⁻¹‖⁻¹` from `x` is a unit. Here we construct its `units` structure. -/ @[simps coe] -def unit_of_nearby (x : Rˣ) (y : R) (h : ∥y - x∥ < ∥(↑x⁻¹ : R)∥⁻¹) : Rˣ := +def unit_of_nearby (x : Rˣ) (y : R) (h : ‖y - x‖ < ‖(↑x⁻¹ : R)‖⁻¹) : Rˣ := units.copy (x.add (y - x : R) h) y (by simp) _ rfl /-- The group of units of a complete normed ring is an open subset of the ring. -/ @@ -70,7 +70,7 @@ begin nontriviality R, apply metric.is_open_iff.mpr, rintros x' ⟨x, rfl⟩, - refine ⟨∥(↑x⁻¹ : R)∥⁻¹, _root_.inv_pos.mpr (units.norm_pos x⁻¹), _⟩, + refine ⟨‖(↑x⁻¹ : R)‖⁻¹, _root_.inv_pos.mpr (units.norm_pos x⁻¹), _⟩, intros y hy, rw [metric.mem_ball, dist_eq_norm] at hy, exact (x.unit_of_nearby y hy).is_unit @@ -98,7 +98,7 @@ namespace normed_ring open_locale classical big_operators open asymptotics filter metric finset ring -lemma inverse_one_sub (t : R) (h : ∥t∥ < 1) : inverse (1 - t) = ↑(units.one_sub t h)⁻¹ := +lemma inverse_one_sub (t : R) (h : ‖t‖ < 1) : inverse (1 - t) = ↑(units.one_sub t h)⁻¹ := by rw [← inverse_unit (units.one_sub t h), units.coe_one_sub] /-- The formula `inverse (x + t) = inverse (1 + x⁻¹ * t) * x⁻¹` holds for `t` sufficiently small. -/ @@ -107,11 +107,11 @@ lemma inverse_add (x : Rˣ) : begin nontriviality R, rw [eventually_iff, metric.mem_nhds_iff], - have hinv : 0 < ∥(↑x⁻¹ : R)∥⁻¹, by cancel_denoms, - use [∥(↑x⁻¹ : R)∥⁻¹, hinv], + have hinv : 0 < ‖(↑x⁻¹ : R)‖⁻¹, by cancel_denoms, + use [‖(↑x⁻¹ : R)‖⁻¹, hinv], intros t ht, simp only [mem_ball, dist_zero_right] at ht, - have ht' : ∥-↑x⁻¹ * t∥ < 1, + have ht' : ‖-↑x⁻¹ * t‖ < 1, { refine lt_of_le_of_lt (norm_mul_le _ _) _, rw norm_neg, refine lt_of_lt_of_le (mul_lt_mul_of_pos_left ht x⁻¹.norm_pos) _, @@ -168,16 +168,16 @@ end lemma inverse_one_sub_norm : (λ t : R, inverse (1 - t)) =O[𝓝 0] (λ t, 1 : R → ℝ) := begin simp only [is_O, is_O_with, eventually_iff, metric.mem_nhds_iff], - refine ⟨∥(1:R)∥ + 1, (2:ℝ)⁻¹, by norm_num, _⟩, + refine ⟨‖(1:R)‖ + 1, (2:ℝ)⁻¹, by norm_num, _⟩, intros t ht, simp only [ball, dist_zero_right, set.mem_set_of_eq] at ht, - have ht' : ∥t∥ < 1, + have ht' : ‖t‖ < 1, { have : (2:ℝ)⁻¹ < 1 := by cancel_denoms, linarith }, simp only [inverse_one_sub t ht', norm_one, mul_one, set.mem_set_of_eq], - change ∥∑' n : ℕ, t ^ n∥ ≤ _, + change ‖∑' n : ℕ, t ^ n‖ ≤ _, have := normed_ring.tsum_geometric_of_norm_lt_1 t ht', - have : (1 - ∥t∥)⁻¹ ≤ 2, + have : (1 - ‖t‖)⁻¹ ≤ 2, { rw ← inv_inv (2:ℝ), refine inv_le_inv_of_le (by norm_num) _, have : (2:ℝ)⁻¹ + (2:ℝ)⁻¹ = 1 := by ring, @@ -190,7 +190,7 @@ lemma inverse_add_norm (x : Rˣ) : (λ t : R, inverse (↑x + t)) =O[𝓝 0] (λ begin simp only [is_O_iff, norm_one, mul_one], cases is_O_iff.mp (@inverse_one_sub_norm R _ _) with C hC, - use C * ∥((x⁻¹:Rˣ):R)∥, + use C * ‖((x⁻¹:Rˣ):R)‖, have hzero : tendsto (λ t, - (↑x⁻¹ : R) * t) (𝓝 0) (𝓝 0), { convert ((mul_left_continuous (-↑x⁻¹ : R)).tendsto 0).comp tendsto_id, simp }, @@ -207,14 +207,14 @@ end is `O(t ^ n)` as `t → 0`. -/ lemma inverse_add_norm_diff_nth_order (x : Rˣ) (n : ℕ) : (λ t : R, inverse (↑x + t) - (∑ i in range n, (- ↑x⁻¹ * t) ^ i) * ↑x⁻¹) =O[𝓝 (0:R)] - (λ t, ∥t∥ ^ n) := + (λ t, ‖t‖ ^ n) := begin by_cases h : n = 0, { simpa [h] using inverse_add_norm x }, have hn : 0 < n := nat.pos_of_ne_zero h, simp [is_O_iff], cases (is_O_iff.mp (inverse_add_norm x)) with C hC, - use C * ∥(1:ℝ)∥ * ∥(↑x⁻¹ : R)∥ ^ n, + use C * ‖(1:ℝ)‖ * ‖(↑x⁻¹ : R)‖ ^ n, have h : eventually_eq (𝓝 (0:R)) (λ t, inverse (↑x + t) - (∑ i in range n, (- ↑x⁻¹ * t) ^ i) * ↑x⁻¹) (λ t, ((- ↑x⁻¹ * t) ^ n) * inverse (x + t)), @@ -227,13 +227,13 @@ begin simp only [neg_mul] at hLHS, rw hLHS, refine le_trans (norm_mul_le _ _ ) _, - have h' : ∥(-(↑x⁻¹ * t)) ^ n∥ ≤ ∥(↑x⁻¹ : R)∥ ^ n * ∥t∥ ^ n, - { calc ∥(-(↑x⁻¹ * t)) ^ n∥ ≤ ∥(-(↑x⁻¹ * t))∥ ^ n : norm_pow_le' _ hn - ... = ∥↑x⁻¹ * t∥ ^ n : by rw norm_neg - ... ≤ (∥(↑x⁻¹ : R)∥ * ∥t∥) ^ n : _ - ... = ∥(↑x⁻¹ : R)∥ ^ n * ∥t∥ ^ n : mul_pow _ _ n, + have h' : ‖(-(↑x⁻¹ * t)) ^ n‖ ≤ ‖(↑x⁻¹ : R)‖ ^ n * ‖t‖ ^ n, + { calc ‖(-(↑x⁻¹ * t)) ^ n‖ ≤ ‖(-(↑x⁻¹ * t))‖ ^ n : norm_pow_le' _ hn + ... = ‖↑x⁻¹ * t‖ ^ n : by rw norm_neg + ... ≤ (‖(↑x⁻¹ : R)‖ * ‖t‖) ^ n : _ + ... = ‖(↑x⁻¹ : R)‖ ^ n * ‖t‖ ^ n : mul_pow _ _ n, exact pow_le_pow_of_le_left (norm_nonneg _) (norm_mul_le ↑x⁻¹ t) n }, - have h'' : 0 ≤ ∥(↑x⁻¹ : R)∥ ^ n * ∥t∥ ^ n, + have h'' : 0 ≤ ‖(↑x⁻¹ : R)‖ ^ n * ‖t‖ ^ n, { refine mul_nonneg _ _; exact pow_nonneg (norm_nonneg _) n }, nlinarith [norm_nonneg (inverse (↑x + t))], @@ -241,14 +241,14 @@ end /-- The function `λ t, inverse (x + t) - x⁻¹` is `O(t)` as `t → 0`. -/ lemma inverse_add_norm_diff_first_order (x : Rˣ) : - (λ t : R, inverse (↑x + t) - ↑x⁻¹) =O[𝓝 0] (λ t, ∥t∥) := + (λ t : R, inverse (↑x + t) - ↑x⁻¹) =O[𝓝 0] (λ t, ‖t‖) := by simpa using inverse_add_norm_diff_nth_order x 1 /-- The function `λ t, inverse (x + t) - x⁻¹ + x⁻¹ * t * x⁻¹` is `O(t ^ 2)` as `t → 0`. -/ lemma inverse_add_norm_diff_second_order (x : Rˣ) : - (λ t : R, inverse (↑x + t) - ↑x⁻¹ + ↑x⁻¹ * t * ↑x⁻¹) =O[𝓝 0] (λ t, ∥t∥ ^ 2) := + (λ t : R, inverse (↑x + t) - ↑x⁻¹ + ↑x⁻¹ * t * ↑x⁻¹) =O[𝓝 0] (λ t, ‖t‖ ^ 2) := begin convert inverse_add_norm_diff_nth_order x 2, ext t, @@ -305,7 +305,7 @@ end units namespace ideal /-- An ideal which contains an element within `1` of `1 : R` is the unit ideal. -/ -lemma eq_top_of_norm_lt_one (I : ideal R) {x : R} (hxI : x ∈ I) (hx : ∥1 - x∥ < 1) : I = ⊤ := +lemma eq_top_of_norm_lt_one (I : ideal R) {x : R} (hxI : x ∈ I) (hx : ‖1 - x‖ < 1) : I = ⊤ := let u := units.one_sub (1 - x) hx in (I.eq_top_iff_one.mpr $ by simpa only [show u.inv * x = 1, by simp] using I.mul_mem_left u.inv hxI) diff --git a/src/analysis/normed_space/weak_dual.lean b/src/analysis/normed_space/weak_dual.lean index 03f09ac772803..f9f756bde4f73 100644 --- a/src/analysis/normed_space/weak_dual.lean +++ b/src/analysis/normed_space/weak_dual.lean @@ -164,7 +164,7 @@ variables (𝕜) weak-star topology is `weak_dual.polar 𝕜 s`. -/ def polar (s : set E) : set (weak_dual 𝕜 E) := to_normed_dual ⁻¹' polar 𝕜 s -lemma polar_def (s : set E) : polar 𝕜 s = {f : weak_dual 𝕜 E | ∀ x ∈ s, ∥f x∥ ≤ 1} := rfl +lemma polar_def (s : set E) : polar 𝕜 s = {f : weak_dual 𝕜 E | ∀ x ∈ s, ‖f x‖ ≤ 1} := rfl /-- The polar `polar 𝕜 s` of a set `s : E` is a closed subset when the weak star topology is used. -/ diff --git a/src/analysis/quaternion.lean b/src/analysis/quaternion.lean index 6ff353e009e1d..3e3000dc5c7fa 100644 --- a/src/analysis/quaternion.lean +++ b/src/analysis/quaternion.lean @@ -48,16 +48,16 @@ inner_product_space.of_core add_left := λ x y z, by simp only [inner_def, add_mul, add_re], smul_left := λ x y r, by simp [inner_def] } -lemma norm_sq_eq_norm_sq (a : ℍ) : norm_sq a = ∥a∥ * ∥a∥ := +lemma norm_sq_eq_norm_sq (a : ℍ) : norm_sq a = ‖a‖ * ‖a‖ := by rw [← inner_self, real_inner_self_eq_norm_mul_norm] instance : norm_one_class ℍ := ⟨by rw [norm_eq_sqrt_real_inner, inner_self, norm_sq.map_one, real.sqrt_one]⟩ -@[simp, norm_cast] lemma norm_coe (a : ℝ) : ∥(a : ℍ)∥ = ∥a∥ := +@[simp, norm_cast] lemma norm_coe (a : ℝ) : ‖(a : ℍ)‖ = ‖a‖ := by rw [norm_eq_sqrt_real_inner, inner_self, norm_sq_coe, real.sqrt_sq_eq_abs, real.norm_eq_abs] -@[simp, norm_cast] lemma nnnorm_coe (a : ℝ) : ∥(a : ℍ)∥₊ = ∥a∥₊ := +@[simp, norm_cast] lemma nnnorm_coe (a : ℝ) : ‖(a : ℍ)‖₊ = ‖a‖₊ := subtype.ext $ norm_coe a noncomputable instance : normed_division_ring ℍ := diff --git a/src/analysis/schwartz_space.lean b/src/analysis/schwartz_space.lean index a17055be490f2..b0ff178f015e6 100644 --- a/src/analysis/schwartz_space.lean +++ b/src/analysis/schwartz_space.lean @@ -18,7 +18,7 @@ functions $f : ℝ^n → ℂ$ such that there exists $C_{αβ} > 0$ with $$|x^α all $x ∈ ℝ^n$ and for all multiindices $α, β$. In mathlib, we use a slightly different approach and define define the Schwartz space as all smooth functions `f : E → F`, where `E` and `F` are real normed vector spaces such that for all -natural numbers `k` and `n` we have uniform bounds `∥x∥^k * ∥iterated_fderiv ℝ n f x∥ < C`. +natural numbers `k` and `n` we have uniform bounds `‖x‖^k * ‖iterated_fderiv ℝ n f x‖ < C`. This approach completely avoids using partial derivatives as well as polynomials. We construct the topology on the Schwartz space by a family of seminorms, which are the best constants in the above estimates, which is by abstract theory from @@ -28,7 +28,7 @@ Schwartz space into a locally convex topological vector space. ## Main definitions * `schwartz_map`: The Schwartz space is the space of smooth functions such that all derivatives -decay faster than any power of `∥x∥`. +decay faster than any power of `‖x‖`. * `schwartz_map.seminorm`: The family of seminorms as described above * `schwartz_map.fderiv_clm`: The differential as a continuous linear map `𝓢(E, F) →L[𝕜] 𝓢(E, E →L[ℝ] F)` @@ -61,11 +61,11 @@ variables [normed_add_comm_group F] [normed_space ℝ F] variables (E F) /-- A function is a Schwartz function if it is smooth and all derivatives decay faster than - any power of `∥x∥`. -/ + any power of `‖x‖`. -/ structure schwartz_map := (to_fun : E → F) (smooth' : cont_diff ℝ ⊤ to_fun) - (decay' : ∀ (k n : ℕ), ∃ (C : ℝ), ∀ x, ∥x∥^k * ∥iterated_fderiv ℝ n to_fun x∥ ≤ C) + (decay' : ∀ (k n : ℕ), ∃ (C : ℝ), ∀ x, ‖x‖^k * ‖iterated_fderiv ℝ n to_fun x‖ ≤ C) localized "notation `𝓢(` E `, ` F `)` := schwartz_map E F" in schwartz_space @@ -84,7 +84,7 @@ instance : has_coe_to_fun 𝓢(E, F) (λ _, E → F) := ⟨λ p, p.to_fun⟩ /-- All derivatives of a Schwartz function are rapidly decaying. -/ lemma decay (f : 𝓢(E, F)) (k n : ℕ) : ∃ (C : ℝ) (hC : 0 < C), - ∀ x, ∥x∥^k * ∥iterated_fderiv ℝ n f x∥ ≤ C := + ∀ x, ‖x‖^k * ‖iterated_fderiv ℝ n f x‖ ≤ C := begin rcases f.decay' k n with ⟨C, hC⟩, exact ⟨max C 1, by positivity, λ x, (hC x).trans (le_max_left _ _)⟩, @@ -105,17 +105,17 @@ lemma smooth (f : 𝓢(E, F)) (n : ℕ∞) : cont_diff ℝ n f := f.smooth'.of_l section aux lemma bounds_nonempty (k n : ℕ) (f : 𝓢(E, F)) : - ∃ (c : ℝ), c ∈ {c : ℝ | 0 ≤ c ∧ ∀ (x : E), ∥x∥^k * ∥iterated_fderiv ℝ n f x∥ ≤ c} := + ∃ (c : ℝ), c ∈ {c : ℝ | 0 ≤ c ∧ ∀ (x : E), ‖x‖^k * ‖iterated_fderiv ℝ n f x‖ ≤ c} := let ⟨M, hMp, hMb⟩ := f.decay k n in ⟨M, le_of_lt hMp, hMb⟩ lemma bounds_bdd_below (k n : ℕ) (f : 𝓢(E, F)) : - bdd_below {c | 0 ≤ c ∧ ∀ x, ∥x∥^k * ∥iterated_fderiv ℝ n f x∥ ≤ c} := + bdd_below {c | 0 ≤ c ∧ ∀ x, ‖x‖^k * ‖iterated_fderiv ℝ n f x‖ ≤ c} := ⟨0, λ _ ⟨hn, _⟩, hn⟩ lemma decay_add_le_aux (k n : ℕ) (f g : 𝓢(E, F)) (x : E) : - ∥x∥^k * ∥iterated_fderiv ℝ n (f+g) x∥ ≤ - ∥x∥^k * ∥iterated_fderiv ℝ n f x∥ - + ∥x∥^k * ∥iterated_fderiv ℝ n g x∥ := + ‖x‖^k * ‖iterated_fderiv ℝ n (f+g) x‖ ≤ + ‖x‖^k * ‖iterated_fderiv ℝ n f x‖ + + ‖x‖^k * ‖iterated_fderiv ℝ n g x‖ := begin rw ←mul_add, refine mul_le_mul_of_nonneg_left _ (by positivity), @@ -124,7 +124,7 @@ begin end lemma decay_neg_aux (k n : ℕ) (f : 𝓢(E, F)) (x : E) : - ∥x∥ ^ k * ∥iterated_fderiv ℝ n (-f) x∥ = ∥x∥ ^ k * ∥iterated_fderiv ℝ n f x∥ := + ‖x‖ ^ k * ‖iterated_fderiv ℝ n (-f) x‖ = ‖x‖ ^ k * ‖iterated_fderiv ℝ n f x‖ := begin nth_rewrite 3 ←norm_neg, congr, @@ -134,9 +134,9 @@ end variables [normed_field 𝕜] [normed_space 𝕜 F] [smul_comm_class ℝ 𝕜 F] lemma decay_smul_aux (k n : ℕ) (f : 𝓢(E, F)) (c : 𝕜) (x : E) : - ∥x∥ ^ k * ∥iterated_fderiv ℝ n (c • f) x∥ = - ∥c∥ * ∥x∥ ^ k * ∥iterated_fderiv ℝ n f x∥ := -by rw [mul_comm (∥c∥), mul_assoc, iterated_fderiv_const_smul_apply (f.smooth _), norm_smul] + ‖x‖ ^ k * ‖iterated_fderiv ℝ n (c • f) x‖ = + ‖c‖ * ‖x‖ ^ k * ‖iterated_fderiv ℝ n f x‖ := +by rw [mul_comm (‖c‖), mul_assoc, iterated_fderiv_const_smul_apply (f.smooth _), norm_smul] end aux @@ -145,18 +145,18 @@ section seminorm_aux /-- Helper definition for the seminorms of the Schwartz space. -/ @[protected] def seminorm_aux (k n : ℕ) (f : 𝓢(E, F)) : ℝ := -Inf {c | 0 ≤ c ∧ ∀ x, ∥x∥^k * ∥iterated_fderiv ℝ n f x∥ ≤ c} +Inf {c | 0 ≤ c ∧ ∀ x, ‖x‖^k * ‖iterated_fderiv ℝ n f x‖ ≤ c} lemma seminorm_aux_nonneg (k n : ℕ) (f : 𝓢(E, F)) : 0 ≤ f.seminorm_aux k n := le_cInf (bounds_nonempty k n f) (λ _ ⟨hx, _⟩, hx) lemma le_seminorm_aux (k n : ℕ) (f : 𝓢(E, F)) (x : E) : - ∥x∥ ^ k * ∥iterated_fderiv ℝ n ⇑f x∥ ≤ f.seminorm_aux k n := + ‖x‖ ^ k * ‖iterated_fderiv ℝ n ⇑f x‖ ≤ f.seminorm_aux k n := le_cInf (bounds_nonempty k n f) (λ y ⟨_, h⟩, h x) /-- If one controls the norm of every `A x`, then one controls the norm of `A`. -/ lemma seminorm_aux_le_bound (k n : ℕ) (f : 𝓢(E, F)) {M : ℝ} (hMp: 0 ≤ M) - (hM : ∀ x, ∥x∥^k * ∥iterated_fderiv ℝ n f x∥ ≤ M) : + (hM : ∀ x, ‖x‖^k * ‖iterated_fderiv ℝ n f x‖ ≤ M) : f.seminorm_aux k n ≤ M := cInf_le (bounds_bdd_below k n f) ⟨hMp, hM⟩ @@ -173,11 +173,11 @@ instance : has_smul 𝕜 𝓢(E, F) := ⟨λ c f, { to_fun := c • f, smooth' := (f.smooth _).const_smul c, decay' := λ k n, begin - refine ⟨f.seminorm_aux k n * (∥c∥+1), λ x, _⟩, - have hc : 0 ≤ ∥c∥ := by positivity, + refine ⟨f.seminorm_aux k n * (‖c‖+1), λ x, _⟩, + have hc : 0 ≤ ‖c‖ := by positivity, refine le_trans _ ((mul_le_mul_of_nonneg_right (f.le_seminorm_aux k n x) hc).trans _), { apply eq.le, - rw [mul_comm _ (∥c∥), ← mul_assoc], + rw [mul_comm _ (‖c‖), ← mul_assoc], exact decay_smul_aux k n f c x }, { apply mul_le_mul_of_nonneg_left _ (f.seminorm_aux_nonneg k n), linarith } @@ -193,7 +193,7 @@ instance [smul_comm_class 𝕜 𝕜' F] : smul_comm_class 𝕜 𝕜' 𝓢(E, F) ⟨λ a b f, ext $ λ x, smul_comm a b (f x)⟩ lemma seminorm_aux_smul_le (k n : ℕ) (c : 𝕜) (f : 𝓢(E, F)) : - (c • f).seminorm_aux k n ≤ ∥c∥ * f.seminorm_aux k n := + (c • f).seminorm_aux k n ≤ ‖c‖ * f.seminorm_aux k n := begin refine (c • f).seminorm_aux_le_bound k n (mul_nonneg (norm_nonneg _) (seminorm_aux_nonneg _ _ _)) (λ x, (decay_smul_aux k n f c x).le.trans _), @@ -334,30 +334,30 @@ def seminorm (k n : ℕ) : seminorm 𝕜 𝓢(E, F) := seminorm.of_smul_le (semi /-- If one controls the seminorm for every `x`, then one controls the seminorm. -/ lemma seminorm_le_bound (k n : ℕ) (f : 𝓢(E, F)) {M : ℝ} (hMp: 0 ≤ M) - (hM : ∀ x, ∥x∥^k * ∥iterated_fderiv ℝ n f x∥ ≤ M) : seminorm 𝕜 k n f ≤ M := + (hM : ∀ x, ‖x‖^k * ‖iterated_fderiv ℝ n f x‖ ≤ M) : seminorm 𝕜 k n f ≤ M := f.seminorm_aux_le_bound k n hMp hM /-- The seminorm controls the Schwartz estimate for any fixed `x`. -/ lemma le_seminorm (k n : ℕ) (f : 𝓢(E, F)) (x : E) : - ∥x∥ ^ k * ∥iterated_fderiv ℝ n f x∥ ≤ seminorm 𝕜 k n f := + ‖x‖ ^ k * ‖iterated_fderiv ℝ n f x‖ ≤ seminorm 𝕜 k n f := f.le_seminorm_aux k n x lemma norm_iterated_fderiv_le_seminorm (f : 𝓢(E, F)) (n : ℕ) (x₀ : E) : - ∥iterated_fderiv ℝ n f x₀∥ ≤ (schwartz_map.seminorm 𝕜 0 n) f := + ‖iterated_fderiv ℝ n f x₀‖ ≤ (schwartz_map.seminorm 𝕜 0 n) f := begin have := schwartz_map.le_seminorm 𝕜 0 n f x₀, rwa [pow_zero, one_mul] at this, end lemma norm_pow_mul_le_seminorm (f : 𝓢(E, F)) (k : ℕ) (x₀ : E) : - ∥x₀∥^k * ∥f x₀∥ ≤ (schwartz_map.seminorm 𝕜 k 0) f := + ‖x₀‖^k * ‖f x₀‖ ≤ (schwartz_map.seminorm 𝕜 k 0) f := begin have := schwartz_map.le_seminorm 𝕜 k 0 f x₀, rwa norm_iterated_fderiv_zero at this, end lemma norm_le_seminorm (f : 𝓢(E, F)) (x₀ : E) : - ∥f x₀∥ ≤ (schwartz_map.seminorm 𝕜 0 0) f := + ‖f x₀‖ ≤ (schwartz_map.seminorm 𝕜 0 0) f := begin have := norm_pow_mul_le_seminorm 𝕜 f 0 x₀, rwa [pow_zero, one_mul] at this, diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index acc777329c7c8..a06e236385dc4 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -43,7 +43,7 @@ variables {R R' 𝕜 𝕜₂ 𝕜₃ E E₂ E₃ F G ι : Type*} semidefinite, positive homogeneous, and subadditive. -/ structure seminorm (𝕜 : Type*) (E : Type*) [semi_normed_ring 𝕜] [add_group E] [has_smul 𝕜 E] extends add_group_seminorm E := -(smul' : ∀ (a : 𝕜) (x : E), to_fun (a • x) = ∥a∥ * to_fun x) +(smul' : ∀ (a : 𝕜) (x : E), to_fun (a • x) = ‖a‖ * to_fun x) attribute [nolint doc_blame] seminorm.to_add_group_seminorm @@ -52,7 +52,7 @@ attribute [nolint doc_blame] seminorm.to_add_group_seminorm You should extend this class when you extend `seminorm`. -/ class seminorm_class (F : Type*) (𝕜 E : out_param $ Type*) [semi_normed_ring 𝕜] [add_group E] [has_smul 𝕜 E] extends add_group_seminorm_class F E := -(map_smul_eq_mul (f : F) (a : 𝕜) (x : E) : f (a • x) = ∥a∥ * f x) +(map_smul_eq_mul (f : F) (a : 𝕜) (x : E) : f (a • x) = ‖a‖ * f x) export seminorm_class (map_smul_eq_mul) @@ -65,7 +65,7 @@ section of `semi_norm_ring 𝕜`. -/ def seminorm.of [semi_normed_ring 𝕜] [add_comm_group E] [module 𝕜 E] (f : E → ℝ) (add_le : ∀ (x y : E), f (x + y) ≤ f x + f y) - (smul : ∀ (a : 𝕜) (x : E), f (a • x) = ∥a∥ * f x) : seminorm 𝕜 E := + (smul : ∀ (a : 𝕜) (x : E), f (a • x) = ‖a‖ * f x) : seminorm 𝕜 E := { to_fun := f, map_zero' := by rw [←zero_smul 𝕜 (0 : E), smul, norm_zero, zero_mul], add_le' := add_le, @@ -76,7 +76,7 @@ def seminorm.of [semi_normed_ring 𝕜] [add_comm_group E] [module 𝕜 E] (f : and an inequality for the scalar multiplication. -/ def seminorm.of_smul_le [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] (f : E → ℝ) (map_zero : f 0 = 0) (add_le : ∀ x y, f (x + y) ≤ f x + f y) - (smul_le : ∀ (r : 𝕜) x, f (r • x) ≤ ∥r∥ * f x) : seminorm 𝕜 E := + (smul_le : ∀ (r : 𝕜) x, f (r • x) ≤ ‖r‖ * f x) : seminorm 𝕜 E := seminorm.of f add_le (λ r x, begin refine le_antisymm (smul_le r x) _, @@ -326,7 +326,7 @@ begin { exact nnreal.coe_pos.mpr ha }, end -lemma norm_sub_map_le_sub (p : seminorm 𝕜 E) (x y : E) : ∥p x - p y∥ ≤ p (x - y) := +lemma norm_sub_map_le_sub (p : seminorm 𝕜 E) (x y : E) : ‖p x - p y‖ ≤ p (x - y) := abs_sub_map_le_sub p x y end module @@ -338,12 +338,12 @@ variables {σ₁₂ : 𝕜 →+* 𝕜₂} [ring_hom_isometric σ₁₂] variables [add_comm_group E] [add_comm_group E₂] [module 𝕜 E] [module 𝕜₂ E₂] lemma comp_smul (p : seminorm 𝕜₂ E₂) (f : E →ₛₗ[σ₁₂] E₂) (c : 𝕜₂) : - p.comp (c • f) = ∥c∥₊ • p.comp f := + p.comp (c • f) = ‖c‖₊ • p.comp f := ext $ λ _, by rw [comp_apply, smul_apply, linear_map.smul_apply, map_smul_eq_mul, nnreal.smul_def, coe_nnnorm, smul_eq_mul, comp_apply] lemma comp_smul_apply (p : seminorm 𝕜₂ E₂) (f : E →ₛₗ[σ₁₂] E₂) (c : 𝕜₂) (x : E) : - p.comp (c • f) x = ∥c∥ * p (f x) := map_smul_eq_mul p _ _ + p.comp (c • f) x = ‖c‖ * p (f x) := map_smul_eq_mul p _ _ end semi_normed_comm_ring @@ -759,8 +759,8 @@ section normed_field variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] (p : seminorm 𝕜 E) {A B : set E} {a : 𝕜} {r : ℝ} {x : E} -lemma smul_ball_zero {p : seminorm 𝕜 E} {k : 𝕜} {r : ℝ} (hk : 0 < ∥k∥) : - k • p.ball 0 r = p.ball 0 (∥k∥ * r) := +lemma smul_ball_zero {p : seminorm 𝕜 E} {k : 𝕜} {r : ℝ} (hk : 0 < ‖k‖) : + k • p.ball 0 r = p.ball 0 (‖k‖ * r) := begin ext, rw [set.mem_smul_set, seminorm.mem_ball_zero], @@ -771,7 +771,7 @@ begin exact (mul_lt_mul_left hk).mpr hy }, refine ⟨k⁻¹ • x, _, _⟩, { rw [seminorm.mem_ball_zero, map_smul_eq_mul, norm_inv, ←(mul_lt_mul_left hk), - ←mul_assoc, ←(div_eq_mul_inv ∥k∥ ∥k∥), div_self (ne_of_gt hk), one_mul], + ←mul_assoc, ←(div_eq_mul_inv ‖k‖ ‖k‖), div_self (ne_of_gt hk), one_mul], exact h}, rw [←smul_assoc, smul_eq_mul, ←div_eq_mul_inv, div_self (norm_pos_iff.mp hk), one_smul], end @@ -786,7 +786,7 @@ begin refine ⟨r₂/r, div_pos hr₂ hr, _⟩, simp_rw set.subset_def, intros a ha x hx, - have ha' : 0 < ∥a∥ := lt_of_lt_of_le (div_pos hr₂ hr) ha, + have ha' : 0 < ‖a‖ := lt_of_lt_of_le (div_pos hr₂ hr) ha, rw [smul_ball_zero ha', p.mem_ball_zero], rw p.mem_ball_zero at hx, rw div_le_iff hr at ha, @@ -800,7 +800,7 @@ begin rintro x, have hxr : 0 ≤ p x / r := by positivity, refine ⟨p x/r, hxr, λ a ha, _⟩, - have ha₀ : 0 < ∥a∥ := hxr.trans_lt ha, + have ha₀ : 0 < ‖a‖ := hxr.trans_lt ha, refine ⟨a⁻¹ • x, _, smul_inv_smul₀ (norm_pos_iff.1 ha₀) x⟩, rwa [mem_ball_zero, map_smul_eq_mul, norm_inv, inv_mul_lt_iff ha₀, ←div_lt_iff hr], end @@ -835,7 +835,7 @@ by { ext, rw [mem_neg, mem_ball, mem_ball, ←neg_add', sub_neg_eq_add, map_neg_ @[simp] lemma smul_ball_preimage (p : seminorm 𝕜 E) (y : E) (r : ℝ) (a : 𝕜) (ha : a ≠ 0) : - ((•) a) ⁻¹' p.ball y r = p.ball (a⁻¹ • y) (r / ∥a∥) := + ((•) a) ⁻¹' p.ball y r = p.ball (a⁻¹ • y) (r / ‖a‖) := set.ext $ λ _, by rw [mem_preimage, mem_ball, mem_ball, lt_div_iff (norm_pos_iff.mpr ha), mul_comm, ←map_smul_eq_mul p, smul_sub, smul_inv_smul₀ ha] @@ -852,7 +852,7 @@ protected lemma convex_on : convex_on ℝ univ p := begin refine ⟨convex_univ, λ x _ y _ a b ha hb hab, _⟩, calc p (a • x + b • y) ≤ p (a • x) + p (b • y) : map_add_le_add p _ _ - ... = ∥a • (1 : 𝕜)∥ * p x + ∥b • (1 : 𝕜)∥ * p y + ... = ‖a • (1 : 𝕜)‖ * p x + ‖b • (1 : 𝕜)‖ * p y : by rw [←map_smul_eq_mul p, ←map_smul_eq_mul p, smul_one_smul, smul_one_smul] ... = a * p x + b * p y : by rw [norm_smul, norm_smul, norm_one, mul_one, mul_one, real.norm_of_nonneg ha, @@ -998,7 +998,7 @@ lemma absorbent_ball_zero (hr : 0 < r) : absorbent 𝕜 (metric.ball (0 : E) r) by { rw ←ball_norm_seminorm 𝕜, exact (norm_seminorm _ _).absorbent_ball_zero hr } /-- Balls containing the origin are absorbent. -/ -lemma absorbent_ball (hx : ∥x∥ < r) : absorbent 𝕜 (metric.ball x r) := +lemma absorbent_ball (hx : ‖x‖ < r) : absorbent 𝕜 (metric.ball x r) := by { rw ←ball_norm_seminorm 𝕜, exact (norm_seminorm _ _).absorbent_ball hx } /-- Balls at the origin are balanced. -/ diff --git a/src/analysis/special_functions/bernstein.lean b/src/analysis/special_functions/bernstein.lean index 3cbea87605898..e84fbc01ca3d0 100644 --- a/src/analysis/special_functions/bernstein.lean +++ b/src/analysis/special_functions/bernstein.lean @@ -220,13 +220,13 @@ begin simp only [metric.nhds_basis_ball.tendsto_right_iff, metric.mem_ball, dist_eq_norm], intros ε h, let δ := δ f ε h, - have nhds_zero := tendsto_const_div_at_top_nhds_0_nat (2 * ∥f∥ * δ ^ (-2 : ℤ)), + have nhds_zero := tendsto_const_div_at_top_nhds_0_nat (2 * ‖f‖ * δ ^ (-2 : ℤ)), filter_upwards [nhds_zero.eventually (gt_mem_nhds (half_pos h)), eventually_gt_at_top 0] with n nh npos', have npos : 0 < (n:ℝ) := by exact_mod_cast npos', -- Two easy inequalities we'll need later: - have w₁ : 0 ≤ 2 * ∥f∥ := mul_nonneg (by norm_num) (norm_nonneg f), - have w₂ : 0 ≤ 2 * ∥f∥ * δ^(-2 : ℤ) := mul_nonneg w₁ pow_minus_two_nonneg, + have w₁ : 0 ≤ 2 * ‖f‖ := mul_nonneg (by norm_num) (norm_nonneg f), + have w₂ : 0 ≤ 2 * ‖f‖ * δ^(-2 : ℤ) := mul_nonneg w₁ pow_minus_two_nonneg, -- As `[0,1]` is compact, it suffices to check the inequality pointwise. rw (continuous_map.norm_lt_iff _ h), intro x, @@ -269,17 +269,17 @@ begin (finset.sum_le_univ_sum_of_nonneg (λ k, bernstein_nonneg)) (le_of_lt (half_pos h)) ... = ε/2 : by rw [bernstein.probability, mul_one] }, - { -- We now turn to working on `Sᶜ`: we control the difference term just using `∥f∥`, + { -- We now turn to working on `Sᶜ`: we control the difference term just using `‖f‖`, -- and then insert a `δ^(-2) * (x - k/n)^2` factor -- (which is at least one because we are not in `S`). calc ∑ k in Sᶜ, |f k/ₙ - f x| * bernstein n k x - ≤ ∑ k in Sᶜ, (2 * ∥f∥) * bernstein n k x + ≤ ∑ k in Sᶜ, (2 * ‖f‖) * bernstein n k x : finset.sum_le_sum (λ k m, mul_le_mul_of_nonneg_right (f.dist_le_two_norm _ _) bernstein_nonneg) - ... = (2 * ∥f∥) * ∑ k in Sᶜ, bernstein n k x + ... = (2 * ‖f‖) * ∑ k in Sᶜ, bernstein n k x : by rw finset.mul_sum - ... ≤ (2 * ∥f∥) * ∑ k in Sᶜ, δ^(-2 : ℤ) * (x - k/ₙ)^2 * bernstein n k x + ... ≤ (2 * ‖f‖) * ∑ k in Sᶜ, δ^(-2 : ℤ) * (x - k/ₙ)^2 * bernstein n k x : mul_le_mul_of_nonneg_left (finset.sum_le_sum (λ k m, begin conv_lhs { rw ←one_mul (bernstein _ _ _), }, @@ -287,19 +287,19 @@ begin (le_of_mem_S_compl m) bernstein_nonneg, end)) w₁ -- Again enlarging the sum from `Sᶜ` to all of `fin (n+1)` - ... ≤ (2 * ∥f∥) * ∑ k : fin (n+1), δ^(-2 : ℤ) * (x - k/ₙ)^2 * bernstein n k x + ... ≤ (2 * ‖f‖) * ∑ k : fin (n+1), δ^(-2 : ℤ) * (x - k/ₙ)^2 * bernstein n k x : mul_le_mul_of_nonneg_left (finset.sum_le_univ_sum_of_nonneg (λ k, mul_nonneg (mul_nonneg pow_minus_two_nonneg (sq_nonneg _)) bernstein_nonneg)) w₁ - ... = (2 * ∥f∥) * δ^(-2 : ℤ) * ∑ k : fin (n+1), (x - k/ₙ)^2 * bernstein n k x + ... = (2 * ‖f‖) * δ^(-2 : ℤ) * ∑ k : fin (n+1), (x - k/ₙ)^2 * bernstein n k x : by conv_rhs { rw [mul_assoc, finset.mul_sum], simp only [←mul_assoc], } -- `bernstein.variance` and `x ∈ [0,1]` gives the uniform bound - ... = (2 * ∥f∥) * δ^(-2 : ℤ) * x * (1-x) / n + ... = (2 * ‖f‖) * δ^(-2 : ℤ) * x * (1-x) / n : by { rw variance npos, ring, } - ... ≤ (2 * ∥f∥) * δ^(-2 : ℤ) / n + ... ≤ (2 * ‖f‖) * δ^(-2 : ℤ) / n : (div_le_div_right npos).mpr $ by refine mul_le_of_le_of_le_one' (mul_le_of_le_one_right w₂ _) _ _ w₂; unit_interval ... < ε/2 : nh, } diff --git a/src/analysis/special_functions/exp.lean b/src/analysis/special_functions/exp.lean index 11de0b639b955..7960020f35ea6 100644 --- a/src/analysis/special_functions/exp.lean +++ b/src/analysis/special_functions/exp.lean @@ -27,36 +27,36 @@ namespace complex variables {z y x : ℝ} -lemma exp_bound_sq (x z : ℂ) (hz : ∥z∥ ≤ 1) : - ∥exp (x + z) - exp x - z • exp x∥ ≤ ∥exp x∥ * ∥z∥ ^ 2 := -calc ∥exp (x + z) - exp x - z * exp x∥ - = ∥exp x * (exp z - 1 - z)∥ : by { congr, rw [exp_add], ring } -... = ∥exp x∥ * ∥exp z - 1 - z∥ : norm_mul _ _ -... ≤ ∥exp x∥ * ∥z∥^2 : mul_le_mul_of_nonneg_left (abs_exp_sub_one_sub_id_le hz) (norm_nonneg _) +lemma exp_bound_sq (x z : ℂ) (hz : ‖z‖ ≤ 1) : + ‖exp (x + z) - exp x - z • exp x‖ ≤ ‖exp x‖ * ‖z‖ ^ 2 := +calc ‖exp (x + z) - exp x - z * exp x‖ + = ‖exp x * (exp z - 1 - z)‖ : by { congr, rw [exp_add], ring } +... = ‖exp x‖ * ‖exp z - 1 - z‖ : norm_mul _ _ +... ≤ ‖exp x‖ * ‖z‖^2 : mul_le_mul_of_nonneg_left (abs_exp_sub_one_sub_id_le hz) (norm_nonneg _) lemma locally_lipschitz_exp {r : ℝ} (hr_nonneg : 0 ≤ r) (hr_le : r ≤ 1) (x y : ℂ) - (hyx : ∥y - x∥ < r) : - ∥exp y - exp x∥ ≤ (1 + r) * ∥exp x∥ * ∥y - x∥ := + (hyx : ‖y - x‖ < r) : + ‖exp y - exp x‖ ≤ (1 + r) * ‖exp x‖ * ‖y - x‖ := begin have hy_eq : y = x + (y - x), by abel, - have hyx_sq_le : ∥y - x∥ ^ 2 ≤ r * ∥y - x∥, + have hyx_sq_le : ‖y - x‖ ^ 2 ≤ r * ‖y - x‖, { rw pow_two, exact mul_le_mul hyx.le le_rfl (norm_nonneg _) hr_nonneg, }, - have h_sq : ∀ z, ∥z∥ ≤ 1 → ∥exp (x + z) - exp x∥ ≤ ∥z∥ * ∥exp x∥ + ∥exp x∥ * ∥z∥ ^ 2, + have h_sq : ∀ z, ‖z‖ ≤ 1 → ‖exp (x + z) - exp x‖ ≤ ‖z‖ * ‖exp x‖ + ‖exp x‖ * ‖z‖ ^ 2, { intros z hz, - have : ∥exp (x + z) - exp x - z • exp x∥ ≤ ∥exp x∥ * ∥z∥ ^ 2, from exp_bound_sq x z hz, + have : ‖exp (x + z) - exp x - z • exp x‖ ≤ ‖exp x‖ * ‖z‖ ^ 2, from exp_bound_sq x z hz, rw [← sub_le_iff_le_add', ← norm_smul z], exact (norm_sub_norm_le _ _).trans this, }, - calc ∥exp y - exp x∥ = ∥exp (x + (y - x)) - exp x∥ : by nth_rewrite 0 hy_eq - ... ≤ ∥y - x∥ * ∥exp x∥ + ∥exp x∥ * ∥y - x∥ ^ 2 : h_sq (y - x) (hyx.le.trans hr_le) - ... ≤ ∥y - x∥ * ∥exp x∥ + ∥exp x∥ * (r * ∥y - x∥) : + calc ‖exp y - exp x‖ = ‖exp (x + (y - x)) - exp x‖ : by nth_rewrite 0 hy_eq + ... ≤ ‖y - x‖ * ‖exp x‖ + ‖exp x‖ * ‖y - x‖ ^ 2 : h_sq (y - x) (hyx.le.trans hr_le) + ... ≤ ‖y - x‖ * ‖exp x‖ + ‖exp x‖ * (r * ‖y - x‖) : add_le_add_left (mul_le_mul le_rfl hyx_sq_le (sq_nonneg _) (norm_nonneg _)) _ - ... = (1 + r) * ∥exp x∥ * ∥y - x∥ : by ring, + ... = (1 + r) * ‖exp x‖ * ‖y - x‖ : by ring, end @[continuity] lemma continuous_exp : continuous exp := continuous_iff_continuous_at.mpr $ - λ x, continuous_at_of_locally_lipschitz zero_lt_one (2 * ∥exp x∥) + λ x, continuous_at_of_locally_lipschitz zero_lt_one (2 * ‖exp x‖) (locally_lipschitz_exp zero_le_one le_rfl x) lemma continuous_on_exp {s : set ℂ} : continuous_on exp s := diff --git a/src/analysis/special_functions/exp_deriv.lean b/src/analysis/special_functions/exp_deriv.lean index 4c88a9c96fef7..07d2b5e7aae93 100644 --- a/src/analysis/special_functions/exp_deriv.lean +++ b/src/analysis/special_functions/exp_deriv.lean @@ -30,7 +30,7 @@ lemma has_deriv_at_exp (x : ℂ) : has_deriv_at exp (exp x) x := begin rw has_deriv_at_iff_is_o_nhds_zero, have : (1 : ℕ) < 2 := by norm_num, - refine (is_O.of_bound (∥exp x∥) _).trans_is_o (is_o_pow_id this), + refine (is_O.of_bound (‖exp x‖) _).trans_is_o (is_o_pow_id this), filter_upwards [metric.ball_mem_nhds (0 : ℂ) zero_lt_one], simp only [metric.mem_ball, dist_zero_right, norm_pow], exact λ z hz, exp_bound_sq x z hz.le, diff --git a/src/analysis/special_functions/gamma.lean b/src/analysis/special_functions/gamma.lean index e43c4c07bbed9..40005eed78a37 100644 --- a/src/analysis/special_functions/gamma.lean +++ b/src/analysis/special_functions/gamma.lean @@ -242,7 +242,7 @@ begin suffices : tendsto (λ X, -X ^ s * (-X).exp : ℝ → ℂ) at_top (𝓝 0), { simpa using tendsto.add (tendsto.const_mul s (tendsto_partial_Gamma hs)) this }, rw tendsto_zero_iff_norm_tendsto_zero, - have : (λ (e : ℝ), ∥-(e:ℂ) ^ s * (-e).exp∥ ) =ᶠ[at_top] (λ (e : ℝ), e ^ s.re * (-1 * e).exp ), + have : (λ (e : ℝ), ‖-(e:ℂ) ^ s * (-e).exp‖ ) =ᶠ[at_top] (λ (e : ℝ), e ^ s.re * (-1 * e).exp ), { refine eventually_eq_of_mem (Ioi_mem_at_top 0) _, intros x hx, dsimp only, rw [norm_eq_abs, map_mul, abs.map_neg, abs_cpow_eq_rpow_re_of_pos hx, @@ -380,7 +380,7 @@ end /-- Absolute convergence of the integral which will give the derivative of the `Γ` function on `1 < re s`. -/ lemma dGamma_integral_abs_convergent (s : ℝ) (hs : 1 < s) : - integrable_on (λ x:ℝ, ∥exp (-x) * log x * x ^ (s-1)∥) (Ioi 0) := + integrable_on (λ x:ℝ, ‖exp (-x) * log x * x ^ (s-1)‖) (Ioi 0) := begin rw [←Ioc_union_Ioi_eq_Ioi (@zero_le_one ℝ _ _ _ _), integrable_on_union], refine ⟨⟨_, _⟩, _⟩, @@ -407,7 +407,7 @@ end /-- A uniform bound for the `s`-derivative of the `Γ` integrand for `s` in vertical strips. -/ lemma loc_unif_bound_dGamma_integrand {t : ℂ} {s1 s2 x : ℝ} (ht1 : s1 ≤ t.re) (ht2: t.re ≤ s2) (hx : 0 < x) : - ∥dGamma_integrand t x∥ ≤ dGamma_integrand_real s1 x + dGamma_integrand_real s2 x := + ‖dGamma_integrand t x‖ ≤ dGamma_integrand_real s1 x + dGamma_integrand_real s2 x := begin rcases le_or_lt 1 x with h|h, { -- case 1 ≤ x @@ -457,7 +457,7 @@ begin rw this, refine continuous_on.mul (cont s) (continuous_at.continuous_on _), exact λ x hx, continuous_of_real.continuous_at.comp (continuous_at_log (mem_Ioi.mp hx).ne'), }, - have h_bound : ∀ᵐ (x : ℝ) ∂μ, ∀ (t : ℂ), t ∈ metric.ball s ε → ∥ dGamma_integrand t x ∥ ≤ bound x, + have h_bound : ∀ᵐ (x : ℝ) ∂μ, ∀ (t : ℂ), t ∈ metric.ball s ε → ‖ dGamma_integrand t x ‖ ≤ bound x, { refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (λ x hx, _)), intros t ht, rw [metric.mem_ball, complex.dist_eq] at ht, diff --git a/src/analysis/special_functions/gaussian.lean b/src/analysis/special_functions/gaussian.lean index f270da52503f3..50838cc92347f 100644 --- a/src/analysis/special_functions/gaussian.lean +++ b/src/analysis/special_functions/gaussian.lean @@ -87,7 +87,7 @@ lemma integrable_exp_neg_mul_sq_iff {b : ℝ} : begin refine ⟨λ h, _, integrable_exp_neg_mul_sq⟩, by_contra' hb, - have : ∫⁻ x:ℝ, 1 ≤ ∫⁻ x:ℝ, ∥exp (-b * x^2)∥₊, + have : ∫⁻ x:ℝ, 1 ≤ ∫⁻ x:ℝ, ‖exp (-b * x^2)‖₊, { apply lintegral_mono (λ x, _), simp only [neg_mul, ennreal.one_le_coe_iff, ← to_nnreal_one, to_nnreal_le_iff_le_coe, real.norm_of_nonneg (exp_pos _).le, coe_nnnorm, one_le_exp_iff, right.nonneg_neg_iff], diff --git a/src/analysis/special_functions/japanese_bracket.lean b/src/analysis/special_functions/japanese_bracket.lean index 4e45a6967bbf4..30ab3a86d24a3 100644 --- a/src/analysis/special_functions/japanese_bracket.lean +++ b/src/analysis/special_functions/japanese_bracket.lean @@ -33,15 +33,15 @@ open asymptotics filter set real measure_theory finite_dimensional variables {E : Type*} [normed_add_comm_group E] -lemma sqrt_one_add_norm_sq_le (x : E) : real.sqrt (1 + ∥x∥^2) ≤ 1 + ∥x∥ := +lemma sqrt_one_add_norm_sq_le (x : E) : real.sqrt (1 + ‖x‖^2) ≤ 1 + ‖x‖ := begin refine le_of_pow_le_pow 2 (by positivity) two_pos _, simp [sq_sqrt (zero_lt_one_add_norm_sq x).le, add_pow_two], end -lemma one_add_norm_le_sqrt_two_mul_sqrt (x : E) : 1 + ∥x∥ ≤ (real.sqrt 2) * sqrt (1 + ∥x∥^2) := +lemma one_add_norm_le_sqrt_two_mul_sqrt (x : E) : 1 + ‖x‖ ≤ (real.sqrt 2) * sqrt (1 + ‖x‖^2) := begin - suffices : (sqrt 2 * sqrt (1 + ∥x∥ ^ 2)) ^ 2 - (1 + ∥x∥) ^ 2 = (1 - ∥x∥) ^2, + suffices : (sqrt 2 * sqrt (1 + ‖x‖ ^ 2)) ^ 2 - (1 + ‖x‖) ^ 2 = (1 - ‖x‖) ^2, { refine le_of_pow_le_pow 2 (by positivity) (by norm_num) _, rw [←sub_nonneg, this], positivity, }, @@ -51,13 +51,13 @@ begin end lemma rpow_neg_one_add_norm_sq_le {r : ℝ} (x : E) (hr : 0 < r) : - (1 + ∥x∥^2)^(-r/2) ≤ 2^(r/2) * (1 + ∥x∥)^(-r) := + (1 + ‖x‖^2)^(-r/2) ≤ 2^(r/2) * (1 + ‖x‖)^(-r) := begin have h1 : 0 ≤ (2 : ℝ) := by positivity, have h3 : 0 < sqrt 2 := by positivity, - have h4 : 0 < 1 + ∥x∥ := by positivity, - have h5 : 0 < sqrt (1 + ∥x∥ ^ 2) := by positivity, - have h6 : 0 < sqrt 2 * sqrt (1 + ∥x∥^2) := mul_pos h3 h5, + have h4 : 0 < 1 + ‖x‖ := by positivity, + have h5 : 0 < sqrt (1 + ‖x‖ ^ 2) := by positivity, + have h6 : 0 < sqrt 2 * sqrt (1 + ‖x‖^2) := mul_pos h3 h5, rw [rpow_div_two_eq_sqrt _ h1, rpow_div_two_eq_sqrt _ (zero_lt_one_add_norm_sq x).le, ←inv_mul_le_iff (rpow_pos_of_pos h3 _), rpow_neg h4.le, rpow_neg (sqrt_nonneg _), ←mul_inv, ←mul_rpow h3.le h5.le, inv_le_inv (rpow_pos_of_pos h6 _) (rpow_pos_of_pos h4 _), @@ -66,7 +66,7 @@ begin end lemma le_rpow_one_add_norm_iff_norm_le {r t : ℝ} (hr : 0 < r) (ht : 0 < t) (x : E) : - t ≤ (1 + ∥x∥) ^ -r ↔ ∥x∥ ≤ t ^ -r⁻¹ - 1 := + t ≤ (1 + ‖x‖) ^ -r ↔ ‖x‖ ≤ t ^ -r⁻¹ - 1 := begin rw [le_sub_iff_add_le', neg_inv], exact (real.le_rpow_inv_iff_of_neg (by positivity) ht (neg_lt_zero.mpr hr)).symm, @@ -111,20 +111,20 @@ end lemma finite_integral_one_add_norm [measure_space E] [borel_space E] [(@volume E _).is_add_haar_measure] {r : ℝ} (hnr : (finrank ℝ E : ℝ) < r) : - ∫⁻ (x : E), ennreal.of_real ((1 + ∥x∥) ^ -r) < ∞ := + ∫⁻ (x : E), ennreal.of_real ((1 + ‖x‖) ^ -r) < ∞ := begin have hr : 0 < r := lt_of_le_of_lt (finrank ℝ E).cast_nonneg hnr, -- We start by applying the layer cake formula - have h_meas : measurable (λ (ω : E), (1 + ∥ω∥) ^ -r) := by measurability, - have h_pos : ∀ x : E, 0 ≤ (1 + ∥x∥) ^ -r := + have h_meas : measurable (λ (ω : E), (1 + ‖ω‖) ^ -r) := by measurability, + have h_pos : ∀ x : E, 0 ≤ (1 + ‖x‖) ^ -r := by { intros x, positivity }, rw lintegral_eq_lintegral_meas_le volume h_pos h_meas, -- We use the first transformation of the integrant to show that we only have to integrate from -- 0 to 1 and from 1 to ∞ have h_int : ∀ (t : ℝ) (ht : t ∈ Ioi (0 : ℝ)), - (volume {a : E | t ≤ (1 + ∥a∥) ^ -r} : ennreal) = + (volume {a : E | t ≤ (1 + ‖a‖) ^ -r} : ennreal) = volume (metric.closed_ball (0 : E) (t^(-r⁻¹) - 1)) := begin intros t ht, @@ -171,11 +171,11 @@ end lemma integrable_one_add_norm [measure_space E] [borel_space E] [(@volume E _).is_add_haar_measure] {r : ℝ} (hnr : (finrank ℝ E : ℝ) < r) : - integrable (λ (x : E), (1 + ∥x∥) ^ -r) := + integrable (λ (x : E), (1 + ‖x‖) ^ -r) := begin refine ⟨by measurability, _⟩, -- Lower Lebesgue integral - have : ∫⁻ (a : E), ∥(1 + ∥a∥) ^ -r∥₊ = ∫⁻ (a : E), ennreal.of_real ((1 + ∥a∥) ^ -r) := + have : ∫⁻ (a : E), ‖(1 + ‖a‖) ^ -r‖₊ = ∫⁻ (a : E), ennreal.of_real ((1 + ‖a‖) ^ -r) := lintegral_nnnorm_eq_of_nonneg (λ _, rpow_nonneg_of_nonneg (by positivity) _), rw [has_finite_integral, this], exact finite_integral_one_add_norm hnr, @@ -183,13 +183,13 @@ end lemma integrable_rpow_neg_one_add_norm_sq [measure_space E] [borel_space E] [(@volume E _).is_add_haar_measure] {r : ℝ} (hnr : (finrank ℝ E : ℝ) < r) : - integrable (λ (x : E), (1 + ∥x∥^2) ^ (-r/2)) := + integrable (λ (x : E), (1 + ‖x‖^2) ^ (-r/2)) := begin have hr : 0 < r := lt_of_le_of_lt (finrank ℝ E).cast_nonneg hnr, refine ((integrable_one_add_norm hnr).const_mul $ 2 ^ (r / 2)).mono (by measurability) (eventually_of_forall $ λ x, _), - have h1 : 0 ≤ (1 + ∥x∥ ^ 2) ^ (-r/2) := by positivity, - have h2 : 0 ≤ (1 + ∥x∥) ^ -r := by positivity, + have h1 : 0 ≤ (1 + ‖x‖ ^ 2) ^ (-r/2) := by positivity, + have h2 : 0 ≤ (1 + ‖x‖) ^ -r := by positivity, have h3 : 0 ≤ (2 : ℝ)^(r/2) := by positivity, simp_rw [norm_mul, norm_eq_abs, abs_of_nonneg h1, abs_of_nonneg h2, abs_of_nonneg h3], exact rpow_neg_one_add_norm_sq_le _ hr, diff --git a/src/analysis/special_functions/log/deriv.lean b/src/analysis/special_functions/log/deriv.lean index 50ba19bde3f8b..434ae117622d4 100644 --- a/src/analysis/special_functions/log/deriv.lean +++ b/src/analysis/special_functions/log/deriv.lean @@ -224,7 +224,7 @@ begin apply_rules [div_le_div, pow_nonneg, abs_nonneg, pow_le_pow_of_le_left] end }, -- third step: apply the mean value inequality - have C : ∥F x - F 0∥ ≤ (|x|^n / (1 - |x|)) * ∥x - 0∥, + have C : ‖F x - F 0‖ ≤ (|x|^n / (1 - |x|)) * ‖x - 0‖, { have : ∀ y ∈ Icc (- |x|) (|x|), differentiable_at ℝ F y, { assume y hy, have : 1 - y ≠ 0 := sub_ne_zero_of_ne (ne_of_gt (lt_of_le_of_lt hy.2 h)), @@ -252,7 +252,7 @@ begin exact tendsto_pow_at_top_nhds_0_of_lt_1 (abs_nonneg _) h }, show summable (λ (n : ℕ), x ^ (n + 1) / (n + 1)), { refine summable_of_norm_bounded _ (summable_geometric_of_lt_1 (abs_nonneg _) h) (λ i, _), - calc ∥x ^ (i + 1) / (i + 1)∥ + calc ‖x ^ (i + 1) / (i + 1)‖ = |x| ^ (i + 1) / (i + 1) : begin have : (0 : ℝ) ≤ i + 1 := le_of_lt (nat.cast_add_one_pos i), diff --git a/src/analysis/special_functions/non_integrable.lean b/src/analysis/special_functions/non_integrable.lean index 07377437342e9..4fc6b9f5ab2bb 100644 --- a/src/analysis/special_functions/non_integrable.lean +++ b/src/analysis/special_functions/non_integrable.lean @@ -48,7 +48,7 @@ is the derivative of `f`, then `g` is not integrable on any interval `a..b` such `[a, b] ∈ l`. -/ lemma not_interval_integrable_of_tendsto_norm_at_top_of_deriv_is_O_filter {f : ℝ → E} {g : ℝ → F} {a b : ℝ} (l : filter ℝ) [ne_bot l] [tendsto_Ixx_class Icc l l] (hl : [a, b] ∈ l) - (hd : ∀ᶠ x in l, differentiable_at ℝ f x) (hf : tendsto (λ x, ∥f x∥) l at_top) + (hd : ∀ᶠ x in l, differentiable_at ℝ f x) (hf : tendsto (λ x, ‖f x‖) l at_top) (hfg : deriv f =O[l] g) : ¬interval_integrable g volume a b := begin @@ -56,49 +56,49 @@ begin obtain ⟨C, hC₀, s, hsl, hsub, hfd, hg⟩ : ∃ (C : ℝ) (hC₀ : 0 ≤ C) (s ∈ l), (∀ (x ∈ s) (y ∈ s), [x, y] ⊆ [a, b]) ∧ (∀ (x ∈ s) (y ∈ s) (z ∈ [x, y]), differentiable_at ℝ f z) ∧ - (∀ (x ∈ s) (y ∈ s) (z ∈ [x, y]), ∥deriv f z∥ ≤ C * ∥g z∥), + (∀ (x ∈ s) (y ∈ s) (z ∈ [x, y]), ‖deriv f z‖ ≤ C * ‖g z‖), { rcases hfg.exists_nonneg with ⟨C, C₀, hC⟩, have h : ∀ᶠ x : ℝ × ℝ in l.prod l, ∀ y ∈ [x.1, x.2], (differentiable_at ℝ f y ∧ - ∥deriv f y∥ ≤ C * ∥g y∥) ∧ y ∈ [a, b], + ‖deriv f y‖ ≤ C * ‖g y‖) ∧ y ∈ [a, b], from (tendsto_fst.interval tendsto_snd).eventually ((hd.and hC.bound).and hl).small_sets, rcases mem_prod_self_iff.1 h with ⟨s, hsl, hs⟩, simp only [prod_subset_iff, mem_set_of_eq] at hs, exact ⟨C, C₀, s, hsl, λ x hx y hy z hz, (hs x hx y hy z hz).2, λ x hx y hy z hz, (hs x hx y hy z hz).1.1, λ x hx y hy z hz, (hs x hx y hy z hz).1.2⟩ }, - replace hgi : interval_integrable (λ x, C * ∥g x∥) volume a b, by convert hgi.norm.smul C, - obtain ⟨c, hc, d, hd, hlt⟩ : ∃ (c ∈ s) (d ∈ s), ∥f c∥ + ∫ y in Ι a b, C * ∥g y∥ < ∥f d∥, + replace hgi : interval_integrable (λ x, C * ‖g x‖) volume a b, by convert hgi.norm.smul C, + obtain ⟨c, hc, d, hd, hlt⟩ : ∃ (c ∈ s) (d ∈ s), ‖f c‖ + ∫ y in Ι a b, C * ‖g y‖ < ‖f d‖, { rcases filter.nonempty_of_mem hsl with ⟨c, hc⟩, - have : ∀ᶠ x in l, ∥f c∥ + ∫ y in Ι a b, C * ∥g y∥ < ∥f x∥, + have : ∀ᶠ x in l, ‖f c‖ + ∫ y in Ι a b, C * ‖g y‖ < ‖f x‖, from hf.eventually (eventually_gt_at_top _), exact ⟨c, hc, (this.and hsl).exists.imp (λ d hd, ⟨hd.2, hd.1⟩)⟩ }, specialize hsub c hc d hd, specialize hfd c hc d hd, - replace hg : ∀ x ∈ Ι c d, ∥deriv f x∥ ≤ C * ∥g x∥, from λ z hz, hg c hc d hd z ⟨hz.1.le, hz.2⟩, - have hg_ae : ∀ᵐ x ∂(volume.restrict (Ι c d)), ∥deriv f x∥ ≤ C * ∥g x∥, + replace hg : ∀ x ∈ Ι c d, ‖deriv f x‖ ≤ C * ‖g x‖, from λ z hz, hg c hc d hd z ⟨hz.1.le, hz.2⟩, + have hg_ae : ∀ᵐ x ∂(volume.restrict (Ι c d)), ‖deriv f x‖ ≤ C * ‖g x‖, from (ae_restrict_mem measurable_set_interval_oc).mono hg, have hsub' : Ι c d ⊆ Ι a b, from interval_oc_subset_interval_oc_of_interval_subset_interval hsub, have hfi : interval_integrable (deriv f) volume c d, from (hgi.mono_set hsub).mono_fun' (ae_strongly_measurable_deriv _ _) hg_ae, refine hlt.not_le (sub_le_iff_le_add'.1 _), - calc ∥f d∥ - ∥f c∥ ≤ ∥f d - f c∥ : norm_sub_norm_le _ _ - ... = ∥∫ x in c..d, deriv f x∥ : congr_arg _ (integral_deriv_eq_sub hfd hfi).symm - ... = ∥∫ x in Ι c d, deriv f x∥ : norm_integral_eq_norm_integral_Ioc _ - ... ≤ ∫ x in Ι c d, ∥deriv f x∥ : norm_integral_le_integral_norm _ - ... ≤ ∫ x in Ι c d, C * ∥g x∥ : + calc ‖f d‖ - ‖f c‖ ≤ ‖f d - f c‖ : norm_sub_norm_le _ _ + ... = ‖∫ x in c..d, deriv f x‖ : congr_arg _ (integral_deriv_eq_sub hfd hfi).symm + ... = ‖∫ x in Ι c d, deriv f x‖ : norm_integral_eq_norm_integral_Ioc _ + ... ≤ ∫ x in Ι c d, ‖deriv f x‖ : norm_integral_le_integral_norm _ + ... ≤ ∫ x in Ι c d, C * ‖g x‖ : set_integral_mono_on hfi.norm.def (hgi.def.mono_set hsub') measurable_set_interval_oc hg - ... ≤ ∫ x in Ι a b, C * ∥g x∥ : + ... ≤ ∫ x in Ι a b, C * ‖g x‖ : set_integral_mono_set hgi.def (ae_of_all _ $ λ x, mul_nonneg hC₀ (norm_nonneg _)) hsub'.eventually_le end /-- If `a ≠ b`, `c ∈ [a, b]`, `f` is differentiable in the neighborhood of `c` within -`[a, b] \ {c}`, `∥f x∥ → ∞` as `x → c` within `[a, b] \ {c}`, and `f' = O(g)` along +`[a, b] \ {c}`, `‖f x‖ → ∞` as `x → c` within `[a, b] \ {c}`, and `f' = O(g)` along `𝓝[[a, b] \ {c}] c`, where `f'` is the derivative of `f`, then `g` is not interval integrable on `a..b`. -/ lemma not_interval_integrable_of_tendsto_norm_at_top_of_deriv_is_O_within_diff_singleton {f : ℝ → E} {g : ℝ → F} {a b c : ℝ} (hne : a ≠ b) (hc : c ∈ [a, b]) (h_deriv : ∀ᶠ x in 𝓝[[a, b] \ {c}] c, differentiable_at ℝ f x) - (h_infty : tendsto (λ x, ∥f x∥) (𝓝[[a, b] \ {c}] c) at_top) + (h_infty : tendsto (λ x, ‖f x‖) (𝓝[[a, b] \ {c}] c) at_top) (hg : deriv f =O[𝓝[[a, b] \ {c}] c] g) : ¬interval_integrable g volume a b := begin @@ -118,13 +118,13 @@ begin (h_deriv.filter_mono this) (h_infty.mono_left this) (hg.mono this), end -/-- If `f` is differentiable in a punctured neighborhood of `c`, `∥f x∥ → ∞` as `x → c` (more +/-- If `f` is differentiable in a punctured neighborhood of `c`, `‖f x‖ → ∞` as `x → c` (more formally, along the filter `𝓝[≠] c`), and `f' = O(g)` along `𝓝[≠] c`, where `f'` is the derivative of `f`, then `g` is not interval integrable on any nontrivial interval `a..b` such that `c ∈ [a, b]`. -/ lemma not_interval_integrable_of_tendsto_norm_at_top_of_deriv_is_O_punctured {f : ℝ → E} {g : ℝ → F} {a b c : ℝ} (h_deriv : ∀ᶠ x in 𝓝[≠] c, differentiable_at ℝ f x) - (h_infty : tendsto (λ x, ∥f x∥) (𝓝[≠] c) at_top) (hg : deriv f =O[𝓝[≠] c] g) + (h_infty : tendsto (λ x, ‖f x‖) (𝓝[≠] c) at_top) (hg : deriv f =O[𝓝[≠] c] g) (hne : a ≠ b) (hc : c ∈ [a, b]) : ¬interval_integrable g volume a b := have 𝓝[[a, b] \ {c}] c ≤ 𝓝[≠] c, from nhds_within_mono _ (inter_subset_right _ _), @@ -140,7 +140,7 @@ begin have A : ∀ᶠ x in 𝓝[≠] c, has_deriv_at (λ x, real.log (x - c)) (x - c)⁻¹ x, { filter_upwards [self_mem_nhds_within] with x hx, simpa using ((has_deriv_at_id x).sub_const c).log (sub_ne_zero.2 hx) }, - have B : tendsto (λ x, ∥real.log (x - c)∥) (𝓝[≠] c) at_top, + have B : tendsto (λ x, ‖real.log (x - c)‖) (𝓝[≠] c) at_top, { refine tendsto_abs_at_bot_at_top.comp (real.tendsto_log_nhds_within_zero.comp _), rw ← sub_self c, exact ((has_deriv_at_id c).sub_const c).tendsto_punctured_nhds one_ne_zero }, diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index cbac7ba763994..ac25bb5888698 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -387,7 +387,7 @@ begin { rw [rpow_def_of_pos (abs_pos.2 hx), log_abs] } end -lemma norm_rpow_of_nonneg {x y : ℝ} (hx_nonneg : 0 ≤ x) : ∥x ^ y∥ = ∥x∥ ^ y := +lemma norm_rpow_of_nonneg {x y : ℝ} (hx_nonneg : 0 ≤ x) : ‖x ^ y‖ = ‖x‖ ^ y := by { simp_rw real.norm_eq_abs, exact abs_rpow_of_nonneg hx_nonneg, } end real diff --git a/src/analysis/specific_limits/normed.lean b/src/analysis/specific_limits/normed.lean index 14ce7e45cac43..f2ea85546c0e6 100644 --- a/src/analysis/specific_limits/normed.lean +++ b/src/analysis/specific_limits/normed.lean @@ -44,12 +44,12 @@ tendsto_norm_zero.inf $ tendsto_principal_principal.2 $ λ x hx, norm_pos_iff.2 namespace normed_field lemma tendsto_norm_inverse_nhds_within_0_at_top {𝕜 : Type*} [normed_field 𝕜] : - tendsto (λ x:𝕜, ∥x⁻¹∥) (𝓝[≠] 0) at_top := + tendsto (λ x:𝕜, ‖x⁻¹‖) (𝓝[≠] 0) at_top := (tendsto_inv_zero_at_top.comp tendsto_norm_zero').congr $ λ x, (norm_inv x).symm lemma tendsto_norm_zpow_nhds_within_0_at_top {𝕜 : Type*} [normed_field 𝕜] {m : ℤ} (hm : m < 0) : - tendsto (λ x : 𝕜, ∥x ^ m∥) (𝓝[≠] 0) at_top := + tendsto (λ x : 𝕜, ‖x ^ m‖) (𝓝[≠] 0) at_top := begin rcases neg_surjective m with ⟨m, rfl⟩, rw neg_lt_zero at hm, lift m to ℕ using hm.le, rw int.coe_nat_pos at hm, @@ -179,7 +179,7 @@ begin suffices : (λ n, n ^ k : ℕ → R) =O[at_top] (λ n : ℕ, (r' ^ k) ^ n), from this.trans_is_o (is_o_pow_pow_of_lt_left (pow_nonneg h0 _) hr'), conv in ((r' ^ _) ^ _) { rw [← pow_mul, mul_comm, pow_mul] }, - suffices : ∀ n : ℕ, ∥(n : R)∥ ≤ (r' - 1)⁻¹ * ∥(1 : R)∥ * ∥r' ^ n∥, + suffices : ∀ n : ℕ, ‖(n : R)‖ ≤ (r' - 1)⁻¹ * ‖(1 : R)‖ * ‖r' ^ n‖, from (is_O_of_le' _ this).pow _, intro n, rw mul_right_comm, refine n.norm_cast_le.trans (mul_le_mul_of_nonneg_right _ (norm_nonneg _)), @@ -191,18 +191,18 @@ lemma is_o_coe_const_pow_of_one_lt {R : Type*} [normed_ring R] {r : ℝ} (hr : 1 (coe : ℕ → R) =o[at_top] (λ n, r ^ n) := by simpa only [pow_one] using @is_o_pow_const_const_pow_of_one_lt R _ 1 _ hr -/-- If `∥r₁∥ < r₂`, then for any naturak `k` we have `n ^ k r₁ ^ n = o (r₂ ^ n)` as `n → ∞`. -/ +/-- If `‖r₁‖ < r₂`, then for any naturak `k` we have `n ^ k r₁ ^ n = o (r₂ ^ n)` as `n → ∞`. -/ lemma is_o_pow_const_mul_const_pow_const_pow_of_norm_lt {R : Type*} [normed_ring R] (k : ℕ) - {r₁ : R} {r₂ : ℝ} (h : ∥r₁∥ < r₂) : + {r₁ : R} {r₂ : ℝ} (h : ‖r₁‖ < r₂) : (λ n, n ^ k * r₁ ^ n : ℕ → R) =o[at_top] (λ n, r₂ ^ n) := begin by_cases h0 : r₁ = 0, { refine (is_o_zero _ _).congr' (mem_at_top_sets.2 $ ⟨1, λ n hn, _⟩) eventually_eq.rfl, simp [zero_pow (zero_lt_one.trans_le hn), h0] }, rw [← ne.def, ← norm_pos_iff] at h0, - have A : (λ n, n ^ k : ℕ → R) =o[at_top] (λ n, (r₂ / ∥r₁∥) ^ n), + have A : (λ n, n ^ k : ℕ → R) =o[at_top] (λ n, (r₂ / ‖r₁‖) ^ n), from is_o_pow_const_const_pow_of_one_lt k ((one_lt_div h0).2 h), - suffices : (λ n, r₁ ^ n) =O[at_top] (λ n, ∥r₁∥ ^ n), + suffices : (λ n, r₁ ^ n) =O[at_top] (λ n, ‖r₁‖ ^ n), by simpa [div_mul_cancel _ (pow_pos h0 _).ne'] using A.mul_is_O this, exact is_O.of_bound 1 (by simpa using eventually_norm_pow_le r₁) end @@ -241,9 +241,9 @@ lemma tendsto_self_mul_const_pow_of_lt_one {r : ℝ} (hr : 0 ≤ r) (h'r : r < 1 tendsto (λ n, n * r ^ n : ℕ → ℝ) at_top (𝓝 0) := by simpa only [pow_one] using tendsto_pow_const_mul_const_pow_of_lt_one 1 hr h'r -/-- In a normed ring, the powers of an element x with `∥x∥ < 1` tend to zero. -/ +/-- In a normed ring, the powers of an element x with `‖x‖ < 1` tend to zero. -/ lemma tendsto_pow_at_top_nhds_0_of_norm_lt_1 {R : Type*} [normed_ring R] {x : R} - (h : ∥x∥ < 1) : tendsto (λ (n : ℕ), x ^ n) at_top (𝓝 0) := + (h : ‖x‖ < 1) : tendsto (λ (n : ℕ), x ^ n) at_top (𝓝 0) := begin apply squeeze_zero_norm' (eventually_norm_pow_le x), exact tendsto_pow_at_top_nhds_0_of_lt_1 (norm_nonneg _) h, @@ -258,7 +258,7 @@ section geometric variables {K : Type*} [normed_field K] {ξ : K} -lemma has_sum_geometric_of_norm_lt_1 (h : ∥ξ∥ < 1) : has_sum (λn:ℕ, ξ ^ n) (1 - ξ)⁻¹ := +lemma has_sum_geometric_of_norm_lt_1 (h : ‖ξ‖ < 1) : has_sum (λn:ℕ, ξ ^ n) (1 - ξ)⁻¹ := begin have xi_ne_one : ξ ≠ 1, by { contrapose! h, simp [h] }, have A : tendsto (λn, (ξ ^ n - 1) * (ξ - 1)⁻¹) at_top (𝓝 ((0 - 1) * (ξ - 1)⁻¹)), @@ -268,10 +268,10 @@ begin { simp [norm_pow, summable_geometric_of_lt_1 (norm_nonneg _) h] } end -lemma summable_geometric_of_norm_lt_1 (h : ∥ξ∥ < 1) : summable (λn:ℕ, ξ ^ n) := +lemma summable_geometric_of_norm_lt_1 (h : ‖ξ‖ < 1) : summable (λn:ℕ, ξ ^ n) := ⟨_, has_sum_geometric_of_norm_lt_1 h⟩ -lemma tsum_geometric_of_norm_lt_1 (h : ∥ξ∥ < 1) : ∑'n:ℕ, ξ ^ n = (1 - ξ)⁻¹ := +lemma tsum_geometric_of_norm_lt_1 (h : ‖ξ‖ < 1) : ∑'n:ℕ, ξ ^ n = (1 - ξ)⁻¹ := (has_sum_geometric_of_norm_lt_1 h).tsum_eq lemma has_sum_geometric_of_abs_lt_1 {r : ℝ} (h : |r| < 1) : has_sum (λn:ℕ, r ^ n) (1 - r)⁻¹ := @@ -285,7 +285,7 @@ tsum_geometric_of_norm_lt_1 h /-- A geometric series in a normed field is summable iff the norm of the common ratio is less than one. -/ -@[simp] lemma summable_geometric_iff_norm_lt_1 : summable (λ n : ℕ, ξ ^ n) ↔ ∥ξ∥ < 1 := +@[simp] lemma summable_geometric_iff_norm_lt_1 : summable (λ n : ℕ, ξ ^ n) ↔ ‖ξ‖ < 1 := begin refine ⟨λ h, _, summable_geometric_of_norm_lt_1⟩, obtain ⟨k : ℕ, hk : dist (ξ ^ k) 0 < 1⟩ := @@ -300,7 +300,7 @@ end geometric section mul_geometric lemma summable_norm_pow_mul_geometric_of_norm_lt_1 {R : Type*} [normed_ring R] - (k : ℕ) {r : R} (hr : ∥r∥ < 1) : summable (λ n : ℕ, ∥(n ^ k * r ^ n : R)∥) := + (k : ℕ) {r : R} (hr : ‖r‖ < 1) : summable (λ n : ℕ, ‖(n ^ k * r ^ n : R)‖) := begin rcases exists_between hr with ⟨r', hrr', h⟩, exact summable_of_is_O_nat (summable_geometric_of_lt_1 ((norm_nonneg _).trans hrr'.le) h) @@ -308,12 +308,12 @@ begin end lemma summable_pow_mul_geometric_of_norm_lt_1 {R : Type*} [normed_ring R] [complete_space R] - (k : ℕ) {r : R} (hr : ∥r∥ < 1) : summable (λ n, n ^ k * r ^ n : ℕ → R) := + (k : ℕ) {r : R} (hr : ‖r‖ < 1) : summable (λ n, n ^ k * r ^ n : ℕ → R) := summable_of_summable_norm $ summable_norm_pow_mul_geometric_of_norm_lt_1 _ hr -/-- If `∥r∥ < 1`, then `∑' n : ℕ, n * r ^ n = r / (1 - r) ^ 2`, `has_sum` version. -/ +/-- If `‖r‖ < 1`, then `∑' n : ℕ, n * r ^ n = r / (1 - r) ^ 2`, `has_sum` version. -/ lemma has_sum_coe_mul_geometric_of_norm_lt_1 {𝕜 : Type*} [normed_field 𝕜] [complete_space 𝕜] - {r : 𝕜} (hr : ∥r∥ < 1) : has_sum (λ n, n * r ^ n : ℕ → 𝕜) (r / (1 - r) ^ 2) := + {r : 𝕜} (hr : ‖r‖ < 1) : has_sum (λ n, n * r ^ n : ℕ → 𝕜) (r / (1 - r) ^ 2) := begin have A : summable (λ n, n * r ^ n : ℕ → 𝕜), by simpa using summable_pow_mul_geometric_of_norm_lt_1 1 hr, @@ -332,9 +332,9 @@ begin div_div] end -/-- If `∥r∥ < 1`, then `∑' n : ℕ, n * r ^ n = r / (1 - r) ^ 2`. -/ +/-- If `‖r‖ < 1`, then `∑' n : ℕ, n * r ^ n = r / (1 - r) ^ 2`. -/ lemma tsum_coe_mul_geometric_of_norm_lt_1 {𝕜 : Type*} [normed_field 𝕜] [complete_space 𝕜] - {r : 𝕜} (hr : ∥r∥ < 1) : + {r : 𝕜} (hr : ‖r‖ < 1) : (∑' n : ℕ, n * r ^ n : 𝕜) = (r / (1 - r) ^ 2) := (has_sum_coe_mul_geometric_of_norm_lt_1 hr).tsum_eq @@ -345,29 +345,29 @@ section summable_le_geometric variables [seminormed_add_comm_group α] {r C : ℝ} {f : ℕ → α} lemma seminormed_add_comm_group.cauchy_seq_of_le_geometric {C : ℝ} {r : ℝ} (hr : r < 1) - {u : ℕ → α} (h : ∀ n, ∥u n - u (n + 1)∥ ≤ C*r^n) : cauchy_seq u := + {u : ℕ → α} (h : ∀ n, ‖u n - u (n + 1)‖ ≤ C*r^n) : cauchy_seq u := cauchy_seq_of_le_geometric r C hr (by simpa [dist_eq_norm] using h) -lemma dist_partial_sum_le_of_le_geometric (hf : ∀n, ∥f n∥ ≤ C * r^n) (n : ℕ) : +lemma dist_partial_sum_le_of_le_geometric (hf : ∀n, ‖f n‖ ≤ C * r^n) (n : ℕ) : dist (∑ i in range n, f i) (∑ i in range (n+1), f i) ≤ C * r ^ n := begin rw [sum_range_succ, dist_eq_norm, ← norm_neg, neg_sub, add_sub_cancel'], exact hf n, end -/-- If `∥f n∥ ≤ C * r ^ n` for all `n : ℕ` and some `r < 1`, then the partial sums of `f` form a +/-- If `‖f n‖ ≤ C * r ^ n` for all `n : ℕ` and some `r < 1`, then the partial sums of `f` form a Cauchy sequence. This lemma does not assume `0 ≤ r` or `0 ≤ C`. -/ -lemma cauchy_seq_finset_of_geometric_bound (hr : r < 1) (hf : ∀n, ∥f n∥ ≤ C * r^n) : +lemma cauchy_seq_finset_of_geometric_bound (hr : r < 1) (hf : ∀n, ‖f n‖ ≤ C * r^n) : cauchy_seq (λ s : finset (ℕ), ∑ x in s, f x) := cauchy_seq_finset_of_norm_bounded _ (aux_has_sum_of_le_geometric hr (dist_partial_sum_le_of_le_geometric hf)).summable hf -/-- If `∥f n∥ ≤ C * r ^ n` for all `n : ℕ` and some `r < 1`, then the partial sums of `f` are within +/-- If `‖f n‖ ≤ C * r ^ n` for all `n : ℕ` and some `r < 1`, then the partial sums of `f` are within distance `C * r ^ n / (1 - r)` of the sum of the series. This lemma does not assume `0 ≤ r` or `0 ≤ C`. -/ -lemma norm_sub_le_of_geometric_bound_of_has_sum (hr : r < 1) (hf : ∀n, ∥f n∥ ≤ C * r^n) +lemma norm_sub_le_of_geometric_bound_of_has_sum (hr : r < 1) (hf : ∀n, ‖f n‖ ≤ C * r^n) {a : α} (ha : has_sum f a) (n : ℕ) : - ∥(∑ x in finset.range n, f x) - a∥ ≤ (C * r ^ n) / (1 - r) := + ‖(∑ x in finset.range n, f x) - a‖ ≤ (C * r ^ n) / (1 - r) := begin rw ← dist_eq_norm, apply dist_le_of_le_geometric_of_tendsto r C hr (dist_partial_sum_le_of_le_geometric hf), @@ -375,24 +375,24 @@ begin end @[simp] lemma dist_partial_sum (u : ℕ → α) (n : ℕ) : - dist (∑ k in range (n + 1), u k) (∑ k in range n, u k) = ∥u n∥ := + dist (∑ k in range (n + 1), u k) (∑ k in range n, u k) = ‖u n‖ := by simp [dist_eq_norm, sum_range_succ] @[simp] lemma dist_partial_sum' (u : ℕ → α) (n : ℕ) : - dist (∑ k in range n, u k) (∑ k in range (n+1), u k) = ∥u n∥ := + dist (∑ k in range n, u k) (∑ k in range (n+1), u k) = ‖u n‖ := by simp [dist_eq_norm', sum_range_succ] lemma cauchy_series_of_le_geometric {C : ℝ} {u : ℕ → α} - {r : ℝ} (hr : r < 1) (h : ∀ n, ∥u n∥ ≤ C*r^n) : cauchy_seq (λ n, ∑ k in range n, u k) := + {r : ℝ} (hr : r < 1) (h : ∀ n, ‖u n‖ ≤ C*r^n) : cauchy_seq (λ n, ∑ k in range n, u k) := cauchy_seq_of_le_geometric r C hr (by simp [h]) lemma normed_add_comm_group.cauchy_series_of_le_geometric' {C : ℝ} {u : ℕ → α} {r : ℝ} (hr : r < 1) - (h : ∀ n, ∥u n∥ ≤ C*r^n) : cauchy_seq (λ n, ∑ k in range (n + 1), u k) := + (h : ∀ n, ‖u n‖ ≤ C*r^n) : cauchy_seq (λ n, ∑ k in range (n + 1), u k) := (cauchy_series_of_le_geometric hr h).comp_tendsto $ tendsto_add_at_top_nat 1 lemma normed_add_comm_group.cauchy_series_of_le_geometric'' {C : ℝ} {u : ℕ → α} {N : ℕ} {r : ℝ} (hr₀ : 0 < r) (hr₁ : r < 1) - (h : ∀ n ≥ N, ∥u n∥ ≤ C*r^n) : cauchy_seq (λ n, ∑ k in range (n + 1), u k) := + (h : ∀ n ≥ N, ‖u n‖ ≤ C*r^n) : cauchy_seq (λ n, ∑ k in range (n + 1), u k) := begin set v : ℕ → α := λ n, if n < N then 0 else u n, have hC : 0 ≤ C, @@ -422,30 +422,30 @@ open normed_space /-- A geometric series in a complete normed ring is summable. Proved above (same name, different namespace) for not-necessarily-complete normed fields. -/ lemma normed_ring.summable_geometric_of_norm_lt_1 - (x : R) (h : ∥x∥ < 1) : summable (λ (n:ℕ), x ^ n) := + (x : R) (h : ‖x‖ < 1) : summable (λ (n:ℕ), x ^ n) := begin - have h1 : summable (λ (n:ℕ), ∥x∥ ^ n) := summable_geometric_of_lt_1 (norm_nonneg _) h, + have h1 : summable (λ (n:ℕ), ‖x‖ ^ n) := summable_geometric_of_lt_1 (norm_nonneg _) h, refine summable_of_norm_bounded_eventually _ h1 _, rw nat.cofinite_eq_at_top, exact eventually_norm_pow_le x, end /-- Bound for the sum of a geometric series in a normed ring. This formula does not assume that the -normed ring satisfies the axiom `∥1∥ = 1`. -/ +normed ring satisfies the axiom `‖1‖ = 1`. -/ lemma normed_ring.tsum_geometric_of_norm_lt_1 - (x : R) (h : ∥x∥ < 1) : ∥∑' n:ℕ, x ^ n∥ ≤ ∥(1:R)∥ - 1 + (1 - ∥x∥)⁻¹ := + (x : R) (h : ‖x‖ < 1) : ‖∑' n:ℕ, x ^ n‖ ≤ ‖(1:R)‖ - 1 + (1 - ‖x‖)⁻¹ := begin rw tsum_eq_zero_add (normed_ring.summable_geometric_of_norm_lt_1 x h), simp only [pow_zero], refine le_trans (norm_add_le _ _) _, - have : ∥∑' b : ℕ, (λ n, x ^ (n + 1)) b∥ ≤ (1 - ∥x∥)⁻¹ - 1, + have : ‖∑' b : ℕ, (λ n, x ^ (n + 1)) b‖ ≤ (1 - ‖x‖)⁻¹ - 1, { refine tsum_of_norm_bounded _ (λ b, norm_pow_le' _ (nat.succ_pos b)), convert (has_sum_nat_add_iff' 1).mpr (has_sum_geometric_of_lt_1 (norm_nonneg x) h), simp }, linarith end -lemma geom_series_mul_neg (x : R) (h : ∥x∥ < 1) : +lemma geom_series_mul_neg (x : R) (h : ‖x‖ < 1) : (∑' i:ℕ, x ^ i) * (1 - x) = 1 := begin have := ((normed_ring.summable_geometric_of_norm_lt_1 x h).has_sum.mul_right (1 - x)), @@ -457,7 +457,7 @@ begin rw [←geom_sum_mul_neg, finset.sum_mul], end -lemma mul_neg_geom_series (x : R) (h : ∥x∥ < 1) : +lemma mul_neg_geom_series (x : R) (h : ‖x‖ < 1) : (1 - x) * ∑' i:ℕ, x ^ i = 1 := begin have := (normed_ring.summable_geometric_of_norm_lt_1 x h).has_sum.mul_left (1 - x), @@ -476,13 +476,13 @@ end normed_ring_geometric lemma summable_of_ratio_norm_eventually_le {α : Type*} [seminormed_add_comm_group α] [complete_space α] {f : ℕ → α} {r : ℝ} (hr₁ : r < 1) - (h : ∀ᶠ n in at_top, ∥f (n+1)∥ ≤ r * ∥f n∥) : summable f := + (h : ∀ᶠ n in at_top, ‖f (n+1)‖ ≤ r * ‖f n‖) : summable f := begin by_cases hr₀ : 0 ≤ r, { rw eventually_at_top at h, rcases h with ⟨N, hN⟩, rw ← @summable_nat_add_iff α _ _ _ _ N, - refine summable_of_norm_bounded (λ n, ∥f N∥ * r^n) + refine summable_of_norm_bounded (λ n, ‖f N‖ * r^n) (summable.mul_left _ $ summable_geometric_of_lt_1 hr₀ hr₁) (λ n, _), conv_rhs {rw [mul_comm, ← zero_add N]}, refine le_geom hr₀ n (λ i _, _), @@ -498,7 +498,7 @@ end lemma summable_of_ratio_test_tendsto_lt_one {α : Type*} [normed_add_comm_group α] [complete_space α] {f : ℕ → α} {l : ℝ} (hl₁ : l < 1) (hf : ∀ᶠ n in at_top, f n ≠ 0) - (h : tendsto (λ n, ∥f (n+1)∥/∥f n∥) at_top (𝓝 l)) : summable f := + (h : tendsto (λ n, ‖f (n+1)‖/‖f n‖) at_top (𝓝 l)) : summable f := begin rcases exists_between hl₁ with ⟨r, hr₀, hr₁⟩, refine summable_of_ratio_norm_eventually_le hr₁ _, @@ -507,8 +507,8 @@ begin end lemma not_summable_of_ratio_norm_eventually_ge {α : Type*} [seminormed_add_comm_group α] - {f : ℕ → α} {r : ℝ} (hr : 1 < r) (hf : ∃ᶠ n in at_top, ∥f n∥ ≠ 0) - (h : ∀ᶠ n in at_top, r * ∥f n∥ ≤ ∥f (n+1)∥) : ¬ summable f := + {f : ℕ → α} {r : ℝ} (hr : 1 < r) (hf : ∃ᶠ n in at_top, ‖f n‖ ≠ 0) + (h : ∀ᶠ n in at_top, r * ‖f n‖ ≤ ‖f (n+1)‖) : ¬ summable f := begin rw eventually_at_top at h, rcases h with ⟨N₀, hN₀⟩, @@ -531,9 +531,9 @@ end lemma not_summable_of_ratio_test_tendsto_gt_one {α : Type*} [seminormed_add_comm_group α] {f : ℕ → α} {l : ℝ} (hl : 1 < l) - (h : tendsto (λ n, ∥f (n+1)∥/∥f n∥) at_top (𝓝 l)) : ¬ summable f := + (h : tendsto (λ n, ‖f (n+1)‖/‖f n‖) at_top (𝓝 l)) : ¬ summable f := begin - have key : ∀ᶠ n in at_top, ∥f n∥ ≠ 0, + have key : ∀ᶠ n in at_top, ‖f n‖ ≠ 0, { filter_upwards [eventually_ge_of_tendsto_gt hl h] with _ hn hc, rw [hc, div_zero] at hn, linarith }, @@ -551,7 +551,7 @@ variables {b : ℝ} {f : ℕ → ℝ} {z : ℕ → E} /-- **Dirichlet's Test** for monotone sequences. -/ theorem monotone.cauchy_seq_series_mul_of_tendsto_zero_of_bounded - (hfa : monotone f) (hf0 : tendsto f at_top (𝓝 0)) (hgb : ∀ n, ∥∑ i in range n, z i∥ ≤ b) : + (hfa : monotone f) (hf0 : tendsto f at_top (𝓝 0)) (hgb : ∀ n, ‖∑ i in range n, z i‖ ≤ b) : cauchy_seq (λ n, ∑ i in range (n + 1), (f i) • z i) := begin simp_rw [finset.sum_range_by_parts _ _ (nat.succ _), sub_eq_add_neg, @@ -569,7 +569,7 @@ end /-- **Dirichlet's test** for antitone sequences. -/ theorem antitone.cauchy_seq_series_mul_of_tendsto_zero_of_bounded - (hfa : antitone f) (hf0 : tendsto f at_top (𝓝 0)) (hzb : ∀ n, ∥∑ i in range n, z i∥ ≤ b) : + (hfa : antitone f) (hf0 : tendsto f at_top (𝓝 0)) (hzb : ∀ n, ‖∑ i in range n, z i‖ ≤ b) : cauchy_seq (λ n, ∑ i in range (n+1), (f i) • z i) := begin have hfa': monotone (λ n, -f n) := λ _ _ hab, neg_le_neg $ hfa hab, @@ -579,7 +579,7 @@ begin simp end -lemma norm_sum_neg_one_pow_le (n : ℕ) : ∥∑ i in range n, (-1 : ℝ) ^ i∥ ≤ 1 := +lemma norm_sum_neg_one_pow_le (n : ℕ) : ‖∑ i in range n, (-1 : ℝ) ^ i‖ ≤ 1 := by { rw [neg_one_geom_sum], split_ifs; norm_num } /-- The **alternating series test** for monotone sequences. @@ -628,18 +628,18 @@ lemma real.summable_pow_div_factorial (x : ℝ) : summable (λ n, x ^ n / n! : ℕ → ℝ) := begin -- We start with trivial extimates - have A : (0 : ℝ) < ⌊∥x∥⌋₊ + 1, from zero_lt_one.trans_le (by simp), - have B : ∥x∥ / (⌊∥x∥⌋₊ + 1) < 1, from (div_lt_one A).2 (nat.lt_floor_add_one _), - -- Then we apply the ratio test. The estimate works for `n ≥ ⌊∥x∥⌋₊`. - suffices : ∀ n ≥ ⌊∥x∥⌋₊, ∥x ^ (n + 1) / (n + 1)!∥ ≤ ∥x∥ / (⌊∥x∥⌋₊ + 1) * ∥x ^ n / ↑n!∥, - from summable_of_ratio_norm_eventually_le B (eventually_at_top.2 ⟨⌊∥x∥⌋₊, this⟩), + have A : (0 : ℝ) < ⌊‖x‖⌋₊ + 1, from zero_lt_one.trans_le (by simp), + have B : ‖x‖ / (⌊‖x‖⌋₊ + 1) < 1, from (div_lt_one A).2 (nat.lt_floor_add_one _), + -- Then we apply the ratio test. The estimate works for `n ≥ ⌊‖x‖⌋₊`. + suffices : ∀ n ≥ ⌊‖x‖⌋₊, ‖x ^ (n + 1) / (n + 1)!‖ ≤ ‖x‖ / (⌊‖x‖⌋₊ + 1) * ‖x ^ n / ↑n!‖, + from summable_of_ratio_norm_eventually_le B (eventually_at_top.2 ⟨⌊‖x‖⌋₊, this⟩), -- Finally, we prove the upper estimate intros n hn, - calc ∥x ^ (n + 1) / (n + 1)!∥ = (∥x∥ / (n + 1)) * ∥x ^ n / n!∥ : + calc ‖x ^ (n + 1) / (n + 1)!‖ = (‖x‖ / (n + 1)) * ‖x ^ n / n!‖ : by rw [pow_succ, nat.factorial_succ, nat.cast_mul, ← div_mul_div_comm, norm_mul, norm_div, real.norm_coe_nat, nat.cast_succ] - ... ≤ (∥x∥ / (⌊∥x∥⌋₊ + 1)) * ∥x ^ n / n!∥ : - by mono* with [0 ≤ ∥x ^ n / n!∥, 0 ≤ ∥x∥]; apply norm_nonneg + ... ≤ (‖x‖ / (⌊‖x‖⌋₊ + 1)) * ‖x ^ n / n!‖ : + by mono* with [0 ≤ ‖x ^ n / n!‖, 0 ≤ ‖x‖]; apply norm_nonneg end lemma real.tendsto_pow_div_factorial_at_top (x : ℝ) : diff --git a/src/combinatorics/additive/behrend.lean b/src/combinatorics/additive/behrend.lean index 4ccc8b2547f5b..aa84a203c8664 100644 --- a/src/combinatorics/additive/behrend.lean +++ b/src/combinatorics/additive/behrend.lean @@ -78,7 +78,7 @@ lemma sphere_zero_subset : sphere n d 0 ⊆ 0 := lemma sphere_subset_box : sphere n d k ⊆ box n d := filter_subset _ _ lemma norm_of_mem_sphere {x : fin n → ℕ} (hx : x ∈ sphere n d k) : - ∥(pi_Lp.equiv 2 _).symm (coe ∘ x : fin n → ℝ)∥ = sqrt k := + ‖(pi_Lp.equiv 2 _).symm (coe ∘ x : fin n → ℝ)‖ = sqrt k := begin rw euclidean_space.norm_eq, dsimp, diff --git a/src/data/complex/is_R_or_C.lean b/src/data/complex/is_R_or_C.lean index e420147d842ab..f188aef992065 100644 --- a/src/data/complex/is_R_or_C.lean +++ b/src/data/complex/is_R_or_C.lean @@ -61,9 +61,9 @@ class is_R_or_C (K : Type*) (conj_re_ax : ∀ z : K, re (conj z) = re z) (conj_im_ax : ∀ z : K, im (conj z) = -(im z)) (conj_I_ax : conj I = -I) -(norm_sq_eq_def_ax : ∀ (z : K), ∥z∥^2 = (re z) * (re z) + (im z) * (im z)) +(norm_sq_eq_def_ax : ∀ (z : K), ‖z‖^2 = (re z) * (re z) + (im z) * (im z)) (mul_im_I_ax : ∀ (z : K), (im z) * im I = im z) -(inv_def_ax : ∀ (z : K), z⁻¹ = conj z * 𝓚 ((∥z∥^2)⁻¹)) +(inv_def_ax : ∀ (z : K), z⁻¹ = conj z * 𝓚 ((‖z‖^2)⁻¹)) (div_I_ax : ∀ (z : K), z / I = -(z * I)) end @@ -103,7 +103,7 @@ is_R_or_C.mul_re_ax @[simp, is_R_or_C_simps] lemma mul_im : ∀ z w : K, im (z * w) = re z * im w + im z * re w := is_R_or_C.mul_im_ax -theorem inv_def (z : K) : z⁻¹ = conj z * ((∥z∥^2)⁻¹:ℝ) := +theorem inv_def (z : K) : z⁻¹ = conj z * ((‖z‖^2)⁻¹:ℝ) := is_R_or_C.inv_def_ax z theorem ext_iff : ∀ {z w : K}, z = w ↔ re z = re w ∧ im z = im w := @@ -177,7 +177,7 @@ by simp only [add_zero, of_real_im, zero_mul, of_real_re, mul_im] @[is_R_or_C_simps] lemma smul_im : ∀ (r : ℝ) (z : K), im (r • z) = r * (im z) := λ r z, by { rw algebra.smul_def, apply of_real_mul_im } -@[simp, is_R_or_C_simps] lemma norm_real (r : ℝ) : ∥(r : K)∥ = ∥r∥ := +@[simp, is_R_or_C_simps] lemma norm_real (r : ℝ) : ‖(r : K)‖ = ‖r‖ := by rw [is_R_or_C.of_real_alg, norm_smul, norm_one, mul_one] /-! ### The imaginary unit, `I` -/ @@ -258,8 +258,8 @@ def norm_sq : K →*₀ ℝ := map_one' := by simp only [one_im, add_zero, mul_one, one_re, mul_zero], map_mul' := λ z w, by { simp only [mul_im, mul_re], ring } } -lemma norm_sq_eq_def {z : K} : ∥z∥^2 = (re z) * (re z) + (im z) * (im z) := norm_sq_eq_def_ax z -lemma norm_sq_eq_def' (z : K) : norm_sq z = ∥z∥^2 := by { rw norm_sq_eq_def, refl } +lemma norm_sq_eq_def {z : K} : ‖z‖^2 = (re z) * (re z) + (im z) * (im z) := norm_sq_eq_def_ax z +lemma norm_sq_eq_def' (z : K) : norm_sq z = ‖z‖^2 := by { rw norm_sq_eq_def, refl } @[is_R_or_C_simps] lemma norm_sq_zero : norm_sq (0 : K) = 0 := norm_sq.map_zero @[is_R_or_C_simps] lemma norm_sq_one : norm_sq (1 : K) = 1 := norm_sq.map_one @@ -331,9 +331,9 @@ lemma norm_sq_sub (z w : K) : norm_sq (z - w) = norm_sq z + norm_sq w - 2 * re ( by simp only [norm_sq_add, sub_eq_add_neg, ring_equiv.map_neg, mul_neg, norm_sq_neg, map_neg] -lemma sqrt_norm_sq_eq_norm {z : K} : real.sqrt (norm_sq z) = ∥z∥ := +lemma sqrt_norm_sq_eq_norm {z : K} : real.sqrt (norm_sq z) = ‖z‖ := begin - have h₂ : ∥z∥ = real.sqrt (∥z∥^2) := (real.sqrt_sq (norm_nonneg z)).symm, + have h₂ : ‖z‖ = real.sqrt (‖z‖^2) := (real.sqrt_sq (norm_nonneg z)).symm, rw [h₂], exact congr_arg real.sqrt (norm_sq_eq_def' z) end @@ -411,11 +411,11 @@ map_inv₀ (@norm_sq K _) z @[simp, is_R_or_C_simps] lemma norm_sq_div (z w : K) : norm_sq (z / w) = norm_sq z / norm_sq w := map_div₀ (@norm_sq K _) z w -@[is_R_or_C_simps] lemma norm_conj {z : K} : ∥conj z∥ = ∥z∥ := +@[is_R_or_C_simps] lemma norm_conj {z : K} : ‖conj z‖ = ‖z‖ := by simp only [←sqrt_norm_sq_eq_norm, norm_sq_conj] @[priority 100] instance : cstar_ring K := -{ norm_star_mul_self := λ x, (norm_mul _ _).trans $ congr_arg (* ∥x∥) norm_conj } +{ norm_star_mul_self := λ x, (norm_mul _ _).trans $ congr_arg (* ‖x‖) norm_conj } /-! ### Cast lemmas -/ @@ -477,17 +477,17 @@ local notation `absK` := @abs K _ by simp only [abs, norm_sq, real.sqrt_mul_self_eq_abs, add_zero, of_real_im, monoid_with_zero_hom.coe_mk, of_real_re, mul_zero] -lemma norm_eq_abs (z : K) : ∥z∥ = absK z := +lemma norm_eq_abs (z : K) : ‖z‖ = absK z := by simp only [abs, norm_sq_eq_def', norm_nonneg, real.sqrt_sq] @[is_R_or_C_simps, norm_cast] -lemma norm_of_real (z : ℝ) : ∥(z : K)∥ = ∥z∥ := +lemma norm_of_real (z : ℝ) : ‖(z : K)‖ = ‖z‖ := by { rw [is_R_or_C.norm_eq_abs, is_R_or_C.abs_of_real, real.norm_eq_abs] } lemma abs_of_nonneg {r : ℝ} (h : 0 ≤ r) : absK r = r := (abs_of_real _).trans (abs_of_nonneg h) -lemma norm_of_nonneg {r : ℝ} (r_nn : 0 ≤ r) : ∥(r : K)∥ = r := +lemma norm_of_nonneg {r : ℝ} (r_nn : 0 ≤ r) : ‖(r : K)‖ = r := by { rw norm_of_real, exact abs_eq_self.mpr r_nn, } lemma abs_of_nat (n : ℕ) : absK n = n := @@ -528,10 +528,10 @@ by rw [mul_self_le_mul_self_iff (_root_.abs_nonneg (im z)) (abs_nonneg _), abs_mul_abs_self, mul_self_abs]; apply im_sq_le_norm_sq -lemma norm_re_le_norm (z : K) : ∥re z∥ ≤ ∥z∥ := +lemma norm_re_le_norm (z : K) : ‖re z‖ ≤ ‖z‖ := by { rw [is_R_or_C.norm_eq_abs, real.norm_eq_abs], exact is_R_or_C.abs_re_le_abs _, } -lemma norm_im_le_norm (z : K) : ∥im z∥ ≤ ∥z∥ := +lemma norm_im_le_norm (z : K) : ‖im z‖ ≤ ‖z‖ := by { rw [is_R_or_C.norm_eq_abs, real.norm_eq_abs], exact is_R_or_C.abs_im_le_abs _, } lemma re_le_abs (z : K) : re z ≤ abs z := @@ -786,7 +786,7 @@ noncomputable def re_clm : K →L[ℝ] ℝ := linear_map.mk_continuous re_lm 1 $ by { simp only [norm_eq_abs, re_lm_coe, one_mul, abs_to_real], exact abs_re_le_abs, } -@[simp, is_R_or_C_simps] lemma re_clm_norm : ∥(re_clm : K →L[ℝ] ℝ)∥ = 1 := +@[simp, is_R_or_C_simps] lemma re_clm_norm : ‖(re_clm : K →L[ℝ] ℝ)‖ = 1 := begin apply le_antisymm (linear_map.mk_continuous_norm_le _ zero_le_one _), convert continuous_linear_map.ratio_le_op_norm _ (1 : K), @@ -843,7 +843,7 @@ noncomputable def conj_cle : K ≃L[ℝ] K := @conj_lie K _ @[simp, is_R_or_C_simps] lemma conj_cle_apply : (conj_cle : K → K) = conj := rfl -@[simp, is_R_or_C_simps] lemma conj_cle_norm : ∥(@conj_cle K _ : K →L[ℝ] K)∥ = 1 := +@[simp, is_R_or_C_simps] lemma conj_cle_norm : ‖(@conj_cle K _ : K →L[ℝ] K)‖ = 1 := (@conj_lie K _).to_linear_isometry.norm_to_continuous_linear_map @[priority 100] @@ -870,7 +870,7 @@ noncomputable def of_real_clm : ℝ →L[ℝ] K := of_real_li.to_continuous_line @[simp, is_R_or_C_simps] lemma of_real_clm_apply : (of_real_clm : ℝ → K) = coe := rfl -@[simp, is_R_or_C_simps] lemma of_real_clm_norm : ∥(of_real_clm : ℝ →L[ℝ] K)∥ = 1 := +@[simp, is_R_or_C_simps] lemma of_real_clm_norm : ‖(of_real_clm : ℝ →L[ℝ] K)‖ = 1 := linear_isometry.norm_to_continuous_linear_map of_real_li @[continuity] lemma continuous_of_real : continuous (coe : ℝ → K) := of_real_li.continuous diff --git a/src/geometry/euclidean/angle/oriented/basic.lean b/src/geometry/euclidean/angle/oriented/basic.lean index 58df4f36d7329..b7ae5c7cd9bec 100644 --- a/src/geometry/euclidean/angle/oriented/basic.lean +++ b/src/geometry/euclidean/angle/oriented/basic.lean @@ -335,7 +335,7 @@ lemma oangle_ne_zero_and_ne_pi_iff_linear_independent {x y : V} : by rw [←not_or_distrib, ←not_iff_not, not_not, oangle_eq_zero_or_eq_pi_iff_not_linear_independent] /-- Two vectors are equal if and only if they have equal norms and zero angle between them. -/ -lemma eq_iff_norm_eq_and_oangle_eq_zero (x y : V) : x = y ↔ ∥x∥ = ∥y∥ ∧ o.oangle x y = 0 := +lemma eq_iff_norm_eq_and_oangle_eq_zero (x y : V) : x = y ↔ ‖x‖ = ‖y‖ ∧ o.oangle x y = 0 := begin rw oangle_eq_zero_iff_same_ray, split, @@ -345,7 +345,7 @@ begin { simp }, rintros ⟨h₁, h₂⟩, obtain ⟨r, hr, rfl⟩ := h₂.exists_nonneg_right hy, - have : ∥y∥ ≠ 0 := by simpa using hy, + have : ‖y‖ ≠ 0 := by simpa using hy, obtain rfl : r = 1, { apply mul_right_cancel₀ this, simpa [norm_smul, _root_.abs_of_nonneg hr] using h₁ }, @@ -353,12 +353,12 @@ begin end /-- Two vectors with equal norms are equal if and only if they have zero angle between them. -/ -lemma eq_iff_oangle_eq_zero_of_norm_eq {x y : V} (h : ∥x∥ = ∥y∥) : x = y ↔ o.oangle x y = 0 := +lemma eq_iff_oangle_eq_zero_of_norm_eq {x y : V} (h : ‖x‖ = ‖y‖) : x = y ↔ o.oangle x y = 0 := ⟨λ he, ((o.eq_iff_norm_eq_and_oangle_eq_zero x y).1 he).2, λ ha, (o.eq_iff_norm_eq_and_oangle_eq_zero x y).2 ⟨h, ha⟩⟩ /-- Two vectors with zero angle between them are equal if and only if they have equal norms. -/ -lemma eq_iff_norm_eq_of_oangle_eq_zero {x y : V} (h : o.oangle x y = 0) : x = y ↔ ∥x∥ = ∥y∥ := +lemma eq_iff_norm_eq_of_oangle_eq_zero {x y : V} (h : o.oangle x y = 0) : x = y ↔ ‖x‖ = ‖y‖ := ⟨λ he, ((o.eq_iff_norm_eq_and_oangle_eq_zero x y).1 he).1, λ hn, (o.eq_iff_norm_eq_and_oangle_eq_zero x y).2 ⟨hn, h⟩⟩ @@ -370,9 +370,9 @@ begin simp_rw [oangle], rw [←complex.arg_mul_coe_angle, o.kahler_mul y x z], congr' 1, - convert complex.arg_real_mul _ (_ : 0 < ∥y∥ ^ 2) using 2, + convert complex.arg_real_mul _ (_ : 0 < ‖y‖ ^ 2) using 2, { norm_cast }, - { have : 0 < ∥y∥ := by simpa using hy, + { have : 0 < ‖y‖ := by simpa using hy, positivity }, { exact o.kahler_ne_zero hx hy, }, { exact o.kahler_ne_zero hy hz } @@ -419,13 +419,13 @@ sum of the angles of a triangle. -/ by simp_rw [←oangle_neg_left_eq_neg_right, o.oangle_add_cyc3_neg_left hx hy hz] /-- Pons asinorum, oriented vector angle form. -/ -lemma oangle_sub_eq_oangle_sub_rev_of_norm_eq {x y : V} (h : ∥x∥ = ∥y∥) : +lemma oangle_sub_eq_oangle_sub_rev_of_norm_eq {x y : V} (h : ‖x‖ = ‖y‖) : o.oangle x (x - y) = o.oangle (y - x) y := by simp [oangle, h] /-- The angle at the apex of an isosceles triangle is `π` minus twice a base angle, oriented vector angle form. -/ -lemma oangle_eq_pi_sub_two_zsmul_oangle_sub_of_norm_eq {x y : V} (hn : x ≠ y) (h : ∥x∥ = ∥y∥) : +lemma oangle_eq_pi_sub_two_zsmul_oangle_sub_of_norm_eq {x y : V} (hn : x ≠ y) (h : ‖x‖ = ‖y‖) : o.oangle y x = π - (2 : ℤ) • o.oangle (y - x) y := begin rw two_zsmul, @@ -685,7 +685,7 @@ by rw [←rotation_eq_self_iff, eq_comm] /-- Rotating a vector by the angle to another vector gives the second vector if and only if the norms are equal. -/ @[simp] lemma rotation_oangle_eq_iff_norm_eq (x y : V) : - o.rotation (o.oangle x y) x = y ↔ ∥x∥ = ∥y∥ := + o.rotation (o.oangle x y) x = y ↔ ‖x‖ = ‖y‖ := begin split, { intro h, @@ -698,7 +698,7 @@ end /-- The angle between two nonzero vectors is `θ` if and only if the second vector is the first rotated by `θ` and scaled by the ratio of the norms. -/ lemma oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) - (θ : real.angle) : o.oangle x y = θ ↔ y = (∥y∥ / ∥x∥) • o.rotation θ x := + (θ : real.angle) : o.oangle x y = θ ↔ y = (‖y‖ / ‖x‖) • o.rotation θ x := begin have hp := div_pos (norm_pos_iff.2 hy) (norm_pos_iff.2 hx), split, @@ -718,7 +718,7 @@ begin split, { intro h, rw o.oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero hx hy at h, - exact ⟨∥y∥ / ∥x∥, div_pos (norm_pos_iff.2 hy) (norm_pos_iff.2 hx), h⟩ }, + exact ⟨‖y‖ / ‖x‖, div_pos (norm_pos_iff.2 hy) (norm_pos_iff.2 hx), h⟩ }, { rintro ⟨r, hr, rfl⟩, rw [o.oangle_smul_right_of_pos _ _ hr, o.oangle_rotation_self_right hx] } end @@ -728,7 +728,7 @@ is the first rotated by `θ` and scaled by the ratio of the norms, or `θ` and a vectors are zero. -/ lemma oangle_eq_iff_eq_norm_div_norm_smul_rotation_or_eq_zero {x y : V} (θ : real.angle) : o.oangle x y = θ ↔ - (x ≠ 0 ∧ y ≠ 0 ∧ y = (∥y∥ / ∥x∥) • o.rotation θ x) ∨ (θ = 0 ∧ (x = 0 ∨ y = 0)) := + (x ≠ 0 ∧ y ≠ 0 ∧ y = (‖y‖ / ‖x‖) • o.rotation θ x) ∨ (θ = 0 ∧ (x = 0 ∨ y = 0)) := begin by_cases hx : x = 0, { simp [hx, eq_comm] }, @@ -830,12 +830,12 @@ linear_isometry_equiv.ext $ by simp [rotation_apply] /-- The inner product of two vectors is the product of the norms and the cosine of the oriented angle between the vectors. -/ lemma inner_eq_norm_mul_norm_mul_cos_oangle (x y : V) : - ⟪x, y⟫ = ∥x∥ * ∥y∥ * real.angle.cos (o.oangle x y) := + ⟪x, y⟫ = ‖x‖ * ‖y‖ * real.angle.cos (o.oangle x y) := begin by_cases hx : x = 0, { simp [hx] }, by_cases hy : y = 0, { simp [hy] }, - have : ∥x∥ ≠ 0 := by simpa using hx, - have : ∥y∥ ≠ 0 := by simpa using hy, + have : ‖x‖ ≠ 0 := by simpa using hx, + have : ‖y‖ ≠ 0 := by simpa using hy, rw [oangle, real.angle.cos_coe, complex.cos_arg, o.abs_kahler], { simp only [kahler_apply_apply, real_smul, add_re, of_real_re, mul_re, I_re, of_real_im], field_simp, @@ -846,7 +846,7 @@ end /-- The cosine of the oriented angle between two nonzero vectors is the inner product divided by the product of the norms. -/ lemma cos_oangle_eq_inner_div_norm_mul_norm {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - real.angle.cos (o.oangle x y) = ⟪x, y⟫ / (∥x∥ * ∥y∥) := + real.angle.cos (o.oangle x y) = ⟪x, y⟫ / (‖x‖ * ‖y‖) := begin rw o.inner_eq_norm_mul_norm_mul_cos_oangle, field_simp [norm_ne_zero_iff.2 hx, norm_ne_zero_iff.2 hy], diff --git a/src/geometry/euclidean/angle/sphere.lean b/src/geometry/euclidean/angle/sphere.lean index 1a00515a37289..b897a4d2e2f70 100644 --- a/src/geometry/euclidean/angle/sphere.lean +++ b/src/geometry/euclidean/angle/sphere.lean @@ -25,7 +25,7 @@ variables [fact (finrank ℝ V = 2)] (o : orientation ℝ V (fin 2)) /-- Angle at center of a circle equals twice angle at circumference, oriented vector angle form. -/ lemma oangle_eq_two_zsmul_oangle_sub_of_norm_eq {x y z : V} (hxyne : x ≠ y) (hxzne : x ≠ z) - (hxy : ∥x∥ = ∥y∥) (hxz : ∥x∥ = ∥z∥) : o.oangle y z = (2 : ℤ) • o.oangle (y - x) (z - x) := + (hxy : ‖x‖ = ‖y‖) (hxz : ‖x‖ = ‖z‖) : o.oangle y z = (2 : ℤ) • o.oangle (y - x) (z - x) := begin have hy : y ≠ 0, { rintro rfl, @@ -48,7 +48,7 @@ end /-- Angle at center of a circle equals twice angle at circumference, oriented vector angle form with radius specified. -/ lemma oangle_eq_two_zsmul_oangle_sub_of_norm_eq_real {x y z : V} (hxyne : x ≠ y) (hxzne : x ≠ z) - {r : ℝ} (hx : ∥x∥ = r) (hy : ∥y∥ = r) (hz : ∥z∥ = r) : + {r : ℝ} (hx : ‖x‖ = r) (hy : ‖y‖ = r) (hz : ‖z‖ = r) : o.oangle y z = (2 : ℤ) • o.oangle (y - x) (z - x) := o.oangle_eq_two_zsmul_oangle_sub_of_norm_eq hxyne hxzne (hy.symm ▸ hx) (hz.symm ▸ hx) @@ -56,8 +56,8 @@ o.oangle_eq_two_zsmul_oangle_sub_of_norm_eq hxyne hxzne (hy.symm ▸ hx) (hz.sym a cyclic quadrilateral add to π", for oriented angles mod π (for which those are the same result), represented here as equality of twice the angles. -/ lemma two_zsmul_oangle_sub_eq_two_zsmul_oangle_sub_of_norm_eq {x₁ x₂ y z : V} (hx₁yne : x₁ ≠ y) - (hx₁zne : x₁ ≠ z) (hx₂yne : x₂ ≠ y) (hx₂zne : x₂ ≠ z) {r : ℝ} (hx₁ : ∥x₁∥ = r) (hx₂ : ∥x₂∥ = r) - (hy : ∥y∥ = r) (hz : ∥z∥ = r) : + (hx₁zne : x₁ ≠ z) (hx₂yne : x₂ ≠ y) (hx₂zne : x₂ ≠ z) {r : ℝ} (hx₁ : ‖x₁‖ = r) (hx₂ : ‖x₂‖ = r) + (hy : ‖y‖ = r) (hz : ‖z‖ = r) : (2 : ℤ) • o.oangle (y - x₁) (z - x₁) = (2 : ℤ) • o.oangle (y - x₂) (z - x₂) := o.oangle_eq_two_zsmul_oangle_sub_of_norm_eq_real hx₁yne hx₁zne hx₁ hy hz ▸ o.oangle_eq_two_zsmul_oangle_sub_of_norm_eq_real hx₂yne hx₂zne hx₂ hy hz diff --git a/src/geometry/euclidean/angle/unoriented/basic.lean b/src/geometry/euclidean/angle/unoriented/basic.lean index 63ae34d4a793e..b361b70d0be13 100644 --- a/src/geometry/euclidean/angle/unoriented/basic.lean +++ b/src/geometry/euclidean/angle/unoriented/basic.lean @@ -30,7 +30,7 @@ variables {V : Type*} [inner_product_space ℝ V] /-- The undirected angle between two vectors. If either vector is 0, this is π/2. See `orientation.oangle` for the corresponding oriented angle definition. -/ -def angle (x y : V) : ℝ := real.arccos (⟪x, y⟫ / (∥x∥ * ∥y∥)) +def angle (x y : V) : ℝ := real.arccos (⟪x, y⟫ / (‖x‖ * ‖y‖)) lemma continuous_at_angle {x : V × V} (hx1 : x.1 ≠ 0) (hx2 : x.2 ≠ 0) : continuous_at (λ y : V × V, angle y.1 y.2) x := @@ -42,7 +42,7 @@ lemma angle_smul_smul {c : ℝ} (hc : c ≠ 0) (x y : V) : angle (c • x) (c • y) = angle x y := have c * c ≠ 0, from mul_ne_zero hc hc, by rw [angle, angle, real_inner_smul_left, inner_smul_right, norm_smul, norm_smul, real.norm_eq_abs, - mul_mul_mul_comm _ (∥x∥), abs_mul_abs_self, ← mul_assoc c c, mul_div_mul_left _ _ this] + mul_mul_mul_comm _ (‖x‖), abs_mul_abs_self, ← mul_assoc c c, mul_div_mul_left _ _ this] @[simp] lemma _root_.linear_isometry.angle_map {E F : Type*} [inner_product_space ℝ E] [inner_product_space ℝ F] (f : E →ₗᵢ[ℝ] F) (u v : E) : @@ -72,7 +72,7 @@ lemma conformal_at.preserves_angle {E F : Type*} let ⟨f₁, h₁, c⟩ := H in h₁.unique h ▸ is_conformal_map.preserves_angle c u v /-- The cosine of the angle between two vectors. -/ -lemma cos_angle (x y : V) : real.cos (angle x y) = ⟪x, y⟫ / (∥x∥ * ∥y∥) := +lemma cos_angle (x y : V) : real.cos (angle x y) = ⟪x, y⟫ / (‖x‖ * ‖y‖) := real.cos_arccos (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).1 (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).2 @@ -167,7 +167,7 @@ by rw [angle_comm, angle_smul_right_of_neg y x hr, angle_comm] /-- The cosine of the angle between two vectors, multiplied by the product of their norms. -/ -lemma cos_angle_mul_norm_mul_norm (x y : V) : real.cos (angle x y) * (∥x∥ * ∥y∥) = ⟪x, y⟫ := +lemma cos_angle_mul_norm_mul_norm (x y : V) : real.cos (angle x y) * (‖x‖ * ‖y‖) = ⟪x, y⟫ := begin rw [cos_angle, div_mul_cancel_of_imp], simp [or_imp_distrib] { contextual := tt }, @@ -175,7 +175,7 @@ end /-- The sine of the angle between two vectors, multiplied by the product of their norms. -/ -lemma sin_angle_mul_norm_mul_norm (x y : V) : real.sin (angle x y) * (∥x∥ * ∥y∥) = +lemma sin_angle_mul_norm_mul_norm (x y : V) : real.sin (angle x y) * (‖x‖ * ‖y‖) = real.sqrt (⟪x, x⟫ * ⟪y, y⟫ - ⟪x, y⟫ * ⟪x, y⟫) := begin unfold angle, @@ -185,8 +185,8 @@ begin real.sqrt_mul_self (mul_nonneg (norm_nonneg x) (norm_nonneg y)), real_inner_self_eq_norm_mul_norm, real_inner_self_eq_norm_mul_norm], - by_cases h : (∥x∥ * ∥y∥) = 0, - { rw [(show ∥x∥ * ∥x∥ * (∥y∥ * ∥y∥) = (∥x∥ * ∥y∥) * (∥x∥ * ∥y∥), by ring), h, mul_zero, mul_zero, + by_cases h : (‖x‖ * ‖y‖) = 0, + { rw [(show ‖x‖ * ‖x‖ * (‖y‖ * ‖y‖) = (‖x‖ * ‖y‖) * (‖x‖ * ‖y‖), by ring), h, mul_zero, mul_zero, zero_sub], cases eq_zero_or_eq_zero_of_mul_eq_zero h with hx hy, { rw norm_eq_zero at hx, @@ -229,36 +229,36 @@ iff.symm $ by simp [angle, or_imp_distrib] { contextual := tt } /-- If the angle between two vectors is π, the inner product equals the negative product of the norms. -/ -lemma inner_eq_neg_mul_norm_of_angle_eq_pi {x y : V} (h : angle x y = π) : ⟪x, y⟫ = - (∥x∥ * ∥y∥) := +lemma inner_eq_neg_mul_norm_of_angle_eq_pi {x y : V} (h : angle x y = π) : ⟪x, y⟫ = - (‖x‖ * ‖y‖) := by simp [← cos_angle_mul_norm_mul_norm, h] /-- If the angle between two vectors is 0, the inner product equals the product of the norms. -/ -lemma inner_eq_mul_norm_of_angle_eq_zero {x y : V} (h : angle x y = 0) : ⟪x, y⟫ = ∥x∥ * ∥y∥ := +lemma inner_eq_mul_norm_of_angle_eq_zero {x y : V} (h : angle x y = 0) : ⟪x, y⟫ = ‖x‖ * ‖y‖ := by simp [← cos_angle_mul_norm_mul_norm, h] /-- The inner product of two non-zero vectors equals the negative product of their norms if and only if the angle between the two vectors is π. -/ lemma inner_eq_neg_mul_norm_iff_angle_eq_pi {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - ⟪x, y⟫ = - (∥x∥ * ∥y∥) ↔ angle x y = π := + ⟪x, y⟫ = - (‖x‖ * ‖y‖) ↔ angle x y = π := begin refine ⟨λ h, _, inner_eq_neg_mul_norm_of_angle_eq_pi⟩, - have h₁ : (∥x∥ * ∥y∥) ≠ 0 := (mul_pos (norm_pos_iff.mpr hx) (norm_pos_iff.mpr hy)).ne', + have h₁ : (‖x‖ * ‖y‖) ≠ 0 := (mul_pos (norm_pos_iff.mpr hx) (norm_pos_iff.mpr hy)).ne', rw [angle, h, neg_div, div_self h₁, real.arccos_neg_one], end /-- The inner product of two non-zero vectors equals the product of their norms if and only if the angle between the two vectors is 0. -/ lemma inner_eq_mul_norm_iff_angle_eq_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - ⟪x, y⟫ = ∥x∥ * ∥y∥ ↔ angle x y = 0 := + ⟪x, y⟫ = ‖x‖ * ‖y‖ ↔ angle x y = 0 := begin refine ⟨λ h, _, inner_eq_mul_norm_of_angle_eq_zero⟩, - have h₁ : (∥x∥ * ∥y∥) ≠ 0 := (mul_pos (norm_pos_iff.mpr hx) (norm_pos_iff.mpr hy)).ne', + have h₁ : (‖x‖ * ‖y‖) ≠ 0 := (mul_pos (norm_pos_iff.mpr hx) (norm_pos_iff.mpr hy)).ne', rw [angle, h, div_self h₁, real.arccos_one], end /-- If the angle between two vectors is π, the norm of their difference equals the sum of their norms. -/ -lemma norm_sub_eq_add_norm_of_angle_eq_pi {x y : V} (h : angle x y = π) : ∥x - y∥ = ∥x∥ + ∥y∥ := +lemma norm_sub_eq_add_norm_of_angle_eq_pi {x y : V} (h : angle x y = π) : ‖x - y‖ = ‖x‖ + ‖y‖ := begin rw ← sq_eq_sq (norm_nonneg (x - y)) (add_nonneg (norm_nonneg x) (norm_nonneg y)), rw [norm_sub_pow_two_real, inner_eq_neg_mul_norm_of_angle_eq_pi h], @@ -267,7 +267,7 @@ end /-- If the angle between two vectors is 0, the norm of their sum equals the sum of their norms. -/ -lemma norm_add_eq_add_norm_of_angle_eq_zero {x y : V} (h : angle x y = 0) : ∥x + y∥ = ∥x∥ + ∥y∥ := +lemma norm_add_eq_add_norm_of_angle_eq_zero {x y : V} (h : angle x y = 0) : ‖x + y‖ = ‖x‖ + ‖y‖ := begin rw ← sq_eq_sq (norm_nonneg (x + y)) (add_nonneg (norm_nonneg x) (norm_nonneg y)), rw [norm_add_pow_two_real, inner_eq_mul_norm_of_angle_eq_zero h], @@ -277,56 +277,56 @@ end /-- If the angle between two vectors is 0, the norm of their difference equals the absolute value of the difference of their norms. -/ lemma norm_sub_eq_abs_sub_norm_of_angle_eq_zero {x y : V} (h : angle x y = 0) : - ∥x - y∥ = |∥x∥ - ∥y∥| := + ‖x - y‖ = |‖x‖ - ‖y‖| := begin - rw [← sq_eq_sq (norm_nonneg (x - y)) (abs_nonneg (∥x∥ - ∥y∥)), - norm_sub_pow_two_real, inner_eq_mul_norm_of_angle_eq_zero h, sq_abs (∥x∥ - ∥y∥)], + rw [← sq_eq_sq (norm_nonneg (x - y)) (abs_nonneg (‖x‖ - ‖y‖)), + norm_sub_pow_two_real, inner_eq_mul_norm_of_angle_eq_zero h, sq_abs (‖x‖ - ‖y‖)], ring, end /-- The norm of the difference of two non-zero vectors equals the sum of their norms if and only the angle between the two vectors is π. -/ lemma norm_sub_eq_add_norm_iff_angle_eq_pi {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - ∥x - y∥ = ∥x∥ + ∥y∥ ↔ angle x y = π := + ‖x - y‖ = ‖x‖ + ‖y‖ ↔ angle x y = π := begin refine ⟨λ h, _, norm_sub_eq_add_norm_of_angle_eq_pi⟩, rw ← inner_eq_neg_mul_norm_iff_angle_eq_pi hx hy, obtain ⟨hxy₁, hxy₂⟩ := ⟨norm_nonneg (x - y), add_nonneg (norm_nonneg x) (norm_nonneg y)⟩, rw [← sq_eq_sq hxy₁ hxy₂, norm_sub_pow_two_real] at h, - calc ⟪x, y⟫ = (∥x∥ ^ 2 + ∥y∥ ^ 2 - (∥x∥ + ∥y∥) ^ 2) / 2 : by linarith - ... = -(∥x∥ * ∥y∥) : by ring, + calc ⟪x, y⟫ = (‖x‖ ^ 2 + ‖y‖ ^ 2 - (‖x‖ + ‖y‖) ^ 2) / 2 : by linarith + ... = -(‖x‖ * ‖y‖) : by ring, end /-- The norm of the sum of two non-zero vectors equals the sum of their norms if and only the angle between the two vectors is 0. -/ lemma norm_add_eq_add_norm_iff_angle_eq_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - ∥x + y∥ = ∥x∥ + ∥y∥ ↔ angle x y = 0 := + ‖x + y‖ = ‖x‖ + ‖y‖ ↔ angle x y = 0 := begin refine ⟨λ h, _, norm_add_eq_add_norm_of_angle_eq_zero⟩, rw ← inner_eq_mul_norm_iff_angle_eq_zero hx hy, obtain ⟨hxy₁, hxy₂⟩ := ⟨norm_nonneg (x + y), add_nonneg (norm_nonneg x) (norm_nonneg y)⟩, rw [← sq_eq_sq hxy₁ hxy₂, norm_add_pow_two_real] at h, - calc ⟪x, y⟫ = ((∥x∥ + ∥y∥) ^ 2 - ∥x∥ ^ 2 - ∥y∥ ^ 2)/ 2 : by linarith - ... = ∥x∥ * ∥y∥ : by ring, + calc ⟪x, y⟫ = ((‖x‖ + ‖y‖) ^ 2 - ‖x‖ ^ 2 - ‖y‖ ^ 2)/ 2 : by linarith + ... = ‖x‖ * ‖y‖ : by ring, end /-- The norm of the difference of two non-zero vectors equals the absolute value of the difference of their norms if and only the angle between the two vectors is 0. -/ lemma norm_sub_eq_abs_sub_norm_iff_angle_eq_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - ∥x - y∥ = |∥x∥ - ∥y∥| ↔ angle x y = 0 := + ‖x - y‖ = |‖x‖ - ‖y‖| ↔ angle x y = 0 := begin refine ⟨λ h, _, norm_sub_eq_abs_sub_norm_of_angle_eq_zero⟩, rw ← inner_eq_mul_norm_iff_angle_eq_zero hx hy, - have h1 : ∥x - y∥ ^ 2 = (∥x∥ - ∥y∥) ^ 2, { rw h, exact sq_abs (∥x∥ - ∥y∥) }, + have h1 : ‖x - y‖ ^ 2 = (‖x‖ - ‖y‖) ^ 2, { rw h, exact sq_abs (‖x‖ - ‖y‖) }, rw norm_sub_pow_two_real at h1, - calc ⟪x, y⟫ = ((∥x∥ + ∥y∥) ^ 2 - ∥x∥ ^ 2 - ∥y∥ ^ 2)/ 2 : by linarith - ... = ∥x∥ * ∥y∥ : by ring, + calc ⟪x, y⟫ = ((‖x‖ + ‖y‖) ^ 2 - ‖x‖ ^ 2 - ‖y‖ ^ 2)/ 2 : by linarith + ... = ‖x‖ * ‖y‖ : by ring, end /-- The norm of the sum of two vectors equals the norm of their difference if and only if the angle between them is π/2. -/ lemma norm_add_eq_norm_sub_iff_angle_eq_pi_div_two (x y : V) : - ∥x + y∥ = ∥x - y∥ ↔ angle x y = π / 2 := + ‖x + y‖ = ‖x - y‖ ↔ angle x y = π / 2 := begin rw [← sq_eq_sq (norm_nonneg (x + y)) (norm_nonneg (x - y)), ← inner_eq_zero_iff_angle_eq_pi_div_two x y, norm_add_pow_two_real, norm_sub_pow_two_real], diff --git a/src/geometry/euclidean/angle/unoriented/right_angle.lean b/src/geometry/euclidean/angle/unoriented/right_angle.lean index ef0d5a0f5be2f..21e7df7841621 100644 --- a/src/geometry/euclidean/angle/unoriented/right_angle.lean +++ b/src/geometry/euclidean/angle/unoriented/right_angle.lean @@ -36,7 +36,7 @@ variables {V : Type*} [inner_product_space ℝ V] /-- Pythagorean theorem, if-and-only-if vector angle form. -/ lemma norm_add_sq_eq_norm_sq_add_norm_sq_iff_angle_eq_pi_div_two (x y : V) : - ∥x + y∥ * ∥x + y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ ↔ angle x y = π / 2 := + ‖x + y‖ * ‖x + y‖ = ‖x‖ * ‖x‖ + ‖y‖ * ‖y‖ ↔ angle x y = π / 2 := begin rw norm_add_sq_eq_norm_sq_add_norm_sq_iff_real_inner_eq_zero, exact inner_eq_zero_iff_angle_eq_pi_div_two x y @@ -44,12 +44,12 @@ end /-- Pythagorean theorem, vector angle form. -/ lemma norm_add_sq_eq_norm_sq_add_norm_sq' (x y : V) (h : angle x y = π / 2) : - ∥x + y∥ * ∥x + y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ := + ‖x + y‖ * ‖x + y‖ = ‖x‖ * ‖x‖ + ‖y‖ * ‖y‖ := (norm_add_sq_eq_norm_sq_add_norm_sq_iff_angle_eq_pi_div_two x y).2 h /-- Pythagorean theorem, subtracting vectors, if-and-only-if vector angle form. -/ lemma norm_sub_sq_eq_norm_sq_add_norm_sq_iff_angle_eq_pi_div_two (x y : V) : - ∥x - y∥ * ∥x - y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ ↔ angle x y = π / 2 := + ‖x - y‖ * ‖x - y‖ = ‖x‖ * ‖x‖ + ‖y‖ * ‖y‖ ↔ angle x y = π / 2 := begin rw norm_sub_sq_eq_norm_sq_add_norm_sq_iff_real_inner_eq_zero, exact inner_eq_zero_iff_angle_eq_pi_div_two x y @@ -57,23 +57,23 @@ end /-- Pythagorean theorem, subtracting vectors, vector angle form. -/ lemma norm_sub_sq_eq_norm_sq_add_norm_sq' (x y : V) (h : angle x y = π / 2) : - ∥x - y∥ * ∥x - y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ := + ‖x - y‖ * ‖x - y‖ = ‖x‖ * ‖x‖ + ‖y‖ * ‖y‖ := (norm_sub_sq_eq_norm_sq_add_norm_sq_iff_angle_eq_pi_div_two x y).2 h /-- An angle in a right-angled triangle expressed using `arccos`. -/ lemma angle_add_eq_arccos_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) : - angle x (x + y) = real.arccos (∥x∥ / ∥x + y∥) := + angle x (x + y) = real.arccos (‖x‖ / ‖x + y‖) := begin rw [angle, inner_add_right, h, add_zero, real_inner_self_eq_norm_mul_norm], - by_cases hx : ∥x∥ = 0, { simp [hx] }, + by_cases hx : ‖x‖ = 0, { simp [hx] }, rw [div_mul_eq_div_div, mul_self_div_self] end /-- An angle in a right-angled triangle expressed using `arcsin`. -/ lemma angle_add_eq_arcsin_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) (h0 : x ≠ 0 ∨ y ≠ 0) : - angle x (x + y) = real.arcsin (∥y∥ / ∥x + y∥) := + angle x (x + y) = real.arcsin (‖y‖ / ‖x + y‖) := begin - have hxy : ∥x + y∥ ^ 2 ≠ 0, + have hxy : ‖x + y‖ ^ 2 ≠ 0, { rw [pow_two, norm_add_sq_eq_norm_sq_add_norm_sq_real h, ne_comm], refine ne_of_lt _, rcases h0 with h0 | h0, @@ -91,13 +91,13 @@ end /-- An angle in a right-angled triangle expressed using `arctan`. -/ lemma angle_add_eq_arctan_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) (h0 : x ≠ 0) : - angle x (x + y) = real.arctan (∥y∥ / ∥x∥) := + angle x (x + y) = real.arctan (‖y‖ / ‖x‖) := begin rw [angle_add_eq_arcsin_of_inner_eq_zero h (or.inl h0), real.arctan_eq_arcsin, ←div_mul_eq_div_div, norm_add_eq_sqrt_iff_real_inner_eq_zero.2 h], nth_rewrite 2 [←real.sqrt_sq (norm_nonneg x)], rw [←real.sqrt_mul (sq_nonneg _), div_pow, pow_two, pow_two, mul_add, mul_one, mul_div, - mul_comm (∥x∥ * ∥x∥), ←mul_div, div_self (mul_self_pos.2 (norm_ne_zero_iff.2 h0)).ne', + mul_comm (‖x‖ * ‖x‖), ←mul_div, div_self (mul_self_pos.2 (norm_ne_zero_iff.2 h0)).ne', mul_one] end @@ -136,7 +136,7 @@ end /-- The cosine of an angle in a right-angled triangle as a ratio of sides. -/ lemma cos_angle_add_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) : - real.cos (angle x (x + y)) = ∥x∥ / ∥x + y∥ := + real.cos (angle x (x + y)) = ‖x‖ / ‖x + y‖ := begin rw [angle_add_eq_arccos_of_inner_eq_zero h, real.cos_arccos (le_trans (by norm_num) (div_nonneg (norm_nonneg _) (norm_nonneg _))) @@ -148,7 +148,7 @@ end /-- The sine of an angle in a right-angled triangle as a ratio of sides. -/ lemma sin_angle_add_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) (h0 : x ≠ 0 ∨ y ≠ 0) : - real.sin (angle x (x + y)) = ∥y∥ / ∥x + y∥ := + real.sin (angle x (x + y)) = ‖y‖ / ‖x + y‖ := begin rw [angle_add_eq_arcsin_of_inner_eq_zero h h0, real.sin_arcsin (le_trans (by norm_num) (div_nonneg (norm_nonneg _) (norm_nonneg _))) @@ -160,7 +160,7 @@ end /-- The tangent of an angle in a right-angled triangle as a ratio of sides. -/ lemma tan_angle_add_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) : - real.tan (angle x (x + y)) = ∥y∥ / ∥x∥ := + real.tan (angle x (x + y)) = ‖y‖ / ‖x‖ := begin by_cases h0 : x = 0, { simp [h0] }, rw [angle_add_eq_arctan_of_inner_eq_zero h h0, real.tan_arctan] @@ -169,12 +169,12 @@ end /-- The cosine of an angle in a right-angled triangle multiplied by the hypotenuse equals the adjacent side. -/ lemma cos_angle_add_mul_norm_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) : - real.cos (angle x (x + y)) * ∥x + y∥ = ∥x∥ := + real.cos (angle x (x + y)) * ‖x + y‖ = ‖x‖ := begin rw cos_angle_add_of_inner_eq_zero h, - by_cases hxy : ∥x + y∥ = 0, + by_cases hxy : ‖x + y‖ = 0, { have h' := norm_add_sq_eq_norm_sq_add_norm_sq_real h, - rw [hxy, zero_mul, eq_comm, add_eq_zero_iff' (mul_self_nonneg ∥x∥) (mul_self_nonneg ∥y∥), + rw [hxy, zero_mul, eq_comm, add_eq_zero_iff' (mul_self_nonneg ‖x‖) (mul_self_nonneg ‖y‖), mul_self_eq_zero] at h', simp [h'.1] }, { exact div_mul_cancel _ hxy } @@ -183,7 +183,7 @@ end /-- The sine of an angle in a right-angled triangle multiplied by the hypotenuse equals the opposite side. -/ lemma sin_angle_add_mul_norm_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) : - real.sin (angle x (x + y)) * ∥x + y∥ = ∥y∥ := + real.sin (angle x (x + y)) * ‖x + y‖ = ‖y‖ := begin by_cases h0 : x = 0 ∧ y = 0, { simp [h0] }, rw not_and_distrib at h0, @@ -200,7 +200,7 @@ end /-- The tangent of an angle in a right-angled triangle multiplied by the adjacent side equals the opposite side. -/ lemma tan_angle_add_mul_norm_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) (h0 : x ≠ 0 ∨ y = 0) : - real.tan (angle x (x + y)) * ∥x∥ = ∥y∥ := + real.tan (angle x (x + y)) * ‖x‖ = ‖y‖ := begin rw [tan_angle_add_of_inner_eq_zero h], rcases h0 with h0 | h0; simp [h0] @@ -209,7 +209,7 @@ end /-- A side of a right-angled triangle divided by the cosine of the adjacent angle equals the hypotenuse. -/ lemma norm_div_cos_angle_add_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) (h0 : x ≠ 0 ∨ y = 0) : - ∥x∥ / real.cos (angle x (x + y)) = ∥x + y∥ := + ‖x‖ / real.cos (angle x (x + y)) = ‖x + y‖ := begin rw cos_angle_add_of_inner_eq_zero h, rcases h0 with h0 | h0, @@ -221,7 +221,7 @@ end /-- A side of a right-angled triangle divided by the sine of the opposite angle equals the hypotenuse. -/ lemma norm_div_sin_angle_add_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) (h0 : x = 0 ∨ y ≠ 0) : - ∥y∥ / real.sin (angle x (x + y)) = ∥x + y∥ := + ‖y‖ / real.sin (angle x (x + y)) = ‖x + y‖ := begin rcases h0 with h0 | h0, { simp [h0] }, rw [sin_angle_add_of_inner_eq_zero h (or.inr h0), div_div_eq_mul_div, mul_comm, div_eq_mul_inv, @@ -231,7 +231,7 @@ end /-- A side of a right-angled triangle divided by the tangent of the opposite angle equals the adjacent side. -/ lemma norm_div_tan_angle_add_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) (h0 : x = 0 ∨ y ≠ 0) : - ∥y∥ / real.tan (angle x (x + y)) = ∥x∥ := + ‖y‖ / real.tan (angle x (x + y)) = ‖x‖ := begin rw tan_angle_add_of_inner_eq_zero h, rcases h0 with h0 | h0, @@ -242,7 +242,7 @@ end /-- An angle in a right-angled triangle expressed using `arccos`, version subtracting vectors. -/ lemma angle_sub_eq_arccos_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) : - angle x (x - y) = real.arccos (∥x∥ / ∥x - y∥) := + angle x (x - y) = real.arccos (‖x‖ / ‖x - y‖) := begin rw [←neg_eq_zero, ←inner_neg_right] at h, rw [sub_eq_add_neg, angle_add_eq_arccos_of_inner_eq_zero h] @@ -250,7 +250,7 @@ end /-- An angle in a right-angled triangle expressed using `arcsin`, version subtracting vectors. -/ lemma angle_sub_eq_arcsin_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) (h0 : x ≠ 0 ∨ y ≠ 0) : - angle x (x - y) = real.arcsin (∥y∥ / ∥x - y∥) := + angle x (x - y) = real.arcsin (‖y‖ / ‖x - y‖) := begin rw [←neg_eq_zero, ←inner_neg_right] at h, nth_rewrite 1 ←neg_ne_zero at h0, @@ -259,7 +259,7 @@ end /-- An angle in a right-angled triangle expressed using `arctan`, version subtracting vectors. -/ lemma angle_sub_eq_arctan_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) (h0 : x ≠ 0) : - angle x (x - y) = real.arctan (∥y∥ / ∥x∥) := + angle x (x - y) = real.arctan (‖y‖ / ‖x‖) := begin rw [←neg_eq_zero, ←inner_neg_right] at h, rw [sub_eq_add_neg, angle_add_eq_arctan_of_inner_eq_zero h h0, norm_neg] @@ -298,7 +298,7 @@ end /-- The cosine of an angle in a right-angled triangle as a ratio of sides, version subtracting vectors. -/ lemma cos_angle_sub_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) : - real.cos (angle x (x - y)) = ∥x∥ / ∥x - y∥ := + real.cos (angle x (x - y)) = ‖x‖ / ‖x - y‖ := begin rw [←neg_eq_zero, ←inner_neg_right] at h, rw [sub_eq_add_neg, cos_angle_add_of_inner_eq_zero h] @@ -307,7 +307,7 @@ end /-- The sine of an angle in a right-angled triangle as a ratio of sides, version subtracting vectors. -/ lemma sin_angle_sub_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) (h0 : x ≠ 0 ∨ y ≠ 0) : - real.sin (angle x (x - y)) = ∥y∥ / ∥x - y∥ := + real.sin (angle x (x - y)) = ‖y‖ / ‖x - y‖ := begin rw [←neg_eq_zero, ←inner_neg_right] at h, nth_rewrite 1 ←neg_ne_zero at h0, @@ -317,7 +317,7 @@ end /-- The tangent of an angle in a right-angled triangle as a ratio of sides, version subtracting vectors. -/ lemma tan_angle_sub_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) : - real.tan (angle x (x - y)) = ∥y∥ / ∥x∥ := + real.tan (angle x (x - y)) = ‖y‖ / ‖x‖ := begin rw [←neg_eq_zero, ←inner_neg_right] at h, rw [sub_eq_add_neg, tan_angle_add_of_inner_eq_zero h, norm_neg] @@ -326,7 +326,7 @@ end /-- The cosine of an angle in a right-angled triangle multiplied by the hypotenuse equals the adjacent side, version subtracting vectors. -/ lemma cos_angle_sub_mul_norm_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) : - real.cos (angle x (x - y)) * ∥x - y∥ = ∥x∥ := + real.cos (angle x (x - y)) * ‖x - y‖ = ‖x‖ := begin rw [←neg_eq_zero, ←inner_neg_right] at h, rw [sub_eq_add_neg, cos_angle_add_mul_norm_of_inner_eq_zero h] @@ -335,7 +335,7 @@ end /-- The sine of an angle in a right-angled triangle multiplied by the hypotenuse equals the opposite side, version subtracting vectors. -/ lemma sin_angle_sub_mul_norm_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) : - real.sin (angle x (x - y)) * ∥x - y∥ = ∥y∥ := + real.sin (angle x (x - y)) * ‖x - y‖ = ‖y‖ := begin rw [←neg_eq_zero, ←inner_neg_right] at h, rw [sub_eq_add_neg, sin_angle_add_mul_norm_of_inner_eq_zero h, norm_neg] @@ -344,7 +344,7 @@ end /-- The tangent of an angle in a right-angled triangle multiplied by the adjacent side equals the opposite side, version subtracting vectors. -/ lemma tan_angle_sub_mul_norm_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) (h0 : x ≠ 0 ∨ y = 0) : - real.tan (angle x (x - y)) * ∥x∥ = ∥y∥ := + real.tan (angle x (x - y)) * ‖x‖ = ‖y‖ := begin rw [←neg_eq_zero, ←inner_neg_right] at h, rw ←neg_eq_zero at h0, @@ -354,7 +354,7 @@ end /-- A side of a right-angled triangle divided by the cosine of the adjacent angle equals the hypotenuse, version subtracting vectors. -/ lemma norm_div_cos_angle_sub_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) (h0 : x ≠ 0 ∨ y = 0) : - ∥x∥ / real.cos (angle x (x - y)) = ∥x - y∥ := + ‖x‖ / real.cos (angle x (x - y)) = ‖x - y‖ := begin rw [←neg_eq_zero, ←inner_neg_right] at h, rw ←neg_eq_zero at h0, @@ -364,7 +364,7 @@ end /-- A side of a right-angled triangle divided by the sine of the opposite angle equals the hypotenuse, version subtracting vectors. -/ lemma norm_div_sin_angle_sub_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) (h0 : x = 0 ∨ y ≠ 0) : - ∥y∥ / real.sin (angle x (x - y)) = ∥x - y∥ := + ‖y‖ / real.sin (angle x (x - y)) = ‖x - y‖ := begin rw [←neg_eq_zero, ←inner_neg_right] at h, rw ←neg_ne_zero at h0, @@ -374,7 +374,7 @@ end /-- A side of a right-angled triangle divided by the tangent of the opposite angle equals the adjacent side, version subtracting vectors. -/ lemma norm_div_tan_angle_sub_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) (h0 : x = 0 ∨ y ≠ 0) : - ∥y∥ / real.tan (angle x (x - y)) = ∥x∥ := + ‖y‖ / real.tan (angle x (x - y)) = ‖x‖ := begin rw [←neg_eq_zero, ←inner_neg_right] at h, rw ←neg_ne_zero at h0, diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index eed88f2749b19..57d4ced013447 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -498,18 +498,18 @@ lemma dist_sq_smul_orthogonal_vadd_smul_orthogonal_vadd {s : affine_subspace ℝ {p1 p2 : P} (hp1 : p1 ∈ s) (hp2 : p2 ∈ s) (r1 r2 : ℝ) {v : V} (hv : v ∈ s.directionᗮ) : dist (r1 • v +ᵥ p1) (r2 • v +ᵥ p2) * dist (r1 • v +ᵥ p1) (r2 • v +ᵥ p2) = - dist p1 p2 * dist p1 p2 + (r1 - r2) * (r1 - r2) * (∥v∥ * ∥v∥) := + dist p1 p2 * dist p1 p2 + (r1 - r2) * (r1 - r2) * (‖v‖ * ‖v‖) := calc dist (r1 • v +ᵥ p1) (r2 • v +ᵥ p2) * dist (r1 • v +ᵥ p1) (r2 • v +ᵥ p2) - = ∥(p1 -ᵥ p2) + (r1 - r2) • v∥ * ∥(p1 -ᵥ p2) + (r1 - r2) • v∥ + = ‖(p1 -ᵥ p2) + (r1 - r2) • v‖ * ‖(p1 -ᵥ p2) + (r1 - r2) • v‖ : by rw [dist_eq_norm_vsub V (r1 • v +ᵥ p1), vsub_vadd_eq_vsub_sub, vadd_vsub_assoc, sub_smul, add_comm, add_sub_assoc] -... = ∥p1 -ᵥ p2∥ * ∥p1 -ᵥ p2∥ + ∥(r1 - r2) • v∥ * ∥(r1 - r2) • v∥ +... = ‖p1 -ᵥ p2‖ * ‖p1 -ᵥ p2‖ + ‖(r1 - r2) • v‖ * ‖(r1 - r2) • v‖ : norm_add_sq_eq_norm_sq_add_norm_sq_real (submodule.inner_right_of_mem_orthogonal (vsub_mem_direction hp1 hp2) (submodule.smul_mem _ _ hv)) -... = ∥(p1 -ᵥ p2 : V)∥ * ∥(p1 -ᵥ p2 : V)∥ + |r1 - r2| * |r1 - r2| * ∥v∥ * ∥v∥ +... = ‖(p1 -ᵥ p2 : V)‖ * ‖(p1 -ᵥ p2 : V)‖ + |r1 - r2| * |r1 - r2| * ‖v‖ * ‖v‖ : by { rw [norm_smul, real.norm_eq_abs], ring } -... = dist p1 p2 * dist p1 p2 + (r1 - r2) * (r1 - r2) * (∥v∥ * ∥v∥) +... = dist p1 p2 * dist p1 p2 + (r1 - r2) * (r1 - r2) * (‖v‖ * ‖v‖) : by { rw [dist_eq_norm_vsub V p1, abs_mul_abs_self, mul_assoc] } /-- Reflection in an affine subspace, which is expected to be nonempty diff --git a/src/geometry/euclidean/sphere.lean b/src/geometry/euclidean/sphere.lean index 516e092f4e428..938a557c9c451 100644 --- a/src/geometry/euclidean/sphere.lean +++ b/src/geometry/euclidean/sphere.lean @@ -54,8 +54,8 @@ which are used to deduce corresponding results for Euclidean affine spaces. -/ lemma mul_norm_eq_abs_sub_sq_norm {x y z : V} - (h₁ : ∃ k : ℝ, k ≠ 1 ∧ x + y = k • (x - y)) (h₂ : ∥z - y∥ = ∥z + y∥) : - ∥x - y∥ * ∥x + y∥ = |∥z + y∥ ^ 2 - ∥z - x∥ ^ 2| := + (h₁ : ∃ k : ℝ, k ≠ 1 ∧ x + y = k • (x - y)) (h₂ : ‖z - y‖ = ‖z + y‖) : + ‖x - y‖ * ‖x + y‖ = |‖z + y‖ ^ 2 - ‖z - x‖ ^ 2| := begin obtain ⟨k, hk_ne_one, hk⟩ := h₁, let r := (k - 1)⁻¹ * (k + 1), @@ -74,14 +74,14 @@ begin have hzx : ⟪z, x⟫ = 0 := by rw [hxy, inner_smul_right, hzy, mul_zero], - calc ∥x - y∥ * ∥x + y∥ - = ∥(r - 1) • y∥ * ∥(r + 1) • y∥ : by simp [sub_smul, add_smul, hxy] - ... = ∥r - 1∥ * ∥y∥ * (∥r + 1∥ * ∥y∥) : by simp_rw [norm_smul] - ... = ∥r - 1∥ * ∥r + 1∥ * ∥y∥ ^ 2 : by ring - ... = |(r - 1) * (r + 1) * ∥y∥ ^ 2| : by simp [abs_mul] - ... = |r ^ 2 * ∥y∥ ^ 2 - ∥y∥ ^ 2| : by ring_nf - ... = |∥x∥ ^ 2 - ∥y∥ ^ 2| : by simp [hxy, norm_smul, mul_pow, sq_abs] - ... = |∥z + y∥ ^ 2 - ∥z - x∥ ^ 2| : by simp [norm_add_sq_real, norm_sub_sq_real, + calc ‖x - y‖ * ‖x + y‖ + = ‖(r - 1) • y‖ * ‖(r + 1) • y‖ : by simp [sub_smul, add_smul, hxy] + ... = ‖r - 1‖ * ‖y‖ * (‖r + 1‖ * ‖y‖) : by simp_rw [norm_smul] + ... = ‖r - 1‖ * ‖r + 1‖ * ‖y‖ ^ 2 : by ring + ... = |(r - 1) * (r + 1) * ‖y‖ ^ 2| : by simp [abs_mul] + ... = |r ^ 2 * ‖y‖ ^ 2 - ‖y‖ ^ 2| : by ring_nf + ... = |‖x‖ ^ 2 - ‖y‖ ^ 2| : by simp [hxy, norm_smul, mul_pow, sq_abs] + ... = |‖z + y‖ ^ 2 - ‖z - x‖ ^ 2| : by simp [norm_add_sq_real, norm_sub_sq_real, hzy, hzx, abs_sub_comm], end diff --git a/src/geometry/euclidean/triangle.lean b/src/geometry/euclidean/triangle.lean index 741ed3f0d7126..1c707d9e7227d 100644 --- a/src/geometry/euclidean/triangle.lean +++ b/src/geometry/euclidean/triangle.lean @@ -54,15 +54,15 @@ variables {V : Type*} [inner_product_space ℝ V] /-- Law of cosines (cosine rule), vector angle form. -/ lemma norm_sub_sq_eq_norm_sq_add_norm_sq_sub_two_mul_norm_mul_norm_mul_cos_angle (x y : V) : - ∥x - y∥ * ∥x - y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ - 2 * ∥x∥ * ∥y∥ * real.cos (angle x y) := -by rw [(show 2 * ∥x∥ * ∥y∥ * real.cos (angle x y) = - 2 * (real.cos (angle x y) * (∥x∥ * ∥y∥)), by ring), + ‖x - y‖ * ‖x - y‖ = ‖x‖ * ‖x‖ + ‖y‖ * ‖y‖ - 2 * ‖x‖ * ‖y‖ * real.cos (angle x y) := +by rw [(show 2 * ‖x‖ * ‖y‖ * real.cos (angle x y) = + 2 * (real.cos (angle x y) * (‖x‖ * ‖y‖)), by ring), cos_angle_mul_norm_mul_norm, ←real_inner_self_eq_norm_mul_norm, ←real_inner_self_eq_norm_mul_norm, ←real_inner_self_eq_norm_mul_norm, real_inner_sub_sub_self, sub_add_eq_add_sub] /-- Pons asinorum, vector angle form. -/ -lemma angle_sub_eq_angle_sub_rev_of_norm_eq {x y : V} (h : ∥x∥ = ∥y∥) : +lemma angle_sub_eq_angle_sub_rev_of_norm_eq {x y : V} (h : ‖x‖ = ‖y‖) : angle x (x - y) = angle y (y - x) := begin refine real.inj_on_cos ⟨angle_nonneg _ _, angle_le_pi _ _⟩ ⟨angle_nonneg _ _, angle_le_pi _ _⟩ _, @@ -73,14 +73,14 @@ end /-- Converse of pons asinorum, vector angle form. -/ lemma norm_eq_of_angle_sub_eq_angle_sub_rev_of_angle_ne_pi {x y : V} - (h : angle x (x - y) = angle y (y - x)) (hpi : angle x y ≠ π) : ∥x∥ = ∥y∥ := + (h : angle x (x - y) = angle y (y - x)) (hpi : angle x y ≠ π) : ‖x‖ = ‖y‖ := begin replace h := real.arccos_inj_on (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x (x - y))) (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one y (y - x))) h, by_cases hxy : x = y, { rw hxy }, - { rw [←norm_neg (y - x), neg_sub, mul_comm, mul_comm ∥y∥, div_eq_mul_inv, div_eq_mul_inv, + { rw [←norm_neg (y - x), neg_sub, mul_comm, mul_comm ‖y‖, div_eq_mul_inv, div_eq_mul_inv, mul_inv_rev, mul_inv_rev, ←mul_assoc, ←mul_assoc] at h, replace h := mul_right_cancel₀ (inv_ne_zero (λ hz, hxy (eq_of_sub_eq_zero (norm_eq_zero.1 hz)))) h, @@ -112,17 +112,17 @@ begin { rw [hxy, angle_self hy], simp }, { rw [real.cos_add, cos_angle, cos_angle, cos_angle], - have hxn : ∥x∥ ≠ 0 := (λ h, hx (norm_eq_zero.1 h)), - have hyn : ∥y∥ ≠ 0 := (λ h, hy (norm_eq_zero.1 h)), - have hxyn : ∥x - y∥ ≠ 0 := (λ h, hxy (eq_of_sub_eq_zero (norm_eq_zero.1 h))), + have hxn : ‖x‖ ≠ 0 := (λ h, hx (norm_eq_zero.1 h)), + have hyn : ‖y‖ ≠ 0 := (λ h, hy (norm_eq_zero.1 h)), + have hxyn : ‖x - y‖ ≠ 0 := (λ h, hxy (eq_of_sub_eq_zero (norm_eq_zero.1 h))), apply mul_right_cancel₀ hxn, apply mul_right_cancel₀ hyn, apply mul_right_cancel₀ hxyn, apply mul_right_cancel₀ hxyn, have H1 : real.sin (angle x (x - y)) * real.sin (angle y (y - x)) * - ∥x∥ * ∥y∥ * ∥x - y∥ * ∥x - y∥ = - (real.sin (angle x (x - y)) * (∥x∥ * ∥x - y∥)) * - (real.sin (angle y (y - x)) * (∥y∥ * ∥x - y∥)), { ring }, + ‖x‖ * ‖y‖ * ‖x - y‖ * ‖x - y‖ = + (real.sin (angle x (x - y)) * (‖x‖ * ‖x - y‖)) * + (real.sin (angle y (y - x)) * (‖y‖ * ‖x - y‖)), { ring }, have H2 : ⟪x, x⟫ * (⟪x, x⟫ - ⟪x, y⟫ - (⟪x, y⟫ - ⟪y, y⟫)) - (⟪x, x⟫ - ⟪x, y⟫) * (⟪x, x⟫ - ⟪x, y⟫) = ⟪x, x⟫ * ⟪y, y⟫ - ⟪x, y⟫ * ⟪x, y⟫, { ring }, @@ -149,19 +149,19 @@ begin { rw [hxy, angle_self hy], simp }, { rw [real.sin_add, cos_angle, cos_angle], - have hxn : ∥x∥ ≠ 0 := (λ h, hx (norm_eq_zero.1 h)), - have hyn : ∥y∥ ≠ 0 := (λ h, hy (norm_eq_zero.1 h)), - have hxyn : ∥x - y∥ ≠ 0 := (λ h, hxy (eq_of_sub_eq_zero (norm_eq_zero.1 h))), + have hxn : ‖x‖ ≠ 0 := (λ h, hx (norm_eq_zero.1 h)), + have hyn : ‖y‖ ≠ 0 := (λ h, hy (norm_eq_zero.1 h)), + have hxyn : ‖x - y‖ ≠ 0 := (λ h, hxy (eq_of_sub_eq_zero (norm_eq_zero.1 h))), apply mul_right_cancel₀ hxn, apply mul_right_cancel₀ hyn, apply mul_right_cancel₀ hxyn, apply mul_right_cancel₀ hxyn, - have H1 : real.sin (angle x (x - y)) * (⟪y, y - x⟫ / (∥y∥ * ∥y - x∥)) * ∥x∥ * ∥y∥ * ∥x - y∥ = - real.sin (angle x (x - y)) * (∥x∥ * ∥x - y∥) * - (⟪y, y - x⟫ / (∥y∥ * ∥y - x∥)) * ∥y∥, { ring }, - have H2 : ⟪x, x - y⟫ / (∥x∥ * ∥y - x∥) * real.sin (angle y (y - x)) * ∥x∥ * ∥y∥ * ∥y - x∥ = - ⟪x, x - y⟫ / (∥x∥ * ∥y - x∥) * - (real.sin (angle y (y - x)) * (∥y∥ * ∥y - x∥)) * ∥x∥, { ring }, + have H1 : real.sin (angle x (x - y)) * (⟪y, y - x⟫ / (‖y‖ * ‖y - x‖)) * ‖x‖ * ‖y‖ * ‖x - y‖ = + real.sin (angle x (x - y)) * (‖x‖ * ‖x - y‖) * + (⟪y, y - x⟫ / (‖y‖ * ‖y - x‖)) * ‖y‖, { ring }, + have H2 : ⟪x, x - y⟫ / (‖x‖ * ‖y - x‖) * real.sin (angle y (y - x)) * ‖x‖ * ‖y‖ * ‖y - x‖ = + ⟪x, x - y⟫ / (‖x‖ * ‖y - x‖) * + (real.sin (angle y (y - x)) * (‖y‖ * ‖y - x‖)) * ‖x‖, { ring }, have H3 : ⟪x, x⟫ * (⟪x, x⟫ - ⟪x, y⟫ - (⟪x, y⟫ - ⟪y, y⟫)) - (⟪x, x⟫ - ⟪x, y⟫) * (⟪x, x⟫ - ⟪x, y⟫) = ⟪x, x⟫ * ⟪y, y⟫ - ⟪x, y⟫ * ⟪x, y⟫, { ring }, diff --git a/src/geometry/manifold/complex.lean b/src/geometry/manifold/complex.lean index 506672ce6f86f..bbd2e7ddfff65 100644 --- a/src/geometry/manifold/complex.lean +++ b/src/geometry/manifold/complex.lean @@ -86,7 +86,7 @@ begin -- `f` pulled back by the chart at `p` has a local max at `chart_at E p p` have hf'' : is_local_max (norm ∘ f ∘ (chart_at E p).symm) (chart_at E p p), { refine filter.eventually_of_mem key₁ (λ z hz, _), - refine (hp₀ ((chart_at E p).symm z) hz).trans (_ : ∥f p₀∥ ≤ ∥f _∥), + refine (hp₀ ((chart_at E p).symm z) hz).trans (_ : ‖f p₀‖ ≤ ‖f _‖), rw [← hp, local_homeomorph.left_inv _ (mem_chart_source E p)] }, -- so by the maximum principle `f` is equal to `f p` near `p` obtain ⟨U, hU, hUf⟩ := (complex.eventually_eq_of_is_local_max_norm hf' hf'').exists_mem, diff --git a/src/geometry/manifold/instances/sphere.lean b/src/geometry/manifold/instances/sphere.lean index cc216dc0997ef..c7a7a668c29bf 100644 --- a/src/geometry/manifold/instances/sphere.lean +++ b/src/geometry/manifold/instances/sphere.lean @@ -104,25 +104,25 @@ projection. This is a map from the orthogonal complement of a unit vector `v` i space `E` to `E`; we will later prove that it takes values in the unit sphere. For most purposes, use `stereo_inv_fun`, not `stereo_inv_fun_aux`. -/ -def stereo_inv_fun_aux (w : E) : E := (∥w∥ ^ 2 + 4)⁻¹ • ((4:ℝ) • w + (∥w∥ ^ 2 - 4) • v) +def stereo_inv_fun_aux (w : E) : E := (‖w‖ ^ 2 + 4)⁻¹ • ((4:ℝ) • w + (‖w‖ ^ 2 - 4) • v) variables {v} @[simp] lemma stereo_inv_fun_aux_apply (w : E) : - stereo_inv_fun_aux v w = (∥w∥ ^ 2 + 4)⁻¹ • ((4:ℝ) • w + (∥w∥ ^ 2 - 4) • v) := + stereo_inv_fun_aux v w = (‖w‖ ^ 2 + 4)⁻¹ • ((4:ℝ) • w + (‖w‖ ^ 2 - 4) • v) := rfl -lemma stereo_inv_fun_aux_mem (hv : ∥v∥ = 1) {w : E} (hw : w ∈ (ℝ ∙ v)ᗮ) : +lemma stereo_inv_fun_aux_mem (hv : ‖v‖ = 1) {w : E} (hw : w ∈ (ℝ ∙ v)ᗮ) : stereo_inv_fun_aux v w ∈ (sphere (0:E) 1) := begin - have h₁ : 0 ≤ ∥w∥ ^ 2 + 4 := by nlinarith, - suffices : ∥(4:ℝ) • w + (∥w∥ ^ 2 - 4) • v∥ = ∥w∥ ^ 2 + 4, - { have h₂ : ∥w∥ ^ 2 + 4 ≠ 0 := by nlinarith, + have h₁ : 0 ≤ ‖w‖ ^ 2 + 4 := by nlinarith, + suffices : ‖(4:ℝ) • w + (‖w‖ ^ 2 - 4) • v‖ = ‖w‖ ^ 2 + 4, + { have h₂ : ‖w‖ ^ 2 + 4 ≠ 0 := by nlinarith, simp only [mem_sphere_zero_iff_norm, norm_smul, real.norm_eq_abs, abs_inv, this, abs_of_nonneg h₁, stereo_inv_fun_aux_apply], field_simp }, - suffices : ∥(4:ℝ) • w + (∥w∥ ^ 2 - 4) • v∥ ^ 2 = (∥w∥ ^ 2 + 4) ^ 2, - { have h₃ : 0 ≤ ∥stereo_inv_fun_aux v w∥ := norm_nonneg _, + suffices : ‖(4:ℝ) • w + (‖w‖ ^ 2 - 4) • v‖ ^ 2 = (‖w‖ ^ 2 + 4) ^ 2, + { have h₃ : 0 ≤ ‖stereo_inv_fun_aux v w‖ := norm_nonneg _, simpa [h₁, h₃, -one_pow] using this }, simp [norm_add_sq_real, norm_smul, inner_smul_left, inner_smul_right, inner_left_of_mem_orthogonal_singleton _ hw, mul_pow, real.norm_eq_abs, hv], @@ -132,12 +132,12 @@ end lemma has_fderiv_at_stereo_inv_fun_aux (v : E) : has_fderiv_at (stereo_inv_fun_aux v) (continuous_linear_map.id ℝ E) 0 := begin - have h₀ : has_fderiv_at (λ w : E, ∥w∥ ^ 2) (0 : E →L[ℝ] ℝ) 0, + have h₀ : has_fderiv_at (λ w : E, ‖w‖ ^ 2) (0 : E →L[ℝ] ℝ) 0, { convert (has_strict_fderiv_at_norm_sq _).has_fderiv_at, simp }, - have h₁ : has_fderiv_at (λ w : E, (∥w∥ ^ 2 + 4)⁻¹) (0 : E →L[ℝ] ℝ) 0, + have h₁ : has_fderiv_at (λ w : E, (‖w‖ ^ 2 + 4)⁻¹) (0 : E →L[ℝ] ℝ) 0, { convert (has_fderiv_at_inv _).comp _ (h₀.add (has_fderiv_at_const 4 0)); simp }, - have h₂ : has_fderiv_at (λ w, (4:ℝ) • w + (∥w∥ ^ 2 - 4) • v) + have h₂ : has_fderiv_at (λ w, (4:ℝ) • w + (‖w‖ ^ 2 - 4) • v) ((4:ℝ) • continuous_linear_map.id ℝ E) 0, { convert ((has_fderiv_at_const (4:ℝ) 0).smul (has_fderiv_at_id 0)).add ((h₀.sub (has_fderiv_at_const (4:ℝ) 0)).smul (has_fderiv_at_const v 0)), @@ -161,12 +161,12 @@ end lemma cont_diff_stereo_inv_fun_aux : cont_diff ℝ ⊤ (stereo_inv_fun_aux v) := begin - have h₀ : cont_diff ℝ ⊤ (λ w : E, ∥w∥ ^ 2) := cont_diff_norm_sq, - have h₁ : cont_diff ℝ ⊤ (λ w : E, (∥w∥ ^ 2 + 4)⁻¹), + have h₀ : cont_diff ℝ ⊤ (λ w : E, ‖w‖ ^ 2) := cont_diff_norm_sq, + have h₁ : cont_diff ℝ ⊤ (λ w : E, (‖w‖ ^ 2 + 4)⁻¹), { refine (h₀.add cont_diff_const).inv _, intros x, nlinarith }, - have h₂ : cont_diff ℝ ⊤ (λ w, (4:ℝ) • w + (∥w∥ ^ 2 - 4) • v), + have h₂ : cont_diff ℝ ⊤ (λ w, (4:ℝ) • w + (‖w‖ ^ 2 - 4) • v), { refine (cont_diff_const.smul cont_diff_id).add _, refine (h₀.sub cont_diff_const).smul cont_diff_const }, exact h₁.smul h₂ @@ -174,20 +174,20 @@ end /-- Stereographic projection, reverse direction. This is a map from the orthogonal complement of a unit vector `v` in an inner product space `E` to the unit sphere in `E`. -/ -def stereo_inv_fun (hv : ∥v∥ = 1) (w : (ℝ ∙ v)ᗮ) : sphere (0:E) 1 := +def stereo_inv_fun (hv : ‖v‖ = 1) (w : (ℝ ∙ v)ᗮ) : sphere (0:E) 1 := ⟨stereo_inv_fun_aux v (w:E), stereo_inv_fun_aux_mem hv w.2⟩ -@[simp] lemma stereo_inv_fun_apply (hv : ∥v∥ = 1) (w : (ℝ ∙ v)ᗮ) : - (stereo_inv_fun hv w : E) = (∥w∥ ^ 2 + 4)⁻¹ • ((4:ℝ) • w + (∥w∥ ^ 2 - 4) • v) := +@[simp] lemma stereo_inv_fun_apply (hv : ‖v‖ = 1) (w : (ℝ ∙ v)ᗮ) : + (stereo_inv_fun hv w : E) = (‖w‖ ^ 2 + 4)⁻¹ • ((4:ℝ) • w + (‖w‖ ^ 2 - 4) • v) := rfl -lemma stereo_inv_fun_ne_north_pole (hv : ∥v∥ = 1) (w : (ℝ ∙ v)ᗮ) : +lemma stereo_inv_fun_ne_north_pole (hv : ‖v‖ = 1) (w : (ℝ ∙ v)ᗮ) : stereo_inv_fun hv w ≠ (⟨v, by simp [hv]⟩ : sphere (0:E) 1) := begin refine subtype.ne_of_val_ne _, rw ← inner_lt_one_iff_real_of_norm_one _ hv, { have hw : ⟪v, w⟫_ℝ = 0 := inner_right_of_mem_orthogonal_singleton v w.2, - have hw' : (∥(w:E)∥ ^ 2 + 4)⁻¹ * (∥(w:E)∥ ^ 2 - 4) < 1, + have hw' : (‖(w:E)‖ ^ 2 + 4)⁻¹ * (‖(w:E)‖ ^ 2 - 4) < 1, { refine (inv_mul_lt_iff' _).mpr _, { nlinarith }, linarith }, @@ -196,12 +196,12 @@ begin { simpa using stereo_inv_fun_aux_mem hv w.2 } end -lemma continuous_stereo_inv_fun (hv : ∥v∥ = 1) : continuous (stereo_inv_fun hv) := +lemma continuous_stereo_inv_fun (hv : ‖v‖ = 1) : continuous (stereo_inv_fun hv) := continuous_induced_rng.2 (cont_diff_stereo_inv_fun_aux.continuous.comp continuous_subtype_coe) variables [complete_space E] -lemma stereo_left_inv (hv : ∥v∥ = 1) {x : sphere (0:E) 1} (hx : (x:E) ≠ v) : +lemma stereo_left_inv (hv : ‖v‖ = 1) {x : sphere (0:E) 1} (hx : (x:E) ≠ v) : stereo_inv_fun hv (stereo_to_fun v x) = x := begin ext, @@ -213,7 +213,7 @@ begin { convert eq_sum_orthogonal_projection_self_orthogonal_complement (ℝ ∙ v) x, exact (orthogonal_projection_unit_singleton ℝ hv x).symm }, have hvy : ⟪v, y⟫_ℝ = 0 := inner_right_of_mem_orthogonal_singleton v y.2, - have pythag : 1 = a ^ 2 + ∥y∥ ^ 2, + have pythag : 1 = a ^ 2 + ‖y‖ ^ 2, { have hvy' : ⟪a • v, y⟫_ℝ = 0 := by simp [inner_smul_left, hvy], convert norm_add_sq_eq_norm_sq_add_norm_sq_of_inner_eq_zero _ _ hvy' using 2, { simp [← split] }, @@ -223,19 +223,19 @@ begin have ha : 1 - a ≠ 0, { have : a < 1 := (inner_lt_one_iff_real_of_norm_one hv (by simp)).mpr hx.symm, linarith }, - have : 2 ^ 2 * ∥y∥ ^ 2 + 4 * (1 - a) ^ 2 ≠ 0, + have : 2 ^ 2 * ‖y‖ ^ 2 + 4 * (1 - a) ^ 2 ≠ 0, { refine ne_of_gt _, have := norm_nonneg (y:E), have : 0 < (1 - a) ^ 2 := sq_pos_of_ne_zero (1 - a) ha, nlinarith }, -- the core of the problem is these two algebraic identities: - have h₁ : (2 ^ 2 / (1 - a) ^ 2 * ∥y∥ ^ 2 + 4)⁻¹ * 4 * (2 / (1 - a)) = 1, + have h₁ : (2 ^ 2 / (1 - a) ^ 2 * ‖y‖ ^ 2 + 4)⁻¹ * 4 * (2 / (1 - a)) = 1, { field_simp, simp only [submodule.coe_norm] at *, nlinarith }, - have h₂ : (2 ^ 2 / (1 - a) ^ 2 * ∥y∥ ^ 2 + 4)⁻¹ * (2 ^ 2 / (1 - a) ^ 2 * ∥y∥ ^ 2 - 4) = a, + have h₂ : (2 ^ 2 / (1 - a) ^ 2 * ‖y‖ ^ 2 + 4)⁻¹ * (2 ^ 2 / (1 - a) ^ 2 * ‖y‖ ^ 2 - 4) = a, { field_simp, - transitivity (1 - a) ^ 2 * (a * (2 ^ 2 * ∥y∥ ^ 2 + 4 * (1 - a) ^ 2)), + transitivity (1 - a) ^ 2 * (a * (2 ^ 2 * ‖y‖ ^ 2 + 4 * (1 - a) ^ 2)), { congr, simp only [submodule.coe_norm] at *, nlinarith }, @@ -248,11 +248,11 @@ begin { simp [split, add_comm] } end -lemma stereo_right_inv (hv : ∥v∥ = 1) (w : (ℝ ∙ v)ᗮ) : +lemma stereo_right_inv (hv : ‖v‖ = 1) (w : (ℝ ∙ v)ᗮ) : stereo_to_fun v (stereo_inv_fun hv w) = w := begin - have : 2 / (1 - (∥(w:E)∥ ^ 2 + 4)⁻¹ * (∥(w:E)∥ ^ 2 - 4)) * (∥(w:E)∥ ^ 2 + 4)⁻¹ * 4 = 1, - { have : ∥(w:E)∥ ^ 2 + 4 ≠ 0 := by nlinarith, + have : 2 / (1 - (‖(w:E)‖ ^ 2 + 4)⁻¹ * (‖(w:E)‖ ^ 2 - 4)) * (‖(w:E)‖ ^ 2 + 4)⁻¹ * 4 = 1, + { have : ‖(w:E)‖ ^ 2 + 4 ≠ 0 := by nlinarith, have : (4:ℝ) + 4 ≠ 0 := by nlinarith, field_simp, ring }, @@ -270,7 +270,7 @@ end /-- Stereographic projection from the unit sphere in `E`, centred at a unit vector `v` in `E`; this is the version as a local homeomorphism. -/ -def stereographic (hv : ∥v∥ = 1) : local_homeomorph (sphere (0:E) 1) (ℝ ∙ v)ᗮ := +def stereographic (hv : ‖v‖ = 1) : local_homeomorph (sphere (0:E) 1) (ℝ ∙ v)ᗮ := { to_fun := (stereo_to_fun v) ∘ coe, inv_fun := stereo_inv_fun hv, source := {⟨v, by simp [hv]⟩}ᶜ, @@ -285,15 +285,15 @@ def stereographic (hv : ∥v∥ = 1) : local_homeomorph (sphere (0:E) 1) (ℝ (λ w h, h ∘ subtype.ext ∘ eq.symm ∘ (inner_eq_norm_mul_iff_of_norm_one hv (by simp)).mp), continuous_inv_fun := (continuous_stereo_inv_fun hv).continuous_on } -lemma stereographic_apply (hv : ∥v∥ = 1) (x : sphere (0 : E) 1) : +lemma stereographic_apply (hv : ‖v‖ = 1) (x : sphere (0 : E) 1) : stereographic hv x = (2 / ((1:ℝ) - inner v x)) • orthogonal_projection (ℝ ∙ v)ᗮ x := rfl -@[simp] lemma stereographic_source (hv : ∥v∥ = 1) : +@[simp] lemma stereographic_source (hv : ‖v‖ = 1) : (stereographic hv).source = {⟨v, by simp [hv]⟩}ᶜ := rfl -@[simp] lemma stereographic_target (hv : ∥v∥ = 1) : (stereographic hv).target = set.univ := rfl +@[simp] lemma stereographic_target (hv : ‖v‖ = 1) : (stereographic hv).target = set.univ := rfl @[simp] lemma stereographic_apply_neg (v : sphere (0:E) 1) : stereographic (norm_eq_of_mem_sphere v) (-v) = 0 := @@ -370,8 +370,8 @@ lemma stereographic'_symm_apply {n : ℕ} [fact (finrank ℝ E = n + 1)] let U : (ℝ ∙ (v:E))ᗮ ≃ₗᵢ[ℝ] euclidean_space ℝ (fin n) := (orthonormal_basis.from_orthogonal_span_singleton n (ne_zero_of_mem_unit_sphere v)).repr in - ((∥(U.symm x : E)∥ ^ 2 + 4)⁻¹ • (4 : ℝ) • (U.symm x : E) + - (∥(U.symm x : E)∥ ^ 2 + 4)⁻¹ • (∥(U.symm x : E)∥ ^ 2 - 4) • v) := + ((‖(U.symm x : E)‖ ^ 2 + 4)⁻¹ • (4 : ℝ) • (U.symm x : E) + + (‖(U.symm x : E)‖ ^ 2 + 4)⁻¹ • (‖(U.symm x : E)‖ ^ 2 - 4) • v) := by simp [real_inner_comm, stereographic, stereographic', ← submodule.coe_norm] /-! ### Smooth manifold structure on the sphere -/ @@ -439,7 +439,7 @@ begin convert (H₁.of_le le_top).comp' H₂ using 1, ext x, have hfxv : f x = -↑v ↔ ⟪f x, -↑v⟫_ℝ = 1, - { have hfx : ∥f x∥ = 1 := by simpa using hf' x, + { have hfx : ‖f x‖ = 1 := by simpa using hf' x, rw inner_eq_norm_mul_iff_of_norm_one hfx, exact norm_eq_of_mem_sphere (-v) }, dsimp [chart_at], diff --git a/src/information_theory/hamming.lean b/src/information_theory/hamming.lean index 01ff0c312fa65..c1aacfc128e44 100644 --- a/src/information_theory/hamming.lean +++ b/src/information_theory/hamming.lean @@ -293,13 +293,13 @@ instance : metric_space (hamming β) := instance [Π i, has_zero (β i)] : has_norm (hamming β) := ⟨λ x, hamming_norm (of_hamming x)⟩ @[simp, push_cast] lemma norm_eq_hamming_norm [Π i, has_zero (β i)] (x : hamming β) : - ∥x∥ = hamming_norm (of_hamming x) := rfl + ‖x‖ = hamming_norm (of_hamming x) := rfl instance [Π i, add_comm_group (β i)] : seminormed_add_comm_group (hamming β) := { dist_eq := by { push_cast, exact_mod_cast hamming_dist_eq_hamming_norm }, ..pi.add_comm_group } @[simp, push_cast] lemma nnnorm_eq_hamming_norm [Π i, add_comm_group (β i)] (x : hamming β) : - ∥x∥₊ = hamming_norm (of_hamming x) := rfl + ‖x‖₊ = hamming_norm (of_hamming x) := rfl instance [Π i, add_comm_group (β i)] : normed_add_comm_group (hamming β) := { ..hamming.seminormed_add_comm_group } diff --git a/src/linear_algebra/affine_space/affine_subspace.lean b/src/linear_algebra/affine_space/affine_subspace.lean index 478ccdcd6d8d5..3afebb52a98d1 100644 --- a/src/linear_algebra/affine_space/affine_subspace.lean +++ b/src/linear_algebra/affine_space/affine_subspace.lean @@ -22,7 +22,7 @@ This file defines affine subspaces (over modules) and the affine span of a set o various lemmas relating to the set of vectors in the `direction`, and relating the lattice structure on affine subspaces to that on their directions. -* `affine_subspace.parallel`, notation `‖`, gives the property of two affine subspaces being +* `affine_subspace.parallel`, notation `∥`, gives the property of two affine subspaces being parallel (one being a translate of the other). * `affine_span` gives the affine subspace spanned by a set of points, with `vector_span` giving its direction. `affine_span` is defined @@ -1551,9 +1551,9 @@ def parallel (s₁ s₂ : affine_subspace k P) : Prop := /- The notation should logically be U+2225 PARALLEL TO, but that is used globally for norms at present, and norms and parallelism are both widely used in geometry, so use U+2016 DOUBLE VERTICAL LINE (which is logically more appropriate for norms) instead here to avoid conflict. -/ -localized "infix (name := affine_subspace.parallel) ` ‖ `:50 := affine_subspace.parallel" in affine +localized "infix (name := affine_subspace.parallel) ` ∥ `:50 := affine_subspace.parallel" in affine -@[symm] lemma parallel.symm {s₁ s₂ : affine_subspace k P} (h : s₁ ‖ s₂) : s₂ ‖ s₁ := +@[symm] lemma parallel.symm {s₁ s₂ : affine_subspace k P} (h : s₁ ∥ s₂) : s₂ ∥ s₁ := begin rcases h with ⟨v, rfl⟩, refine ⟨-v, _⟩, @@ -1561,14 +1561,14 @@ begin coe_refl_to_affine_map, map_id] end -lemma parallel_comm {s₁ s₂ : affine_subspace k P} : s₁ ‖ s₂ ↔ s₂ ‖ s₁ := +lemma parallel_comm {s₁ s₂ : affine_subspace k P} : s₁ ∥ s₂ ↔ s₂ ∥ s₁ := ⟨parallel.symm, parallel.symm⟩ -@[refl] lemma parallel.refl (s : affine_subspace k P) : s ‖ s := +@[refl] lemma parallel.refl (s : affine_subspace k P) : s ∥ s := ⟨0, by simp⟩ -@[trans] lemma parallel.trans {s₁ s₂ s₃ : affine_subspace k P} (h₁₂ : s₁ ‖ s₂) (h₂₃ : s₂ ‖ s₃) : - s₁ ‖ s₃ := +@[trans] lemma parallel.trans {s₁ s₂ s₃ : affine_subspace k P} (h₁₂ : s₁ ∥ s₂) (h₂₃ : s₂ ∥ s₃) : + s₁ ∥ s₃ := begin rcases h₁₂ with ⟨v₁₂, rfl⟩, rcases h₂₃ with ⟨v₂₃, rfl⟩, @@ -1576,7 +1576,7 @@ begin rw [map_map, ←coe_trans_to_affine_map, ←const_vadd_add] end -lemma parallel.direction_eq {s₁ s₂ : affine_subspace k P} (h : s₁ ‖ s₂) : +lemma parallel.direction_eq {s₁ s₂ : affine_subspace k P} (h : s₁ ∥ s₂) : s₁.direction = s₂.direction := begin rcases h with ⟨v, rfl⟩, @@ -1584,7 +1584,7 @@ begin end @[simp] lemma parallel_bot_iff_eq_bot {s : affine_subspace k P} : - s ‖ ⊥ ↔ s = ⊥ := + s ∥ ⊥ ↔ s = ⊥ := begin refine ⟨λ h, _, λ h, h ▸ parallel.refl _⟩, rcases h with ⟨v, h⟩, @@ -1592,11 +1592,11 @@ begin end @[simp] lemma bot_parallel_iff_eq_bot {s : affine_subspace k P} : - ⊥ ‖ s ↔ s = ⊥ := + ⊥ ∥ s ↔ s = ⊥ := by rw [parallel_comm, parallel_bot_iff_eq_bot] lemma parallel_iff_direction_eq_and_eq_bot_iff_eq_bot {s₁ s₂ : affine_subspace k P} : - s₁ ‖ s₂ ↔ s₁.direction = s₂.direction ∧ (s₁ = ⊥ ↔ s₂ = ⊥) := + s₁ ∥ s₂ ↔ s₁.direction = s₂.direction ∧ (s₁ = ⊥ ↔ s₂ = ⊥) := begin refine ⟨λ h, ⟨h.direction_eq, _, _⟩, λ h, _⟩, { rintro rfl, exact bot_parallel_iff_eq_bot.1 h }, diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index ae47ff79ff557..04f2a970078c0 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -1863,26 +1863,26 @@ lemma measurable_nnnorm : measurable (nnnorm : α → ℝ≥0) := continuous_nnnorm.measurable @[measurability] -lemma measurable.nnnorm {f : β → α} (hf : measurable f) : measurable (λ a, ∥f a∥₊) := +lemma measurable.nnnorm {f : β → α} (hf : measurable f) : measurable (λ a, ‖f a‖₊) := measurable_nnnorm.comp hf @[measurability] lemma ae_measurable.nnnorm {f : β → α} {μ : measure β} (hf : ae_measurable f μ) : - ae_measurable (λ a, ∥f a∥₊) μ := + ae_measurable (λ a, ‖f a‖₊) μ := measurable_nnnorm.comp_ae_measurable hf @[measurability] -lemma measurable_ennnorm : measurable (λ x : α, (∥x∥₊ : ℝ≥0∞)) := +lemma measurable_ennnorm : measurable (λ x : α, (‖x‖₊ : ℝ≥0∞)) := measurable_nnnorm.coe_nnreal_ennreal @[measurability] lemma measurable.ennnorm {f : β → α} (hf : measurable f) : - measurable (λ a, (∥f a∥₊ : ℝ≥0∞)) := + measurable (λ a, (‖f a‖₊ : ℝ≥0∞)) := hf.nnnorm.coe_nnreal_ennreal @[measurability] lemma ae_measurable.ennnorm {f : β → α} {μ : measure β} (hf : ae_measurable f μ) : - ae_measurable (λ a, (∥f a∥₊ : ℝ≥0∞)) μ := + ae_measurable (λ a, (‖f a‖₊ : ℝ≥0∞)) μ := measurable_ennnorm.comp_ae_measurable hf end normed_add_comm_group diff --git a/src/measure_theory/constructions/prod.lean b/src/measure_theory/constructions/prod.lean index f9b1f98845397..6bb0900ec5f03 100644 --- a/src/measure_theory/constructions/prod.lean +++ b/src/measure_theory/constructions/prod.lean @@ -37,7 +37,7 @@ We also prove Tonelli's theorem and Fubini's theorem. is measurable. * `measure_theory.integrable_prod_iff` states that a binary function is integrable iff both * `y ↦ f (x, y)` is integrable for almost every `x`, and - * the function `x ↦ ∫ ∥f (x, y)∥ dy` is integrable. + * the function `x ↦ ∫ ‖f (x, y)‖ dy` is integrable. * `measure_theory.integral_prod`: Fubini's theorem. It states that for a integrable function `α × β → E` (where `E` is a second countable Banach space) we have `∫ z, f z ∂(μ.prod ν) = ∫ x, ∫ y, f (x, y) ∂ν ∂μ`. This theorem has the same variants as @@ -280,7 +280,7 @@ begin simp_rw [s', simple_func.coe_comp], exact simple_func.norm_approx_on_zero_le _ _ (x, y) n }, simp only [f', hfx, simple_func.integral_eq_integral _ (this _), indicator_of_mem, mem_set_of_eq], - refine tendsto_integral_of_dominated_convergence (λ y, ∥f x y∥ + ∥f x y∥) + refine tendsto_integral_of_dominated_convergence (λ y, ‖f x y‖ + ‖f x y‖) (λ n, (s' n x).ae_strongly_measurable) (hfx.norm.add hfx.norm) _ _, { exact λ n, eventually_of_forall (λ y, simple_func.norm_approx_on_zero_le _ _ (x, y) n) }, { refine eventually_of_forall (λ y, simple_func.tendsto_approx_on _ _ _), @@ -838,10 +838,10 @@ lemma integrable_swap_iff [sigma_finite μ] ⦃f : α × β → E⦄ : lemma has_finite_integral_prod_iff ⦃f : α × β → E⦄ (h1f : strongly_measurable f) : has_finite_integral f (μ.prod ν) ↔ (∀ᵐ x ∂ μ, has_finite_integral (λ y, f (x, y)) ν) ∧ - has_finite_integral (λ x, ∫ y, ∥f (x, y)∥ ∂ν) μ := + has_finite_integral (λ x, ∫ y, ‖f (x, y)‖ ∂ν) μ := begin simp only [has_finite_integral, lintegral_prod_of_measurable _ h1f.ennnorm], - have : ∀ x, ∀ᵐ y ∂ν, 0 ≤ ∥f (x, y)∥ := λ x, eventually_of_forall (λ y, norm_nonneg _), + have : ∀ x, ∀ᵐ y ∂ν, 0 ≤ ‖f (x, y)‖ := λ x, eventually_of_forall (λ y, norm_nonneg _), simp_rw [integral_eq_lintegral_of_nonneg_ae (this _) (h1f.norm.comp_measurable measurable_prod_mk_left).ae_strongly_measurable, ennnorm_eq_of_real to_real_nonneg, of_real_norm_eq_coe_nnnorm], @@ -857,7 +857,7 @@ end lemma has_finite_integral_prod_iff' ⦃f : α × β → E⦄ (h1f : ae_strongly_measurable f (μ.prod ν)) : has_finite_integral f (μ.prod ν) ↔ (∀ᵐ x ∂ μ, has_finite_integral (λ y, f (x, y)) ν) ∧ - has_finite_integral (λ x, ∫ y, ∥f (x, y)∥ ∂ν) μ := + has_finite_integral (λ x, ∫ y, ‖f (x, y)‖ ∂ν) μ := begin rw [has_finite_integral_congr h1f.ae_eq_mk, has_finite_integral_prod_iff h1f.strongly_measurable_mk], @@ -873,19 +873,19 @@ begin end /-- A binary function is integrable if the function `y ↦ f (x, y)` is integrable for almost every - `x` and the function `x ↦ ∫ ∥f (x, y)∥ dy` is integrable. -/ + `x` and the function `x ↦ ∫ ‖f (x, y)‖ dy` is integrable. -/ lemma integrable_prod_iff ⦃f : α × β → E⦄ (h1f : ae_strongly_measurable f (μ.prod ν)) : integrable f (μ.prod ν) ↔ - (∀ᵐ x ∂ μ, integrable (λ y, f (x, y)) ν) ∧ integrable (λ x, ∫ y, ∥f (x, y)∥ ∂ν) μ := + (∀ᵐ x ∂ μ, integrable (λ y, f (x, y)) ν) ∧ integrable (λ x, ∫ y, ‖f (x, y)‖ ∂ν) μ := by simp [integrable, h1f, has_finite_integral_prod_iff', h1f.norm.integral_prod_right', h1f.prod_mk_left] /-- A binary function is integrable if the function `x ↦ f (x, y)` is integrable for almost every - `y` and the function `y ↦ ∫ ∥f (x, y)∥ dx` is integrable. -/ + `y` and the function `y ↦ ∫ ‖f (x, y)‖ dx` is integrable. -/ lemma integrable_prod_iff' [sigma_finite μ] ⦃f : α × β → E⦄ (h1f : ae_strongly_measurable f (μ.prod ν)) : integrable f (μ.prod ν) ↔ - (∀ᵐ y ∂ ν, integrable (λ x, f (x, y)) μ) ∧ integrable (λ y, ∫ x, ∥f (x, y)∥ ∂μ) ν := + (∀ᵐ y ∂ ν, integrable (λ x, f (x, y)) μ) ∧ integrable (λ y, ∫ x, ‖f (x, y)‖ ∂μ) ν := by { convert integrable_prod_iff (h1f.prod_swap) using 1, rw [integrable_swap_iff] } lemma integrable.prod_left_ae [sigma_finite μ] ⦃f : α × β → E⦄ @@ -897,11 +897,11 @@ lemma integrable.prod_right_ae [sigma_finite μ] ⦃f : α × β → E⦄ hf.swap.prod_left_ae lemma integrable.integral_norm_prod_left ⦃f : α × β → E⦄ - (hf : integrable f (μ.prod ν)) : integrable (λ x, ∫ y, ∥f (x, y)∥ ∂ν) μ := + (hf : integrable f (μ.prod ν)) : integrable (λ x, ∫ y, ‖f (x, y)‖ ∂ν) μ := ((integrable_prod_iff hf.ae_strongly_measurable).mp hf).2 lemma integrable.integral_norm_prod_right [sigma_finite μ] ⦃f : α × β → E⦄ - (hf : integrable f (μ.prod ν)) : integrable (λ y, ∫ x, ∥f (x, y)∥ ∂μ) ν := + (hf : integrable f (μ.prod ν)) : integrable (λ y, ∫ x, ‖f (x, y)‖ ∂μ) ν := hf.swap.integral_norm_prod_left lemma integrable_prod_mul {f : α → ℝ} {g : β → ℝ} (hf : integrable f μ) (hg : integrable g ν) : @@ -1011,14 +1011,14 @@ begin rw [continuous_iff_continuous_at], intro g, refine tendsto_integral_of_L1 _ (L1.integrable_coe_fn g).integral_prod_left (eventually_of_forall $ λ h, (L1.integrable_coe_fn h).integral_prod_left) _, - simp_rw [← lintegral_fn_integral_sub (λ x, (∥x∥₊ : ℝ≥0∞)) (L1.integrable_coe_fn _) + simp_rw [← lintegral_fn_integral_sub (λ x, (‖x‖₊ : ℝ≥0∞)) (L1.integrable_coe_fn _) (L1.integrable_coe_fn g)], refine tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds _ (λ i, zero_le _) _, - { exact λ i, ∫⁻ x, ∫⁻ y, ∥i (x, y) - g (x, y)∥₊ ∂ν ∂μ }, + { exact λ i, ∫⁻ x, ∫⁻ y, ‖i (x, y) - g (x, y)‖₊ ∂ν ∂μ }, swap, { exact λ i, lintegral_mono (λ x, ennnorm_integral_le_lintegral_ennnorm _) }, show tendsto (λ (i : α × β →₁[μ.prod ν] E), - ∫⁻ x, ∫⁻ (y : β), ∥i (x, y) - g (x, y)∥₊ ∂ν ∂μ) (𝓝 g) (𝓝 0), - have : ∀ (i : α × β →₁[μ.prod ν] E), measurable (λ z, (∥i z - g z∥₊ : ℝ≥0∞)) := + ∫⁻ x, ∫⁻ (y : β), ‖i (x, y) - g (x, y)‖₊ ∂ν ∂μ) (𝓝 g) (𝓝 0), + have : ∀ (i : α × β →₁[μ.prod ν] E), measurable (λ z, (‖i z - g z‖₊ : ℝ≥0∞)) := λ i, ((Lp.strongly_measurable i).sub (Lp.strongly_measurable g)).ennnorm, simp_rw [← lintegral_prod_of_measurable _ (this _), ← L1.of_real_norm_sub_eq_lintegral, ← of_real_zero], diff --git a/src/measure_theory/covering/besicovitch_vector_space.lean b/src/measure_theory/covering/besicovitch_vector_space.lean index b8da20e0c2eb4..8c4b787747f5c 100644 --- a/src/measure_theory/covering/besicovitch_vector_space.lean +++ b/src/measure_theory/covering/besicovitch_vector_space.lean @@ -124,7 +124,7 @@ end satellite_config /-- The maximum cardinality of a `1`-separated set in the ball of radius `2`. This is also the optimal number of families in the Besicovitch covering theorem. -/ def multiplicity (E : Type*) [normed_add_comm_group E] := -Sup {N | ∃ s : finset E, s.card = N ∧ (∀ c ∈ s, ∥c∥ ≤ 2) ∧ (∀ c ∈ s, ∀ d ∈ s, c ≠ d → 1 ≤ ∥c - d∥)} +Sup {N | ∃ s : finset E, s.card = N ∧ (∀ c ∈ s, ‖c‖ ≤ 2) ∧ (∀ c ∈ s, ∀ d ∈ s, c ≠ d → 1 ≤ ‖c - d‖)} section variables [normed_space ℝ E] [finite_dimensional ℝ E] @@ -133,7 +133,7 @@ variables [normed_space ℝ E] [finite_dimensional ℝ E] useful to show that the supremum in the definition of `besicovitch.multiplicity E` is well behaved. -/ lemma card_le_of_separated - (s : finset E) (hs : ∀ c ∈ s, ∥c∥ ≤ 2) (h : ∀ (c ∈ s) (d ∈ s), c ≠ d → 1 ≤ ∥c - d∥) : + (s : finset E) (hs : ∀ c ∈ s, ‖c‖ ≤ 2) (h : ∀ (c ∈ s) (d ∈ s), c ≠ d → 1 ≤ ‖c - d‖) : s.card ≤ 5 ^ (finrank ℝ E) := begin /- We consider balls of radius `1/2` around the points in `s`. They are disjoint, and all @@ -187,7 +187,7 @@ begin end lemma card_le_multiplicity - {s : finset E} (hs : ∀ c ∈ s, ∥c∥ ≤ 2) (h's : ∀ (c ∈ s) (d ∈ s), c ≠ d → 1 ≤ ∥c - d∥) : + {s : finset E} (hs : ∀ c ∈ s, ‖c‖ ≤ 2) (h's : ∀ (c ∈ s) (d ∈ s), c ≠ d → 1 ≤ ‖c - d‖) : s.card ≤ multiplicity E := begin apply le_cSup, @@ -202,8 +202,8 @@ variable (E) /-- If `δ` is small enough, a `(1-δ)`-separated set in the ball of radius `2` also has cardinality at most `multiplicity E`. -/ -lemma exists_good_δ : ∃ (δ : ℝ), 0 < δ ∧ δ < 1 ∧ ∀ (s : finset E), (∀ c ∈ s, ∥c∥ ≤ 2) → - (∀ (c ∈ s) (d ∈ s), c ≠ d → 1 - δ ≤ ∥c - d∥) → s.card ≤ multiplicity E := +lemma exists_good_δ : ∃ (δ : ℝ), 0 < δ ∧ δ < 1 ∧ ∀ (s : finset E), (∀ c ∈ s, ‖c‖ ≤ 2) → + (∀ (c ∈ s) (d ∈ s), c ≠ d → 1 - δ ≤ ‖c - d‖) → s.card ≤ multiplicity E := begin /- This follows from a compactness argument: otherwise, one could extract a converging subsequence, to obtain a `1`-separated set in the ball of radius `2` with cardinality @@ -212,8 +212,8 @@ begin classical, by_contra' h, set N := multiplicity E + 1 with hN, - have : ∀ (δ : ℝ), 0 < δ → ∃ f : fin N → E, (∀ (i : fin N), ∥f i∥ ≤ 2) - ∧ (∀ i j, i ≠ j → 1 - δ ≤ ∥f i - f j∥), + have : ∀ (δ : ℝ), 0 < δ → ∃ f : fin N → E, (∀ (i : fin N), ‖f i‖ ≤ 2) + ∧ (∀ i j, i ≠ j → 1 - δ ≤ ‖f i - f j‖), { assume δ hδ, rcases lt_or_le δ 1 with hδ'|hδ', { rcases h δ hδ hδ' with ⟨s, hs, h's, s_card⟩, @@ -228,7 +228,7 @@ begin -- in the image are separated by `1 - δ`. choose! F hF using this, -- Choose a converging subsequence when `δ → 0`. - have : ∃ f : fin N → E, (∀ (i : fin N), ∥f i∥ ≤ 2) ∧ (∀ i j, i ≠ j → 1 ≤ ∥f i - f j∥), + have : ∃ f : fin N → E, (∀ (i : fin N), ‖f i‖ ≤ 2) ∧ (∀ i j, i ≠ j → 1 ≤ ‖f i - f j‖), { obtain ⟨u, u_mono, zero_lt_u, hu⟩ : ∃ (u : ℕ → ℝ), (∀ (m n : ℕ), m < n → u n < u m) ∧ (∀ (n : ℕ), 0 < u n) ∧ filter.tendsto u filter.at_top (𝓝 0) := exists_seq_strict_anti_tendsto (0 : ℝ), @@ -242,7 +242,7 @@ begin refine ⟨f, λ i, _, λ i j hij, _⟩, { simp only [pi_norm_le_iff_of_nonneg zero_le_two, mem_closed_ball, dist_zero_right] at fmem, exact fmem i }, - { have A : tendsto (λ n, ∥F (u (φ n)) i - F (u (φ n)) j∥) at_top (𝓝 (∥f i - f j∥)) := + { have A : tendsto (λ n, ‖F (u (φ n)) i - F (u (φ n)) j‖) at_top (𝓝 (‖f i - f j‖)) := ((hf.apply i).sub (hf.apply j)).norm, have B : tendsto (λ n, 1 - u (φ n)) at_top (𝓝 (1 - 0)) := tendsto_const_nhds.sub (hu.comp φ_mono.tendsto_at_top), @@ -253,16 +253,16 @@ begin have finj : function.injective f, { assume i j hij, by_contra, - have : 1 ≤ ∥f i - f j∥ := h'f i j h, + have : 1 ≤ ‖f i - f j‖ := h'f i j h, simp only [hij, norm_zero, sub_self] at this, exact lt_irrefl _ (this.trans_lt zero_lt_one) }, let s := finset.image f finset.univ, have s_card : s.card = N, by { rw finset.card_image_of_injective _ finj, exact finset.card_fin N }, - have hs : ∀ c ∈ s, ∥c∥ ≤ 2, + have hs : ∀ c ∈ s, ‖c‖ ≤ 2, by simp only [hf, forall_apply_eq_imp_iff', forall_const, forall_exists_index, finset.mem_univ, finset.mem_image], - have h's : ∀ (c ∈ s) (d ∈ s), c ≠ d → 1 ≤ ∥c - d∥, + have h's : ∀ (c ∈ s) (d ∈ s), c ≠ d → 1 ≤ ‖c - d‖, { simp only [s, forall_apply_eq_imp_iff', forall_exists_index, finset.mem_univ, finset.mem_image, ne.def, exists_true_left, forall_apply_eq_imp_iff', forall_true_left], assume i j hij, @@ -289,29 +289,29 @@ by { dsimp [good_τ, good_δ], linarith [(exists_good_δ E).some_spec.1] } variable {E} -lemma card_le_multiplicity_of_δ {s : finset E} (hs : ∀ c ∈ s, ∥c∥ ≤ 2) - (h's : ∀ (c ∈ s) (d ∈ s), c ≠ d → 1 - good_δ E ≤ ∥c - d∥) : +lemma card_le_multiplicity_of_δ {s : finset E} (hs : ∀ c ∈ s, ‖c‖ ≤ 2) + (h's : ∀ (c ∈ s) (d ∈ s), c ≠ d → 1 - good_δ E ≤ ‖c - d‖) : s.card ≤ multiplicity E := (classical.some_spec (exists_good_δ E)).2.2 s hs h's -lemma le_multiplicity_of_δ_of_fin {n : ℕ} (f : fin n → E) (h : ∀ i, ∥f i∥ ≤ 2) - (h' : ∀ i j, i ≠ j → 1 - good_δ E ≤ ∥f i - f j∥) : +lemma le_multiplicity_of_δ_of_fin {n : ℕ} (f : fin n → E) (h : ∀ i, ‖f i‖ ≤ 2) + (h' : ∀ i j, i ≠ j → 1 - good_δ E ≤ ‖f i - f j‖) : n ≤ multiplicity E := begin classical, have finj : function.injective f, { assume i j hij, by_contra, - have : 1 - good_δ E ≤ ∥f i - f j∥ := h' i j h, + have : 1 - good_δ E ≤ ‖f i - f j‖ := h' i j h, simp only [hij, norm_zero, sub_self] at this, linarith [good_δ_lt_one E] }, let s := finset.image f finset.univ, have s_card : s.card = n, by { rw finset.card_image_of_injective _ finj, exact finset.card_fin n }, - have hs : ∀ c ∈ s, ∥c∥ ≤ 2, + have hs : ∀ c ∈ s, ‖c‖ ≤ 2, by simp only [h, forall_apply_eq_imp_iff', forall_const, forall_exists_index, finset.mem_univ, finset.mem_image, implies_true_iff], - have h's : ∀ (c ∈ s) (d ∈ s), c ≠ d → 1 - good_δ E ≤ ∥c - d∥, + have h's : ∀ (c ∈ s) (d ∈ s), c ≠ d → 1 - good_δ E ≤ ‖c - d‖, { simp only [s, forall_apply_eq_imp_iff', forall_exists_index, finset.mem_univ, finset.mem_image, ne.def, exists_true_left, forall_apply_eq_imp_iff', forall_true_left], assume i j hij, @@ -331,23 +331,23 @@ namespace satellite_config We prove that the number of points in a satellite configuration is bounded by the maximal number of `1`-separated points in the ball of radius `2`. For this, start from a satellite congifuration `c`. Without loss of generality, one can assume that the last ball is centered at `0` and of -radius `1`. Define `c' i = c i` if `∥c i∥ ≤ 2`, and `c' i = (2/∥c i∥) • c i` if `∥c i∥ > 2`. +radius `1`. Define `c' i = c i` if `‖c i‖ ≤ 2`, and `c' i = (2/‖c i‖) • c i` if `‖c i‖ > 2`. It turns out that these points are `1 - δ`-separated, where `δ` is arbitrarily small if `τ` is close enough to `1`. The number of such configurations is bounded by `multiplicity E` if `δ` is suitably small. To check that the points `c' i` are `1 - δ`-separated, one treats separately the cases where -both `∥c i∥` and `∥c j∥` are `≤ 2`, where one of them is `≤ 2` and the other one is `> 2`, and +both `‖c i‖` and `‖c j‖` are `≤ 2`, where one of them is `≤ 2` and the other one is `> 2`, and where both of them are `> 2`. -/ lemma exists_normalized_aux1 {N : ℕ} {τ : ℝ} (a : satellite_config E N τ) (lastr : a.r (last N) = 1) (hτ : 1 ≤ τ) (δ : ℝ) (hδ1 : τ ≤ 1 + δ / 4) (hδ2 : δ ≤ 1) (i j : fin N.succ) (inej : i ≠ j) : - 1 - δ ≤ ∥a.c i - a.c j∥ := + 1 - δ ≤ ‖a.c i - a.c j‖ := begin - have ah : ∀ i j, i ≠ j → (a.r i ≤ ∥a.c i - a.c j∥ ∧ a.r j ≤ τ * a.r i) ∨ - (a.r j ≤ ∥a.c j - a.c i∥ ∧ a.r i ≤ τ * a.r j), + have ah : ∀ i j, i ≠ j → (a.r i ≤ ‖a.c i - a.c j‖ ∧ a.r j ≤ τ * a.r i) ∨ + (a.r j ≤ ‖a.c j - a.c i‖ ∧ a.r i ≤ τ * a.r j), by simpa only [dist_eq_norm] using a.h, have δnonneg : 0 ≤ δ := by linarith only [hτ, hδ1], have D : 0 ≤ 1 - δ / 4, by linarith only [hδ2], @@ -358,7 +358,7 @@ begin ... ≤ 1 : (by linarith only [sq_nonneg δ]), have J : 1 - δ ≤ 1 - δ / 4, by linarith only [δnonneg], have K : 1 - δ / 4 ≤ τ⁻¹, by { rw [inv_eq_one_div, le_div_iff τpos], exact I }, - suffices L : τ⁻¹ ≤ ∥a.c i - a.c j∥, by linarith only [J, K, L], + suffices L : τ⁻¹ ≤ ‖a.c i - a.c j‖, by linarith only [J, K, L], have hτ' : ∀ k, τ⁻¹ ≤ a.r k, { assume k, rw [inv_eq_one_div, div_le_iff τpos, ← lastr, mul_comm], @@ -376,16 +376,16 @@ variable [normed_space ℝ E] lemma exists_normalized_aux2 {N : ℕ} {τ : ℝ} (a : satellite_config E N τ) (lastc : a.c (last N) = 0) (lastr : a.r (last N) = 1) (hτ : 1 ≤ τ) (δ : ℝ) (hδ1 : τ ≤ 1 + δ / 4) (hδ2 : δ ≤ 1) - (i j : fin N.succ) (inej : i ≠ j) (hi : ∥a.c i∥ ≤ 2) (hj : 2 < ∥a.c j∥) : - 1 - δ ≤ ∥a.c i - (2 / ∥a.c j∥) • a.c j∥ := + (i j : fin N.succ) (inej : i ≠ j) (hi : ‖a.c i‖ ≤ 2) (hj : 2 < ‖a.c j‖) : + 1 - δ ≤ ‖a.c i - (2 / ‖a.c j‖) • a.c j‖ := begin - have ah : ∀ i j, i ≠ j → (a.r i ≤ ∥a.c i - a.c j∥ ∧ a.r j ≤ τ * a.r i) ∨ - (a.r j ≤ ∥a.c j - a.c i∥ ∧ a.r i ≤ τ * a.r j), + have ah : ∀ i j, i ≠ j → (a.r i ≤ ‖a.c i - a.c j‖ ∧ a.r j ≤ τ * a.r i) ∨ + (a.r j ≤ ‖a.c j - a.c i‖ ∧ a.r i ≤ τ * a.r j), by simpa only [dist_eq_norm] using a.h, have δnonneg : 0 ≤ δ := by linarith only [hτ, hδ1], have D : 0 ≤ 1 - δ / 4, by linarith only [hδ2], have τpos : 0 < τ := _root_.zero_lt_one.trans_le hτ, - have hcrj : ∥a.c j∥ ≤ a.r j + 1, + have hcrj : ‖a.c j‖ ≤ a.r j + 1, by simpa only [lastc, lastr, dist_zero_right] using a.inter' j, have I : a.r i ≤ 2, { rcases lt_or_le i (last N) with H|H, @@ -398,7 +398,7 @@ begin (1 - δ / 4) * τ ≤ (1 - δ / 4) * (1 + δ / 4) : mul_le_mul_of_nonneg_left hδ1 D ... = 1 - δ^2 / 16 : by ring ... ≤ 1 : (by linarith only [sq_nonneg δ]), - have A : a.r j - δ ≤ ∥a.c i - a.c j∥, + have A : a.r j - δ ≤ ‖a.c i - a.c j‖, { rcases ah j i inej.symm with H|H, { rw norm_sub_rev, linarith [H.1] }, have C : a.r j ≤ 4 := calc a.r j ≤ τ * a.r i : H.2 @@ -414,14 +414,14 @@ begin ... ≤ (1 - δ / 4) * (τ * a.r i) : mul_le_mul_of_nonneg_left (H.2) D ... ≤ 1 * a.r i : by { rw [← mul_assoc], apply mul_le_mul_of_nonneg_right J (a.rpos _).le } - ... ≤ ∥a.c i - a.c j∥ : by { rw [one_mul], exact H.1 } }, - set d := (2 / ∥a.c j∥) • a.c j with hd, - have : a.r j - δ ≤ ∥a.c i - d∥ + (a.r j - 1) := calc - a.r j - δ ≤ ∥a.c i - a.c j∥ : A - ... ≤ ∥a.c i - d∥ + ∥d - a.c j∥ : by simp only [← dist_eq_norm, dist_triangle] - ... ≤ ∥a.c i - d∥ + (a.r j - 1) : begin + ... ≤ ‖a.c i - a.c j‖ : by { rw [one_mul], exact H.1 } }, + set d := (2 / ‖a.c j‖) • a.c j with hd, + have : a.r j - δ ≤ ‖a.c i - d‖ + (a.r j - 1) := calc + a.r j - δ ≤ ‖a.c i - a.c j‖ : A + ... ≤ ‖a.c i - d‖ + ‖d - a.c j‖ : by simp only [← dist_eq_norm, dist_triangle] + ... ≤ ‖a.c i - d‖ + (a.r j - 1) : begin apply add_le_add_left, - have A : 0 ≤ 1 - 2 / ∥a.c j∥, by simpa [div_le_iff (zero_le_two.trans_lt hj)] using hj.le, + have A : 0 ≤ 1 - 2 / ‖a.c j‖, by simpa [div_le_iff (zero_le_two.trans_lt hj)] using hj.le, rw [← one_smul ℝ (a.c j), hd, ← sub_smul, norm_smul, norm_sub_rev, real.norm_eq_abs, abs_of_nonneg A, sub_mul], field_simp [(zero_le_two.trans_lt hj).ne'], @@ -433,17 +433,17 @@ end lemma exists_normalized_aux3 {N : ℕ} {τ : ℝ} (a : satellite_config E N τ) (lastc : a.c (last N) = 0) (lastr : a.r (last N) = 1) (hτ : 1 ≤ τ) (δ : ℝ) (hδ1 : τ ≤ 1 + δ / 4) - (i j : fin N.succ) (inej : i ≠ j) (hi : 2 < ∥a.c i∥) (hij : ∥a.c i∥ ≤ ∥a.c j∥) : - 1 - δ ≤ ∥(2 / ∥a.c i∥) • a.c i - (2 / ∥a.c j∥) • a.c j∥ := + (i j : fin N.succ) (inej : i ≠ j) (hi : 2 < ‖a.c i‖) (hij : ‖a.c i‖ ≤ ‖a.c j‖) : + 1 - δ ≤ ‖(2 / ‖a.c i‖) • a.c i - (2 / ‖a.c j‖) • a.c j‖ := begin - have ah : ∀ i j, i ≠ j → (a.r i ≤ ∥a.c i - a.c j∥ ∧ a.r j ≤ τ * a.r i) ∨ - (a.r j ≤ ∥a.c j - a.c i∥ ∧ a.r i ≤ τ * a.r j), + have ah : ∀ i j, i ≠ j → (a.r i ≤ ‖a.c i - a.c j‖ ∧ a.r j ≤ τ * a.r i) ∨ + (a.r j ≤ ‖a.c j - a.c i‖ ∧ a.r i ≤ τ * a.r j), by simpa only [dist_eq_norm] using a.h, have δnonneg : 0 ≤ δ := by linarith only [hτ, hδ1], have τpos : 0 < τ := _root_.zero_lt_one.trans_le hτ, - have hcrj : ∥a.c j∥ ≤ a.r j + 1, + have hcrj : ‖a.c j‖ ≤ a.r j + 1, by simpa only [lastc, lastr, dist_zero_right] using a.inter' j, - have A : a.r i ≤ ∥a.c i∥, + have A : a.r i ≤ ‖a.c i‖, { have : i < last N, { apply lt_top_iff_ne_top.2, assume iN, @@ -452,25 +452,25 @@ begin exact lt_irrefl _ (zero_le_two.trans_lt hi) }, convert (a.hlast i this).1, rw [dist_eq_norm, lastc, sub_zero] }, - have hj : 2 < ∥a.c j∥ := hi.trans_le hij, - set s := ∥a.c i∥ with hs, + have hj : 2 < ‖a.c j‖ := hi.trans_le hij, + set s := ‖a.c i‖ with hs, have spos : 0 < s := zero_lt_two.trans hi, - set d := (s/∥a.c j∥) • a.c j with hd, - have I : ∥a.c j - a.c i∥ ≤ ∥a.c j∥ - s + ∥d - a.c i∥ := calc - ∥a.c j - a.c i∥ ≤ ∥a.c j - d∥ + ∥d - a.c i∥ : by simp [← dist_eq_norm, dist_triangle] - ... = ∥a.c j∥ - ∥a.c i∥ + ∥d - a.c i∥ : begin + set d := (s/‖a.c j‖) • a.c j with hd, + have I : ‖a.c j - a.c i‖ ≤ ‖a.c j‖ - s + ‖d - a.c i‖ := calc + ‖a.c j - a.c i‖ ≤ ‖a.c j - d‖ + ‖d - a.c i‖ : by simp [← dist_eq_norm, dist_triangle] + ... = ‖a.c j‖ - ‖a.c i‖ + ‖d - a.c i‖ : begin nth_rewrite 0 ← one_smul ℝ (a.c j), rw [add_left_inj, hd, ← sub_smul, norm_smul, real.norm_eq_abs, abs_of_nonneg, sub_mul, one_mul, div_mul_cancel _ (zero_le_two.trans_lt hj).ne'], rwa [sub_nonneg, div_le_iff (zero_lt_two.trans hj), one_mul], end, - have J : a.r j - ∥a.c j - a.c i∥ ≤ s / 2 * δ := calc - a.r j - ∥a.c j - a.c i∥ ≤ s * (τ - 1) : begin + have J : a.r j - ‖a.c j - a.c i‖ ≤ s / 2 * δ := calc + a.r j - ‖a.c j - a.c i‖ ≤ s * (τ - 1) : begin rcases ah j i inej.symm with H|H, - { calc a.r j - ∥a.c j - a.c i∥ ≤ 0 : sub_nonpos.2 H.1 + { calc a.r j - ‖a.c j - a.c i‖ ≤ 0 : sub_nonpos.2 H.1 ... ≤ s * (τ - 1) : mul_nonneg spos.le (sub_nonneg.2 hτ) }, { rw norm_sub_rev at H, - calc a.r j - ∥a.c j - a.c i∥ ≤ τ * a.r i - a.r i : sub_le_sub H.2 H.1 + calc a.r j - ‖a.c j - a.c i‖ ≤ τ * a.r i - a.r i : sub_le_sub H.2 H.1 ... = a.r i * (τ - 1) : by ring ... ≤ s * (τ - 1) : mul_le_mul_of_nonneg_right A (sub_nonneg.2 hτ) } end @@ -478,9 +478,9 @@ begin ... = s / 2 * δ : by ring, have invs_nonneg : 0 ≤ 2 / s := (div_nonneg zero_le_two (zero_le_two.trans hi.le)), calc 1 - δ = (2 / s) * (s / 2 - (s / 2) * δ) : by { field_simp [spos.ne'], ring } - ... ≤ (2 / s) * ∥d - a.c i∥ : + ... ≤ (2 / s) * ‖d - a.c i‖ : mul_le_mul_of_nonneg_left (by linarith only [hcrj, I, J, hi]) invs_nonneg - ... = ∥(2 / s) • a.c i - (2 / ∥a.c j∥) • a.c j∥ : begin + ... = ‖(2 / s) • a.c i - (2 / ‖a.c j‖) • a.c j‖ : begin conv_lhs { rw [norm_sub_rev, ← abs_of_nonneg invs_nonneg] }, rw [← real.norm_eq_abs, ← norm_smul, smul_sub, hd, smul_smul], congr' 3, @@ -491,33 +491,33 @@ end lemma exists_normalized {N : ℕ} {τ : ℝ} (a : satellite_config E N τ) (lastc : a.c (last N) = 0) (lastr : a.r (last N) = 1) (hτ : 1 ≤ τ) (δ : ℝ) (hδ1 : τ ≤ 1 + δ / 4) (hδ2 : δ ≤ 1) : - ∃ (c' : fin N.succ → E), (∀ n, ∥c' n∥ ≤ 2) ∧ (∀ i j, i ≠ j → 1 - δ ≤ ∥c' i - c' j∥) := + ∃ (c' : fin N.succ → E), (∀ n, ‖c' n‖ ≤ 2) ∧ (∀ i j, i ≠ j → 1 - δ ≤ ‖c' i - c' j‖) := begin - let c' : fin N.succ → E := λ i, if ∥a.c i∥ ≤ 2 then a.c i else (2 / ∥a.c i∥) • a.c i, - have norm_c'_le : ∀ i, ∥c' i∥ ≤ 2, + let c' : fin N.succ → E := λ i, if ‖a.c i‖ ≤ 2 then a.c i else (2 / ‖a.c i‖) • a.c i, + have norm_c'_le : ∀ i, ‖c' i‖ ≤ 2, { assume i, simp only [c'], split_ifs, { exact h }, - by_cases hi : ∥a.c i∥ = 0; + by_cases hi : ‖a.c i‖ = 0; field_simp [norm_smul, hi] }, refine ⟨c', λ n, norm_c'_le n, λ i j inej, _⟩, - -- up to exchanging `i` and `j`, one can assume `∥c i∥ ≤ ∥c j∥`. - wlog hij : ∥a.c i∥ ≤ ∥a.c j∥ := le_total (∥a.c i∥) (∥a.c j∥) using [i j, j i] tactic.skip, swap, + -- up to exchanging `i` and `j`, one can assume `‖c i‖ ≤ ‖c j‖`. + wlog hij : ‖a.c i‖ ≤ ‖a.c j‖ := le_total (‖a.c i‖) (‖a.c j‖) using [i j, j i] tactic.skip, swap, { assume i_ne_j, rw norm_sub_rev, exact this i_ne_j.symm }, - rcases le_or_lt (∥a.c j∥) 2 with Hj|Hj, - -- case `∥c j∥ ≤ 2` (and therefore also `∥c i∥ ≤ 2`) + rcases le_or_lt (‖a.c j‖) 2 with Hj|Hj, + -- case `‖c j‖ ≤ 2` (and therefore also `‖c i‖ ≤ 2`) { simp_rw [c', Hj, hij.trans Hj, if_true], exact exists_normalized_aux1 a lastr hτ δ hδ1 hδ2 i j inej }, - -- case `2 < ∥c j∥` - { have H'j : (∥a.c j∥ ≤ 2) ↔ false, by simpa only [not_le, iff_false] using Hj, - rcases le_or_lt (∥a.c i∥) 2 with Hi|Hi, - { -- case `∥c i∥ ≤ 2` + -- case `2 < ‖c j‖` + { have H'j : (‖a.c j‖ ≤ 2) ↔ false, by simpa only [not_le, iff_false] using Hj, + rcases le_or_lt (‖a.c i‖) 2 with Hi|Hi, + { -- case `‖c i‖ ≤ 2` simp_rw [c', Hi, if_true, H'j, if_false], exact exists_normalized_aux2 a lastc lastr hτ δ hδ1 hδ2 i j inej Hi Hj }, - { -- case `2 < ∥c i∥` - have H'i : (∥a.c i∥ ≤ 2) ↔ false, by simpa only [not_le, iff_false] using Hi, + { -- case `2 < ‖c i‖` + have H'i : (‖a.c i‖ ≤ 2) ↔ false, by simpa only [not_le, iff_false] using Hi, simp_rw [c', H'i, if_false, H'j, if_false], exact exists_normalized_aux3 a lastc lastr hτ δ hδ1 i j inej Hi hij } } end diff --git a/src/measure_theory/covering/density_theorem.lean b/src/measure_theory/covering/density_theorem.lean index cd37a11be4cb2..4e5cc91bedc99 100644 --- a/src/measure_theory/covering/density_theorem.lean +++ b/src/measure_theory/covering/density_theorem.lean @@ -156,7 +156,7 @@ lemma ae_tendsto_average_norm_sub {f : α → E} (hf : integrable f μ) (K : ℝ ∀ᵐ x ∂μ, ∀ {ι : Type*} {l : filter ι} (w : ι → α) (δ : ι → ℝ) (δlim : tendsto δ l (𝓝[>] 0)) (xmem : ∀ᶠ j in l, x ∈ closed_ball (w j) (K * δ j)), - tendsto (λ j, ⨍ y in closed_ball (w j) (δ j), ∥f y - f x∥ ∂μ) l (𝓝 0) := + tendsto (λ j, ⨍ y in closed_ball (w j) (δ j), ‖f y - f x‖ ∂μ) l (𝓝 0) := by filter_upwards [(vitali_family μ K).ae_tendsto_average_norm_sub hf] with x hx ι l w δ δlim xmem using hx.comp (tendsto_closed_ball_filter_at μ _ _ δlim xmem) diff --git a/src/measure_theory/covering/differentiation.lean b/src/measure_theory/covering/differentiation.lean index 3903a4bc70342..9babd8246a002 100644 --- a/src/measure_theory/covering/differentiation.lean +++ b/src/measure_theory/covering/differentiation.lean @@ -32,7 +32,7 @@ derived: * `vitali_family.ae_tendsto_measure_inter_div` states that, for almost every point `x ∈ s`, then `μ (s ∩ a) / μ a` tends to `1` as `a` shrinks to `x` along a Vitali family. * `vitali_family.ae_tendsto_average_norm_sub` states that, for almost every point `x`, then the - average of `y ↦ ∥f y - f x∥` on `a` tends to `0` as `a` shrinks to `x` along a Vitali family. + average of `y ↦ ‖f y - f x‖` on `a` tends to `0` as `a` shrinks to `x` along a Vitali family. ## Sketch of proof @@ -789,9 +789,9 @@ end lemma ae_tendsto_lintegral_nnnorm_sub_div' {f : α → E} (hf : integrable f μ) (h'f : strongly_measurable f) : - ∀ᵐ x ∂μ, tendsto (λ a, (∫⁻ y in a, ∥f y - f x∥₊ ∂μ) / μ a) (v.filter_at x) (𝓝 0) := + ∀ᵐ x ∂μ, tendsto (λ a, (∫⁻ y in a, ‖f y - f x‖₊ ∂μ) / μ a) (v.filter_at x) (𝓝 0) := begin - /- For every `c`, then `(∫⁻ y in a, ∥f y - c∥₊ ∂μ) / μ a` tends almost everywhere to `∥f x - c∥`. + /- For every `c`, then `(∫⁻ y in a, ‖f y - c‖₊ ∂μ) / μ a` tends almost everywhere to `‖f x - c‖`. We apply this to a countable set of `c` which is dense in the range of `f`, to deduce the desired convergence. A minor technical inconvenience is that constants are not integrable, so to apply previous lemmas @@ -800,16 +800,16 @@ begin let A := measure_theory.measure.finite_spanning_sets_in_open μ, rcases h'f.is_separable_range with ⟨t, t_count, ht⟩, have main : ∀ᵐ x ∂μ, ∀ (n : ℕ) (c : E) (hc : c ∈ t), - tendsto (λ a, (∫⁻ y in a, ∥f y - (A.set n).indicator (λ y, c) y∥₊ ∂μ) / μ a) - (v.filter_at x) (𝓝 (∥f x - (A.set n).indicator (λ y, c) x∥₊)), + tendsto (λ a, (∫⁻ y in a, ‖f y - (A.set n).indicator (λ y, c) y‖₊ ∂μ) / μ a) + (v.filter_at x) (𝓝 (‖f x - (A.set n).indicator (λ y, c) x‖₊)), { simp_rw [ae_all_iff, ae_ball_iff t_count], assume n c hc, apply ae_tendsto_lintegral_div', { refine (h'f.sub _).ennnorm, exact strongly_measurable_const.indicator (is_open.measurable_set (A.set_mem n)) }, { apply ne_of_lt, - calc ∫⁻ y, ↑∥f y - (A.set n).indicator (λ (y : α), c) y∥₊ ∂μ - ≤ ∫⁻ y, (∥f y∥₊ + ∥(A.set n).indicator (λ (y : α), c) y∥₊) ∂μ : + calc ∫⁻ y, ↑‖f y - (A.set n).indicator (λ (y : α), c) y‖₊ ∂μ + ≤ ∫⁻ y, (‖f y‖₊ + ‖(A.set n).indicator (λ (y : α), c) y‖₊) ∂μ : begin apply lintegral_mono, assume x, @@ -817,7 +817,7 @@ begin rw ← ennreal.coe_add, exact ennreal.coe_le_coe.2 (nnnorm_sub_le _ _), end - ... = ∫⁻ y, ∥f y∥₊ ∂μ + ∫⁻ y, ∥(A.set n).indicator (λ (y : α), c) y∥₊ ∂μ : + ... = ∫⁻ y, ‖f y‖₊ ∂μ + ∫⁻ y, ‖(A.set n).indicator (λ (y : α), c) y‖₊ ∂μ : lintegral_add_left h'f.ennnorm _ ... < ∞ + ∞ : begin @@ -827,8 +827,8 @@ begin exact ennreal.add_lt_add hf.2 I.2, end } }, filter_upwards [main, v.ae_eventually_measure_pos] with x hx h'x, - have M : ∀ c ∈ t, tendsto (λ a, (∫⁻ y in a, ∥f y - c∥₊ ∂μ) / μ a) - (v.filter_at x) (𝓝 (∥f x - c∥₊)), + have M : ∀ c ∈ t, tendsto (λ a, (∫⁻ y in a, ‖f y - c‖₊ ∂μ) / μ a) + (v.filter_at x) (𝓝 (‖f x - c‖₊)), { assume c hc, obtain ⟨n, xn⟩ : ∃ n, x ∈ A.set n, by simpa [← A.spanning] using mem_univ x, specialize hx n c hc, @@ -843,20 +843,20 @@ begin assume hy, simp only [ha hy, indicator_of_mem] }, apply ennreal.tendsto_nhds_zero.2 (λ ε εpos, _), - obtain ⟨c, ct, xc⟩ : ∃ c ∈ t, (∥f x - c∥₊ : ℝ≥0∞) < ε / 2, + obtain ⟨c, ct, xc⟩ : ∃ c ∈ t, (‖f x - c‖₊ : ℝ≥0∞) < ε / 2, { simp_rw ← edist_eq_coe_nnnorm_sub, have : f x ∈ closure t, from ht (mem_range_self _), exact emetric.mem_closure_iff.1 this (ε / 2) (ennreal.half_pos (ne_of_gt εpos)) }, filter_upwards [(tendsto_order.1 (M c ct)).2 (ε / 2) xc, h'x, v.eventually_measure_lt_top x] with a ha h'a h''a, apply ennreal.div_le_of_le_mul, - calc ∫⁻ y in a, ∥f y - f x∥₊ ∂μ - ≤ ∫⁻ y in a, ∥f y - c∥₊ + ∥f x - c∥₊ ∂μ : + calc ∫⁻ y in a, ‖f y - f x‖₊ ∂μ + ≤ ∫⁻ y in a, ‖f y - c‖₊ + ‖f x - c‖₊ ∂μ : begin apply lintegral_mono (λ x, _), simpa only [← edist_eq_coe_nnnorm_sub] using edist_triangle_right _ _ _, end - ... = ∫⁻ y in a, ∥f y - c∥₊ ∂μ + ∫⁻ y in a, ∥f x - c∥₊ ∂μ : + ... = ∫⁻ y in a, ‖f y - c‖₊ ∂μ + ∫⁻ y in a, ‖f x - c‖₊ ∂μ : lintegral_add_right _ measurable_const ... ≤ ε / 2 * μ a + ε / 2 * μ a : begin @@ -870,7 +870,7 @@ begin end lemma ae_tendsto_lintegral_nnnorm_sub_div {f : α → E} (hf : integrable f μ) : - ∀ᵐ x ∂μ, tendsto (λ a, (∫⁻ y in a, ∥f y - f x∥₊ ∂μ) / μ a) (v.filter_at x) (𝓝 0) := + ∀ᵐ x ∂μ, tendsto (λ a, (∫⁻ y in a, ‖f y - f x‖₊ ∂μ) / μ a) (v.filter_at x) (𝓝 0) := begin have I : integrable (hf.1.mk f) μ, from hf.congr hf.1.ae_eq_mk, filter_upwards [v.ae_tendsto_lintegral_nnnorm_sub_div' I hf.1.strongly_measurable_mk, @@ -885,9 +885,9 @@ begin end /-- *Lebesgue differentiation theorem*: for almost every point `x`, the -average of `∥f y - f x∥` on `a` tends to `0` as `a` shrinks to `x` along a Vitali family.-/ +average of `‖f y - f x‖` on `a` tends to `0` as `a` shrinks to `x` along a Vitali family.-/ lemma ae_tendsto_average_norm_sub {f : α → E} (hf : integrable f μ) : - ∀ᵐ x ∂μ, tendsto (λ a, ⨍ y in a, ∥f y - f x∥ ∂μ) (v.filter_at x) (𝓝 0) := + ∀ᵐ x ∂μ, tendsto (λ a, ⨍ y in a, ‖f y - f x‖ ∂μ) (v.filter_at x) (𝓝 0) := begin filter_upwards [v.ae_tendsto_lintegral_nnnorm_sub_div hf, v.ae_eventually_measure_pos] with x hx h'x, @@ -896,7 +896,7 @@ begin apply tendsto.congr' _ this, filter_upwards [h'x, v.eventually_measure_lt_top x] with a ha h'a, simp only [function.comp_app, ennreal.to_real_div, set_average_eq, div_eq_inv_mul], - have A : integrable_on (λ y, (∥f y - f x∥₊ : ℝ)) a μ, + have A : integrable_on (λ y, (‖f y - f x‖₊ : ℝ)) a μ, { simp_rw [coe_nnnorm], exact (hf.integrable_on.sub (integrable_on_const.2 (or.inr h'a))).norm }, rw [lintegral_coe_eq_integral _ A, ennreal.to_real_of_real], diff --git a/src/measure_theory/function/ae_eq_of_integral.lean b/src/measure_theory/function/ae_eq_of_integral.lean index 077b2508b4c7d..b847808f2a0ce 100644 --- a/src/measure_theory/function/ae_eq_of_integral.lean +++ b/src/measure_theory/function/ae_eq_of_integral.lean @@ -77,28 +77,28 @@ lemma ae_eq_zero_of_forall_dual_of_is_separable [normed_add_comm_group E] [norme begin rcases ht with ⟨d, d_count, hd⟩, haveI : encodable d := d_count.to_encodable, - have : ∀ (x : d), ∃ g : E →L[𝕜] 𝕜, ∥g∥ ≤ 1 ∧ g x = ∥(x : E)∥ := λ x, exists_dual_vector'' 𝕜 x, + have : ∀ (x : d), ∃ g : E →L[𝕜] 𝕜, ‖g‖ ≤ 1 ∧ g x = ‖(x : E)‖ := λ x, exists_dual_vector'' 𝕜 x, choose s hs using this, have A : ∀ (a : E), a ∈ t → (∀ x, ⟪a, s x⟫ = (0 : 𝕜)) → a = 0, { assume a hat ha, contrapose! ha, - have a_pos : 0 < ∥a∥, by simp only [ha, norm_pos_iff, ne.def, not_false_iff], + have a_pos : 0 < ‖a‖, by simp only [ha, norm_pos_iff, ne.def, not_false_iff], have a_mem : a ∈ closure d := hd hat, - obtain ⟨x, hx⟩ : ∃ (x : d), dist a x < ∥a∥ / 2, - { rcases metric.mem_closure_iff.1 a_mem (∥a∥/2) (half_pos a_pos) with ⟨x, h'x, hx⟩, + obtain ⟨x, hx⟩ : ∃ (x : d), dist a x < ‖a‖ / 2, + { rcases metric.mem_closure_iff.1 a_mem (‖a‖/2) (half_pos a_pos) with ⟨x, h'x, hx⟩, exact ⟨⟨x, h'x⟩, hx⟩ }, use x, - have I : ∥a∥/2 < ∥(x : E)∥, - { have : ∥a∥ ≤ ∥(x : E)∥ + ∥a - x∥ := norm_le_insert' _ _, - have : ∥a - x∥ < ∥a∥/2, by rwa dist_eq_norm at hx, + have I : ‖a‖/2 < ‖(x : E)‖, + { have : ‖a‖ ≤ ‖(x : E)‖ + ‖a - x‖ := norm_le_insert' _ _, + have : ‖a - x‖ < ‖a‖/2, by rwa dist_eq_norm at hx, linarith }, assume h, - apply lt_irrefl (∥s x x∥), - calc ∥s x x∥ = ∥s x (x - a)∥ : by simp only [h, sub_zero, continuous_linear_map.map_sub] - ... ≤ 1 * ∥(x : E) - a∥ : continuous_linear_map.le_of_op_norm_le _ (hs x).1 _ - ... < ∥a∥ / 2 : by { rw [one_mul], rwa dist_eq_norm' at hx } - ... < ∥(x : E)∥ : I - ... = ∥s x x∥ : by rw [(hs x).2, is_R_or_C.norm_coe_norm] }, + apply lt_irrefl (‖s x x‖), + calc ‖s x x‖ = ‖s x (x - a)‖ : by simp only [h, sub_zero, continuous_linear_map.map_sub] + ... ≤ 1 * ‖(x : E) - a‖ : continuous_linear_map.le_of_op_norm_le _ (hs x).1 _ + ... < ‖a‖ / 2 : by { rw [one_mul], rwa dist_eq_norm' at hx } + ... < ‖(x : E)‖ : I + ... = ‖s x x‖ : by rw [(hs x).2, is_R_or_C.norm_coe_norm] }, have hfs : ∀ (y : d), ∀ᵐ x ∂μ, ⟪f x, s y⟫ = (0 : 𝕜), from λ y, hf (s y), have hf' : ∀ᵐ x ∂μ, ∀ (y : d), ⟪f x, s y⟫ = (0 : 𝕜), by rwa ae_all_iff, filter_upwards [hf', h't] with x hx h'x, @@ -252,7 +252,7 @@ begin have mus : μ s < ∞, { let c : ℝ≥0 := ⟨|b|, abs_nonneg _⟩, have c_pos : (c : ℝ≥0∞) ≠ 0, by simpa using hb_neg.ne, - calc μ s ≤ μ {x | (c : ℝ≥0∞) ≤ ∥f x∥₊} : + calc μ s ≤ μ {x | (c : ℝ≥0∞) ≤ ‖f x‖₊} : begin apply measure_mono, assume x hx, @@ -260,7 +260,7 @@ begin simpa only [nnnorm, abs_of_neg hb_neg, abs_of_neg (hx.trans_lt hb_neg), real.norm_eq_abs, subtype.mk_le_mk, neg_le_neg_iff, set.mem_set_of_eq, ennreal.coe_le_coe] using hx, end - ... ≤ (∫⁻ x, ∥f x∥₊ ∂μ) / c : + ... ≤ (∫⁻ x, ‖f x‖₊ ∂μ) / c : meas_ge_le_lintegral_div hfm.ae_measurable.ennnorm c_pos ennreal.coe_ne_top ... < ∞ : ennreal.div_lt_top (ne_of_lt hf.2) c_pos }, have h_int_gt : ∫ x in s, f x ∂μ ≤ b * (μ s).to_real, diff --git a/src/measure_theory/function/conditional_expectation/basic.lean b/src/measure_theory/function/conditional_expectation/basic.lean index ed3df4fb13910..4a71620469317 100644 --- a/src/measure_theory/function/conditional_expectation/basic.lean +++ b/src/measure_theory/function/conditional_expectation/basic.lean @@ -462,7 +462,7 @@ end /-- `Lp_meas_subgroup_to_Lp_trim` preserves the norm. -/ lemma Lp_meas_subgroup_to_Lp_trim_norm_map [hp : fact (1 ≤ p)] (hm : m ≤ m0) (f : Lp_meas_subgroup F m p μ) : - ∥Lp_meas_subgroup_to_Lp_trim F p μ hm f∥ = ∥f∥ := + ‖Lp_meas_subgroup_to_Lp_trim F p μ hm f‖ = ‖f‖ := begin rw [Lp.norm_def, snorm_trim hm (Lp.strongly_measurable _), snorm_congr_ae (Lp_meas_subgroup_to_Lp_trim_ae_eq hm _), Lp_meas_subgroup_coe, ← Lp.norm_def], @@ -848,13 +848,13 @@ variables {m m0 : measurable_space α} {μ : measure α} {s : set α} /-- Let `m` be a sub-σ-algebra of `m0`, `f` a `m0`-measurable function and `g` a `m`-measurable function, such that their integrals coincide on `m`-measurable sets with finite measure. -Then `∫ x in s, ∥g x∥ ∂μ ≤ ∫ x in s, ∥f x∥ ∂μ` on all `m`-measurable sets with finite measure. -/ +Then `∫ x in s, ‖g x‖ ∂μ ≤ ∫ x in s, ‖f x‖ ∂μ` on all `m`-measurable sets with finite measure. -/ lemma integral_norm_le_of_forall_fin_meas_integral_eq (hm : m ≤ m0) {f g : α → ℝ} (hf : strongly_measurable f) (hfi : integrable_on f s μ) (hg : strongly_measurable[m] g) (hgi : integrable_on g s μ) (hgf : ∀ t, measurable_set[m] t → μ t < ∞ → ∫ x in t, g x ∂μ = ∫ x in t, f x ∂μ) (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) : - ∫ x in s, ∥g x∥ ∂μ ≤ ∫ x in s, ∥f x∥ ∂μ := + ∫ x in s, ‖g x‖ ∂μ ≤ ∫ x in s, ‖f x‖ ∂μ := begin rw [integral_norm_eq_pos_sub_neg (hg.mono hm) hgi, integral_norm_eq_pos_sub_neg hf hfi], have h_meas_nonneg_g : measurable_set[m] {x | 0 ≤ g x}, @@ -884,14 +884,14 @@ end /-- Let `m` be a sub-σ-algebra of `m0`, `f` a `m0`-measurable function and `g` a `m`-measurable function, such that their integrals coincide on `m`-measurable sets with finite measure. -Then `∫⁻ x in s, ∥g x∥₊ ∂μ ≤ ∫⁻ x in s, ∥f x∥₊ ∂μ` on all `m`-measurable sets with finite +Then `∫⁻ x in s, ‖g x‖₊ ∂μ ≤ ∫⁻ x in s, ‖f x‖₊ ∂μ` on all `m`-measurable sets with finite measure. -/ lemma lintegral_nnnorm_le_of_forall_fin_meas_integral_eq (hm : m ≤ m0) {f g : α → ℝ} (hf : strongly_measurable f) (hfi : integrable_on f s μ) (hg : strongly_measurable[m] g) (hgi : integrable_on g s μ) (hgf : ∀ t, measurable_set[m] t → μ t < ∞ → ∫ x in t, g x ∂μ = ∫ x in t, f x ∂μ) (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) : - ∫⁻ x in s, ∥g x∥₊ ∂μ ≤ ∫⁻ x in s, ∥f x∥₊ ∂μ := + ∫⁻ x in s, ‖g x‖₊ ∂μ ≤ ∫⁻ x in s, ‖f x‖₊ ∂μ := begin rw [← of_real_integral_norm_eq_lintegral_nnnorm hfi, ← of_real_integral_norm_eq_lintegral_nnnorm hgi, ennreal.of_real_le_of_real_iff], @@ -935,10 +935,10 @@ lemma integrable_condexp_L2_of_is_finite_measure (hm : m ≤ m0) [is_finite_meas integrable (condexp_L2 𝕜 hm f) μ := integrable_on_univ.mp $ integrable_on_condexp_L2_of_measure_ne_top hm (measure_ne_top _ _) f -lemma norm_condexp_L2_le_one (hm : m ≤ m0) : ∥@condexp_L2 α E 𝕜 _ _ _ _ _ μ hm∥ ≤ 1 := +lemma norm_condexp_L2_le_one (hm : m ≤ m0) : ‖@condexp_L2 α E 𝕜 _ _ _ _ _ μ hm‖ ≤ 1 := by { haveI : fact (m ≤ m0) := ⟨hm⟩, exact orthogonal_projection_norm_le _, } -lemma norm_condexp_L2_le (hm : m ≤ m0) (f : α →₂[μ] E) : ∥condexp_L2 𝕜 hm f∥ ≤ ∥f∥ := +lemma norm_condexp_L2_le (hm : m ≤ m0) (f : α →₂[μ] E) : ‖condexp_L2 𝕜 hm f‖ ≤ ‖f‖ := ((@condexp_L2 _ E 𝕜 _ _ _ _ _ μ hm).le_op_norm f).trans (mul_le_of_le_one_left (norm_nonneg _) (norm_condexp_L2_le_one hm)) @@ -951,7 +951,7 @@ begin end lemma norm_condexp_L2_coe_le (hm : m ≤ m0) (f : α →₂[μ] E) : - ∥(condexp_L2 𝕜 hm f : α →₂[μ] E)∥ ≤ ∥f∥ := + ‖(condexp_L2 𝕜 hm f : α →₂[μ] E)‖ ≤ ‖f‖ := begin rw [Lp.norm_def, Lp.norm_def, ← Lp_meas_coe], refine (ennreal.to_real_le_to_real _ (Lp.snorm_ne_top _)).mpr (snorm_condexp_L2_le hm f), @@ -1003,15 +1003,15 @@ begin end lemma lintegral_nnnorm_condexp_L2_le (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) (f : Lp ℝ 2 μ) : - ∫⁻ x in s, ∥condexp_L2 ℝ hm f x∥₊ ∂μ ≤ ∫⁻ x in s, ∥f x∥₊ ∂μ := + ∫⁻ x in s, ‖condexp_L2 ℝ hm f x‖₊ ∂μ ≤ ∫⁻ x in s, ‖f x‖₊ ∂μ := begin let h_meas := Lp_meas.ae_strongly_measurable' (condexp_L2 ℝ hm f), let g := h_meas.some, have hg_meas : strongly_measurable[m] g, from h_meas.some_spec.1, have hg_eq : g =ᵐ[μ] condexp_L2 ℝ hm f, from h_meas.some_spec.2.symm, have hg_eq_restrict : g =ᵐ[μ.restrict s] condexp_L2 ℝ hm f, from ae_restrict_of_ae hg_eq, - have hg_nnnorm_eq : (λ x, (∥g x∥₊ : ℝ≥0∞)) - =ᵐ[μ.restrict s] (λ x, (∥condexp_L2 ℝ hm f x∥₊ : ℝ≥0∞)), + have hg_nnnorm_eq : (λ x, (‖g x‖₊ : ℝ≥0∞)) + =ᵐ[μ.restrict s] (λ x, (‖condexp_L2 ℝ hm f x‖₊ : ℝ≥0∞)), { refine hg_eq_restrict.mono (λ x hx, _), dsimp only, rw hx, }, @@ -1031,7 +1031,7 @@ lemma condexp_L2_ae_eq_zero_of_ae_eq_zero (hs : measurable_set[m] s) (hμs : μ {f : Lp ℝ 2 μ} (hf : f =ᵐ[μ.restrict s] 0) : condexp_L2 ℝ hm f =ᵐ[μ.restrict s] 0 := begin - suffices h_nnnorm_eq_zero : ∫⁻ x in s, ∥condexp_L2 ℝ hm f x∥₊ ∂μ = 0, + suffices h_nnnorm_eq_zero : ∫⁻ x in s, ‖condexp_L2 ℝ hm f x‖₊ ∂μ = 0, { rw lintegral_eq_zero_iff at h_nnnorm_eq_zero, refine h_nnnorm_eq_zero.mono (λ x hx, _), dsimp only at hx, @@ -1052,10 +1052,10 @@ end lemma lintegral_nnnorm_condexp_L2_indicator_le_real (hs : measurable_set s) (hμs : μ s ≠ ∞) (ht : measurable_set[m] t) (hμt : μ t ≠ ∞) : - ∫⁻ a in t, ∥condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a∥₊ ∂μ ≤ μ (s ∩ t) := + ∫⁻ a in t, ‖condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a‖₊ ∂μ ≤ μ (s ∩ t) := begin refine (lintegral_nnnorm_condexp_L2_le ht hμt _).trans (le_of_eq _), - have h_eq : ∫⁻ x in t, ∥(indicator_const_Lp 2 hs hμs (1 : ℝ)) x∥₊ ∂μ + have h_eq : ∫⁻ x in t, ‖(indicator_const_Lp 2 hs hμs (1 : ℝ)) x‖₊ ∂μ = ∫⁻ x in t, s.indicator (λ x, (1 : ℝ≥0∞)) x ∂μ, { refine lintegral_congr_ae (ae_restrict_of_ae _), refine (@indicator_const_Lp_coe_fn _ _ _ 2 _ _ _ hs hμs (1 : ℝ)).mono (λ x hx, _), @@ -1185,25 +1185,25 @@ variables {𝕜} lemma set_lintegral_nnnorm_condexp_L2_indicator_le (hm : m ≤ m0) (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : E') {t : set α} (ht : measurable_set[m] t) (hμt : μ t ≠ ∞) : - ∫⁻ a in t, ∥condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x) a∥₊ ∂μ ≤ μ (s ∩ t) * ∥x∥₊ := -calc ∫⁻ a in t, ∥condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x) a∥₊ ∂μ - = ∫⁻ a in t, ∥(condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a) • x∥₊ ∂μ : + ∫⁻ a in t, ‖condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x) a‖₊ ∂μ ≤ μ (s ∩ t) * ‖x‖₊ := +calc ∫⁻ a in t, ‖condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x) a‖₊ ∂μ + = ∫⁻ a in t, ‖(condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a) • x‖₊ ∂μ : set_lintegral_congr_fun (hm t ht) ((condexp_L2_indicator_ae_eq_smul 𝕜 hm hs hμs x).mono (λ a ha hat, by rw ha)) -... = ∫⁻ a in t, ∥condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a∥₊ ∂μ * ∥x∥₊ : +... = ∫⁻ a in t, ‖condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a‖₊ ∂μ * ‖x‖₊ : begin simp_rw [nnnorm_smul, ennreal.coe_mul], rw [lintegral_mul_const, Lp_meas_coe], exact (Lp.strongly_measurable _).ennnorm end -... ≤ μ (s ∩ t) * ∥x∥₊ : +... ≤ μ (s ∩ t) * ‖x‖₊ : ennreal.mul_le_mul (lintegral_nnnorm_condexp_L2_indicator_le_real hs hμs ht hμt) le_rfl lemma lintegral_nnnorm_condexp_L2_indicator_le (hm : m ≤ m0) (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : E') [sigma_finite (μ.trim hm)] : - ∫⁻ a, ∥condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x) a∥₊ ∂μ ≤ μ s * ∥x∥₊ := + ∫⁻ a, ‖condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x) a‖₊ ∂μ ≤ μ s * ‖x‖₊ := begin - refine lintegral_le_of_forall_fin_meas_le' hm (μ s * ∥x∥₊) _ (λ t ht hμt, _), + refine lintegral_le_of_forall_fin_meas_le' hm (μ s * ‖x‖₊) _ (λ t ht hμt, _), { rw Lp_meas_coe, exact (Lp.ae_strongly_measurable _).ennnorm }, refine (set_lintegral_nnnorm_condexp_L2_indicator_le hm hs hμs x ht hμt).trans _, @@ -1217,7 +1217,7 @@ lemma integrable_condexp_L2_indicator (hm : m ≤ m0) [sigma_finite (μ.trim hm) (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : E') : integrable (condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x)) μ := begin - refine integrable_of_forall_fin_meas_le' hm (μ s * ∥x∥₊) + refine integrable_of_forall_fin_meas_le' hm (μ s * ‖x‖₊) (ennreal.mul_lt_top hμs ennreal.coe_ne_top) _ _, { rw Lp_meas_coe, exact Lp.ae_strongly_measurable _, }, { refine λ t ht hμt, (set_lintegral_nnnorm_condexp_L2_indicator_le hm hs hμs x ht hμt).trans _, @@ -1272,25 +1272,25 @@ lemma condexp_ind_smul_ae_eq_smul (hm : m ≤ m0) (hs : measurable_set s) (hμs lemma set_lintegral_nnnorm_condexp_ind_smul_le (hm : m ≤ m0) (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) {t : set α} (ht : measurable_set[m] t) (hμt : μ t ≠ ∞) : - ∫⁻ a in t, ∥condexp_ind_smul hm hs hμs x a∥₊ ∂μ ≤ μ (s ∩ t) * ∥x∥₊ := -calc ∫⁻ a in t, ∥condexp_ind_smul hm hs hμs x a∥₊ ∂μ - = ∫⁻ a in t, ∥condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a • x∥₊ ∂μ : + ∫⁻ a in t, ‖condexp_ind_smul hm hs hμs x a‖₊ ∂μ ≤ μ (s ∩ t) * ‖x‖₊ := +calc ∫⁻ a in t, ‖condexp_ind_smul hm hs hμs x a‖₊ ∂μ + = ∫⁻ a in t, ‖condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a • x‖₊ ∂μ : set_lintegral_congr_fun (hm t ht) ((condexp_ind_smul_ae_eq_smul hm hs hμs x).mono (λ a ha hat, by rw ha )) -... = ∫⁻ a in t, ∥condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a∥₊ ∂μ * ∥x∥₊ : +... = ∫⁻ a in t, ‖condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a‖₊ ∂μ * ‖x‖₊ : begin simp_rw [nnnorm_smul, ennreal.coe_mul], rw [lintegral_mul_const, Lp_meas_coe], exact (Lp.strongly_measurable _).ennnorm end -... ≤ μ (s ∩ t) * ∥x∥₊ : +... ≤ μ (s ∩ t) * ‖x‖₊ : ennreal.mul_le_mul (lintegral_nnnorm_condexp_L2_indicator_le_real hs hμs ht hμt) le_rfl lemma lintegral_nnnorm_condexp_ind_smul_le (hm : m ≤ m0) (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) [sigma_finite (μ.trim hm)] : - ∫⁻ a, ∥condexp_ind_smul hm hs hμs x a∥₊ ∂μ ≤ μ s * ∥x∥₊ := + ∫⁻ a, ‖condexp_ind_smul hm hs hμs x a‖₊ ∂μ ≤ μ s * ‖x‖₊ := begin - refine lintegral_le_of_forall_fin_meas_le' hm (μ s * ∥x∥₊) _ (λ t ht hμt, _), + refine lintegral_le_of_forall_fin_meas_le' hm (μ s * ‖x‖₊) _ (λ t ht hμt, _), { exact (Lp.ae_strongly_measurable _).ennnorm }, refine (set_lintegral_nnnorm_condexp_ind_smul_le hm hs hμs x ht hμt).trans _, refine ennreal.mul_le_mul _ le_rfl, @@ -1303,7 +1303,7 @@ lemma integrable_condexp_ind_smul (hm : m ≤ m0) [sigma_finite (μ.trim hm)] (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : integrable (condexp_ind_smul hm hs hμs x) μ := begin - refine integrable_of_forall_fin_meas_le' hm (μ s * ∥x∥₊) + refine integrable_of_forall_fin_meas_le' hm (μ s * ‖x‖₊) (ennreal.mul_lt_top hμs ennreal.coe_ne_top) _ _, { exact Lp.ae_strongly_measurable _, }, { refine λ t ht hμt, (set_lintegral_nnnorm_condexp_ind_smul_le hm hs hμs x ht hμt).trans _, @@ -1444,17 +1444,17 @@ begin end lemma norm_condexp_ind_L1_fin_le (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : - ∥condexp_ind_L1_fin hm hs hμs x∥ ≤ (μ s).to_real * ∥x∥ := + ‖condexp_ind_L1_fin hm hs hμs x‖ ≤ (μ s).to_real * ‖x‖ := begin - have : 0 ≤ ∫ (a : α), ∥condexp_ind_L1_fin hm hs hμs x a∥ ∂μ, + have : 0 ≤ ∫ (a : α), ‖condexp_ind_L1_fin hm hs hμs x a‖ ∂μ, from integral_nonneg (λ a, norm_nonneg _), rw [L1.norm_eq_integral_norm, ← ennreal.to_real_of_real (norm_nonneg x), ← ennreal.to_real_mul, ← ennreal.to_real_of_real this, ennreal.to_real_le_to_real ennreal.of_real_ne_top (ennreal.mul_ne_top hμs ennreal.of_real_ne_top), of_real_integral_norm_eq_lintegral_nnnorm], swap, { rw [← mem_ℒp_one_iff_integrable], exact Lp.mem_ℒp _, }, - have h_eq : ∫⁻ a, ∥condexp_ind_L1_fin hm hs hμs x a∥₊ ∂μ - = ∫⁻ a, ∥condexp_ind_smul hm hs hμs x a∥₊ ∂μ, + have h_eq : ∫⁻ a, ‖condexp_ind_L1_fin hm hs hμs x a‖₊ ∂μ + = ∫⁻ a, ‖condexp_ind_smul hm hs hμs x a‖₊ ∂μ, { refine lintegral_congr_ae _, refine (condexp_ind_L1_fin_ae_eq_condexp_ind_smul hm hs hμs x).mono (λ z hz, _), dsimp only, @@ -1550,7 +1550,7 @@ begin end lemma norm_condexp_ind_L1_le (x : G) : - ∥condexp_ind_L1 hm μ s x∥ ≤ (μ s).to_real * ∥x∥ := + ‖condexp_ind_L1 hm μ s x‖ ≤ (μ s).to_real * ‖x‖ := begin by_cases hs : measurable_set s, swap, {simp_rw condexp_ind_L1_of_not_measurable_set hs, rw Lp.norm_zero, @@ -1617,10 +1617,10 @@ lemma condexp_ind_smul' [normed_space ℝ F] [smul_comm_class ℝ 𝕜 F] (c : condexp_ind hm μ s (c • x) = c • condexp_ind hm μ s x := condexp_ind_L1_smul' c x -lemma norm_condexp_ind_apply_le (x : G) : ∥condexp_ind hm μ s x∥ ≤ (μ s).to_real * ∥x∥ := +lemma norm_condexp_ind_apply_le (x : G) : ‖condexp_ind hm μ s x‖ ≤ (μ s).to_real * ‖x‖ := norm_condexp_ind_L1_le x -lemma norm_condexp_ind_le : ∥(condexp_ind hm μ s : G →L[ℝ] α →₁[μ] G)∥ ≤ (μ s).to_real := +lemma norm_condexp_ind_le : ‖(condexp_ind hm μ s : G →L[ℝ] α →₁[μ] G)‖ ≤ (μ s).to_real := continuous_linear_map.op_norm_le_bound _ ennreal.to_real_nonneg norm_condexp_ind_apply_le lemma condexp_ind_disjoint_union_apply (hs : measurable_set s) (ht : measurable_set t) @@ -2230,7 +2230,7 @@ end lemma tendsto_condexp_L1_of_dominated_convergence (hm : m ≤ m0) [sigma_finite (μ.trim hm)] {fs : ℕ → α → F'} {f : α → F'} (bound_fs : α → ℝ) (hfs_meas : ∀ n, ae_strongly_measurable (fs n) μ) (h_int_bound_fs : integrable bound_fs μ) - (hfs_bound : ∀ n, ∀ᵐ x ∂μ, ∥fs n x∥ ≤ bound_fs x) + (hfs_bound : ∀ n, ∀ᵐ x ∂μ, ‖fs n x‖ ≤ bound_fs x) (hfs : ∀ᵐ x ∂μ, tendsto (λ n, fs n x) at_top (𝓝 (f x))) : tendsto (λ n, condexp_L1 hm μ (fs n)) at_top (𝓝 (condexp_L1 hm μ f)) := tendsto_set_to_fun_of_dominated_convergence _ bound_fs hfs_meas h_int_bound_fs hfs_bound hfs @@ -2244,8 +2244,8 @@ lemma tendsto_condexp_unique (fs gs : ℕ → α → F') (f g : α → F') (hgs : ∀ᵐ x ∂μ, tendsto (λ n, gs n x) at_top (𝓝 (g x))) (bound_fs : α → ℝ) (h_int_bound_fs : integrable bound_fs μ) (bound_gs : α → ℝ) (h_int_bound_gs : integrable bound_gs μ) - (hfs_bound : ∀ n, ∀ᵐ x ∂μ, ∥fs n x∥ ≤ bound_fs x) - (hgs_bound : ∀ n, ∀ᵐ x ∂μ, ∥gs n x∥ ≤ bound_gs x) + (hfs_bound : ∀ n, ∀ᵐ x ∂μ, ‖fs n x‖ ≤ bound_fs x) + (hgs_bound : ∀ n, ∀ᵐ x ∂μ, ‖gs n x‖ ≤ bound_gs x) (hfg : ∀ n, μ[fs n | m] =ᵐ[μ] μ[gs n | m]) : μ[f | m] =ᵐ[μ] μ[g | m] := begin diff --git a/src/measure_theory/function/conditional_expectation/real.lean b/src/measure_theory/function/conditional_expectation/real.lean index 85bd4f776aab5..a93d5a547b9e7 100644 --- a/src/measure_theory/function/conditional_expectation/real.lean +++ b/src/measure_theory/function/conditional_expectation/real.lean @@ -196,7 +196,7 @@ lemma integrable.uniform_integrable_condexp {ι : Type*} [is_finite_measure μ] {g : α → ℝ} (hint : integrable g μ) {ℱ : ι → measurable_space α} (hℱ : ∀ i, ℱ i ≤ m0) : uniform_integrable (λ i, μ[g | ℱ i]) 1 μ := begin - have hmeas : ∀ n, ∀ C, measurable_set {x | C ≤ ∥μ[g | ℱ n] x∥₊} := + have hmeas : ∀ n, ∀ C, measurable_set {x | C ≤ ‖μ[g | ℱ n] x‖₊} := λ n C, measurable_set_le measurable_const (strongly_measurable_condexp.mono (hℱ n)).measurable.nnnorm, have hg : mem_ℒp g 1 μ := mem_ℒp_one_iff_integrable.2 hint, @@ -213,7 +213,7 @@ begin set C : ℝ≥0 := ⟨δ, hδ.le⟩⁻¹ * (snorm g 1 μ).to_nnreal with hC, have hCpos : 0 < C := mul_pos (nnreal.inv_pos.2 hδ) (ennreal.to_nnreal_pos hne hg.snorm_lt_top.ne), - have : ∀ n, μ {x : α | C ≤ ∥μ[g | ℱ n] x∥₊} ≤ ennreal.of_real δ, + have : ∀ n, μ {x : α | C ≤ ‖μ[g | ℱ n] x‖₊} ≤ ennreal.of_real δ, { intro n, have := mul_meas_ge_le_pow_snorm' μ one_ne_zero ennreal.one_ne_top ((@strongly_measurable_condexp _ _ _ _ _ (ℱ n) _ μ g).mono @@ -228,8 +228,8 @@ begin ennreal.coe_to_nnreal hg.snorm_lt_top.ne, ← mul_assoc, ← ennreal.of_real_eq_coe_nnreal, ← ennreal.of_real_mul hδ.le, mul_inv_cancel hδ.ne.symm, ennreal.of_real_one, one_mul], exact snorm_one_condexp_le_snorm _ }, - refine ⟨C, λ n, le_trans _ (h {x : α | C ≤ ∥μ[g | ℱ n] x∥₊} (hmeas n C) (this n))⟩, - have hmeasℱ : measurable_set[ℱ n] {x : α | C ≤ ∥μ[g | ℱ n] x∥₊} := + refine ⟨C, λ n, le_trans _ (h {x : α | C ≤ ‖μ[g | ℱ n] x‖₊} (hmeas n C) (this n))⟩, + have hmeasℱ : measurable_set[ℱ n] {x : α | C ≤ ‖μ[g | ℱ n] x‖₊} := @measurable_set_le _ _ _ _ _ (ℱ n) _ _ _ _ _ measurable_const (@measurable.nnnorm _ _ _ _ _ (ℱ n) _ strongly_measurable_condexp.measurable), rw ← snorm_congr_ae (condexp_indicator hint hmeasℱ), @@ -269,7 +269,7 @@ end lemma condexp_strongly_measurable_mul_of_bound (hm : m ≤ m0) [is_finite_measure μ] {f g : α → ℝ} (hf : strongly_measurable[m] f) (hg : integrable g μ) (c : ℝ) - (hf_bound : ∀ᵐ x ∂μ, ∥f x∥ ≤ c) : + (hf_bound : ∀ᵐ x ∂μ, ‖f x‖ ≤ c) : μ[f * g | m] =ᵐ[μ] f * μ[g | m] := begin let fs := hf.approx_bounded c, @@ -279,9 +279,9 @@ begin { simp only [hμ, ae_zero], }, haveI : μ.ae.ne_bot, by simp only [hμ, ae_ne_bot, ne.def, not_false_iff], have hc : 0 ≤ c, - { have h_exists : ∃ x, ∥f x∥ ≤ c := eventually.exists hf_bound, + { have h_exists : ∃ x, ‖f x‖ ≤ c := eventually.exists hf_bound, exact (norm_nonneg _).trans h_exists.some_spec, }, - have hfs_bound : ∀ n x, ∥fs n x∥ ≤ c := hf.norm_approx_bounded_le hc, + have hfs_bound : ∀ n x, ‖fs n x‖ ≤ c := hf.norm_approx_bounded_le hc, have hn_eq : ∀ n, μ[fs n * g | m] =ᵐ[μ] fs n * μ[g | m], from λ n, condexp_strongly_measurable_simple_func_mul hm _ hg, have : μ[f * μ[g|m]|m] = f * μ[g|m], @@ -289,7 +289,7 @@ begin exact integrable_condexp.bdd_mul' ((hf.mono hm).ae_strongly_measurable) hf_bound, }, rw ← this, refine tendsto_condexp_unique (λ n x, fs n x * g x) (λ n x, fs n x * μ[g|m] x) (f * g) - (f * μ[g|m]) _ _ _ _ (λ x, c * ∥g x∥) _ (λ x, c * ∥μ[g|m] x∥) _ _ _ _, + (f * μ[g|m]) _ _ _ _ (λ x, c * ‖g x‖) _ (λ x, c * ‖μ[g|m] x‖) _ _ _ _, { exact λ n, hg.bdd_mul' ((simple_func.strongly_measurable (fs n)).mono hm).ae_strongly_measurable (eventually_of_forall (hfs_bound n)), }, @@ -322,7 +322,7 @@ end lemma condexp_strongly_measurable_mul_of_bound₀ (hm : m ≤ m0) [is_finite_measure μ] {f g : α → ℝ} (hf : ae_strongly_measurable' m f μ) (hg : integrable g μ) (c : ℝ) - (hf_bound : ∀ᵐ x ∂μ, ∥f x∥ ≤ c) : + (hf_bound : ∀ᵐ x ∂μ, ‖f x‖ ≤ c) : μ[f * g | m] =ᵐ[μ] f * μ[g | m] := begin have : μ[f * g | m] =ᵐ[μ] μ[hf.mk f * g | m], diff --git a/src/measure_theory/function/continuous_map_dense.lean b/src/measure_theory/function/continuous_map_dense.lean index dd99888f8f17e..9db92100eb208 100644 --- a/src/measure_theory/function/continuous_map_dense.lean +++ b/src/measure_theory/function/continuous_map_dense.lean @@ -80,8 +80,8 @@ begin intros ε hε, -- A little bit of pre-emptive work, to find `η : ℝ≥0` which will be a margin small enough for -- our purposes - obtain ⟨η, hη_pos, hη_le⟩ : ∃ η, 0 < η ∧ (↑(∥bit0 (∥c∥)∥₊ * (2 * η) ^ (1 / p.to_real)) : ℝ) ≤ ε, - { have : filter.tendsto (λ x : ℝ≥0, ∥bit0 (∥c∥)∥₊ * (2 * x) ^ (1 / p.to_real)) (𝓝 0) (𝓝 0), + obtain ⟨η, hη_pos, hη_le⟩ : ∃ η, 0 < η ∧ (↑(‖bit0 (‖c‖)‖₊ * (2 * η) ^ (1 / p.to_real)) : ℝ) ≤ ε, + { have : filter.tendsto (λ x : ℝ≥0, ‖bit0 (‖c‖)‖₊ * (2 * x) ^ (1 / p.to_real)) (𝓝 0) (𝓝 0), { have : filter.tendsto (λ x : ℝ≥0, 2 * x) (𝓝 0) (𝓝 (2 * 0)) := filter.tendsto_id.const_mul 2, convert ((nnreal.continuous_at_rpow_const (or.inr hp₀')).tendsto.comp this).const_mul _, simp [hp₀''.ne'] }, @@ -118,24 +118,24 @@ begin exists_continuous_zero_one_of_closed u_open.is_closed_compl F_closed this, -- Multiply this by `c` to get a continuous approximation to the function `f`; the key point is -- that this is pointwise bounded by the indicator of the set `u \ F` - have g_norm : ∀ x, ∥g x∥ = g x := λ x, by rw [real.norm_eq_abs, abs_of_nonneg (hg_range x).1], - have gc_bd : ∀ x, ∥g x • c - s.indicator (λ x, c) x∥ ≤ ∥(u \ F).indicator (λ x, bit0 ∥c∥) x∥, + have g_norm : ∀ x, ‖g x‖ = g x := λ x, by rw [real.norm_eq_abs, abs_of_nonneg (hg_range x).1], + have gc_bd : ∀ x, ‖g x • c - s.indicator (λ x, c) x‖ ≤ ‖(u \ F).indicator (λ x, bit0 ‖c‖) x‖, { intros x, by_cases hu : x ∈ u, { rw ← set.diff_union_of_subset (Fs.trans su) at hu, cases hu with hFu hF, { refine (norm_sub_le _ _).trans _, refine (add_le_add_left (norm_indicator_le_norm_self (λ x, c) x) _).trans _, - have h₀ : g x * ∥c∥ + ∥c∥ ≤ 2 * ∥c∥, + have h₀ : g x * ‖c‖ + ‖c‖ ≤ 2 * ‖c‖, { nlinarith [(hg_range x).1, (hg_range x).2, norm_nonneg c] }, - have h₁ : (2:ℝ) * ∥c∥ = bit0 (∥c∥) := by simpa using add_mul (1:ℝ) 1 (∥c∥), + have h₁ : (2:ℝ) * ‖c‖ = bit0 (‖c‖) := by simpa using add_mul (1:ℝ) 1 (‖c‖), simp [hFu, norm_smul, h₀, ← h₁, g_norm x] }, { simp [hgF hF, Fs hF] } }, { have : x ∉ s := λ h, hu (su h), simp [hgu hu, this] } }, -- The rest is basically just `ennreal`-arithmetic have gc_snorm : snorm ((λ x, g x • c) - s.indicator (λ x, c)) p μ - ≤ (↑(∥bit0 (∥c∥)∥₊ * (2 * η) ^ (1 / p.to_real)) : ℝ≥0∞), + ≤ (↑(‖bit0 (‖c‖)‖₊ * (2 * η) ^ (1 / p.to_real)) : ℝ≥0∞), { refine (snorm_mono_ae (filter.eventually_of_forall gc_bd)).trans _, rw snorm_indicator_const (u_open.sdiff F_closed).measurable_set hp₀.ne' hp, push_cast [← ennreal.coe_rpow_of_nonneg _ hp₀'], @@ -153,9 +153,9 @@ begin rw [simple_func.coe_indicator_const, indicator_const_Lp, ← mem_ℒp.to_Lp_sub, Lp.norm_to_Lp], exact ennreal.to_real_le_coe_of_le_coe gc_snorm }, { rw [set_like.mem_coe, mem_bounded_continuous_function_iff], - refine ⟨bounded_continuous_function.of_normed_add_comm_group _ gc_cont (∥c∥) _, rfl⟩, + refine ⟨bounded_continuous_function.of_normed_add_comm_group _ gc_cont (‖c‖) _, rfl⟩, intros x, - have h₀ : g x * ∥c∥ ≤ ∥c∥, + have h₀ : g x * ‖c‖ ≤ ‖c‖, { nlinarith [(hg_range x).1, (hg_range x).2, norm_nonneg c] }, simp [norm_smul, g_norm x, h₀] }, end diff --git a/src/measure_theory/function/convergence_in_measure.lean b/src/measure_theory/function/convergence_in_measure.lean index 188b4f5bd627a..403082e262ea7 100644 --- a/src/measure_theory/function/convergence_in_measure.lean +++ b/src/measure_theory/function/convergence_in_measure.lean @@ -53,7 +53,7 @@ def tendsto_in_measure [has_dist E] {m : measurable_space α} lemma tendsto_in_measure_iff_norm [seminormed_add_comm_group E] {l : filter ι} {f : ι → α → E} {g : α → E} : tendsto_in_measure μ f l g - ↔ ∀ ε (hε : 0 < ε), tendsto (λ i, μ {x | ε ≤ ∥f i x - g x∥}) l (𝓝 0) := + ↔ ∀ ε (hε : 0 < ε), tendsto (λ i, μ {x | ε ≤ ‖f i x - g x‖}) l (𝓝 0) := by simp_rw [tendsto_in_measure, dist_eq_norm] namespace tendsto_in_measure @@ -340,7 +340,7 @@ begin refine hfg.mono (λ n hn, _), simp only [true_and, gt_iff_lt, ge_iff_le, zero_tsub, zero_le, zero_add, set.mem_Icc, pi.sub_apply] at *, - have : ess_sup (λ (x : α), (∥f n x - g x∥₊ : ℝ≥0∞)) μ < ennreal.of_real δ := + have : ess_sup (λ (x : α), (‖f n x - g x‖₊ : ℝ≥0∞)) μ < ennreal.of_real δ := lt_of_le_of_lt hn (ennreal.half_lt_self (ennreal.of_real_pos.2 hδ).ne.symm ennreal.of_real_lt_top.ne), refine ((le_of_eq _).trans (ae_lt_of_ess_sup_lt this).le).trans hε.le, diff --git a/src/measure_theory/function/jacobian.lean b/src/measure_theory/function/jacobian.lean index 6798ad18d1a78..389f7a7893e4a 100644 --- a/src/measure_theory/function/jacobian.lean +++ b/src/measure_theory/function/jacobian.lean @@ -133,7 +133,7 @@ begin -- `M n z` is the set of points `x` such that `f y - f x` is close to `f' z (y - x)` for `y` -- in the ball of radius `u n` around `x`. let M : ℕ → T → set E := λ n z, {x | x ∈ s ∧ - ∀ y ∈ s ∩ ball x (u n), ∥f y - f x - f' z (y - x)∥ ≤ r (f' z) * ∥y - x∥}, + ∀ y ∈ s ∩ ball x (u n), ‖f y - f x - f' z (y - x)‖ ≤ r (f' z) * ‖y - x‖}, -- As `f` is differentiable everywhere on `s`, the sets `M n z` cover `s` by design. have s_subset : ∀ x ∈ s, ∃ (n : ℕ) (z : T), x ∈ M n z, { assume x xs, @@ -143,30 +143,30 @@ begin refine mem_Union.2 ⟨⟨x, xs⟩, _⟩, simpa only [mem_ball, subtype.coe_mk, dist_self] using (rpos (f' x)).bot_lt }, rwa mem_Union₂ at this }, - obtain ⟨ε, εpos, hε⟩ : ∃ (ε : ℝ), 0 < ε ∧ ∥f' x - f' z∥ + ε ≤ r (f' z), - { refine ⟨r (f' z) - ∥f' x - f' z∥, _, le_of_eq (by abel)⟩, + obtain ⟨ε, εpos, hε⟩ : ∃ (ε : ℝ), 0 < ε ∧ ‖f' x - f' z‖ + ε ≤ r (f' z), + { refine ⟨r (f' z) - ‖f' x - f' z‖, _, le_of_eq (by abel)⟩, simpa only [sub_pos] using mem_ball_iff_norm.mp hz }, obtain ⟨δ, δpos, hδ⟩ : ∃ (δ : ℝ) (H : 0 < δ), - ball x δ ∩ s ⊆ {y | ∥f y - f x - (f' x) (y - x)∥ ≤ ε * ∥y - x∥} := + ball x δ ∩ s ⊆ {y | ‖f y - f x - (f' x) (y - x)‖ ≤ ε * ‖y - x‖} := metric.mem_nhds_within_iff.1 (is_o.def (hf' x xs) εpos), obtain ⟨n, hn⟩ : ∃ n, u n < δ := ((tendsto_order.1 u_lim).2 _ δpos).exists, refine ⟨n, ⟨z, zT⟩, ⟨xs, _⟩⟩, assume y hy, - calc ∥f y - f x - (f' z) (y - x)∥ - = ∥(f y - f x - (f' x) (y - x)) + (f' x - f' z) (y - x)∥ : + calc ‖f y - f x - (f' z) (y - x)‖ + = ‖(f y - f x - (f' x) (y - x)) + (f' x - f' z) (y - x)‖ : begin congr' 1, simp only [continuous_linear_map.coe_sub', map_sub, pi.sub_apply], abel, end - ... ≤ ∥f y - f x - (f' x) (y - x)∥ + ∥(f' x - f' z) (y - x)∥ : norm_add_le _ _ - ... ≤ ε * ∥y - x∥ + ∥f' x - f' z∥ * ∥y - x∥ : + ... ≤ ‖f y - f x - (f' x) (y - x)‖ + ‖(f' x - f' z) (y - x)‖ : norm_add_le _ _ + ... ≤ ε * ‖y - x‖ + ‖f' x - f' z‖ * ‖y - x‖ : begin refine add_le_add (hδ _) (continuous_linear_map.le_op_norm _ _), rw inter_comm, exact inter_subset_inter_right _ (ball_subset_ball hn.le) hy, end - ... ≤ r (f' z) * ∥y - x∥ : + ... ≤ r (f' z) * ‖y - x‖ : begin rw [← add_mul, add_comm], exact mul_le_mul_of_nonneg_right hε (norm_nonneg _), @@ -178,8 +178,8 @@ begin refine ⟨xs, λ y hy, _⟩, obtain ⟨a, aM, a_lim⟩ : ∃ (a : ℕ → E), (∀ k, a k ∈ M n z) ∧ tendsto a at_top (𝓝 x) := mem_closure_iff_seq_limit.1 hx, - have L1 : tendsto (λ (k : ℕ), ∥f y - f (a k) - (f' z) (y - a k)∥) at_top - (𝓝 ∥f y - f x - (f' z) (y - x)∥), + have L1 : tendsto (λ (k : ℕ), ‖f y - f (a k) - (f' z) (y - a k)‖) at_top + (𝓝 ‖f y - f x - (f' z) (y - x)‖), { apply tendsto.norm, have L : tendsto (λ k, f (a k)) at_top (𝓝 (f x)), { apply (hf' x xs).continuous_within_at.tendsto.comp, @@ -187,9 +187,9 @@ begin exact eventually_of_forall (λ k, (aM k).1) }, apply tendsto.sub (tendsto_const_nhds.sub L), exact ((f' z).continuous.tendsto _).comp (tendsto_const_nhds.sub a_lim) }, - have L2 : tendsto (λ (k : ℕ), (r (f' z) : ℝ) * ∥y - a k∥) at_top (𝓝 (r (f' z) * ∥y - x∥)) := + have L2 : tendsto (λ (k : ℕ), (r (f' z) : ℝ) * ‖y - a k‖) at_top (𝓝 (r (f' z) * ‖y - x‖)) := (tendsto_const_nhds.sub a_lim).norm.const_mul _, - have I : ∀ᶠ k in at_top, ∥f y - f (a k) - (f' z) (y - a k)∥ ≤ r (f' z) * ∥y - a k∥, + have I : ∀ᶠ k in at_top, ‖f y - f (a k) - (f' z) (y - a k)‖ ≤ r (f' z) * ‖y - a k‖, { have L : tendsto (λ k, dist y (a k)) at_top (𝓝 (dist y x)) := tendsto_const_nhds.dist a_lim, filter_upwards [(tendsto_order.1 L).2 _ hy.2], assume k hk, @@ -321,8 +321,8 @@ begin { apply mem_image_of_mem, simpa only [dist_eq_norm, mem_closed_ball, mem_closed_ball_zero_iff] using zr }, { rw [mem_closed_ball_iff_norm, add_sub_cancel], - calc ∥f z - f x - A (z - x)∥ - ≤ δ * ∥z - x∥ : hf _ zs _ xs + calc ‖f z - f x - A (z - x)‖ + ≤ δ * ‖z - x‖ : hf _ zs _ xs ... ≤ ε * r : mul_le_mul (le_of_lt hδ) (mem_closed_ball_iff_norm.1 zr) (norm_nonneg _) εpos.le }, { simp only [map_sub, pi.sub_apply], @@ -410,17 +410,17 @@ begin rcases (this.and self_mem_nhds_within).exists with ⟨δ₀, h, h'⟩, exact ⟨δ₀, h', h⟩, }, -- record smallness conditions for `δ` that will be needed to apply `hδ₀` below. - have L1 : ∀ᶠ δ in 𝓝 (0 : ℝ≥0), subsingleton E ∨ δ < ∥(B.symm : E →L[ℝ] E)∥₊⁻¹, + have L1 : ∀ᶠ δ in 𝓝 (0 : ℝ≥0), subsingleton E ∨ δ < ‖(B.symm : E →L[ℝ] E)‖₊⁻¹, { by_cases (subsingleton E), { simp only [h, true_or, eventually_const] }, simp only [h, false_or], apply Iio_mem_nhds, simpa only [h, false_or, nnreal.inv_pos] using B.subsingleton_or_nnnorm_symm_pos }, have L2 : ∀ᶠ δ in 𝓝 (0 : ℝ≥0), - ∥(B.symm : E →L[ℝ] E)∥₊ * (∥(B.symm : E →L[ℝ] E)∥₊⁻¹ - δ)⁻¹ * δ < δ₀, - { have : tendsto (λ δ, ∥(B.symm : E →L[ℝ] E)∥₊ * (∥(B.symm : E →L[ℝ] E)∥₊⁻¹ - δ)⁻¹ * δ) - (𝓝 0) (𝓝 (∥(B.symm : E →L[ℝ] E)∥₊ * (∥(B.symm : E →L[ℝ] E)∥₊⁻¹ - 0)⁻¹ * 0)), - { rcases eq_or_ne (∥(B.symm : E →L[ℝ] E)∥₊) 0 with H|H, + ‖(B.symm : E →L[ℝ] E)‖₊ * (‖(B.symm : E →L[ℝ] E)‖₊⁻¹ - δ)⁻¹ * δ < δ₀, + { have : tendsto (λ δ, ‖(B.symm : E →L[ℝ] E)‖₊ * (‖(B.symm : E →L[ℝ] E)‖₊⁻¹ - δ)⁻¹ * δ) + (𝓝 0) (𝓝 (‖(B.symm : E →L[ℝ] E)‖₊ * (‖(B.symm : E →L[ℝ] E)‖₊⁻¹ - 0)⁻¹ * 0)), + { rcases eq_or_ne (‖(B.symm : E →L[ℝ] E)‖₊) 0 with H|H, { simpa only [H, zero_mul] using tendsto_const_nhds }, refine tendsto.mul (tendsto_const_nhds.mul _) tendsto_id, refine (tendsto.sub tendsto_const_nhds tendsto_id).inv₀ _, @@ -447,12 +447,12 @@ begin end /-- If a differentiable function `f` is approximated by a linear map `A` on a set `s`, up to `δ`, -then at almost every `x` in `s` one has `∥f' x - A∥ ≤ δ`. -/ +then at almost every `x` in `s` one has `‖f' x - A‖ ≤ δ`. -/ lemma _root_.approximates_linear_on.norm_fderiv_sub_le {A : E →L[ℝ] E} {δ : ℝ≥0} (hf : approximates_linear_on f A s δ) (hs : measurable_set s) (f' : E → E →L[ℝ] E) (hf' : ∀ x ∈ s, has_fderiv_within_at f (f' x) s x) : - ∀ᵐ x ∂(μ.restrict s), ∥f' x - A∥₊ ≤ δ := + ∀ᵐ x ∂(μ.restrict s), ‖f' x - A‖₊ ≤ δ := begin /- The conclusion will hold at the Lebesgue density points of `s` (which have full measure). At such a point `x`, for any `z` and any `ε > 0` one has for small `r` @@ -464,11 +464,11 @@ begin assume x hx xs, -- consider an arbitrary vector `z`. apply continuous_linear_map.op_norm_le_bound _ δ.2 (λ z, _), - -- to show that `∥(f' x - A) z∥ ≤ δ ∥z∥`, it suffices to do it up to some error that vanishes + -- to show that `‖(f' x - A) z‖ ≤ δ ‖z‖`, it suffices to do it up to some error that vanishes -- asymptotically in terms of `ε > 0`. - suffices H : ∀ ε, 0 < ε → ∥(f' x - A) z∥ ≤ (δ + ε) * (∥z∥ + ε) + ∥(f' x - A)∥ * ε, - { have : tendsto (λ (ε : ℝ), ((δ : ℝ) + ε) * (∥z∥ + ε) + ∥(f' x - A)∥ * ε) (𝓝[>] 0) - (𝓝 ((δ + 0) * (∥z∥ + 0) + ∥(f' x - A)∥ * 0)) := + suffices H : ∀ ε, 0 < ε → ‖(f' x - A) z‖ ≤ (δ + ε) * (‖z‖ + ε) + ‖(f' x - A)‖ * ε, + { have : tendsto (λ (ε : ℝ), ((δ : ℝ) + ε) * (‖z‖ + ε) + ‖(f' x - A)‖ * ε) (𝓝[>] 0) + (𝓝 ((δ + 0) * (‖z‖ + 0) + ‖(f' x - A)‖ * 0)) := tendsto.mono_left (continuous.tendsto (by continuity) 0) nhds_within_le_nhds, simp only [add_zero, mul_zero] at this, apply le_of_tendsto_of_tendsto tendsto_const_nhds this, @@ -482,7 +482,7 @@ begin eventually_nonempty_inter_smul_of_density_one μ s x hx _ measurable_set_closed_ball (measure_closed_ball_pos μ z εpos).ne', obtain ⟨ρ, ρpos, hρ⟩ : - ∃ ρ > 0, ball x ρ ∩ s ⊆ {y : E | ∥f y - f x - (f' x) (y - x)∥ ≤ ε * ∥y - x∥} := + ∃ ρ > 0, ball x ρ ∩ s ⊆ {y : E | ‖f y - f x - (f' x) (y - x)‖ ≤ ε * ‖y - x‖} := mem_nhds_within_iff.1 (is_o.def (hf' x xs) εpos), -- for small enough `r`, the rescaled ball `r • closed_ball z ε` is included in the set where -- `f y - f x` is well approximated by `f' x (y - x)`. @@ -496,47 +496,47 @@ begin { simp only [mem_smul_set, image_add_left, mem_preimage, singleton_add] at hy, rcases hy with ⟨a, az, ha⟩, exact ⟨a, az, by simp only [ha, add_neg_cancel_left]⟩ }, - have norm_a : ∥a∥ ≤ ∥z∥ + ε := calc - ∥a∥ = ∥z + (a - z)∥ : by simp only [add_sub_cancel'_right] - ... ≤ ∥z∥ + ∥a - z∥ : norm_add_le _ _ - ... ≤ ∥z∥ + ε : add_le_add_left (mem_closed_ball_iff_norm.1 az) _, + have norm_a : ‖a‖ ≤ ‖z‖ + ε := calc + ‖a‖ = ‖z + (a - z)‖ : by simp only [add_sub_cancel'_right] + ... ≤ ‖z‖ + ‖a - z‖ : norm_add_le _ _ + ... ≤ ‖z‖ + ε : add_le_add_left (mem_closed_ball_iff_norm.1 az) _, -- use the approximation properties to control `(f' x - A) a`, and then `(f' x - A) z` as `z` is -- close to `a`. - have I : r * ∥(f' x - A) a∥ ≤ r * (δ + ε) * (∥z∥ + ε) := calc - r * ∥(f' x - A) a∥ = ∥(f' x - A) (r • a)∥ : + have I : r * ‖(f' x - A) a‖ ≤ r * (δ + ε) * (‖z‖ + ε) := calc + r * ‖(f' x - A) a‖ = ‖(f' x - A) (r • a)‖ : by simp only [continuous_linear_map.map_smul, norm_smul, real.norm_eq_abs, abs_of_nonneg rpos.le] - ... = ∥(f y - f x - A (y - x)) - - (f y - f x - (f' x) (y - x))∥ : + ... = ‖(f y - f x - A (y - x)) - + (f y - f x - (f' x) (y - x))‖ : begin congr' 1, simp only [ya, add_sub_cancel', sub_sub_sub_cancel_left, continuous_linear_map.coe_sub', eq_self_iff_true, sub_left_inj, pi.sub_apply, continuous_linear_map.map_smul, smul_sub], end - ... ≤ ∥f y - f x - A (y - x)∥ + - ∥f y - f x - (f' x) (y - x)∥ : norm_sub_le _ _ - ... ≤ δ * ∥y - x∥ + ε * ∥y - x∥ : + ... ≤ ‖f y - f x - A (y - x)‖ + + ‖f y - f x - (f' x) (y - x)‖ : norm_sub_le _ _ + ... ≤ δ * ‖y - x‖ + ε * ‖y - x‖ : add_le_add (hf _ ys _ xs) (hρ ⟨rρ hy, ys⟩) - ... = r * (δ + ε) * ∥a∥ : + ... = r * (δ + ε) * ‖a‖ : by { simp only [ya, add_sub_cancel', norm_smul, real.norm_eq_abs, abs_of_nonneg rpos.le], ring } - ... ≤ r * (δ + ε) * (∥z∥ + ε) : + ... ≤ r * (δ + ε) * (‖z‖ + ε) : mul_le_mul_of_nonneg_left norm_a (mul_nonneg rpos.le (add_nonneg δ.2 εpos.le)), - show ∥(f' x - A) z∥ ≤ (δ + ε) * (∥z∥ + ε) + ∥(f' x - A)∥ * ε, from calc - ∥(f' x - A) z∥ = ∥(f' x - A) a + (f' x - A) (z - a)∥ : + show ‖(f' x - A) z‖ ≤ (δ + ε) * (‖z‖ + ε) + ‖(f' x - A)‖ * ε, from calc + ‖(f' x - A) z‖ = ‖(f' x - A) a + (f' x - A) (z - a)‖ : begin congr' 1, simp only [continuous_linear_map.coe_sub', map_sub, pi.sub_apply], abel end - ... ≤ ∥(f' x - A) a∥ + ∥(f' x - A) (z - a)∥ : norm_add_le _ _ - ... ≤ (δ + ε) * (∥z∥ + ε) + ∥f' x - A∥ * ∥z - a∥ : + ... ≤ ‖(f' x - A) a‖ + ‖(f' x - A) (z - a)‖ : norm_add_le _ _ + ... ≤ (δ + ε) * (‖z‖ + ε) + ‖f' x - A‖ * ‖z - a‖ : begin apply add_le_add, { rw mul_assoc at I, exact (mul_le_mul_left rpos).1 I }, { apply continuous_linear_map.le_op_norm } end - ... ≤ (δ + ε) * (∥z∥ + ε) + ∥f' x - A∥ * ε : add_le_add le_rfl + ... ≤ (δ + ε) * (‖z‖ + ε) + ‖f' x - A‖ * ε : add_le_add le_rfl (mul_le_mul_of_nonneg_left (mem_closed_ball_iff_norm'.1 az) (norm_nonneg _)), end @@ -740,7 +740,7 @@ begin refine ae_sum_iff.2 (λ n, _), -- on almost all `s ∩ t n`, `f' x` is close to `A n` thanks to -- `approximates_linear_on.norm_fderiv_sub_le`. - have E₁ : ∀ᵐ (x : E) ∂μ.restrict (s ∩ t n), ∥f' x - A n∥₊ ≤ δ := + have E₁ : ∀ᵐ (x : E) ∂μ.restrict (s ∩ t n), ‖f' x - A n‖₊ ≤ δ := (ht n).norm_fderiv_sub_le μ (hs.inter (t_meas n)) f' (λ x hx, (hf' x hx.1).mono (inter_subset_left _ _)), -- moreover, `g x` is equal to `A n` there. @@ -812,7 +812,7 @@ begin `A n` (and where `f'` is almost everywhere close to `A n`), and then use that `f` expands the measure of such a set by at most `(A n).det + ε`. -/ have : ∀ (A : E →L[ℝ] E), ∃ (δ : ℝ≥0), 0 < δ ∧ - (∀ (B : E →L[ℝ] E), ∥B - A∥ ≤ δ → |B.det - A.det| ≤ ε) ∧ + (∀ (B : E →L[ℝ] E), ‖B - A‖ ≤ δ → |B.det - A.det| ≤ ε) ∧ ∀ (t : set E) (g : E → E) (hf : approximates_linear_on g A t δ), μ (g '' t) ≤ (ennreal.of_real (|A.det|) + ε) * μ t, { assume A, @@ -830,7 +830,7 @@ begin rw ← real.dist_eq, apply (hδ' B _).le, rw dist_eq_norm, - calc ∥B - A∥ ≤ (min δ δ'' : ℝ≥0) : hB + calc ‖B - A‖ ≤ (min δ δ'' : ℝ≥0) : hB ... ≤ δ'' : by simp only [le_refl, nnreal.coe_min, min_le_iff, or_true] ... < δ' : half_lt_self δ'pos }, { assume t g htg, @@ -957,7 +957,7 @@ begin well-approximated by linear maps `A n` (and where `f'` is almost everywhere close to `A n`), and then use that `f` expands the measure of such a set by at least `(A n).det - ε`. -/ have : ∀ (A : E →L[ℝ] E), ∃ (δ : ℝ≥0), 0 < δ ∧ - (∀ (B : E →L[ℝ] E), ∥B - A∥ ≤ δ → |B.det - A.det| ≤ ε) ∧ + (∀ (B : E →L[ℝ] E), ‖B - A‖ ≤ δ → |B.det - A.det| ≤ ε) ∧ ∀ (t : set E) (g : E → E) (hf : approximates_linear_on g A t δ), ennreal.of_real (|A.det|) * μ t ≤ μ (g '' t) + ε * μ t, { assume A, @@ -965,7 +965,7 @@ begin ∃ (δ' : ℝ) (H : 0 < δ'), ∀ B, dist B A < δ' → dist B.det A.det < ↑ε := continuous_at_iff.1 continuous_linear_map.continuous_det.continuous_at ε εpos, let δ'' : ℝ≥0 := ⟨δ' / 2, (half_pos δ'pos).le⟩, - have I'' : ∀ (B : E →L[ℝ] E), ∥B - A∥ ≤ ↑δ'' → |B.det - A.det| ≤ ↑ε, + have I'' : ∀ (B : E →L[ℝ] E), ‖B - A‖ ≤ ↑δ'' → |B.det - A.det| ≤ ↑ε, { assume B hB, rw ← real.dist_eq, apply (hδ' B _).le, diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index 4f6ac00fd7914..80b6a243cb362 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -29,7 +29,7 @@ classes of integrable functions, already defined as a special case of `L^p` spac ## Main definitions * Let `f : α → β` be a function, where `α` is a `measure_space` and `β` a `normed_add_comm_group`. - Then `has_finite_integral f` means `(∫⁻ a, ∥f a∥₊) < ∞`. + Then `has_finite_integral f` means `(∫⁻ a, ‖f a‖₊) < ∞`. * If `β` is moreover a `measurable_space` then `f` is called `integrable` if `f` is `measurable` and `has_finite_integral f` holds. @@ -60,11 +60,11 @@ namespace measure_theory /-! ### Some results about the Lebesgue integral involving a normed group -/ lemma lintegral_nnnorm_eq_lintegral_edist (f : α → β) : - ∫⁻ a, ∥f a∥₊ ∂μ = ∫⁻ a, edist (f a) 0 ∂μ := + ∫⁻ a, ‖f a‖₊ ∂μ = ∫⁻ a, edist (f a) 0 ∂μ := by simp only [edist_eq_coe_nnnorm] lemma lintegral_norm_eq_lintegral_edist (f : α → β) : - ∫⁻ a, (ennreal.of_real ∥f a∥) ∂μ = ∫⁻ a, edist (f a) 0 ∂μ := + ∫⁻ a, (ennreal.of_real ‖f a‖) ∂μ = ∫⁻ a, edist (f a) 0 ∂μ := by simp only [of_real_norm_eq_coe_nnnorm, edist_eq_coe_nnnorm] lemma lintegral_edist_triangle {f g h : α → β} @@ -76,31 +76,31 @@ begin apply edist_triangle_right end -lemma lintegral_nnnorm_zero : ∫⁻ a : α, ∥(0 : β)∥₊ ∂μ = 0 := by simp +lemma lintegral_nnnorm_zero : ∫⁻ a : α, ‖(0 : β)‖₊ ∂μ = 0 := by simp lemma lintegral_nnnorm_add_left {f : α → β} (hf : ae_strongly_measurable f μ) (g : α → γ) : - ∫⁻ a, ∥f a∥₊ + ∥g a∥₊ ∂μ = ∫⁻ a, ∥f a∥₊ ∂μ + ∫⁻ a, ∥g a∥₊ ∂μ := + ∫⁻ a, ‖f a‖₊ + ‖g a‖₊ ∂μ = ∫⁻ a, ‖f a‖₊ ∂μ + ∫⁻ a, ‖g a‖₊ ∂μ := lintegral_add_left' hf.ennnorm _ lemma lintegral_nnnorm_add_right (f : α → β) {g : α → γ} (hg : ae_strongly_measurable g μ) : - ∫⁻ a, ∥f a∥₊ + ∥g a∥₊ ∂μ = ∫⁻ a, ∥f a∥₊ ∂μ + ∫⁻ a, ∥g a∥₊ ∂μ := + ∫⁻ a, ‖f a‖₊ + ‖g a‖₊ ∂μ = ∫⁻ a, ‖f a‖₊ ∂μ + ∫⁻ a, ‖g a‖₊ ∂μ := lintegral_add_right' _ hg.ennnorm lemma lintegral_nnnorm_neg {f : α → β} : - ∫⁻ a, ∥(-f) a∥₊ ∂μ = ∫⁻ a, ∥f a∥₊ ∂μ := + ∫⁻ a, ‖(-f) a‖₊ ∂μ = ∫⁻ a, ‖f a‖₊ ∂μ := by simp only [pi.neg_apply, nnnorm_neg] /-! ### The predicate `has_finite_integral` -/ -/-- `has_finite_integral f μ` means that the integral `∫⁻ a, ∥f a∥ ∂μ` is finite. +/-- `has_finite_integral f μ` means that the integral `∫⁻ a, ‖f a‖ ∂μ` is finite. `has_finite_integral f` means `has_finite_integral f volume`. -/ def has_finite_integral {m : measurable_space α} (f : α → β) (μ : measure α . volume_tac) : Prop := -∫⁻ a, ∥f a∥₊ ∂μ < ∞ +∫⁻ a, ‖f a‖₊ ∂μ < ∞ lemma has_finite_integral_iff_norm (f : α → β) : - has_finite_integral f μ ↔ ∫⁻ a, (ennreal.of_real ∥f a∥) ∂μ < ∞ := + has_finite_integral f μ ↔ ∫⁻ a, (ennreal.of_real ‖f a‖) ∂μ < ∞ := by simp only [has_finite_integral, of_real_norm_eq_coe_nnnorm] lemma has_finite_integral_iff_edist (f : α → β) : @@ -116,24 +116,24 @@ lemma has_finite_integral_iff_of_nnreal {f : α → ℝ≥0} : by simp [has_finite_integral_iff_norm] lemma has_finite_integral.mono {f : α → β} {g : α → γ} (hg : has_finite_integral g μ) - (h : ∀ᵐ a ∂μ, ∥f a∥ ≤ ∥g a∥) : has_finite_integral f μ := + (h : ∀ᵐ a ∂μ, ‖f a‖ ≤ ‖g a‖) : has_finite_integral f μ := begin simp only [has_finite_integral_iff_norm] at *, - calc ∫⁻ a, (ennreal.of_real ∥f a∥) ∂μ ≤ ∫⁻ (a : α), (ennreal.of_real ∥g a∥) ∂μ : + calc ∫⁻ a, (ennreal.of_real ‖f a‖) ∂μ ≤ ∫⁻ (a : α), (ennreal.of_real ‖g a‖) ∂μ : lintegral_mono_ae (h.mono $ assume a h, of_real_le_of_real h) ... < ∞ : hg end lemma has_finite_integral.mono' {f : α → β} {g : α → ℝ} (hg : has_finite_integral g μ) - (h : ∀ᵐ a ∂μ, ∥f a∥ ≤ g a) : has_finite_integral f μ := + (h : ∀ᵐ a ∂μ, ‖f a‖ ≤ g a) : has_finite_integral f μ := hg.mono $ h.mono $ λ x hx, le_trans hx (le_abs_self _) lemma has_finite_integral.congr' {f : α → β} {g : α → γ} (hf : has_finite_integral f μ) - (h : ∀ᵐ a ∂μ, ∥f a∥ = ∥g a∥) : + (h : ∀ᵐ a ∂μ, ‖f a‖ = ‖g a‖) : has_finite_integral g μ := hf.mono $ eventually_eq.le $ eventually_eq.symm h -lemma has_finite_integral_congr' {f : α → β} {g : α → γ} (h : ∀ᵐ a ∂μ, ∥f a∥ = ∥g a∥) : +lemma has_finite_integral_congr' {f : α → β} {g : α → γ} (h : ∀ᵐ a ∂μ, ‖f a‖ = ‖g a‖) : has_finite_integral f μ ↔ has_finite_integral g μ := ⟨λ hf, hf.congr' h, λ hg, hg.congr' $ eventually_eq.symm h⟩ @@ -154,7 +154,7 @@ lemma has_finite_integral_const [is_finite_measure μ] (c : β) : has_finite_integral_const_iff.2 (or.inr $ measure_lt_top _ _) lemma has_finite_integral_of_bounded [is_finite_measure μ] {f : α → β} {C : ℝ} - (hC : ∀ᵐ a ∂μ, ∥f a∥ ≤ C) : has_finite_integral f μ := + (hC : ∀ᵐ a ∂μ, ‖f a‖ ≤ C) : has_finite_integral f μ := (has_finite_integral_const C).mono' hC lemma has_finite_integral.mono_measure {f : α → β} (h : has_finite_integral f ν) (hμ : μ ≤ ν) : @@ -205,20 +205,20 @@ by simpa [has_finite_integral] using hfi ⟨λ h, neg_neg f ▸ h.neg, has_finite_integral.neg⟩ lemma has_finite_integral.norm {f : α → β} (hfi : has_finite_integral f μ) : - has_finite_integral (λa, ∥f a∥) μ := -have eq : (λa, (nnnorm ∥f a∥ : ℝ≥0∞)) = λa, (∥f a∥₊ : ℝ≥0∞), + has_finite_integral (λa, ‖f a‖) μ := +have eq : (λa, (nnnorm ‖f a‖ : ℝ≥0∞)) = λa, (‖f a‖₊ : ℝ≥0∞), by { funext, rw nnnorm_norm }, by { rwa [has_finite_integral, eq] } lemma has_finite_integral_norm_iff (f : α → β) : - has_finite_integral (λa, ∥f a∥) μ ↔ has_finite_integral f μ := + has_finite_integral (λa, ‖f a‖) μ ↔ has_finite_integral f μ := has_finite_integral_congr' $ eventually_of_forall $ λ x, norm_norm (f x) lemma has_finite_integral_to_real_of_lintegral_ne_top {f : α → ℝ≥0∞} (hf : ∫⁻ x, f x ∂μ ≠ ∞) : has_finite_integral (λ x, (f x).to_real) μ := begin - have : ∀ x, (∥(f x).to_real∥₊ : ℝ≥0∞) = + have : ∀ x, (‖(f x).to_real‖₊ : ℝ≥0∞) = @coe ℝ≥0 ℝ≥0∞ _ (⟨(f x).to_real, ennreal.to_real_nonneg⟩ : ℝ≥0), { intro x, rw real.nnnorm_of_nonneg }, simp_rw [has_finite_integral, this], @@ -240,18 +240,18 @@ section dominated_convergence variables {F : ℕ → α → β} {f : α → β} {bound : α → ℝ} -lemma all_ae_of_real_F_le_bound (h : ∀ n, ∀ᵐ a ∂μ, ∥F n a∥ ≤ bound a) : - ∀ n, ∀ᵐ a ∂μ, ennreal.of_real ∥F n a∥ ≤ ennreal.of_real (bound a) := +lemma all_ae_of_real_F_le_bound (h : ∀ n, ∀ᵐ a ∂μ, ‖F n a‖ ≤ bound a) : + ∀ n, ∀ᵐ a ∂μ, ennreal.of_real ‖F n a‖ ≤ ennreal.of_real (bound a) := λn, (h n).mono $ λ a h, ennreal.of_real_le_of_real h lemma all_ae_tendsto_of_real_norm (h : ∀ᵐ a ∂μ, tendsto (λ n, F n a) at_top $ 𝓝 $ f a) : - ∀ᵐ a ∂μ, tendsto (λn, ennreal.of_real ∥F n a∥) at_top $ 𝓝 $ ennreal.of_real ∥f a∥ := + ∀ᵐ a ∂μ, tendsto (λn, ennreal.of_real ‖F n a‖) at_top $ 𝓝 $ ennreal.of_real ‖f a‖ := h.mono $ λ a h, tendsto_of_real $ tendsto.comp (continuous.tendsto continuous_norm _) h -lemma all_ae_of_real_f_le_bound (h_bound : ∀ n, ∀ᵐ a ∂μ, ∥F n a∥ ≤ bound a) +lemma all_ae_of_real_f_le_bound (h_bound : ∀ n, ∀ᵐ a ∂μ, ‖F n a‖ ≤ bound a) (h_lim : ∀ᵐ a ∂μ, tendsto (λ n, F n a) at_top (𝓝 (f a))) : - ∀ᵐ a ∂μ, ennreal.of_real ∥f a∥ ≤ ennreal.of_real (bound a) := + ∀ᵐ a ∂μ, ennreal.of_real ‖f a‖ ≤ ennreal.of_real (bound a) := begin have F_le_bound := all_ae_of_real_F_le_bound h_bound, rw ← ae_all_iff at F_le_bound, @@ -262,14 +262,14 @@ end lemma has_finite_integral_of_dominated_convergence {F : ℕ → α → β} {f : α → β} {bound : α → ℝ} (bound_has_finite_integral : has_finite_integral bound μ) - (h_bound : ∀ n, ∀ᵐ a ∂μ, ∥F n a∥ ≤ bound a) + (h_bound : ∀ n, ∀ᵐ a ∂μ, ‖F n a‖ ≤ bound a) (h_lim : ∀ᵐ a ∂μ, tendsto (λ n, F n a) at_top (𝓝 (f a))) : has_finite_integral f μ := -/- `∥F n a∥ ≤ bound a` and `∥F n a∥ --> ∥f a∥` implies `∥f a∥ ≤ bound a`, - and so `∫ ∥f∥ ≤ ∫ bound < ∞` since `bound` is has_finite_integral -/ +/- `‖F n a‖ ≤ bound a` and `‖F n a‖ --> ‖f a‖` implies `‖f a‖ ≤ bound a`, + and so `∫ ‖f‖ ≤ ∫ bound < ∞` since `bound` is has_finite_integral -/ begin rw has_finite_integral_iff_norm, - calc ∫⁻ a, (ennreal.of_real ∥f a∥) ∂μ ≤ ∫⁻ a, ennreal.of_real (bound a) ∂μ : + calc ∫⁻ a, (ennreal.of_real ‖f a‖) ∂μ ≤ ∫⁻ a, ennreal.of_real (bound a) ∂μ : lintegral_mono_ae $ all_ae_of_real_f_le_bound h_bound h_lim ... < ∞ : begin @@ -283,20 +283,20 @@ lemma tendsto_lintegral_norm_of_dominated_convergence {F : ℕ → α → β} {f : α → β} {bound : α → ℝ} (F_measurable : ∀ n, ae_strongly_measurable (F n) μ) (bound_has_finite_integral : has_finite_integral bound μ) - (h_bound : ∀ n, ∀ᵐ a ∂μ, ∥F n a∥ ≤ bound a) + (h_bound : ∀ n, ∀ᵐ a ∂μ, ‖F n a‖ ≤ bound a) (h_lim : ∀ᵐ a ∂μ, tendsto (λ n, F n a) at_top (𝓝 (f a))) : - tendsto (λn, ∫⁻ a, (ennreal.of_real ∥F n a - f a∥) ∂μ) at_top (𝓝 0) := + tendsto (λn, ∫⁻ a, (ennreal.of_real ‖F n a - f a‖) ∂μ) at_top (𝓝 0) := have f_measurable : ae_strongly_measurable f μ := ae_strongly_measurable_of_tendsto_ae _ F_measurable h_lim, let b := λ a, 2 * ennreal.of_real (bound a) in -/- `∥F n a∥ ≤ bound a` and `F n a --> f a` implies `∥f a∥ ≤ bound a`, and thus by the - triangle inequality, have `∥F n a - f a∥ ≤ 2 * (bound a). -/ -have hb : ∀ n, ∀ᵐ a ∂μ, ennreal.of_real ∥F n a - f a∥ ≤ b a, +/- `‖F n a‖ ≤ bound a` and `F n a --> f a` implies `‖f a‖ ≤ bound a`, and thus by the + triangle inequality, have `‖F n a - f a‖ ≤ 2 * (bound a). -/ +have hb : ∀ n, ∀ᵐ a ∂μ, ennreal.of_real ‖F n a - f a‖ ≤ b a, begin assume n, filter_upwards [all_ae_of_real_F_le_bound h_bound n, all_ae_of_real_f_le_bound h_bound h_lim] with a h₁ h₂, - calc ennreal.of_real ∥F n a - f a∥ ≤ (ennreal.of_real ∥F n a∥) + (ennreal.of_real ∥f a∥) : + calc ennreal.of_real ‖F n a - f a‖ ≤ (ennreal.of_real ‖F n a‖) + (ennreal.of_real ‖f a‖) : begin rw [← ennreal.of_real_add], apply of_real_le_of_real, @@ -305,21 +305,21 @@ begin ... ≤ (ennreal.of_real (bound a)) + (ennreal.of_real (bound a)) : add_le_add h₁ h₂ ... = b a : by rw ← two_mul end, -/- On the other hand, `F n a --> f a` implies that `∥F n a - f a∥ --> 0` -/ -have h : ∀ᵐ a ∂μ, tendsto (λ n, ennreal.of_real ∥F n a - f a∥) at_top (𝓝 0), +/- On the other hand, `F n a --> f a` implies that `‖F n a - f a‖ --> 0` -/ +have h : ∀ᵐ a ∂μ, tendsto (λ n, ennreal.of_real ‖F n a - f a‖) at_top (𝓝 0), begin rw ← ennreal.of_real_zero, refine h_lim.mono (λ a h, (continuous_of_real.tendsto _).comp _), rwa ← tendsto_iff_norm_tendsto_zero end, /- Therefore, by the dominated convergence theorem for nonnegative integration, have - ` ∫ ∥f a - F n a∥ --> 0 ` -/ + ` ∫ ‖f a - F n a‖ --> 0 ` -/ begin - suffices h : tendsto (λn, ∫⁻ a, (ennreal.of_real ∥F n a - f a∥) ∂μ) at_top (𝓝 (∫⁻ (a:α), 0 ∂μ)), + suffices h : tendsto (λn, ∫⁻ a, (ennreal.of_real ‖F n a - f a‖) ∂μ) at_top (𝓝 (∫⁻ (a:α), 0 ∂μ)), { rwa lintegral_zero at h }, -- Using the dominated convergence theorem. refine tendsto_lintegral_of_dominated_convergence' _ _ hb _ _, - -- Show `λa, ∥f a - F n a∥` is almost everywhere measurable for all `n` + -- Show `λa, ‖f a - F n a‖` is almost everywhere measurable for all `n` { exact λ n, measurable_of_real.comp_ae_measurable ((F_measurable n).sub f_measurable).norm.ae_measurable }, -- Show `2 * bound` is has_finite_integral @@ -328,7 +328,7 @@ begin by { rw lintegral_const_mul', exact coe_ne_top } ... ≠ ∞ : mul_ne_top coe_ne_top bound_has_finite_integral.ne }, filter_upwards [h_bound 0] with _ h using le_trans (norm_nonneg _) h }, - -- Show `∥f a - F n a∥ --> 0` + -- Show `‖f a - F n a‖ --> 0` { exact h } end @@ -356,7 +356,7 @@ lemma has_finite_integral.smul (c : 𝕜) {f : α → β} : has_finite_integral begin simp only [has_finite_integral], assume hfi, calc - ∫⁻ (a : α), ∥c • f a∥₊ ∂μ = ∫⁻ (a : α), (∥c∥₊) * ∥f a∥₊ ∂μ : + ∫⁻ (a : α), ‖c • f a‖₊ ∂μ = ∫⁻ (a : α), (‖c‖₊) * ‖f a‖₊ ∂μ : by simp only [nnnorm_smul, ennreal.coe_mul] ... < ∞ : begin @@ -388,7 +388,7 @@ end normed_space -- variables [measurable_space β] [measurable_space γ] [measurable_space δ] -/-- `integrable f μ` means that `f` is measurable and that the integral `∫⁻ a, ∥f a∥ ∂μ` is finite. +/-- `integrable f μ` means that `f` is measurable and that the integral `∫⁻ a, ‖f a‖ ∂μ` is finite. `integrable f` means `integrable f volume`. -/ def integrable {α} {m : measurable_space α} (f : α → β) (μ : measure α . volume_tac) : Prop := ae_strongly_measurable f μ ∧ has_finite_integral f μ @@ -409,22 +409,22 @@ lemma integrable.has_finite_integral {f : α → β} (hf : integrable f μ) : ha hf.2 lemma integrable.mono {f : α → β} {g : α → γ} - (hg : integrable g μ) (hf : ae_strongly_measurable f μ) (h : ∀ᵐ a ∂μ, ∥f a∥ ≤ ∥g a∥) : + (hg : integrable g μ) (hf : ae_strongly_measurable f μ) (h : ∀ᵐ a ∂μ, ‖f a‖ ≤ ‖g a‖) : integrable f μ := ⟨hf, hg.has_finite_integral.mono h⟩ lemma integrable.mono' {f : α → β} {g : α → ℝ} - (hg : integrable g μ) (hf : ae_strongly_measurable f μ) (h : ∀ᵐ a ∂μ, ∥f a∥ ≤ g a) : + (hg : integrable g μ) (hf : ae_strongly_measurable f μ) (h : ∀ᵐ a ∂μ, ‖f a‖ ≤ g a) : integrable f μ := ⟨hf, hg.has_finite_integral.mono' h⟩ lemma integrable.congr' {f : α → β} {g : α → γ} - (hf : integrable f μ) (hg : ae_strongly_measurable g μ) (h : ∀ᵐ a ∂μ, ∥f a∥ = ∥g a∥) : + (hf : integrable f μ) (hg : ae_strongly_measurable g μ) (h : ∀ᵐ a ∂μ, ‖f a‖ = ‖g a‖) : integrable g μ := ⟨hg, hf.has_finite_integral.congr' h⟩ lemma integrable_congr' {f : α → β} {g : α → γ} - (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (h : ∀ᵐ a ∂μ, ∥f a∥ = ∥g a∥) : + (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (h : ∀ᵐ a ∂μ, ‖f a‖ = ‖g a‖) : integrable f μ ↔ integrable g μ := ⟨λ h2f, h2f.congr' hg h, λ h2g, h2g.congr' hf $ eventually_eq.symm h⟩ @@ -447,7 +447,7 @@ integrable_const_iff.2 $ or.inr $ measure_lt_top _ _ lemma mem_ℒp.integrable_norm_rpow {f : α → β} {p : ℝ≥0∞} (hf : mem_ℒp f p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) : - integrable (λ (x : α), ∥f x∥ ^ p.to_real) μ := + integrable (λ (x : α), ‖f x‖ ^ p.to_real) μ := begin rw ← mem_ℒp_one_iff_integrable, exact hf.norm_rpow hp_ne_zero hp_ne_top, @@ -455,7 +455,7 @@ end lemma mem_ℒp.integrable_norm_rpow' [is_finite_measure μ] {f : α → β} {p : ℝ≥0∞} (hf : mem_ℒp f p μ) : - integrable (λ (x : α), ∥f x∥ ^ p.to_real) μ := + integrable (λ (x : α), ‖f x‖ ^ p.to_real) μ := begin by_cases h_zero : p = 0, { simp [h_zero, integrable_const] }, @@ -574,7 +574,7 @@ variables {α β μ} lemma integrable.add' {f g : α → β} (hf : integrable f μ) (hg : integrable g μ) : has_finite_integral (f + g) μ := -calc ∫⁻ a, ∥f a + g a∥₊ ∂μ ≤ ∫⁻ a, ∥f a∥₊ + ∥g a∥₊ ∂μ : +calc ∫⁻ a, ‖f a + g a‖₊ ∂μ ≤ ∫⁻ a, ‖f a‖₊ + ‖g a‖₊ ∂μ : lintegral_mono (λ a, by exact_mod_cast nnnorm_add_le _ _) ... = _ : lintegral_nnnorm_add_left hf.ae_strongly_measurable _ ... < ∞ : add_lt_top.2 ⟨hf.has_finite_integral, hg.has_finite_integral⟩ @@ -604,7 +604,7 @@ lemma integrable.sub {f g : α → β} by simpa only [sub_eq_add_neg] using hf.add hg.neg lemma integrable.norm {f : α → β} (hf : integrable f μ) : - integrable (λ a, ∥f a∥) μ := + integrable (λ a, ‖f a‖) μ := ⟨hf.ae_strongly_measurable.norm, hf.has_finite_integral.norm⟩ lemma integrable.inf {β} [normed_lattice_add_comm_group β] {f g : α → β} @@ -623,7 +623,7 @@ by { rw ← mem_ℒp_one_iff_integrable at hf ⊢, exact hf.abs, } lemma integrable.bdd_mul {F : Type*} [normed_division_ring F] {f g : α → F} (hint : integrable g μ) (hm : ae_strongly_measurable f μ) - (hfbdd : ∃ C, ∀ x, ∥f x∥ ≤ C) : + (hfbdd : ∃ C, ∀ x, ‖f x‖ ≤ C) : integrable (λ x, f x * g x) μ := begin casesI is_empty_or_nonempty α with hα hα, @@ -632,7 +632,7 @@ begin { refine ⟨hm.mul hint.1, _⟩, obtain ⟨C, hC⟩ := hfbdd, have hCnonneg : 0 ≤ C := le_trans (norm_nonneg _) (hC hα.some), - have : (λ x, ∥f x * g x∥₊) ≤ λ x, ⟨C, hCnonneg⟩ * ∥g x∥₊, + have : (λ x, ‖f x * g x‖₊) ≤ λ x, ⟨C, hCnonneg⟩ * ‖g x‖₊, { intro x, simp only [nnnorm_mul], exact mul_le_mul_of_nonneg_right (hC x) (zero_le _) }, @@ -643,21 +643,21 @@ begin end lemma integrable_norm_iff {f : α → β} (hf : ae_strongly_measurable f μ) : - integrable (λa, ∥f a∥) μ ↔ integrable f μ := + integrable (λa, ‖f a‖) μ ↔ integrable f μ := by simp_rw [integrable, and_iff_right hf, and_iff_right hf.norm, has_finite_integral_norm_iff] lemma integrable_of_norm_sub_le {f₀ f₁ : α → β} {g : α → ℝ} (hf₁_m : ae_strongly_measurable f₁ μ) (hf₀_i : integrable f₀ μ) (hg_i : integrable g μ) - (h : ∀ᵐ a ∂μ, ∥f₀ a - f₁ a∥ ≤ g a) : + (h : ∀ᵐ a ∂μ, ‖f₀ a - f₁ a‖ ≤ g a) : integrable f₁ μ := begin - have : ∀ᵐ a ∂μ, ∥f₁ a∥ ≤ ∥f₀ a∥ + g a, + have : ∀ᵐ a ∂μ, ‖f₁ a‖ ≤ ‖f₀ a‖ + g a, { apply h.mono, intros a ha, - calc ∥f₁ a∥ ≤ ∥f₀ a∥ + ∥f₀ a - f₁ a∥ : norm_le_insert _ _ - ... ≤ ∥f₀ a∥ + g a : add_le_add_left ha _ }, + calc ‖f₁ a‖ ≤ ‖f₀ a‖ + ‖f₀ a - f₁ a‖ : norm_le_insert _ _ + ... ≤ ‖f₀ a‖ + g a : add_le_add_left ha _ }, exact integrable.mono' (hf₀_i.norm.add hg_i) hf₁_m this end @@ -665,8 +665,8 @@ lemma integrable.prod_mk {f : α → β} {g : α → γ} (hf : integrable f μ) integrable (λ x, (f x, g x)) μ := ⟨hf.ae_strongly_measurable.prod_mk hg.ae_strongly_measurable, (hf.norm.add' hg.norm).mono $ eventually_of_forall $ λ x, - calc max ∥f x∥ ∥g x∥ ≤ ∥f x∥ + ∥g x∥ : max_le_add_of_nonneg (norm_nonneg _) (norm_nonneg _) - ... ≤ ∥(∥f x∥ + ∥g x∥)∥ : le_abs_self _⟩ + calc max ‖f x‖ ‖g x‖ ≤ ‖f x‖ + ‖g x‖ : max_le_add_of_nonneg (norm_nonneg _) (norm_nonneg _) + ... ≤ ‖(‖f x‖ + ‖g x‖)‖ : le_abs_self _⟩ lemma mem_ℒp.integrable {q : ℝ≥0∞} (hq1 : 1 ≤ q) {f : α → β} [is_finite_measure μ] (hfq : mem_ℒp f q μ) : integrable f μ := @@ -833,7 +833,7 @@ noncomputable def with_density_smul_li {f : α → ℝ≥0} (f_meas : measurable apply lintegral_congr_ae, filter_upwards [(mem_ℒ1_smul_of_L1_with_density f_meas u).coe_fn_to_Lp] with x hx, rw [hx, pi.mul_apply], - change ↑∥(f x : ℝ) • u x∥₊ = ↑(f x) * ↑∥u x∥₊, + change ↑‖(f x : ℝ) • u x‖₊ = ↑(f x) * ↑‖u x‖₊, simp only [nnnorm_smul, nnreal.nnnorm_eq, ennreal.coe_mul], end } @@ -903,7 +903,7 @@ lemma integrable.div_const {f : α → ℝ} (h : integrable f μ) (c : ℝ) : by simp_rw [div_eq_mul_inv, h.mul_const] lemma integrable.bdd_mul' {f g : α → ℝ} {c : ℝ} (hg : integrable g μ) - (hf : ae_strongly_measurable f μ) (hf_bound : ∀ᵐ x ∂μ, ∥f x∥ ≤ c) : + (hf : ae_strongly_measurable f μ) (hf_bound : ∀ᵐ x ∂μ, ‖f x‖ ≤ c) : integrable (λ x, f x * g x) μ := begin refine integrable.mono' (hg.norm.smul c) (hf.mul hg.1) _, @@ -994,13 +994,13 @@ variables {E : Type*} {m0 : measurable_space α} [normed_add_comm_group E] lemma integrable_of_forall_fin_meas_le' {μ : measure α} (hm : m ≤ m0) [sigma_finite (μ.trim hm)] (C : ℝ≥0∞) (hC : C < ∞) {f : α → E} (hf_meas : ae_strongly_measurable f μ) - (hf : ∀ s, measurable_set[m] s → μ s ≠ ∞ → ∫⁻ x in s, ∥f x∥₊ ∂μ ≤ C) : + (hf : ∀ s, measurable_set[m] s → μ s ≠ ∞ → ∫⁻ x in s, ‖f x‖₊ ∂μ ≤ C) : integrable f μ := ⟨hf_meas, (lintegral_le_of_forall_fin_meas_le' hm C hf_meas.ennnorm hf).trans_lt hC⟩ lemma integrable_of_forall_fin_meas_le [sigma_finite μ] (C : ℝ≥0∞) (hC : C < ∞) {f : α → E} (hf_meas : ae_strongly_measurable f μ) - (hf : ∀ s : set α, measurable_set s → μ s ≠ ∞ → ∫⁻ x in s, ∥f x∥₊ ∂μ ≤ C) : + (hf : ∀ s : set α, measurable_set s → μ s ≠ ∞ → ∫⁻ x in s, ‖f x‖₊ ∂μ ≤ C) : integrable f μ := @integrable_of_forall_fin_meas_le' _ _ _ _ _ _ _ (by rwa trim_eq_self) C hC _ hf_meas hf @@ -1100,14 +1100,14 @@ lemma dist_def (f g : α →₁[μ] β) : by { simp [Lp.dist_def, snorm, snorm'], simp [edist_eq_coe_nnnorm_sub] } lemma norm_def (f : α →₁[μ] β) : - ∥f∥ = (∫⁻ a, ∥f a∥₊ ∂μ).to_real := + ‖f‖ = (∫⁻ a, ‖f a‖₊ ∂μ).to_real := by { simp [Lp.norm_def, snorm, snorm'] } /-- Computing the norm of a difference between two L¹-functions. Note that this is not a special case of `norm_def` since `(f - g) x` and `f x - g x` are not equal (but only a.e.-equal). -/ lemma norm_sub_eq_lintegral (f g : α →₁[μ] β) : - ∥f - g∥ = (∫⁻ x, (∥f x - g x∥₊ : ℝ≥0∞) ∂μ).to_real := + ‖f - g‖ = (∫⁻ x, (‖f x - g x‖₊ : ℝ≥0∞) ∂μ).to_real := begin rw [norm_def], congr' 1, @@ -1117,14 +1117,14 @@ begin end lemma of_real_norm_eq_lintegral (f : α →₁[μ] β) : - ennreal.of_real ∥f∥ = ∫⁻ x, (∥f x∥₊ : ℝ≥0∞) ∂μ := + ennreal.of_real ‖f‖ = ∫⁻ x, (‖f x‖₊ : ℝ≥0∞) ∂μ := by { rw [norm_def, ennreal.of_real_to_real], exact ne_of_lt (has_finite_integral_coe_fn f) } /-- Computing the norm of a difference between two L¹-functions. Note that this is not a special case of `of_real_norm_eq_lintegral` since `(f - g) x` and `f x - g x` are not equal (but only a.e.-equal). -/ lemma of_real_norm_sub_eq_lintegral (f g : α →₁[μ] β) : - ennreal.of_real ∥f - g∥ = ∫⁻ x, (∥f x - g x∥₊ : ℝ≥0∞) ∂μ := + ennreal.of_real ‖f - g‖ = ∫⁻ x, (‖f x - g x‖₊ : ℝ≥0∞) ∂μ := begin simp_rw [of_real_norm_eq_lintegral, ← edist_eq_coe_nnnorm], apply lintegral_congr_ae, @@ -1167,11 +1167,11 @@ lemma to_L1_sub (f g : α → β) (hf : integrable f μ) (hg : integrable g μ) to_L1 (f - g) (hf.sub hg) = to_L1 f hf - to_L1 g hg := rfl lemma norm_to_L1 (f : α → β) (hf : integrable f μ) : - ∥hf.to_L1 f∥ = ennreal.to_real (∫⁻ a, edist (f a) 0 ∂μ) := + ‖hf.to_L1 f‖ = ennreal.to_real (∫⁻ a, edist (f a) 0 ∂μ) := by { simp [to_L1, snorm, snorm'], simp [edist_eq_coe_nnnorm] } lemma norm_to_L1_eq_lintegral_norm (f : α → β) (hf : integrable f μ) : - ∥hf.to_L1 f∥ = ennreal.to_real (∫⁻ a, (ennreal.of_real ∥f a∥) ∂μ) := + ‖hf.to_L1 f‖ = ennreal.to_real (∫⁻ a, (ennreal.of_real ‖f a‖) ∂μ) := by { rw [norm_to_L1, lintegral_norm_eq_lintegral_edist] } @[simp] lemma edist_to_L1_to_L1 (f g : α → β) (hf : integrable f μ) (hg : integrable g μ) : @@ -1202,11 +1202,11 @@ variables {E : Type*} [normed_add_comm_group E] lemma measure_theory.integrable.apply_continuous_linear_map {φ : α → H →L[𝕜] E} (φ_int : integrable φ μ) (v : H) : integrable (λ a, φ a v) μ := -(φ_int.norm.mul_const ∥v∥).mono' (φ_int.ae_strongly_measurable.apply_continuous_linear_map v) +(φ_int.norm.mul_const ‖v‖).mono' (φ_int.ae_strongly_measurable.apply_continuous_linear_map v) (eventually_of_forall $ λ a, (φ a).le_op_norm v) lemma continuous_linear_map.integrable_comp {φ : α → H} (L : H →L[𝕜] E) (φ_int : integrable φ μ) : integrable (λ (a : α), L (φ a)) μ := -((integrable.norm φ_int).const_mul ∥L∥).mono' +((integrable.norm φ_int).const_mul ‖L‖).mono' (L.continuous.comp_ae_strongly_measurable φ_int.ae_strongly_measurable) (eventually_of_forall $ λ a, L.le_op_norm (φ a)) diff --git a/src/measure_theory/function/l2_space.lean b/src/measure_theory/function/l2_space.lean index b3a7ca2d8f3b5..3341afa2b63e5 100644 --- a/src/measure_theory/function/l2_space.lean +++ b/src/measure_theory/function/l2_space.lean @@ -37,7 +37,7 @@ by simpa [← mem_ℒp_one_iff_integrable] using h.norm_rpow ennreal.two_ne_zero ennreal.two_ne_top lemma mem_ℒp_two_iff_integrable_sq_norm {f : α → F} (hf : ae_strongly_measurable f μ) : - mem_ℒp f 2 μ ↔ integrable (λ x, ∥f x∥^2) μ := + mem_ℒp f 2 μ ↔ integrable (λ x, ‖f x‖^2) μ := begin rw ← mem_ℒp_one_iff_integrable, convert (mem_ℒp_norm_rpow_iff hf ennreal.two_ne_zero ennreal.two_ne_top).symm, @@ -63,7 +63,7 @@ variables {α E F 𝕜 : Type*} [is_R_or_C 𝕜] [measurable_space α] {μ : mea local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y -lemma snorm_rpow_two_norm_lt_top (f : Lp F 2 μ) : snorm (λ x, ∥f x∥ ^ (2 : ℝ)) 1 μ < ∞ := +lemma snorm_rpow_two_norm_lt_top (f : Lp F 2 μ) : snorm (λ x, ‖f x‖ ^ (2 : ℝ)) 1 μ < ∞ := begin have h_two : ennreal.of_real (2 : ℝ) = 2, by simp [zero_le_one], rw [snorm_norm_rpow f zero_lt_two, one_mul, h_two], @@ -72,8 +72,8 @@ end lemma snorm_inner_lt_top (f g : α →₂[μ] E) : snorm (λ (x : α), ⟪f x, g x⟫) 1 μ < ∞ := begin - have h : ∀ x, is_R_or_C.abs ⟪f x, g x⟫ ≤ ∥f x∥ * ∥g x∥, from λ x, abs_inner_le_norm _ _, - have h' : ∀ x, is_R_or_C.abs ⟪f x, g x⟫ ≤ is_R_or_C.abs (∥f x∥^2 + ∥g x∥^2), + have h : ∀ x, is_R_or_C.abs ⟪f x, g x⟫ ≤ ‖f x‖ * ‖g x‖, from λ x, abs_inner_le_norm _ _, + have h' : ∀ x, is_R_or_C.abs ⟪f x, g x⟫ ≤ is_R_or_C.abs (‖f x‖^2 + ‖g x‖^2), { refine λ x, le_trans (h x) _, rw [is_R_or_C.abs_to_real, abs_eq_self.mpr], swap, { exact add_nonneg (by simp) (by simp), }, @@ -98,7 +98,7 @@ instance : has_inner 𝕜 (α →₂[μ] E) := ⟨λ f g, ∫ a, ⟪f a, g a⟫ lemma inner_def (f g : α →₂[μ] E) : ⟪f, g⟫ = ∫ a : α, ⟪f a, g a⟫ ∂μ := rfl lemma integral_inner_eq_sq_snorm (f : α →₂[μ] E) : - ∫ a, ⟪f a, f a⟫ ∂μ = ennreal.to_real ∫⁻ a, (∥f a∥₊ : ℝ≥0∞) ^ (2:ℝ) ∂μ := + ∫ a, ⟪f a, f a⟫ ∂μ = ennreal.to_real ∫⁻ a, (‖f a‖₊ : ℝ≥0∞) ^ (2:ℝ) ∂μ := begin simp_rw inner_self_eq_norm_sq_to_K, norm_cast, @@ -114,7 +114,7 @@ begin norm_cast, end -private lemma norm_sq_eq_inner' (f : α →₂[μ] E) : ∥f∥ ^ 2 = is_R_or_C.re ⟪f, f⟫ := +private lemma norm_sq_eq_inner' (f : α →₂[μ] E) : ‖f‖ ^ 2 = is_R_or_C.re ⟪f, f⟫ := begin have h_two : (2 : ℝ≥0∞).to_real = 2 := by simp, rw [inner_def, integral_inner_eq_sq_snorm, norm_def, ← ennreal.to_real_pow, is_R_or_C.of_real_re, diff --git a/src/measure_theory/function/locally_integrable.lean b/src/measure_theory/function/locally_integrable.lean index 68bd275ea2647..aa697fde24a6b 100644 --- a/src/measure_theory/function/locally_integrable.lean +++ b/src/measure_theory/function/locally_integrable.lean @@ -66,7 +66,7 @@ lemma integrable_on.mul_continuous_on_of_subset begin rcases is_compact.exists_bound_of_continuous_on hK hg' with ⟨C, hC⟩, rw [integrable_on, ← mem_ℒp_one_iff_integrable] at hg ⊢, - have : ∀ᵐ x ∂(μ.restrict A), ∥g x * g' x∥ ≤ C * ∥g x∥, + have : ∀ᵐ x ∂(μ.restrict A), ‖g x * g' x‖ ≤ C * ‖g x‖, { filter_upwards [ae_restrict_mem hA] with x hx, rw [real.norm_eq_abs, abs_mul, mul_comm, real.norm_eq_abs], apply mul_le_mul_of_nonneg_right (hC x (hAK hx)) (abs_nonneg _), }, diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index 94d613e5e14bd..35987dfa43c2d 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -16,7 +16,7 @@ import topology.continuous_function.compact This file describes properties of almost everywhere strongly measurable functions with finite seminorm, denoted by `snorm f p μ` and defined for `p:ℝ≥0∞` asmm_group (Lp E p μ) := `0` if `p=0`, -`(∫ ∥f a∥^p ∂μ) ^ (1/p)` for `0 < p < ∞` and `ess_sup ∥f∥ μ` for `p=∞`. +`(∫ ‖f a‖^p ∂μ) ^ (1/p)` for `0 < p < ∞` and `ess_sup ‖f‖ μ` for `p=∞`. The Prop-valued `mem_ℒp f p μ` states that a function `f : α → E` has finite seminorm. The space `Lp E p μ` is the subtype of elements of `α →ₘ[μ] E` (see ae_eq_fun) such that @@ -24,9 +24,9 @@ The space `Lp E p μ` is the subtype of elements of `α →ₘ[μ] E` (see ae_eq ## Main definitions -* `snorm' f p μ` : `(∫ ∥f a∥^p ∂μ) ^ (1/p)` for `f : α → F` and `p : ℝ`, where `α` is a measurable +* `snorm' f p μ` : `(∫ ‖f a‖^p ∂μ) ^ (1/p)` for `f : α → F` and `p : ℝ`, where `α` is a measurable space and `F` is a normed group. -* `snorm_ess_sup f μ` : seminorm in `ℒ∞`, equal to the essential supremum `ess_sup ∥f∥ μ`. +* `snorm_ess_sup f μ` : seminorm in `ℒ∞`, equal to the essential supremum `ess_sup ‖f‖ μ`. * `snorm f p μ` : for `p : ℝ≥0∞`, seminorm in `ℒp`, equal to `0` for `p=0`, to `snorm' f p μ` for `0 < p < ∞` and to `snorm_ess_sup f μ` for `p = ∞`. @@ -102,17 +102,17 @@ deduce it for `snorm`, and translate it in terms of `mem_ℒp`. section ℒp_space_definition -/-- `(∫ ∥f a∥^q ∂μ) ^ (1/q)`, which is a seminorm on the space of measurable functions for which +/-- `(∫ ‖f a‖^q ∂μ) ^ (1/q)`, which is a seminorm on the space of measurable functions for which this quantity is finite -/ def snorm' {m : measurable_space α} (f : α → F) (q : ℝ) (μ : measure α) : ℝ≥0∞ := -(∫⁻ a, ∥f a∥₊^q ∂μ) ^ (1/q) +(∫⁻ a, ‖f a‖₊^q ∂μ) ^ (1/q) -/-- seminorm for `ℒ∞`, equal to the essential supremum of `∥f∥`. -/ +/-- seminorm for `ℒ∞`, equal to the essential supremum of `‖f‖`. -/ def snorm_ess_sup {m : measurable_space α} (f : α → F) (μ : measure α) := -ess_sup (λ x, (∥f x∥₊ : ℝ≥0∞)) μ +ess_sup (λ x, (‖f x‖₊ : ℝ≥0∞)) μ -/-- `ℒp` seminorm, equal to `0` for `p=0`, to `(∫ ∥f a∥^p ∂μ) ^ (1/p)` for `0 < p < ∞` and to -`ess_sup ∥f∥ μ` for `p = ∞`. -/ +/-- `ℒp` seminorm, equal to `0` for `p=0`, to `(∫ ‖f a‖^p ∂μ) ^ (1/p)` for `0 < p < ∞` and to +`ess_sup ‖f‖ μ` for `p = ∞`. -/ def snorm {m : measurable_space α} (f : α → F) (p : ℝ≥0∞) (μ : measure α) : ℝ≥0∞ := if p = 0 then 0 else (if p = ∞ then snorm_ess_sup f μ else snorm' f (ennreal.to_real p) μ) @@ -121,16 +121,16 @@ lemma snorm_eq_snorm' (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) {f : α → by simp [snorm, hp_ne_zero, hp_ne_top] lemma snorm_eq_lintegral_rpow_nnnorm (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) {f : α → F} : - snorm f p μ = (∫⁻ x, ∥f x∥₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) := + snorm f p μ = (∫⁻ x, ‖f x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) := by rw [snorm_eq_snorm' hp_ne_zero hp_ne_top, snorm'] -lemma snorm_one_eq_lintegral_nnnorm {f : α → F} : snorm f 1 μ = ∫⁻ x, ∥f x∥₊ ∂μ := +lemma snorm_one_eq_lintegral_nnnorm {f : α → F} : snorm f 1 μ = ∫⁻ x, ‖f x‖₊ ∂μ := by simp_rw [snorm_eq_lintegral_rpow_nnnorm one_ne_zero ennreal.coe_ne_top, ennreal.one_to_real, one_div_one, ennreal.rpow_one] @[simp] lemma snorm_exponent_top {f : α → F} : snorm f ∞ μ = snorm_ess_sup f μ := by simp [snorm] -/-- The property that `f:α→E` is ae strongly measurable and `(∫ ∥f a∥^p ∂μ)^(1/p)` is finite +/-- The property that `f:α→E` is ae strongly measurable and `(∫ ‖f a‖^p ∂μ)^(1/p)` is finite if `p < ∞`, or `ess_sup f < ∞` if `p = ∞`. -/ def mem_ℒp {α} {m : measurable_space α} (f : α → E) (p : ℝ≥0∞) (μ : measure α . volume_tac) : Prop := @@ -140,7 +140,7 @@ lemma mem_ℒp.ae_strongly_measurable {f : α → E} {p : ℝ≥0∞} (h : mem_ ae_strongly_measurable f μ := h.1 lemma lintegral_rpow_nnnorm_eq_rpow_snorm' {f : α → F} (hq0_lt : 0 < q) : - ∫⁻ a, ∥f a∥₊ ^ q ∂μ = (snorm' f q μ) ^ q := + ∫⁻ a, ‖f a‖₊ ^ q ∂μ = (snorm' f q μ) ^ q := begin rw [snorm', ←ennreal.rpow_mul, one_div, inv_mul_cancel, ennreal.rpow_one], exact (ne_of_lt hq0_lt).symm, @@ -156,7 +156,7 @@ lemma mem_ℒp.snorm_ne_top {f : α → E} (hfp : mem_ℒp f p μ) : snorm f p lemma lintegral_rpow_nnnorm_lt_top_of_snorm'_lt_top {f : α → F} (hq0_lt : 0 < q) (hfq : snorm' f q μ < ∞) : - ∫⁻ a, ∥f a∥₊ ^ q ∂μ < ∞ := + ∫⁻ a, ‖f a‖₊ ^ q ∂μ < ∞ := begin rw lintegral_rpow_nnnorm_eq_rpow_snorm' hq0_lt, exact ennreal.rpow_lt_top_of_nonneg (le_of_lt hq0_lt) (ne_of_lt hfq), @@ -164,7 +164,7 @@ end lemma lintegral_rpow_nnnorm_lt_top_of_snorm_lt_top {f : α → F} (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hfp : snorm f p μ < ∞) : - ∫⁻ a, ∥f a∥₊ ^ p.to_real ∂μ < ∞ := + ∫⁻ a, ‖f a‖₊ ^ p.to_real ∂μ < ∞ := begin apply lintegral_rpow_nnnorm_lt_top_of_snorm'_lt_top, { exact ennreal.to_real_pos hp_ne_zero hp_ne_top }, @@ -173,7 +173,7 @@ end lemma snorm_lt_top_iff_lintegral_rpow_nnnorm_lt_top {f : α → F} (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) : - snorm f p μ < ∞ ↔ ∫⁻ a, ∥f a∥₊ ^ p.to_real ∂μ < ∞ := + snorm f p μ < ∞ ↔ ∫⁻ a, ‖f a‖₊ ^ p.to_real ∂μ < ∞ := ⟨lintegral_rpow_nnnorm_lt_top_of_snorm_lt_top hp_ne_zero hp_ne_top, begin intros h, @@ -262,7 +262,7 @@ end zero section const lemma snorm'_const (c : F) (hq_pos : 0 < q) : - snorm' (λ x : α , c) q μ = (∥c∥₊ : ℝ≥0∞) * (μ set.univ) ^ (1/q) := + snorm' (λ x : α , c) q μ = (‖c‖₊ : ℝ≥0∞) * (μ set.univ) ^ (1/q) := begin rw [snorm', lintegral_const, ennreal.mul_rpow_of_nonneg _ _ (by simp [hq_pos.le] : 0 ≤ 1 / q)], congr, @@ -272,7 +272,7 @@ begin end lemma snorm'_const' [is_finite_measure μ] (c : F) (hc_ne_zero : c ≠ 0) (hq_ne_zero : q ≠ 0) : - snorm' (λ x : α , c) q μ = (∥c∥₊ : ℝ≥0∞) * (μ set.univ) ^ (1/q) := + snorm' (λ x : α , c) q μ = (‖c‖₊ : ℝ≥0∞) * (μ set.univ) ^ (1/q) := begin rw [snorm', lintegral_const, ennreal.mul_rpow_of_ne_top _ (measure_ne_top μ set.univ)], { congr, @@ -287,15 +287,15 @@ begin end lemma snorm_ess_sup_const (c : F) (hμ : μ ≠ 0) : - snorm_ess_sup (λ x : α, c) μ = (∥c∥₊ : ℝ≥0∞) := + snorm_ess_sup (λ x : α, c) μ = (‖c‖₊ : ℝ≥0∞) := by rw [snorm_ess_sup, ess_sup_const _ hμ] lemma snorm'_const_of_is_probability_measure (c : F) (hq_pos : 0 < q) [is_probability_measure μ] : - snorm' (λ x : α , c) q μ = (∥c∥₊ : ℝ≥0∞) := + snorm' (λ x : α , c) q μ = (‖c‖₊ : ℝ≥0∞) := by simp [snorm'_const c hq_pos, measure_univ] lemma snorm_const (c : F) (h0 : p ≠ 0) (hμ : μ ≠ 0) : - snorm (λ x : α , c) p μ = (∥c∥₊ : ℝ≥0∞) * (μ set.univ) ^ (1/(ennreal.to_real p)) := + snorm (λ x : α , c) p μ = (‖c‖₊ : ℝ≥0∞) * (μ set.univ) ^ (1/(ennreal.to_real p)) := begin by_cases h_top : p = ∞, { simp [h_top, snorm_ess_sup_const c hμ], }, @@ -303,7 +303,7 @@ begin end lemma snorm_const' (c : F) (h0 : p ≠ 0) (h_top: p ≠ ∞) : - snorm (λ x : α , c) p μ = (∥c∥₊ : ℝ≥0∞) * (μ set.univ) ^ (1/(ennreal.to_real p)) := + snorm (λ x : α , c) p μ = (‖c‖₊ : ℝ≥0∞) * (μ set.univ) ^ (1/(ennreal.to_real p)) := begin simp [snorm_eq_snorm' h0 h_top, snorm'_const, ennreal.to_real_pos h0 h_top], end @@ -359,7 +359,7 @@ end end const -lemma snorm'_mono_ae {f : α → F} {g : α → G} (hq : 0 ≤ q) (h : ∀ᵐ x ∂μ, ∥f x∥ ≤ ∥g x∥) : +lemma snorm'_mono_ae {f : α → F} {g : α → G} (hq : 0 ≤ q) (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : snorm' f q μ ≤ snorm' g q μ := begin rw [snorm'], @@ -368,10 +368,10 @@ begin exact ennreal.rpow_le_rpow (ennreal.coe_le_coe.2 hx) hq end -lemma snorm'_congr_norm_ae {f g : α → F} (hfg : ∀ᵐ x ∂μ, ∥f x∥ = ∥g x∥) : +lemma snorm'_congr_norm_ae {f g : α → F} (hfg : ∀ᵐ x ∂μ, ‖f x‖ = ‖g x‖) : snorm' f q μ = snorm' g q μ := begin - have : (λ x, (∥f x∥₊ ^ q : ℝ≥0∞)) =ᵐ[μ] (λ x, ∥g x∥₊ ^ q), + have : (λ x, (‖f x‖₊ ^ q : ℝ≥0∞)) =ᵐ[μ] (λ x, ‖g x‖₊ ^ q), from hfg.mono (λ x hx, by { simp only [← coe_nnnorm, nnreal.coe_eq] at hx, simp [hx] }), simp only [snorm', lintegral_congr_ae this] end @@ -383,7 +383,7 @@ lemma snorm_ess_sup_congr_ae {f g : α → F} (hfg : f =ᵐ[μ] g) : snorm_ess_sup f μ = snorm_ess_sup g μ := ess_sup_congr_ae (hfg.fun_comp (coe ∘ nnnorm)) -lemma snorm_mono_ae {f : α → F} {g : α → G} (h : ∀ᵐ x ∂μ, ∥f x∥ ≤ ∥g x∥) : +lemma snorm_mono_ae {f : α → F} {g : α → G} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : snorm f p μ ≤ snorm g p μ := begin simp only [snorm], @@ -394,19 +394,19 @@ begin { exact snorm'_mono_ae ennreal.to_real_nonneg h } end -lemma snorm_mono_ae_real {f : α → F} {g : α → ℝ} (h : ∀ᵐ x ∂μ, ∥f x∥ ≤ g x) : +lemma snorm_mono_ae_real {f : α → F} {g : α → ℝ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ g x) : snorm f p μ ≤ snorm g p μ := snorm_mono_ae $ h.mono (λ x hx, hx.trans ((le_abs_self _).trans (real.norm_eq_abs _).symm.le)) -lemma snorm_mono {f : α → F} {g : α → G} (h : ∀ x, ∥f x∥ ≤ ∥g x∥) : +lemma snorm_mono {f : α → F} {g : α → G} (h : ∀ x, ‖f x‖ ≤ ‖g x‖) : snorm f p μ ≤ snorm g p μ := snorm_mono_ae (eventually_of_forall (λ x, h x)) -lemma snorm_mono_real {f : α → F} {g : α → ℝ} (h : ∀ x, ∥f x∥ ≤ g x) : +lemma snorm_mono_real {f : α → F} {g : α → ℝ} (h : ∀ x, ‖f x‖ ≤ g x) : snorm f p μ ≤ snorm g p μ := snorm_mono_ae_real (eventually_of_forall (λ x, h x)) -lemma snorm_ess_sup_le_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ∥f x∥ ≤ C) : +lemma snorm_ess_sup_le_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : snorm_ess_sup f μ ≤ ennreal.of_real C:= begin simp_rw [snorm_ess_sup, ← of_real_norm_eq_coe_nnnorm], @@ -414,11 +414,11 @@ begin exact ennreal.of_real_le_of_real hx, end -lemma snorm_ess_sup_lt_top_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ∥f x∥ ≤ C) : +lemma snorm_ess_sup_lt_top_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : snorm_ess_sup f μ < ∞ := (snorm_ess_sup_le_of_ae_bound hfC).trans_lt ennreal.of_real_lt_top -lemma snorm_le_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ∥f x∥ ≤ C) : +lemma snorm_le_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : snorm f p μ ≤ ((μ set.univ) ^ p.to_real⁻¹) * (ennreal.of_real C) := begin by_cases hμ : μ = 0, @@ -427,25 +427,25 @@ begin by_cases hp : p = 0, { simp [hp] }, have hC : 0 ≤ C, from le_trans (norm_nonneg _) hfC.exists.some_spec, - have hC' : ∥C∥ = C := by rw [real.norm_eq_abs, abs_eq_self.mpr hC], - have : ∀ᵐ x ∂μ, ∥f x∥ ≤ ∥(λ _, C) x∥, from hfC.mono (λ x hx, hx.trans (le_of_eq hC'.symm)), + have hC' : ‖C‖ = C := by rw [real.norm_eq_abs, abs_eq_self.mpr hC], + have : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖(λ _, C) x‖, from hfC.mono (λ x hx, hx.trans (le_of_eq hC'.symm)), convert snorm_mono_ae this, rw [snorm_const _ hp hμ, mul_comm, ← of_real_norm_eq_coe_nnnorm, hC', one_div] end -lemma snorm_congr_norm_ae {f : α → F} {g : α → G} (hfg : ∀ᵐ x ∂μ, ∥f x∥ = ∥g x∥) : +lemma snorm_congr_norm_ae {f : α → F} {g : α → G} (hfg : ∀ᵐ x ∂μ, ‖f x‖ = ‖g x‖) : snorm f p μ = snorm g p μ := le_antisymm (snorm_mono_ae $ eventually_eq.le hfg) (snorm_mono_ae $ (eventually_eq.symm hfg).le) -@[simp] lemma snorm'_norm {f : α → F} : snorm' (λ a, ∥f a∥) q μ = snorm' f q μ := +@[simp] lemma snorm'_norm {f : α → F} : snorm' (λ a, ‖f a‖) q μ = snorm' f q μ := by simp [snorm'] -@[simp] lemma snorm_norm (f : α → F) : snorm (λ x, ∥f x∥) p μ = snorm f p μ := +@[simp] lemma snorm_norm (f : α → F) : snorm (λ x, ‖f x‖) p μ = snorm f p μ := snorm_congr_norm_ae $ eventually_of_forall $ λ x, norm_norm _ lemma snorm'_norm_rpow (f : α → F) (p q : ℝ) (hq_pos : 0 < q) : - snorm' (λ x, ∥f x∥ ^ q) p μ = (snorm' f (p * q) μ) ^ q := + snorm' (λ x, ‖f x‖ ^ q) p μ = (snorm' f (p * q) μ) ^ q := begin simp_rw snorm', rw [← ennreal.rpow_mul, ←one_div_mul_one_div], @@ -459,15 +459,15 @@ begin end lemma snorm_norm_rpow (f : α → F) (hq_pos : 0 < q) : - snorm (λ x, ∥f x∥ ^ q) p μ = (snorm f (p * ennreal.of_real q) μ) ^ q := + snorm (λ x, ‖f x‖ ^ q) p μ = (snorm f (p * ennreal.of_real q) μ) ^ q := begin by_cases h0 : p = 0, { simp [h0, ennreal.zero_rpow_of_pos hq_pos], }, by_cases hp_top : p = ∞, { simp only [hp_top, snorm_exponent_top, ennreal.top_mul, hq_pos.not_le, ennreal.of_real_eq_zero, if_false, snorm_exponent_top, snorm_ess_sup], - have h_rpow : ess_sup (λ (x : α), (∥(∥f x∥ ^ q)∥₊ : ℝ≥0∞)) μ - = ess_sup (λ (x : α), (↑∥f x∥₊) ^ q) μ, + have h_rpow : ess_sup (λ (x : α), (‖(‖f x‖ ^ q)‖₊ : ℝ≥0∞)) μ + = ess_sup (λ (x : α), (↑‖f x‖₊) ^ q) μ, { congr, ext1 x, nth_rewrite 1 ← nnnorm_norm, @@ -479,7 +479,7 @@ begin have h_rpow_mono := ennreal.strict_mono_rpow_of_pos hq_pos, have h_rpow_surj := (ennreal.rpow_left_bijective hq_pos.ne.symm).2, let iso := h_rpow_mono.order_iso_of_surjective _ h_rpow_surj, - exact (iso.ess_sup_apply (λ x, (∥f x∥₊ : ℝ≥0∞)) μ).symm, }, + exact (iso.ess_sup_apply (λ x, (‖f x‖₊ : ℝ≥0∞)) μ).symm, }, rw [snorm_eq_snorm' h0 hp_top, snorm_eq_snorm' _ _], swap, { refine mul_ne_zero h0 _, rwa [ne.def, ennreal.of_real_eq_zero, not_le], }, swap, { exact ennreal.mul_ne_top hp_top ennreal.of_real_ne_top, }, @@ -497,33 +497,33 @@ lemma mem_ℒp.ae_eq {f g : α → E} (hfg : f =ᵐ[μ] g) (hf_Lp : mem_ℒp f p (mem_ℒp_congr_ae hfg).1 hf_Lp lemma mem_ℒp.of_le {f : α → E} {g : α → F} - (hg : mem_ℒp g p μ) (hf : ae_strongly_measurable f μ) (hfg : ∀ᵐ x ∂μ, ∥f x∥ ≤ ∥g x∥) : + (hg : mem_ℒp g p μ) (hf : ae_strongly_measurable f μ) (hfg : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : mem_ℒp f p μ := ⟨hf, (snorm_mono_ae hfg).trans_lt hg.snorm_lt_top⟩ alias mem_ℒp.of_le ← mem_ℒp.mono lemma mem_ℒp.mono' {f : α → E} {g : α → ℝ} (hg : mem_ℒp g p μ) - (hf : ae_strongly_measurable f μ) (h : ∀ᵐ a ∂μ, ∥f a∥ ≤ g a) : mem_ℒp f p μ := + (hf : ae_strongly_measurable f μ) (h : ∀ᵐ a ∂μ, ‖f a‖ ≤ g a) : mem_ℒp f p μ := hg.mono hf $ h.mono $ λ x hx, le_trans hx (le_abs_self _) lemma mem_ℒp.congr_norm {f : α → E} {g : α → F} (hf : mem_ℒp f p μ) - (hg : ae_strongly_measurable g μ) (h : ∀ᵐ a ∂μ, ∥f a∥ = ∥g a∥) : + (hg : ae_strongly_measurable g μ) (h : ∀ᵐ a ∂μ, ‖f a‖ = ‖g a‖) : mem_ℒp g p μ := hf.mono hg $ eventually_eq.le $ eventually_eq.symm h lemma mem_ℒp_congr_norm {f : α → E} {g : α → F} - (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (h : ∀ᵐ a ∂μ, ∥f a∥ = ∥g a∥) : + (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (h : ∀ᵐ a ∂μ, ‖f a‖ = ‖g a‖) : mem_ℒp f p μ ↔ mem_ℒp g p μ := ⟨λ h2f, h2f.congr_norm hg h, λ h2g, h2g.congr_norm hf $ eventually_eq.symm h⟩ lemma mem_ℒp_top_of_bound {f : α → E} (hf : ae_strongly_measurable f μ) (C : ℝ) - (hfC : ∀ᵐ x ∂μ, ∥f x∥ ≤ C) : + (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : mem_ℒp f ∞ μ := ⟨hf, by { rw snorm_exponent_top, exact snorm_ess_sup_lt_top_of_ae_bound hfC, }⟩ lemma mem_ℒp.of_bound [is_finite_measure μ] {f : α → E} (hf : ae_strongly_measurable f μ) - (C : ℝ) (hfC : ∀ᵐ x ∂μ, ∥f x∥ ≤ C) : + (C : ℝ) (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : mem_ℒp f p μ := (mem_ℒp_const C).of_le hf (hfC.mono (λ x hx, le_trans hx (le_abs_self _))) @@ -531,7 +531,7 @@ lemma mem_ℒp.of_bound [is_finite_measure μ] {f : α → E} (hf : ae_strongly_ snorm' f q ν ≤ snorm' f q μ := begin simp_rw snorm', - suffices h_integral_mono : (∫⁻ a, (∥f a∥₊ : ℝ≥0∞) ^ q ∂ν) ≤ ∫⁻ a, ∥f a∥₊ ^ q ∂μ, + suffices h_integral_mono : (∫⁻ a, (‖f a‖₊ : ℝ≥0∞) ^ q ∂ν) ≤ ∫⁻ a, ‖f a‖₊ ^ q ∂μ, from ennreal.rpow_le_rpow h_integral_mono (by simp [hq]), exact lintegral_mono' hμν le_rfl, end @@ -640,11 +640,11 @@ h.mono_measure $ measure.le_add_right $ le_refl _ lemma mem_ℒp.right_of_add_measure {f : α → E} (h : mem_ℒp f p (μ + ν)) : mem_ℒp f p ν := h.mono_measure $ measure.le_add_left $ le_refl _ -lemma mem_ℒp.norm {f : α → E} (h : mem_ℒp f p μ) : mem_ℒp (λ x, ∥f x∥) p μ := +lemma mem_ℒp.norm {f : α → E} (h : mem_ℒp f p μ) : mem_ℒp (λ x, ‖f x‖) p μ := h.of_le h.ae_strongly_measurable.norm (eventually_of_forall (λ x, by simp)) lemma mem_ℒp_norm_iff {f : α → E} (hf : ae_strongly_measurable f μ) : - mem_ℒp (λ x, ∥f x∥) p μ ↔ mem_ℒp f p μ := + mem_ℒp (λ x, ‖f x‖) p μ ↔ mem_ℒp f p μ := ⟨λ h, ⟨hf, by { rw ← snorm_norm, exact h.2, }⟩, λ h, h.norm⟩ lemma snorm'_eq_zero_of_ae_zero {f : α → F} (hq0_lt : 0 < q) (hf_zero : f =ᵐ[μ] 0) : @@ -677,8 +677,8 @@ lemma snorm'_eq_zero_iff (hq0_lt : 0 < q) {f : α → E} (hf : ae_strongly_measu ⟨ae_eq_zero_of_snorm'_eq_zero (le_of_lt hq0_lt) hf, snorm'_eq_zero_of_ae_zero hq0_lt⟩ lemma coe_nnnorm_ae_le_snorm_ess_sup {m : measurable_space α} (f : α → F) (μ : measure α) : - ∀ᵐ x ∂μ, (∥f x∥₊ : ℝ≥0∞) ≤ snorm_ess_sup f μ := -ennreal.ae_le_ess_sup (λ x, (∥f x∥₊ : ℝ≥0∞)) + ∀ᵐ x ∂μ, (‖f x‖₊ : ℝ≥0∞) ≤ snorm_ess_sup f μ := +ennreal.ae_le_ess_sup (λ x, (‖f x‖₊ : ℝ≥0∞)) @[simp] lemma snorm_ess_sup_eq_zero_iff {f : α → F} : snorm_ess_sup f μ = 0 ↔ f =ᵐ[μ] 0 := by simp [eventually_eq, snorm_ess_sup] @@ -695,9 +695,9 @@ end lemma snorm'_add_le {f g : α → E} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (hq1 : 1 ≤ q) : snorm' (f + g) q μ ≤ snorm' f q μ + snorm' g q μ := -calc (∫⁻ a, ↑∥(f + g) a∥₊ ^ q ∂μ) ^ (1 / q) - ≤ (∫⁻ a, (((λ a, (∥f a∥₊ : ℝ≥0∞)) - + (λ a, (∥g a∥₊ : ℝ≥0∞))) a) ^ q ∂μ) ^ (1 / q) : +calc (∫⁻ a, ↑‖(f + g) a‖₊ ^ q ∂μ) ^ (1 / q) + ≤ (∫⁻ a, (((λ a, (‖f a‖₊ : ℝ≥0∞)) + + (λ a, (‖g a‖₊ : ℝ≥0∞))) a) ^ q ∂μ) ^ (1 / q) : begin refine ennreal.rpow_le_rpow _ (by simp [le_trans zero_le_one hq1] : 0 ≤ 1 / q), refine lintegral_mono (λ a, ennreal.rpow_le_rpow _ (le_trans zero_le_one hq1)), @@ -735,11 +735,11 @@ lemma snorm_sub_le calc snorm (f - g) p μ = snorm (f + - g) p μ : by rw sub_eq_add_neg -- We cannot use snorm_add_le on f and (-g) because we don't have `ae_measurable (-g) μ`, since -- we don't suppose `[borel_space E]`. -... = snorm (λ x, ∥f x + - g x∥) p μ : (snorm_norm (f + - g)).symm -... ≤ snorm (λ x, ∥f x∥ + ∥- g x∥) p μ : by +... = snorm (λ x, ‖f x + - g x‖) p μ : (snorm_norm (f + - g)).symm +... ≤ snorm (λ x, ‖f x‖ + ‖- g x‖) p μ : by { refine snorm_mono_real (λ x, _), rw norm_norm, exact norm_add_le _ _, } -... = snorm (λ x, ∥f x∥ + ∥g x∥) p μ : by simp_rw norm_neg -... ≤ snorm (λ x, ∥f x∥) p μ + snorm (λ x, ∥g x∥) p μ : snorm_add_le hf.norm hg.norm hp1 +... = snorm (λ x, ‖f x‖ + ‖g x‖) p μ : by simp_rw norm_neg +... ≤ snorm (λ x, ‖f x‖) p μ + snorm (λ x, ‖g x‖) p μ : snorm_add_le hf.norm hg.norm hp1 ... = snorm f p μ + snorm g p μ : by rw [← snorm_norm f, ← snorm_norm g] lemma snorm_add_lt_top_of_one_le {f g : α → E} (hf : mem_ℒp f p μ) (hg : mem_ℒp g p μ) @@ -750,15 +750,15 @@ lemma snorm'_add_lt_top_of_le_one {f g : α → E} (hf : ae_strongly_measurable f μ) (hf_snorm : snorm' f q μ < ∞) (hg_snorm : snorm' g q μ < ∞) (hq_pos : 0 < q) (hq1 : q ≤ 1) : snorm' (f + g) q μ < ∞ := -calc (∫⁻ a, ↑∥(f + g) a∥₊ ^ q ∂μ) ^ (1 / q) - ≤ (∫⁻ a, (((λ a, (∥f a∥₊ : ℝ≥0∞)) - + (λ a, (∥g a∥₊ : ℝ≥0∞))) a) ^ q ∂μ) ^ (1 / q) : +calc (∫⁻ a, ↑‖(f + g) a‖₊ ^ q ∂μ) ^ (1 / q) + ≤ (∫⁻ a, (((λ a, (‖f a‖₊ : ℝ≥0∞)) + + (λ a, (‖g a‖₊ : ℝ≥0∞))) a) ^ q ∂μ) ^ (1 / q) : begin refine ennreal.rpow_le_rpow _ (by simp [hq_pos.le] : 0 ≤ 1 / q), refine lintegral_mono (λ a, ennreal.rpow_le_rpow _ hq_pos.le), simp [←ennreal.coe_add, -coe_add, nnnorm_add_le], end -... ≤ (∫⁻ a, (∥f a∥₊ : ℝ≥0∞) ^ q + (∥g a∥₊ : ℝ≥0∞) ^ q ∂μ) ^ (1 / q) : +... ≤ (∫⁻ a, (‖f a‖₊ : ℝ≥0∞) ^ q + (‖g a‖₊ : ℝ≥0∞) ^ q ∂μ) ^ (1 / q) : begin refine ennreal.rpow_le_rpow (lintegral_mono (λ a, _)) (by simp [hq_pos.le] : 0 ≤ 1 / q), exact ennreal.rpow_add_le_add_rpow _ _ hq_pos.le hq1, @@ -937,7 +937,7 @@ begin exact le_rfl, }, have hpq : p < q, from lt_of_le_of_ne hpq hpq_eq, let g := λ a : α, (1 : ℝ≥0∞), - have h_rw : ∫⁻ a, ↑∥f a∥₊^p ∂ μ = ∫⁻ a, (∥f a∥₊ * (g a))^p ∂ μ, + have h_rw : ∫⁻ a, ↑‖f a‖₊^p ∂ μ = ∫⁻ a, (‖f a‖₊ * (g a))^p ∂ μ, from lintegral_congr (λ a, by simp), repeat {rw snorm'}, rw h_rw, @@ -946,17 +946,17 @@ begin { field_simp [(ne_of_lt hp0_lt).symm, (ne_of_lt hq0_lt).symm], ring, }, - calc (∫⁻ (a : α), (↑∥f a∥₊ * g a) ^ p ∂μ) ^ (1/p) - ≤ (∫⁻ (a : α), ↑∥f a∥₊ ^ q ∂μ) ^ (1/q) * (∫⁻ (a : α), (g a) ^ r ∂μ) ^ (1/r) : + calc (∫⁻ (a : α), (↑‖f a‖₊ * g a) ^ p ∂μ) ^ (1/p) + ≤ (∫⁻ (a : α), ↑‖f a‖₊ ^ q ∂μ) ^ (1/q) * (∫⁻ (a : α), (g a) ^ r ∂μ) ^ (1/r) : ennreal.lintegral_Lp_mul_le_Lq_mul_Lr hp0_lt hpq hpqr μ hf.ennnorm ae_measurable_const - ... = (∫⁻ (a : α), ↑∥f a∥₊ ^ q ∂μ) ^ (1/q) * μ set.univ ^ (1/p - 1/q) : + ... = (∫⁻ (a : α), ↑‖f a‖₊ ^ q ∂μ) ^ (1/q) * μ set.univ ^ (1/p - 1/q) : by simp [hpqr], end lemma snorm'_le_snorm_ess_sup_mul_rpow_measure_univ (hq_pos : 0 < q) {f : α → F} : snorm' f q μ ≤ snorm_ess_sup f μ * (μ set.univ) ^ (1/q) := begin - have h_le : ∫⁻ (a : α), ↑∥f a∥₊ ^ q ∂μ ≤ ∫⁻ (a : α), (snorm_ess_sup f μ) ^ q ∂μ, + have h_le : ∫⁻ (a : α), ↑‖f a‖₊ ^ q ∂μ ≤ ∫⁻ (a : α), (snorm_ess_sup f μ) ^ q ∂μ, { refine lintegral_mono_ae _, have h_nnnorm_le_snorm_ess_sup := coe_nnnorm_ae_le_snorm_ess_sup f μ, refine h_nnnorm_le_snorm_ess_sup.mono (λ x hx, ennreal.rpow_le_rpow hx (le_of_lt hq_pos)), }, @@ -1038,7 +1038,7 @@ variables (μ) lemma pow_mul_meas_ge_le_snorm {f : α → E} (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf : ae_strongly_measurable f μ) (ε : ℝ≥0∞) : - (ε * μ {x | ε ≤ ∥f x∥₊ ^ p.to_real}) ^ (1 / p.to_real) ≤ snorm f p μ := + (ε * μ {x | ε ≤ ‖f x‖₊ ^ p.to_real}) ^ (1 / p.to_real) ≤ snorm f p μ := begin rw snorm_eq_lintegral_rpow_nnnorm hp_ne_zero hp_ne_top, exact ennreal.rpow_le_rpow (mul_meas_ge_le_lintegral₀ (hf.ennnorm.pow_const _) ε) @@ -1047,13 +1047,13 @@ end lemma mul_meas_ge_le_pow_snorm {f : α → E} (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf : ae_strongly_measurable f μ) (ε : ℝ≥0∞) : - ε * μ {x | ε ≤ ∥f x∥₊ ^ p.to_real} ≤ snorm f p μ ^ p.to_real := + ε * μ {x | ε ≤ ‖f x‖₊ ^ p.to_real} ≤ snorm f p μ ^ p.to_real := begin have : 1 / p.to_real * p.to_real = 1, { refine one_div_mul_cancel _, rw [ne, ennreal.to_real_eq_zero_iff], exact not_or hp_ne_zero hp_ne_top }, - rw [← ennreal.rpow_one (ε * μ {x | ε ≤ ∥f x∥₊ ^ p.to_real}), ← this, ennreal.rpow_mul], + rw [← ennreal.rpow_one (ε * μ {x | ε ≤ ‖f x‖₊ ^ p.to_real}), ← this, ennreal.rpow_mul], exact ennreal.rpow_le_rpow (pow_mul_meas_ge_le_snorm μ hp_ne_zero hp_ne_top hf ε) ennreal.to_real_nonneg, end @@ -1061,7 +1061,7 @@ end /-- A version of Markov's inequality using Lp-norms. -/ lemma mul_meas_ge_le_pow_snorm' {f : α → E} (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf : ae_strongly_measurable f μ) (ε : ℝ≥0∞) : - ε ^ p.to_real * μ {x | ε ≤ ∥f x∥₊} ≤ snorm f p μ ^ p.to_real := + ε ^ p.to_real * μ {x | ε ≤ ‖f x‖₊} ≤ snorm f p μ ^ p.to_real := begin convert mul_meas_ge_le_pow_snorm μ hp_ne_zero hp_ne_top hf (ε ^ p.to_real), ext x, @@ -1070,7 +1070,7 @@ end lemma meas_ge_le_mul_pow_snorm {f : α → E} (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf : ae_strongly_measurable f μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : - μ {x | ε ≤ ∥f x∥₊} ≤ ε⁻¹ ^ p.to_real * snorm f p μ ^ p.to_real := + μ {x | ε ≤ ‖f x‖₊} ≤ ε⁻¹ ^ p.to_real * snorm f p μ ^ p.to_real := begin by_cases ε = ∞, { simp [h] }, @@ -1167,13 +1167,13 @@ section normed_space variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 E] [normed_space 𝕜 F] lemma snorm'_const_smul {f : α → F} (c : 𝕜) (hq_pos : 0 < q) : - snorm' (c • f) q μ = (∥c∥₊ : ℝ≥0∞) * snorm' f q μ := + snorm' (c • f) q μ = (‖c‖₊ : ℝ≥0∞) * snorm' f q μ := begin rw snorm', simp_rw [pi.smul_apply, nnnorm_smul, ennreal.coe_mul, ennreal.mul_rpow_of_nonneg _ _ hq_pos.le], - suffices h_integral : ∫⁻ a, ↑(∥c∥₊) ^ q * ↑∥f a∥₊ ^ q ∂μ - = (∥c∥₊ : ℝ≥0∞)^q * ∫⁻ a, ∥f a∥₊ ^ q ∂μ, + suffices h_integral : ∫⁻ a, ↑(‖c‖₊) ^ q * ↑‖f a‖₊ ^ q ∂μ + = (‖c‖₊ : ℝ≥0∞)^q * ∫⁻ a, ‖f a‖₊ ^ q ∂μ, { apply_fun (λ x, x ^ (1/q)) at h_integral, rw [h_integral, ennreal.mul_rpow_of_nonneg _ _ (by simp [hq_pos.le] : 0 ≤ 1 / q)], congr, @@ -1184,11 +1184,11 @@ begin end lemma snorm_ess_sup_const_smul {f : α → F} (c : 𝕜) : - snorm_ess_sup (c • f) μ = (∥c∥₊ : ℝ≥0∞) * snorm_ess_sup f μ := + snorm_ess_sup (c • f) μ = (‖c‖₊ : ℝ≥0∞) * snorm_ess_sup f μ := by simp_rw [snorm_ess_sup, pi.smul_apply, nnnorm_smul, ennreal.coe_mul, ennreal.ess_sup_const_mul] lemma snorm_const_smul {f : α → F} (c : 𝕜) : - snorm (c • f) p μ = (∥c∥₊ : ℝ≥0∞) * snorm f p μ := + snorm (c • f) p μ = (‖c‖₊ : ℝ≥0∞) * snorm f p μ := begin by_cases h0 : p = 0, { simp [h0], }, @@ -1229,24 +1229,24 @@ begin by_cases hp_zero : p = 0, { simp only [hp_zero, snorm_exponent_zero, mul_zero, le_zero_iff], }, simp_rw [snorm_eq_lintegral_rpow_nnnorm hp_zero hp_top, snorm_exponent_top, snorm_ess_sup], - calc (∫⁻ x, ↑∥(φ • f) x∥₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) - = (∫⁻ x, ↑∥φ x∥₊ ^ p.to_real * ↑∥f x∥₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : + calc (∫⁻ x, ↑‖(φ • f) x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) + = (∫⁻ x, ↑‖φ x‖₊ ^ p.to_real * ↑‖f x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : begin congr, ext1 x, rw [pi.smul_apply', nnnorm_smul, ennreal.coe_mul, ennreal.mul_rpow_of_nonneg _ _ (ennreal.to_real_nonneg)], end - ... ≤ (∫⁻ x, (ess_sup (λ x, ↑∥φ x∥₊) μ) ^ p.to_real * ↑∥f x∥₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : + ... ≤ (∫⁻ x, (ess_sup (λ x, ↑‖φ x‖₊) μ) ^ p.to_real * ↑‖f x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : begin refine ennreal.rpow_le_rpow _ _, swap, { rw one_div_nonneg, exact ennreal.to_real_nonneg, }, refine lintegral_mono_ae _, - filter_upwards [@ennreal.ae_le_ess_sup _ _ μ (λ x, ↑∥φ x∥₊)] with x hx, + filter_upwards [@ennreal.ae_le_ess_sup _ _ μ (λ x, ↑‖φ x‖₊)] with x hx, refine ennreal.mul_le_mul _ le_rfl, exact ennreal.rpow_le_rpow hx ennreal.to_real_nonneg, end - ... = ess_sup (λ x, ↑∥φ x∥₊) μ * (∫⁻ x, ↑∥f x∥₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : + ... = ess_sup (λ x, ↑‖φ x‖₊) μ * (∫⁻ x, ↑‖f x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : begin rw lintegral_const_mul'', swap, { exact hf.nnnorm.ae_measurable.coe_nnreal_ennreal.pow ae_measurable_const, }, @@ -1264,10 +1264,10 @@ lemma snorm_smul_le_snorm_mul_snorm_top (p : ℝ≥0∞) begin rw ← snorm_norm, simp_rw [pi.smul_apply', norm_smul], - have : (λ x, ∥φ x∥ * ∥f x∥) = (λ x, ∥f x∥) • (λ x, ∥φ x∥), + have : (λ x, ‖φ x‖ * ‖f x‖) = (λ x, ‖f x‖) • (λ x, ‖φ x‖), { rw [smul_eq_mul, mul_comm], refl, }, rw this, - have h := snorm_smul_le_snorm_top_mul_snorm p hφ.norm (λ x, ∥f x∥), + have h := snorm_smul_le_snorm_top_mul_snorm p hφ.norm (λ x, ‖f x‖), refine h.trans_eq _, simp_rw snorm_norm, rw mul_comm, @@ -1330,7 +1330,7 @@ end normed_space section monotonicity lemma snorm_le_mul_snorm_aux_of_nonneg {f : α → F} {g : α → G} {c : ℝ} - (h : ∀ᵐ x ∂μ, ∥f x∥ ≤ c * ∥g x∥) (hc : 0 ≤ c) (p : ℝ≥0∞) : + (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) (hc : 0 ≤ c) (p : ℝ≥0∞) : snorm f p μ ≤ (ennreal.of_real c) * snorm g p μ := begin lift c to ℝ≥0 using hc, @@ -1341,7 +1341,7 @@ begin end lemma snorm_le_mul_snorm_aux_of_neg {f : α → F} {g : α → G} {c : ℝ} - (h : ∀ᵐ x ∂μ, ∥f x∥ ≤ c * ∥g x∥) (hc : c < 0) (p : ℝ≥0∞) : + (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) (hc : c < 0) (p : ℝ≥0∞) : snorm f p μ = 0 ∧ snorm g p μ = 0 := begin suffices : f =ᵐ[μ] 0 ∧ g =ᵐ[μ] 0, @@ -1354,7 +1354,7 @@ begin end lemma snorm_le_mul_snorm_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ} - (h : ∀ᵐ x ∂μ, ∥f x∥ ≤ c * ∥g x∥) (p : ℝ≥0∞) : + (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) (p : ℝ≥0∞) : snorm f p μ ≤ (ennreal.of_real c) * snorm g p μ := begin cases le_or_lt 0 c with hc hc, @@ -1363,7 +1363,7 @@ begin end lemma mem_ℒp.of_le_mul {f : α → E} {g : α → F} {c : ℝ} - (hg : mem_ℒp g p μ) (hf : ae_strongly_measurable f μ) (hfg : ∀ᵐ x ∂μ, ∥f x∥ ≤ c * ∥g x∥) : + (hg : mem_ℒp g p μ) (hf : ae_strongly_measurable f μ) (hfg : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) : mem_ℒp f p μ := begin simp only [mem_ℒp, hf, true_and], @@ -1375,7 +1375,7 @@ end monotonicity lemma snorm_indicator_ge_of_bdd_below (hp : p ≠ 0) (hp' : p ≠ ∞) {f : α → F} (C : ℝ≥0) {s : set α} (hs : measurable_set s) - (hf : ∀ᵐ x ∂μ, x ∈ s → C ≤ ∥s.indicator f x∥₊) : + (hf : ∀ᵐ x ∂μ, x ∈ s → C ≤ ‖s.indicator f x‖₊) : C • μ s ^ (1 / p.to_real) ≤ snorm (s.indicator f) p μ := begin rw [ennreal.smul_def, smul_eq_mul, snorm_eq_lintegral_rpow_nnnorm hp hp', @@ -1397,14 +1397,14 @@ variables {𝕜 : Type*} [is_R_or_C 𝕜] {f : α → 𝕜} lemma mem_ℒp.re (hf : mem_ℒp f p μ) : mem_ℒp (λ x, is_R_or_C.re (f x)) p μ := begin - have : ∀ x, ∥is_R_or_C.re (f x)∥ ≤ 1 * ∥f x∥, + have : ∀ x, ‖is_R_or_C.re (f x)‖ ≤ 1 * ‖f x‖, by { intro x, rw one_mul, exact is_R_or_C.norm_re_le_norm (f x), }, exact hf.of_le_mul hf.1.re (eventually_of_forall this), end lemma mem_ℒp.im (hf : mem_ℒp f p μ) : mem_ℒp (λ x, is_R_or_C.im (f x)) p μ := begin - have : ∀ x, ∥is_R_or_C.im (f x)∥ ≤ 1 * ∥f x∥, + have : ∀ x, ‖is_R_or_C.im (f x)‖ ≤ 1 * ‖f x‖, by { intro x, rw one_mul, exact is_R_or_C.norm_im_le_norm (f x), }, exact hf.of_le_mul hf.1.im (eventually_of_forall this), end @@ -1434,7 +1434,7 @@ variables [measurable_space E] [opens_measurable_space E] {R : ℝ≥0} lemma ae_bdd_liminf_at_top_rpow_of_snorm_bdd {p : ℝ≥0∞} {f : ℕ → α → E} (hfmeas : ∀ n, measurable (f n)) (hbdd : ∀ n, snorm (f n) p μ ≤ R) : - ∀ᵐ x ∂μ, liminf (λ n, (∥f n x∥₊ ^ p.to_real : ℝ≥0∞)) at_top < ∞ := + ∀ᵐ x ∂μ, liminf (λ n, (‖f n x‖₊ ^ p.to_real : ℝ≥0∞)) at_top < ∞ := begin by_cases hp0 : p.to_real = 0, { simp only [hp0, ennreal.rpow_zero], @@ -1457,12 +1457,12 @@ end lemma ae_bdd_liminf_at_top_of_snorm_bdd {p : ℝ≥0∞} (hp : p ≠ 0) {f : ℕ → α → E} (hfmeas : ∀ n, measurable (f n)) (hbdd : ∀ n, snorm (f n) p μ ≤ R) : - ∀ᵐ x ∂μ, liminf (λ n, (∥f n x∥₊ : ℝ≥0∞)) at_top < ∞ := + ∀ᵐ x ∂μ, liminf (λ n, (‖f n x‖₊ : ℝ≥0∞)) at_top < ∞ := begin by_cases hp' : p = ∞, { subst hp', simp_rw snorm_exponent_top at hbdd, - have : ∀ n, ∀ᵐ x ∂μ, (∥f n x∥₊ : ℝ≥0∞) < R + 1 := + have : ∀ n, ∀ᵐ x ∂μ, (‖f n x‖₊ : ℝ≥0∞) < R + 1 := λ n, ae_lt_of_ess_sup_lt (lt_of_le_of_lt (hbdd n) $ ennreal.lt_add_right ennreal.coe_ne_top one_ne_zero), rw ← ae_all_iff at this, @@ -1471,14 +1471,14 @@ begin (ennreal.add_lt_top.2 ⟨ennreal.coe_lt_top, ennreal.one_lt_top⟩) }, filter_upwards [ae_bdd_liminf_at_top_rpow_of_snorm_bdd hfmeas hbdd] with x hx, have hppos : 0 < p.to_real := ennreal.to_real_pos hp hp', - have : liminf (λ n, (∥f n x∥₊ ^ p.to_real : ℝ≥0∞)) at_top = - (liminf (λ n, (∥f n x∥₊ : ℝ≥0∞)) at_top)^ p.to_real, - { change liminf (λ n, ennreal.order_iso_rpow p.to_real hppos (∥f n x∥₊ : ℝ≥0∞)) at_top = - ennreal.order_iso_rpow p.to_real hppos (liminf (λ n, (∥f n x∥₊ : ℝ≥0∞)) at_top), + have : liminf (λ n, (‖f n x‖₊ ^ p.to_real : ℝ≥0∞)) at_top = + (liminf (λ n, (‖f n x‖₊ : ℝ≥0∞)) at_top)^ p.to_real, + { change liminf (λ n, ennreal.order_iso_rpow p.to_real hppos (‖f n x‖₊ : ℝ≥0∞)) at_top = + ennreal.order_iso_rpow p.to_real hppos (liminf (λ n, (‖f n x‖₊ : ℝ≥0∞)) at_top), refine (order_iso.liminf_apply (ennreal.order_iso_rpow p.to_real _) _ _ _ _).symm; is_bounded_default }, rw this at hx, - rw [← ennreal.rpow_one (liminf (λ n, ∥f n x∥₊) at_top), ← mul_inv_cancel hppos.ne.symm, + rw [← ennreal.rpow_one (liminf (λ n, ‖f n x‖₊) at_top), ← mul_inv_cancel hppos.ne.symm, ennreal.rpow_mul], exact ennreal.rpow_lt_top_of_nonneg (inv_nonneg.2 hppos.le) hx.ne, end @@ -1609,14 +1609,14 @@ lemma mem_Lp_const (α) {m : measurable_space α} (μ : measure α) (c : E) [is_ instance : has_norm (Lp E p μ) := { norm := λ f, ennreal.to_real (snorm f p μ) } -instance : has_dist (Lp E p μ) := { dist := λ f g, ∥f - g∥} +instance : has_dist (Lp E p μ) := { dist := λ f g, ‖f - g‖} instance : has_edist (Lp E p μ) := { edist := λ f g, snorm (f - g) p μ } -lemma norm_def (f : Lp E p μ) : ∥f∥ = ennreal.to_real (snorm f p μ) := rfl +lemma norm_def (f : Lp E p μ) : ‖f‖ = ennreal.to_real (snorm f p μ) := rfl @[simp] lemma norm_to_Lp (f : α → E) (hf : mem_ℒp f p μ) : - ∥hf.to_Lp f∥ = ennreal.to_real (snorm f p μ) := + ‖hf.to_Lp f‖ = ennreal.to_real (snorm f p μ) := by rw [norm_def, snorm_congr_ae (mem_ℒp.coe_fn_to_Lp hf)] lemma dist_def (f g : Lp E p μ) : dist f g = (snorm (f - g) p μ).to_real := @@ -1637,13 +1637,13 @@ by { rw edist_def, exact snorm_congr_ae (hf.coe_fn_to_Lp.sub hg.coe_fn_to_Lp) } edist (hf.to_Lp f) 0 = snorm f p μ := by { convert edist_to_Lp_to_Lp f 0 hf zero_mem_ℒp, simp } -@[simp] lemma norm_zero : ∥(0 : Lp E p μ)∥ = 0 := +@[simp] lemma norm_zero : ‖(0 : Lp E p μ)‖ = 0 := begin change (snorm ⇑(0 : α →ₘ[μ] E) p μ).to_real = 0, simp [snorm_congr_ae ae_eq_fun.coe_fn_zero, snorm_zero] end -lemma norm_eq_zero_iff {f : Lp E p μ} (hp : 0 < p) : ∥f∥ = 0 ↔ f = 0 := +lemma norm_eq_zero_iff {f : Lp E p μ} (hp : 0 < p) : ‖f‖ = 0 ↔ f = 0 := begin refine ⟨λ hf, _, λ hf, by simp [hf]⟩, rw [norm_def, ennreal.to_real_eq_zero_iff] at hf, @@ -1666,11 +1666,11 @@ begin exact h'a.symm, }, end -@[simp] lemma norm_neg {f : Lp E p μ} : ∥-f∥ = ∥f∥ := +@[simp] lemma norm_neg {f : Lp E p μ} : ‖-f‖ = ‖f‖ := by rw [norm_def, norm_def, snorm_congr_ae (coe_fn_neg _), snorm_neg] lemma norm_le_mul_norm_of_ae_le_mul {c : ℝ} {f : Lp E p μ} {g : Lp F p μ} - (h : ∀ᵐ x ∂μ, ∥f x∥ ≤ c * ∥g x∥) : ∥f∥ ≤ c * ∥g∥ := + (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) : ‖f‖ ≤ c * ‖g‖ := begin by_cases pzero : p = 0, { simp [pzero, norm_def] }, @@ -1686,28 +1686,28 @@ begin simp [this] } end -lemma norm_le_norm_of_ae_le {f : Lp E p μ} {g : Lp F p μ} (h : ∀ᵐ x ∂μ, ∥f x∥ ≤ ∥g x∥) : - ∥f∥ ≤ ∥g∥ := +lemma norm_le_norm_of_ae_le {f : Lp E p μ} {g : Lp F p μ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : + ‖f‖ ≤ ‖g‖ := begin rw [norm_def, norm_def, ennreal.to_real_le_to_real (snorm_ne_top _) (snorm_ne_top _)], exact snorm_mono_ae h end -lemma mem_Lp_of_ae_le_mul {c : ℝ} {f : α →ₘ[μ] E} {g : Lp F p μ} (h : ∀ᵐ x ∂μ, ∥f x∥ ≤ c * ∥g x∥) : +lemma mem_Lp_of_ae_le_mul {c : ℝ} {f : α →ₘ[μ] E} {g : Lp F p μ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) : f ∈ Lp E p μ := mem_Lp_iff_mem_ℒp.2 $ mem_ℒp.of_le_mul (Lp.mem_ℒp g) f.ae_strongly_measurable h -lemma mem_Lp_of_ae_le {f : α →ₘ[μ] E} {g : Lp F p μ} (h : ∀ᵐ x ∂μ, ∥f x∥ ≤ ∥g x∥) : +lemma mem_Lp_of_ae_le {f : α →ₘ[μ] E} {g : Lp F p μ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : f ∈ Lp E p μ := mem_Lp_iff_mem_ℒp.2 $ mem_ℒp.of_le (Lp.mem_ℒp g) f.ae_strongly_measurable h -lemma mem_Lp_of_ae_bound [is_finite_measure μ] {f : α →ₘ[μ] E} (C : ℝ) (hfC : ∀ᵐ x ∂μ, ∥f x∥ ≤ C) : +lemma mem_Lp_of_ae_bound [is_finite_measure μ] {f : α →ₘ[μ] E} (C : ℝ) (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : f ∈ Lp E p μ := mem_Lp_iff_mem_ℒp.2 $ mem_ℒp.of_bound f.ae_strongly_measurable _ hfC lemma norm_le_of_ae_bound [is_finite_measure μ] {f : Lp E p μ} {C : ℝ} (hC : 0 ≤ C) - (hfC : ∀ᵐ x ∂μ, ∥f x∥ ≤ C) : - ∥f∥ ≤ (measure_univ_nnreal μ) ^ (p.to_real)⁻¹ * C := + (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : + ‖f‖ ≤ (measure_univ_nnreal μ) ^ (p.to_real)⁻¹ * C := begin by_cases hμ : μ = 0, { by_cases hp : p.to_real⁻¹ = 0, @@ -1776,7 +1776,7 @@ instance : module 𝕜 (Lp E p μ) := lemma coe_fn_smul (c : 𝕜) (f : Lp E p μ) : ⇑(c • f) =ᵐ[μ] c • f := ae_eq_fun.coe_fn_smul _ _ -lemma norm_const_smul (c : 𝕜) (f : Lp E p μ) : ∥c • f∥ = ∥c∥ * ∥f∥ := +lemma norm_const_smul (c : 𝕜) (f : Lp E p μ) : ‖c • f‖ = ‖c‖ * ‖f‖ := by rw [norm_def, snorm_congr_ae (coe_fn_smul _ _), snorm_const_smul c, ennreal.to_real_mul, ennreal.coe_to_real, coe_nnnorm, norm_def] @@ -1815,7 +1815,7 @@ begin end lemma snorm_ess_sup_indicator_const_le (s : set α) (c : G) : - snorm_ess_sup (s.indicator (λ x : α , c)) μ ≤ ∥c∥₊ := + snorm_ess_sup (s.indicator (λ x : α , c)) μ ≤ ‖c‖₊ := begin by_cases hμ0 : μ = 0, { rw [hμ0, snorm_ess_sup_measure_zero], @@ -1824,7 +1824,7 @@ begin end lemma snorm_ess_sup_indicator_const_eq (s : set α) (c : G) (hμs : μ s ≠ 0) : - snorm_ess_sup (s.indicator (λ x : α , c)) μ = ∥c∥₊ := + snorm_ess_sup (s.indicator (λ x : α , c)) μ = ‖c‖₊ := begin refine le_antisymm (snorm_ess_sup_indicator_const_le s c) _, by_contra' h, @@ -1841,7 +1841,7 @@ lemma snorm_indicator_le {E : Type*} [normed_add_comm_group E] (f : α → E) : snorm (s.indicator f) p μ ≤ snorm f p μ := begin refine snorm_mono_ae (eventually_of_forall (λ x, _)), - suffices : ∥s.indicator f x∥₊ ≤ ∥f x∥₊, + suffices : ‖s.indicator f x‖₊ ≤ ‖f x‖₊, { exact nnreal.coe_mono this }, rw nnnorm_indicator_eq_indicator_nnnorm, exact s.indicator_le_self _ x, @@ -1850,14 +1850,14 @@ end variables {hs} lemma snorm_indicator_const {c : G} (hs : measurable_set s) (hp : p ≠ 0) (hp_top : p ≠ ∞) : - snorm (s.indicator (λ x, c)) p μ = ∥c∥₊ * (μ s) ^ (1 / p.to_real) := + snorm (s.indicator (λ x, c)) p μ = ‖c‖₊ * (μ s) ^ (1 / p.to_real) := begin have hp_pos : 0 < p.to_real, from ennreal.to_real_pos hp hp_top, rw snorm_eq_lintegral_rpow_nnnorm hp hp_top, simp_rw [nnnorm_indicator_eq_indicator_nnnorm, ennreal.coe_indicator], - have h_indicator_pow : (λ a : α, s.indicator (λ (x : α), (∥c∥₊ : ℝ≥0∞)) a ^ p.to_real) - = s.indicator (λ (x : α), ↑∥c∥₊ ^ p.to_real), - { rw set.comp_indicator_const (∥c∥₊ : ℝ≥0∞) (λ x, x ^ p.to_real) _, + have h_indicator_pow : (λ a : α, s.indicator (λ (x : α), (‖c‖₊ : ℝ≥0∞)) a ^ p.to_real) + = s.indicator (λ (x : α), ↑‖c‖₊ ^ p.to_real), + { rw set.comp_indicator_const (‖c‖₊ : ℝ≥0∞) (λ x, x ^ p.to_real) _, simp [hp_pos], }, rw [h_indicator_pow, lintegral_indicator _ hs, set_lintegral_const, ennreal.mul_rpow_of_nonneg], { rw [← ennreal.rpow_mul, mul_one_div_cancel hp_pos.ne.symm, ennreal.rpow_one], }, @@ -1865,7 +1865,7 @@ begin end lemma snorm_indicator_const' {c : G} (hs : measurable_set s) (hμs : μ s ≠ 0) (hp : p ≠ 0) : - snorm (s.indicator (λ _, c)) p μ = ∥c∥₊ * (μ s) ^ (1 / p.to_real) := + snorm (s.indicator (λ _, c)) p μ = ‖c‖₊ * (μ s) ^ (1 / p.to_real) := begin by_cases hp_top : p = ∞, { simp [hp_top, snorm_ess_sup_indicator_const_eq s c hμs], }, @@ -1901,7 +1901,7 @@ begin { simp_rw [hp_top, snorm_exponent_top], exact snorm_ess_sup_indicator_eq_snorm_ess_sup_restrict hs, }, simp_rw snorm_eq_lintegral_rpow_nnnorm hp_zero hp_top, - suffices : ∫⁻ x, ∥s.indicator f x∥₊ ^ p.to_real ∂μ = ∫⁻ x in s, ∥f x∥₊ ^ p.to_real ∂μ, + suffices : ∫⁻ x, ‖s.indicator f x‖₊ ^ p.to_real ∂μ = ∫⁻ x in s, ‖f x‖₊ ^ p.to_real ∂μ, by rw this, rw ← lintegral_indicator _ hs, congr, @@ -1923,7 +1923,7 @@ begin { rw hp_zero, exact mem_ℒp_zero_iff_ae_strongly_measurable.mpr ae_strongly_measurable_const, }, by_cases hp_top : p = ∞, { rw hp_top, - exact mem_ℒp_top_of_bound ae_strongly_measurable_const (∥c∥) + exact mem_ℒp_top_of_bound ae_strongly_measurable_const (‖c‖) (eventually_of_forall (λ x, le_rfl)), }, rw [mem_ℒp_const_iff hp_zero hp_top, measure.restrict_apply_univ], cases hμsc, @@ -1955,18 +1955,18 @@ lemma indicator_const_Lp_coe_fn_nmem : indicator_const_Lp_coe_fn.mono (λ x hx hxs, hx.trans (set.indicator_of_not_mem hxs _)) lemma norm_indicator_const_Lp (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) : - ∥indicator_const_Lp p hs hμs c∥ = ∥c∥ * (μ s).to_real ^ (1 / p.to_real) := + ‖indicator_const_Lp p hs hμs c‖ = ‖c‖ * (μ s).to_real ^ (1 / p.to_real) := by rw [Lp.norm_def, snorm_congr_ae indicator_const_Lp_coe_fn, snorm_indicator_const hs hp_ne_zero hp_ne_top, ennreal.to_real_mul, ennreal.to_real_rpow, ennreal.coe_to_real, coe_nnnorm] -lemma norm_indicator_const_Lp_top (hμs_ne_zero : μ s ≠ 0) : ∥indicator_const_Lp ∞ hs hμs c∥ = ∥c∥ := +lemma norm_indicator_const_Lp_top (hμs_ne_zero : μ s ≠ 0) : ‖indicator_const_Lp ∞ hs hμs c‖ = ‖c‖ := by rw [Lp.norm_def, snorm_congr_ae indicator_const_Lp_coe_fn, snorm_indicator_const' hs hμs_ne_zero ennreal.top_ne_zero, ennreal.top_to_real, div_zero, ennreal.rpow_zero, mul_one, ennreal.coe_to_real, coe_nnnorm] lemma norm_indicator_const_Lp' (hp_pos : p ≠ 0) (hμs_pos : μ s ≠ 0) : - ∥indicator_const_Lp p hs hμs c∥ = ∥c∥ * (μ s).to_real ^ (1 / p.to_real) := + ‖indicator_const_Lp p hs hμs c‖ = ‖c‖ * (μ s).to_real ^ (1 / p.to_real) := begin by_cases hp_top : p = ∞, { rw [hp_top, ennreal.top_to_real, div_zero, real.rpow_zero, mul_one], @@ -2010,7 +2010,7 @@ end indicator_const_Lp lemma mem_ℒp.norm_rpow_div {f : α → E} (hf : mem_ℒp f p μ) (q : ℝ≥0∞) : - mem_ℒp (λ (x : α), ∥f x∥ ^ q.to_real) (p/q) μ := + mem_ℒp (λ (x : α), ‖f x‖ ^ q.to_real) (p/q) μ := begin refine ⟨(hf.1.norm.ae_measurable.pow_const q.to_real).ae_strongly_measurable, _⟩, by_cases q_top : q = ∞, { simp [q_top] }, @@ -2028,7 +2028,7 @@ end lemma mem_ℒp_norm_rpow_iff {q : ℝ≥0∞} {f : α → E} (hf : ae_strongly_measurable f μ) (q_zero : q ≠ 0) (q_top : q ≠ ∞) : - mem_ℒp (λ (x : α), ∥f x∥ ^ q.to_real) (p/q) μ ↔ mem_ℒp f p μ := + mem_ℒp (λ (x : α), ‖f x‖ ^ q.to_real) (p/q) μ ↔ mem_ℒp f p μ := begin refine ⟨λ h, _, λ h, h.norm_rpow_div q⟩, apply (mem_ℒp_norm_iff hf).1, @@ -2043,7 +2043,7 @@ end lemma mem_ℒp.norm_rpow {f : α → E} (hf : mem_ℒp f p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) : - mem_ℒp (λ (x : α), ∥f x∥ ^ p.to_real) 1 μ := + mem_ℒp (λ (x : α), ‖f x‖ ^ p.to_real) 1 μ := begin convert hf.norm_rpow_div p, rw [div_eq_mul_inv, ennreal.mul_inv_cancel hp_ne_zero hp_ne_top], @@ -2069,7 +2069,7 @@ lemma lipschitz_with.comp_mem_ℒp {α E F} {K} [measurable_space α] {μ : meas [normed_add_comm_group E] [normed_add_comm_group F] {f : α → E} {g : E → F} (hg : lipschitz_with K g) (g0 : g 0 = 0) (hL : mem_ℒp f p μ) : mem_ℒp (g ∘ f) p μ := begin - have : ∀ᵐ x ∂μ, ∥g (f x)∥ ≤ K * ∥f x∥, + have : ∀ᵐ x ∂μ, ‖g (f x)‖ ≤ K * ‖f x‖, { apply filter.eventually_of_forall (λ x, _), rw [← dist_zero_right, ← dist_zero_right, ← g0], apply hg.dist_le_mul }, @@ -2081,7 +2081,7 @@ lemma measure_theory.mem_ℒp.of_comp_antilipschitz_with {α E F} {K'} {f : α → E} {g : E → F} (hL : mem_ℒp (g ∘ f) p μ) (hg : uniform_continuous g) (hg' : antilipschitz_with K' g) (g0 : g 0 = 0) : mem_ℒp f p μ := begin - have A : ∀ᵐ x ∂μ, ∥f x∥ ≤ K' * ∥g (f x)∥, + have A : ∀ᵐ x ∂μ, ‖f x‖ ≤ K' * ‖g (f x)‖, { apply filter.eventually_of_forall (λ x, _), rw [← dist_zero_right, ← dist_zero_right, ← g0], apply hg'.le_mul_dist }, @@ -2103,7 +2103,7 @@ defined as an element of `Lp`. -/ def comp_Lp (hg : lipschitz_with c g) (g0 : g 0 = 0) (f : Lp E p μ) : Lp F p μ := ⟨ae_eq_fun.comp g hg.continuous (f : α →ₘ[μ] E), begin - suffices : ∀ᵐ x ∂μ, ∥ae_eq_fun.comp g hg.continuous (f : α →ₘ[μ] E) x∥ ≤ c * ∥f x∥, + suffices : ∀ᵐ x ∂μ, ‖ae_eq_fun.comp g hg.continuous (f : α →ₘ[μ] E) x‖ ≤ c * ‖f x‖, { exact Lp.mem_Lp_of_ae_le_mul this }, filter_upwards [ae_eq_fun.coe_fn_comp g hg.continuous (f : α →ₘ[μ] E)] with a ha, simp only [ha], @@ -2125,7 +2125,7 @@ begin end lemma norm_comp_Lp_sub_le (hg : lipschitz_with c g) (g0 : g 0 = 0) (f f' : Lp E p μ) : - ∥hg.comp_Lp g0 f - hg.comp_Lp g0 f'∥ ≤ c * ∥f - f'∥ := + ‖hg.comp_Lp g0 f - hg.comp_Lp g0 f'‖ ≤ c * ‖f - f'‖ := begin apply Lp.norm_le_mul_norm_of_ae_le_mul, filter_upwards [hg.coe_fn_comp_Lp g0 f, hg.coe_fn_comp_Lp g0 f', @@ -2135,7 +2135,7 @@ begin end lemma norm_comp_Lp_le (hg : lipschitz_with c g) (g0 : g 0 = 0) (f : Lp E p μ) : - ∥hg.comp_Lp g0 f∥ ≤ c * ∥f∥ := + ‖hg.comp_Lp g0 f‖ ≤ c * ‖f‖ := by simpa using hg.norm_comp_Lp_sub_le g0 f 0 lemma lipschitz_with_comp_Lp [fact (1 ≤ p)] (hg : lipschitz_with c g) (g0 : g 0 = 0) : @@ -2215,7 +2215,7 @@ begin refl, end -lemma norm_comp_Lp_le (L : E →L[𝕜] F) (f : Lp E p μ) : ∥L.comp_Lp f∥ ≤ ∥L∥ * ∥f∥ := +lemma norm_comp_Lp_le (L : E →L[𝕜] F) (f : Lp E p μ) : ‖L.comp_Lp f‖ ≤ ‖L‖ * ‖f‖ := lipschitz_with.norm_comp_Lp_le _ _ _ variables (μ p) @@ -2248,7 +2248,7 @@ def comp_Lpₗ (L : E →L[𝕜] F) : (Lp E p μ) →ₗ[𝕜] (Lp F p μ) := * `continuous_linear_map.comp_left_continuous_compact` for continuous functions on compact spaces. -/ def comp_LpL [fact (1 ≤ p)] (L : E →L[𝕜] F) : (Lp E p μ) →L[𝕜] (Lp F p μ) := -linear_map.mk_continuous (L.comp_Lpₗ p μ) ∥L∥ L.norm_comp_Lp_le +linear_map.mk_continuous (L.comp_Lpₗ p μ) ‖L‖ L.norm_comp_Lp_le variables {μ p} @@ -2271,7 +2271,7 @@ lemma smul_comp_LpL_apply [fact (1 ≤ p)] {𝕜'} [normed_field 𝕜'] [normed_ smul_comp_Lp c L f lemma norm_compLpL_le [fact (1 ≤ p)] (L : E →L[𝕜] F) : - ∥L.comp_LpL p μ∥ ≤ ∥L∥ := + ‖L.comp_LpL p μ‖ ≤ ‖L‖ := linear_map.mk_continuous_norm_le _ (norm_nonneg _) _ end continuous_linear_map @@ -2363,15 +2363,15 @@ namespace Lp lemma snorm'_lim_eq_lintegral_liminf {ι} [nonempty ι] [linear_order ι] {f : ι → α → G} {p : ℝ} (hp_nonneg : 0 ≤ p) {f_lim : α → G} (h_lim : ∀ᵐ (x : α) ∂μ, tendsto (λ n, f n x) at_top (𝓝 (f_lim x))) : - snorm' f_lim p μ = (∫⁻ a, at_top.liminf (λ m, (∥f m a∥₊ : ℝ≥0∞)^p) ∂μ) ^ (1/p) := + snorm' f_lim p μ = (∫⁻ a, at_top.liminf (λ m, (‖f m a‖₊ : ℝ≥0∞)^p) ∂μ) ^ (1/p) := begin - suffices h_no_pow : (∫⁻ a, ∥f_lim a∥₊ ^ p ∂μ) - = (∫⁻ a, at_top.liminf (λ m, (∥f m a∥₊ : ℝ≥0∞)^p) ∂μ), + suffices h_no_pow : (∫⁻ a, ‖f_lim a‖₊ ^ p ∂μ) + = (∫⁻ a, at_top.liminf (λ m, (‖f m a‖₊ : ℝ≥0∞)^p) ∂μ), { rw [snorm', h_no_pow], }, refine lintegral_congr_ae (h_lim.mono (λ a ha, _)), rw tendsto.liminf_eq, simp_rw [ennreal.coe_rpow_of_nonneg _ hp_nonneg, ennreal.tendsto_coe], - refine ((nnreal.continuous_rpow_const hp_nonneg).tendsto (∥f_lim a∥₊)).comp _, + refine ((nnreal.continuous_rpow_const hp_nonneg).tendsto (‖f_lim a‖₊)).comp _, exact (continuous_nnnorm.tendsto (f_lim a)).comp ha, end @@ -2396,7 +2396,7 @@ end lemma snorm_exponent_top_lim_eq_ess_sup_liminf {ι} [nonempty ι] [linear_order ι] {f : ι → α → G} {f_lim : α → G} (h_lim : ∀ᵐ (x : α) ∂μ, tendsto (λ n, f n x) at_top (𝓝 (f_lim x))) : - snorm f_lim ∞ μ = ess_sup (λ x, at_top.liminf (λ m, (∥f m x∥₊ : ℝ≥0∞))) μ := + snorm f_lim ∞ μ = ess_sup (λ x, at_top.liminf (λ m, (‖f m x‖₊ : ℝ≥0∞))) μ := begin rw [snorm_exponent_top, snorm_ess_sup], refine ess_sup_congr_ae (h_lim.mono (λ x hx, _)), @@ -2412,7 +2412,7 @@ lemma snorm_exponent_top_lim_le_liminf_snorm_exponent_top {ι} [nonempty ι] [co begin rw snorm_exponent_top_lim_eq_ess_sup_liminf h_lim, simp_rw [snorm_exponent_top, snorm_ess_sup], - exact ennreal.ess_sup_liminf_le (λ n, (λ x, (∥f n x∥₊ : ℝ≥0∞))), + exact ennreal.ess_sup_liminf_le (λ n, (λ x, (‖f n x‖₊ : ℝ≥0∞))), end lemma snorm_lim_le_liminf_snorm {E} [normed_add_comm_group E] @@ -2524,10 +2524,10 @@ end private lemma snorm'_sum_norm_sub_le_tsum_of_cauchy_snorm' {f : ℕ → α → E} (hf : ∀ n, ae_strongly_measurable (f n) μ) {p : ℝ} (hp1 : 1 ≤ p) {B : ℕ → ℝ≥0∞} (h_cau : ∀ (N n m : ℕ), N ≤ n → N ≤ m → snorm' (f n - f m) p μ < B N) (n : ℕ) : - snorm' (λ x, ∑ i in finset.range (n + 1), ∥f (i + 1) x - f i x∥) p μ ≤ ∑' i, B i := + snorm' (λ x, ∑ i in finset.range (n + 1), ‖f (i + 1) x - f i x‖) p μ ≤ ∑' i, B i := begin - let f_norm_diff := λ i x, ∥f (i + 1) x - f i x∥, - have hgf_norm_diff : ∀ n, (λ x, ∑ i in finset.range (n + 1), ∥f (i + 1) x - f i x∥) + let f_norm_diff := λ i x, ‖f (i + 1) x - f i x‖, + have hgf_norm_diff : ∀ n, (λ x, ∑ i in finset.range (n + 1), ‖f (i + 1) x - f i x‖) = ∑ i in finset.range (n + 1), f_norm_diff i, from λ n, funext (λ x, by simp [f_norm_diff]), rw hgf_norm_diff, @@ -2539,8 +2539,8 @@ end private lemma lintegral_rpow_sum_coe_nnnorm_sub_le_rpow_tsum {f : ℕ → α → E} (hf : ∀ n, ae_strongly_measurable (f n) μ) {p : ℝ} (hp1 : 1 ≤ p) {B : ℕ → ℝ≥0∞} (n : ℕ) - (hn : snorm' (λ x, ∑ i in finset.range (n + 1), ∥f (i + 1) x - f i x∥) p μ ≤ ∑' i, B i) : - ∫⁻ a, (∑ i in finset.range (n + 1), ∥f (i + 1) a - f i a∥₊ : ℝ≥0∞)^p ∂μ + (hn : snorm' (λ x, ∑ i in finset.range (n + 1), ‖f (i + 1) x - f i x‖) p μ ≤ ∑' i, B i) : + ∫⁻ a, (∑ i in finset.range (n + 1), ‖f (i + 1) a - f i a‖₊ : ℝ≥0∞)^p ∂μ ≤ (∑' i, B i) ^ p := begin have hp_pos : 0 < p := zero_lt_one.trans_le hp1, @@ -2548,8 +2548,8 @@ begin one_div_one_div p], simp_rw snorm' at hn, have h_nnnorm_nonneg : - (λ a, (∥∑ i in finset.range (n + 1), ∥f (i + 1) a - f i a∥∥₊ : ℝ≥0∞) ^ p) - = λ a, (∑ i in finset.range (n + 1), (∥f (i + 1) a - f i a∥₊ : ℝ≥0∞)) ^ p, + (λ a, (‖∑ i in finset.range (n + 1), ‖f (i + 1) a - f i a‖‖₊ : ℝ≥0∞) ^ p) + = λ a, (∑ i in finset.range (n + 1), (‖f (i + 1) a - f i a‖₊ : ℝ≥0∞)) ^ p, { ext1 a, congr, simp_rw ←of_real_norm_eq_coe_nnnorm, @@ -2557,19 +2557,19 @@ begin { rw real.norm_of_nonneg _, exact finset.sum_nonneg (λ x hx, norm_nonneg _), }, { exact λ x hx, norm_nonneg _, }, }, - change (∫⁻ a, (λ x, ↑∥∑ i in finset.range (n + 1), ∥f (i+1) x - f i x∥∥₊^p) a ∂μ)^(1/p) + change (∫⁻ a, (λ x, ↑‖∑ i in finset.range (n + 1), ‖f (i+1) x - f i x‖‖₊^p) a ∂μ)^(1/p) ≤ ∑' i, B i at hn, rwa h_nnnorm_nonneg at hn, end private lemma lintegral_rpow_tsum_coe_nnnorm_sub_le_tsum {f : ℕ → α → E} (hf : ∀ n, ae_strongly_measurable (f n) μ) {p : ℝ} (hp1 : 1 ≤ p) {B : ℕ → ℝ≥0∞} - (h : ∀ n, ∫⁻ a, (∑ i in finset.range (n + 1), ∥f (i + 1) a - f i a∥₊ : ℝ≥0∞)^p ∂μ + (h : ∀ n, ∫⁻ a, (∑ i in finset.range (n + 1), ‖f (i + 1) a - f i a‖₊ : ℝ≥0∞)^p ∂μ ≤ (∑' i, B i) ^ p) : - (∫⁻ a, (∑' i, ∥f (i + 1) a - f i a∥₊ : ℝ≥0∞)^p ∂μ) ^ (1/p) ≤ ∑' i, B i := + (∫⁻ a, (∑' i, ‖f (i + 1) a - f i a‖₊ : ℝ≥0∞)^p ∂μ) ^ (1/p) ≤ ∑' i, B i := begin have hp_pos : 0 < p := zero_lt_one.trans_le hp1, - suffices h_pow : ∫⁻ a, (∑' i, ∥f (i + 1) a - f i a∥₊ : ℝ≥0∞)^p ∂μ ≤ (∑' i, B i) ^ p, + suffices h_pow : ∫⁻ a, (∑' i, ‖f (i + 1) a - f i a‖₊ : ℝ≥0∞)^p ∂μ ≤ (∑' i, B i) ^ p, by rwa [←ennreal.le_rpow_one_div_iff (by simp [hp_pos] : 0 < 1 / p), one_div_one_div], have h_tsum_1 : ∀ g : ℕ → ℝ≥0∞, ∑' i, g i = at_top.liminf (λ n, ∑ i in finset.range (n + 1), g i), @@ -2577,8 +2577,8 @@ begin simp_rw h_tsum_1 _, rw ← h_tsum_1, have h_liminf_pow : ∫⁻ a, at_top.liminf (λ n, ∑ i in finset.range (n + 1), - (∥f (i + 1) a - f i a∥₊))^p ∂μ - = ∫⁻ a, at_top.liminf (λ n, (∑ i in finset.range (n + 1), (∥f (i + 1) a - f i a∥₊))^p) ∂μ, + (‖f (i + 1) a - f i a‖₊))^p ∂μ + = ∫⁻ a, at_top.liminf (λ n, (∑ i in finset.range (n + 1), (‖f (i + 1) a - f i a‖₊))^p) ∂μ, { refine lintegral_congr (λ x, _), have h_rpow_mono := ennreal.strict_mono_rpow_of_pos (zero_lt_one.trans_le hp1), have h_rpow_surj := (ennreal.rpow_left_bijective hp_pos.ne.symm).2, @@ -2594,16 +2594,16 @@ end private lemma tsum_nnnorm_sub_ae_lt_top {f : ℕ → α → E} (hf : ∀ n, ae_strongly_measurable (f n) μ) {p : ℝ} (hp1 : 1 ≤ p) {B : ℕ → ℝ≥0∞} (hB : ∑' i, B i ≠ ∞) - (h : (∫⁻ a, (∑' i, ∥f (i + 1) a - f i a∥₊ : ℝ≥0∞)^p ∂μ) ^ (1/p) ≤ ∑' i, B i) : - ∀ᵐ x ∂μ, (∑' i, ∥f (i + 1) x - f i x∥₊ : ℝ≥0∞) < ∞ := + (h : (∫⁻ a, (∑' i, ‖f (i + 1) a - f i a‖₊ : ℝ≥0∞)^p ∂μ) ^ (1/p) ≤ ∑' i, B i) : + ∀ᵐ x ∂μ, (∑' i, ‖f (i + 1) x - f i x‖₊ : ℝ≥0∞) < ∞ := begin have hp_pos : 0 < p := zero_lt_one.trans_le hp1, - have h_integral : ∫⁻ a, (∑' i, ∥f (i + 1) a - f i a∥₊ : ℝ≥0∞)^p ∂μ < ∞, + have h_integral : ∫⁻ a, (∑' i, ‖f (i + 1) a - f i a‖₊ : ℝ≥0∞)^p ∂μ < ∞, { have h_tsum_lt_top : (∑' i, B i) ^ p < ∞, from ennreal.rpow_lt_top_of_nonneg hp_pos.le hB, refine lt_of_le_of_lt _ h_tsum_lt_top, rwa [←ennreal.le_rpow_one_div_iff (by simp [hp_pos] : 0 < 1 / p), one_div_one_div] at h, }, - have rpow_ae_lt_top : ∀ᵐ x ∂μ, (∑' i, ∥f (i + 1) x - f i x∥₊ : ℝ≥0∞)^p < ∞, + have rpow_ae_lt_top : ∀ᵐ x ∂μ, (∑' i, ‖f (i + 1) x - f i x‖₊ : ℝ≥0∞)^p < ∞, { refine ae_lt_top' (ae_measurable.pow_const _ _) h_integral.ne, exact ae_measurable.ennreal_tsum (λ n, ((hf (n+1)).sub (hf n)).ennnorm), }, refine rpow_ae_lt_top.mono (λ x hx, _), @@ -2617,15 +2617,15 @@ lemma ae_tendsto_of_cauchy_snorm' [complete_space E] {f : ℕ → α → E} {p : ∀ᵐ x ∂μ, ∃ l : E, at_top.tendsto (λ n, f n x) (𝓝 l) := begin have h_summable : ∀ᵐ x ∂μ, summable (λ (i : ℕ), f (i + 1) x - f i x), - { have h1 : ∀ n, snorm' (λ x, ∑ i in finset.range (n + 1), ∥f (i + 1) x - f i x∥) p μ + { have h1 : ∀ n, snorm' (λ x, ∑ i in finset.range (n + 1), ‖f (i + 1) x - f i x‖) p μ ≤ ∑' i, B i, from snorm'_sum_norm_sub_le_tsum_of_cauchy_snorm' hf hp1 h_cau, - have h2 : ∀ n, ∫⁻ a, (∑ i in finset.range (n + 1), ∥f (i + 1) a - f i a∥₊ : ℝ≥0∞)^p ∂μ + have h2 : ∀ n, ∫⁻ a, (∑ i in finset.range (n + 1), ‖f (i + 1) a - f i a‖₊ : ℝ≥0∞)^p ∂μ ≤ (∑' i, B i) ^ p, from λ n, lintegral_rpow_sum_coe_nnnorm_sub_le_rpow_tsum hf hp1 n (h1 n), - have h3 : (∫⁻ a, (∑' i, ∥f (i + 1) a - f i a∥₊ : ℝ≥0∞)^p ∂μ) ^ (1/p) ≤ ∑' i, B i, + have h3 : (∫⁻ a, (∑' i, ‖f (i + 1) a - f i a‖₊ : ℝ≥0∞)^p ∂μ) ^ (1/p) ≤ ∑' i, B i, from lintegral_rpow_tsum_coe_nnnorm_sub_le_tsum hf hp1 h2, - have h4 : ∀ᵐ x ∂μ, (∑' i, ∥f (i + 1) x - f i x∥₊ : ℝ≥0∞) < ∞, + have h4 : ∀ᵐ x ∂μ, (∑' i, ‖f (i + 1) x - f i x‖₊ : ℝ≥0∞) < ∞, from tsum_nnnorm_sub_ae_lt_top hf hp1 hB h3, exact h4.mono (λ x hx, summable_of_summable_nnnorm (ennreal.tsum_coe_ne_top_iff_summable.mp (lt_top_iff_ne_top.mp hx))), }, @@ -2653,7 +2653,7 @@ lemma ae_tendsto_of_cauchy_snorm [complete_space E] {f : ℕ → α → E} begin by_cases hp_top : p = ∞, { simp_rw [hp_top] at *, - have h_cau_ae : ∀ᵐ x ∂μ, ∀ N n m, N ≤ n → N ≤ m → (∥(f n - f m) x∥₊ : ℝ≥0∞) < B N, + have h_cau_ae : ∀ᵐ x ∂μ, ∀ N n m, N ≤ n → N ≤ m → (‖(f n - f m) x‖₊ : ℝ≥0∞) < B N, { simp_rw ae_all_iff, exact λ N n m hnN hmN, ae_lt_of_ess_sup_lt (h_cau N n m hnN hmN), }, simp_rw [snorm_exponent_top, snorm_ess_sup] at h_cau, @@ -2787,7 +2787,7 @@ variables [is_finite_measure μ] lemma mem_Lp (f : α →ᵇ E) : f.to_continuous_map.to_ae_eq_fun μ ∈ Lp E p μ := begin - refine Lp.mem_Lp_of_ae_bound (∥f∥) _, + refine Lp.mem_Lp_of_ae_bound (‖f‖) _, filter_upwards [f.to_continuous_map.coe_fn_to_ae_eq_fun μ] with x _, convert f.norm_coe_le_norm x end @@ -2795,8 +2795,8 @@ end /-- The `Lp`-norm of a bounded continuous function is at most a constant (depending on the measure of the whole space) times its sup-norm. -/ lemma Lp_norm_le (f : α →ᵇ E) : - ∥(⟨f.to_continuous_map.to_ae_eq_fun μ, mem_Lp f⟩ : Lp E p μ)∥ - ≤ (measure_univ_nnreal μ) ^ (p.to_real)⁻¹ * ∥f∥ := + ‖(⟨f.to_continuous_map.to_ae_eq_fun μ, mem_Lp f⟩ : Lp E p μ)‖ + ≤ (measure_univ_nnreal μ) ^ (p.to_real)⁻¹ * ‖f‖ := begin apply Lp.norm_le_of_ae_bound (norm_nonneg f), { refine (f.to_continuous_map.coe_fn_to_ae_eq_fun μ).mono _, @@ -2854,7 +2854,7 @@ lemma coe_fn_to_Lp [normed_field 𝕜] [normed_space 𝕜 E] [fact (1 ≤ p)] (f ae_eq_fun.coe_fn_mk f _ lemma to_Lp_norm_le [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] [fact (1 ≤ p)] : - ∥(to_Lp p μ 𝕜 : (α →ᵇ E) →L[𝕜] (Lp E p μ))∥ ≤ (measure_univ_nnreal μ) ^ (p.to_real)⁻¹ := + ‖(to_Lp p μ 𝕜 : (α →ᵇ E) →L[𝕜] (Lp E p μ))‖ ≤ (measure_univ_nnreal μ) ^ (p.to_real)⁻¹ := linear_map.mk_continuous_norm_le _ ((measure_univ_nnreal μ) ^ (p.to_real)⁻¹).coe_nonneg _ end bounded_continuous_function @@ -2909,13 +2909,13 @@ rfl variables [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] lemma to_Lp_norm_eq_to_Lp_norm_coe : - ∥(to_Lp p μ 𝕜 : C(α, E) →L[𝕜] (Lp E p μ))∥ - = ∥(bounded_continuous_function.to_Lp p μ 𝕜 : (α →ᵇ E) →L[𝕜] (Lp E p μ))∥ := + ‖(to_Lp p μ 𝕜 : C(α, E) →L[𝕜] (Lp E p μ))‖ + = ‖(bounded_continuous_function.to_Lp p μ 𝕜 : (α →ᵇ E) →L[𝕜] (Lp E p μ))‖ := continuous_linear_map.op_norm_comp_linear_isometry_equiv _ _ /-- Bound for the operator norm of `continuous_map.to_Lp`. -/ lemma to_Lp_norm_le : - ∥(to_Lp p μ 𝕜 : C(α, E) →L[𝕜] (Lp E p μ))∥ ≤ (measure_univ_nnreal μ) ^ (p.to_real)⁻¹ := + ‖(to_Lp p μ 𝕜 : C(α, E) →L[𝕜] (Lp E p μ))‖ ≤ (measure_univ_nnreal μ) ^ (p.to_real)⁻¹ := by { rw to_Lp_norm_eq_to_Lp_norm_coe, exact bounded_continuous_function.to_Lp_norm_le μ } end continuous_map @@ -2928,26 +2928,26 @@ namespace Lp lemma pow_mul_meas_ge_le_norm (f : Lp E p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (ε : ℝ≥0∞) : - (ε * μ {x | ε ≤ ∥f x∥₊ ^ p.to_real}) ^ (1 / p.to_real) ≤ (ennreal.of_real ∥f∥) := + (ε * μ {x | ε ≤ ‖f x‖₊ ^ p.to_real}) ^ (1 / p.to_real) ≤ (ennreal.of_real ‖f‖) := (ennreal.of_real_to_real (snorm_ne_top f)).symm ▸ pow_mul_meas_ge_le_snorm μ hp_ne_zero hp_ne_top (Lp.ae_strongly_measurable f) ε lemma mul_meas_ge_le_pow_norm (f : Lp E p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (ε : ℝ≥0∞) : - ε * μ {x | ε ≤ ∥f x∥₊ ^ p.to_real} ≤ (ennreal.of_real ∥f∥) ^ p.to_real := + ε * μ {x | ε ≤ ‖f x‖₊ ^ p.to_real} ≤ (ennreal.of_real ‖f‖) ^ p.to_real := (ennreal.of_real_to_real (snorm_ne_top f)).symm ▸ mul_meas_ge_le_pow_snorm μ hp_ne_zero hp_ne_top (Lp.ae_strongly_measurable f) ε /-- A version of Markov's inequality with elements of Lp. -/ lemma mul_meas_ge_le_pow_norm' (f : Lp E p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (ε : ℝ≥0∞) : - ε ^ p.to_real * μ {x | ε ≤ ∥f x∥₊} ≤ (ennreal.of_real ∥f∥) ^ p.to_real := + ε ^ p.to_real * μ {x | ε ≤ ‖f x‖₊} ≤ (ennreal.of_real ‖f‖) ^ p.to_real := (ennreal.of_real_to_real (snorm_ne_top f)).symm ▸ mul_meas_ge_le_pow_snorm' μ hp_ne_zero hp_ne_top (Lp.ae_strongly_measurable f) ε lemma meas_ge_le_mul_pow_norm (f : Lp E p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) {ε : ℝ≥0∞} (hε : ε ≠ 0) : - μ {x | ε ≤ ∥f x∥₊} ≤ ε⁻¹ ^ p.to_real * (ennreal.of_real ∥f∥) ^ p.to_real := + μ {x | ε ≤ ‖f x‖₊} ≤ ε⁻¹ ^ p.to_real * (ennreal.of_real ‖f‖) ^ p.to_real := (ennreal.of_real_to_real (snorm_ne_top f)).symm ▸ meas_ge_le_mul_pow_snorm μ hp_ne_zero hp_ne_top (Lp.ae_strongly_measurable f) hε diff --git a/src/measure_theory/function/simple_func_dense_lp.lean b/src/measure_theory/function/simple_func_dense_lp.lean index d37cbc611c38d..8eb99c82022c8 100644 --- a/src/measure_theory/function/simple_func_dense_lp.lean +++ b/src/measure_theory/function/simple_func_dense_lp.lean @@ -58,7 +58,7 @@ variables [measurable_space β] [measurable_space E] [normed_add_comm_group E] lemma nnnorm_approx_on_le [opens_measurable_space E] {f : β → E} (hf : measurable f) {s : set E} {y₀ : E} (h₀ : y₀ ∈ s) [separable_space s] (x : β) (n : ℕ) : - ∥approx_on f hf s y₀ h₀ n x - f x∥₊ ≤ ∥f x - y₀∥₊ := + ‖approx_on f hf s y₀ h₀ n x - f x‖₊ ≤ ‖f x - y₀‖₊ := begin have := edist_approx_on_le hf h₀ x n, rw edist_comm y₀ at this, @@ -68,7 +68,7 @@ end lemma norm_approx_on_y₀_le [opens_measurable_space E] {f : β → E} (hf : measurable f) {s : set E} {y₀ : E} (h₀ : y₀ ∈ s) [separable_space s] (x : β) (n : ℕ) : - ∥approx_on f hf s y₀ h₀ n x - y₀∥ ≤ ∥f x - y₀∥ + ∥f x - y₀∥ := + ‖approx_on f hf s y₀ h₀ n x - y₀‖ ≤ ‖f x - y₀‖ + ‖f x - y₀‖ := begin have := edist_approx_on_y0_le hf h₀ x n, repeat { rw [edist_comm y₀, edist_eq_coe_nnnorm_sub] at this }, @@ -77,7 +77,7 @@ end lemma norm_approx_on_zero_le [opens_measurable_space E] {f : β → E} (hf : measurable f) {s : set E} (h₀ : (0 : E) ∈ s) [separable_space s] (x : β) (n : ℕ) : - ∥approx_on f hf s 0 h₀ n x∥ ≤ ∥f x∥ + ∥f x∥ := + ‖approx_on f hf s 0 h₀ n x‖ ≤ ‖f x‖ + ‖f x‖ := begin have := edist_approx_on_y0_le hf h₀ x n, simp [edist_comm (0 : E), edist_eq_coe_nnnorm] at this, @@ -93,29 +93,29 @@ begin by_cases hp_zero : p = 0, { simpa only [hp_zero, snorm_exponent_zero] using tendsto_const_nhds }, have hp : 0 < p.to_real := to_real_pos hp_zero hp_ne_top, - suffices : tendsto (λ n, ∫⁻ x, ∥approx_on f hf s y₀ h₀ n x - f x∥₊ ^ p.to_real ∂μ) at_top (𝓝 0), + suffices : tendsto (λ n, ∫⁻ x, ‖approx_on f hf s y₀ h₀ n x - f x‖₊ ^ p.to_real ∂μ) at_top (𝓝 0), { simp only [snorm_eq_lintegral_rpow_nnnorm hp_zero hp_ne_top], convert continuous_rpow_const.continuous_at.tendsto.comp this; simp [_root_.inv_pos.mpr hp] }, -- We simply check the conditions of the Dominated Convergence Theorem: -- (1) The function "`p`-th power of distance between `f` and the approximation" is measurable - have hF_meas : ∀ n, measurable (λ x, (∥approx_on f hf s y₀ h₀ n x - f x∥₊ : ℝ≥0∞) ^ p.to_real), + have hF_meas : ∀ n, measurable (λ x, (‖approx_on f hf s y₀ h₀ n x - f x‖₊ : ℝ≥0∞) ^ p.to_real), { simpa only [← edist_eq_coe_nnnorm_sub] using λ n, (approx_on f hf s y₀ h₀ n).measurable_bind (λ y x, (edist y (f x)) ^ p.to_real) (λ y, (measurable_edist_right.comp hf).pow_const p.to_real) }, -- (2) The functions "`p`-th power of distance between `f` and the approximation" are uniformly - -- bounded, at any given point, by `λ x, ∥f x - y₀∥ ^ p.to_real` - have h_bound : ∀ n, (λ x, (∥approx_on f hf s y₀ h₀ n x - f x∥₊ : ℝ≥0∞) ^ p.to_real) - ≤ᵐ[μ] (λ x, ∥f x - y₀∥₊ ^ p.to_real), + -- bounded, at any given point, by `λ x, ‖f x - y₀‖ ^ p.to_real` + have h_bound : ∀ n, (λ x, (‖approx_on f hf s y₀ h₀ n x - f x‖₊ : ℝ≥0∞) ^ p.to_real) + ≤ᵐ[μ] (λ x, ‖f x - y₀‖₊ ^ p.to_real), { exact λ n, eventually_of_forall (λ x, rpow_le_rpow (coe_mono (nnnorm_approx_on_le hf h₀ x n)) to_real_nonneg) }, - -- (3) The bounding function `λ x, ∥f x - y₀∥ ^ p.to_real` has finite integral - have h_fin : ∫⁻ (a : β), ∥f a - y₀∥₊ ^ p.to_real ∂μ ≠ ⊤, + -- (3) The bounding function `λ x, ‖f x - y₀‖ ^ p.to_real` has finite integral + have h_fin : ∫⁻ (a : β), ‖f a - y₀‖₊ ^ p.to_real ∂μ ≠ ⊤, from (lintegral_rpow_nnnorm_lt_top_of_snorm_lt_top hp_zero hp_ne_top hi).ne, -- (4) The functions "`p`-th power of distance between `f` and the approximation" tend pointwise -- to zero have h_lim : ∀ᵐ (a : β) ∂μ, - tendsto (λ n, (∥approx_on f hf s y₀ h₀ n a - f a∥₊ : ℝ≥0∞) ^ p.to_real) at_top (𝓝 0), + tendsto (λ n, (‖approx_on f hf s y₀ h₀ n a - f a‖₊ : ℝ≥0∞) ^ p.to_real) at_top (𝓝 0), { filter_upwards [hμ] with a ha, have : tendsto (λ n, (approx_on f hf s y₀ h₀ n) a - f a) at_top (𝓝 (f a - f a)), { exact (tendsto_approx_on hf h₀ ha).sub tendsto_const_nhds }, @@ -137,8 +137,8 @@ begin convert snorm_add_lt_top this hi₀, ext x, simp }, - have hf' : mem_ℒp (λ x, ∥f x - y₀∥) p μ, - { have h_meas : measurable (λ x, ∥f x - y₀∥), + have hf' : mem_ℒp (λ x, ‖f x - y₀‖) p μ, + { have h_meas : measurable (λ x, ‖f x - y₀‖), { simp only [← dist_eq_norm], exact (continuous_id.dist continuous_const).measurable.comp fmeas }, refine ⟨h_meas.ae_measurable.ae_strongly_measurable, _⟩, @@ -146,14 +146,14 @@ begin convert snorm_add_lt_top hf hi₀.neg, ext x, simp [sub_eq_add_neg] }, - have : ∀ᵐ x ∂μ, ∥approx_on f fmeas s y₀ h₀ n x - y₀∥ ≤ ∥(∥f x - y₀∥ + ∥f x - y₀∥)∥, + have : ∀ᵐ x ∂μ, ‖approx_on f fmeas s y₀ h₀ n x - y₀‖ ≤ ‖(‖f x - y₀‖ + ‖f x - y₀‖)‖, { refine eventually_of_forall _, intros x, convert norm_approx_on_y₀_le fmeas h₀ x n, rw [real.norm_eq_abs, abs_of_nonneg], exact add_nonneg (norm_nonneg _) (norm_nonneg _) }, calc snorm (λ x, approx_on f fmeas s y₀ h₀ n x - y₀) p μ - ≤ snorm (λ x, ∥f x - y₀∥ + ∥f x - y₀∥) p μ : snorm_mono_ae this + ≤ snorm (λ x, ‖f x - y₀‖ + ‖f x - y₀‖) p μ : snorm_mono_ae this ... < ⊤ : snorm_add_lt_top hf' hf', end @@ -197,7 +197,7 @@ variables [measurable_space E] [normed_add_comm_group E] lemma tendsto_approx_on_L1_nnnorm [opens_measurable_space E] {f : β → E} (hf : measurable f) {s : set E} {y₀ : E} (h₀ : y₀ ∈ s) [separable_space s] {μ : measure β} (hμ : ∀ᵐ x ∂μ, f x ∈ closure s) (hi : has_finite_integral (λ x, f x - y₀) μ) : - tendsto (λ n, ∫⁻ x, ∥approx_on f hf s y₀ h₀ n x - f x∥₊ ∂μ) at_top (𝓝 0) := + tendsto (λ n, ∫⁻ x, ‖approx_on f hf s y₀ h₀ n x - f x‖₊ ∂μ) at_top (𝓝 0) := by simpa [snorm_one_eq_lintegral_nnnorm] using tendsto_approx_on_Lp_snorm hf h₀ one_ne_top hμ (by simpa [snorm_one_eq_lintegral_nnnorm] using hi) @@ -214,7 +214,7 @@ end lemma tendsto_approx_on_range_L1_nnnorm [opens_measurable_space E] {f : β → E} {μ : measure β} [separable_space (range f ∪ {0} : set E)] (fmeas : measurable f) (hf : integrable f μ) : - tendsto (λ n, ∫⁻ x, ∥approx_on f fmeas (range f ∪ {0}) 0 (by simp) n x - f x∥₊ ∂μ) + tendsto (λ n, ∫⁻ x, ‖approx_on f fmeas (range f ∪ {0}) 0 (by simp) n x - f x‖₊ ∂μ) at_top (𝓝 0) := begin apply tendsto_approx_on_L1_nnnorm fmeas, @@ -248,8 +248,8 @@ A simple function `f : α →ₛ E` into a normed group `E` verifies, for a meas `mem_ℒp f p μ ↔ integrable f μ ↔ f.fin_meas_supp μ ↔ ∀ y ≠ 0, μ (f ⁻¹' {y}) < ∞`. -/ -lemma exists_forall_norm_le (f : α →ₛ F) : ∃ C, ∀ x, ∥f x∥ ≤ C := -exists_forall_le (f.map (λ x, ∥x∥)) +lemma exists_forall_norm_le (f : α →ₛ F) : ∃ C, ∀ x, ‖f x‖ ≤ C := +exists_forall_le (f.map (λ x, ‖x‖)) lemma mem_ℒp_zero (f : α →ₛ E) (μ : measure α) : mem_ℒp f 0 μ := mem_ℒp_zero_iff_ae_strongly_measurable.mpr f.ae_strongly_measurable @@ -259,8 +259,8 @@ let ⟨C, hfC⟩ := f.exists_forall_norm_le in mem_ℒp_top_of_bound f.ae_strongly_measurable C $ eventually_of_forall hfC protected lemma snorm'_eq {p : ℝ} (f : α →ₛ F) (μ : measure α) : - snorm' f p μ = (∑ y in f.range, (∥y∥₊ : ℝ≥0∞) ^ p * μ (f ⁻¹' {y})) ^ (1/p) := -have h_map : (λ a, (∥f a∥₊ : ℝ≥0∞) ^ p) = f.map (λ a : F, (∥a∥₊ : ℝ≥0∞) ^ p), by simp, + snorm' f p μ = (∑ y in f.range, (‖y‖₊ : ℝ≥0∞) ^ p * μ (f ⁻¹' {y})) ^ (1/p) := +have h_map : (λ a, (‖f a‖₊ : ℝ≥0∞) ^ p) = f.map (λ a : F, (‖a‖₊ : ℝ≥0∞) ^ p), by simp, by rw [snorm', h_map, lintegral_eq_lintegral, map_lintegral] lemma measure_preimage_lt_top_of_mem_ℒp (hp_pos : p ≠ 0) (hp_ne_top : p ≠ ∞) (f : α →ₛ E) @@ -489,7 +489,7 @@ lemma to_Lp_smul (f : α →ₛ E) (hf : mem_ℒp f p μ) (c : 𝕜) : to_Lp (c • f) (hf.const_smul c) = c • to_Lp f hf := rfl lemma norm_to_Lp [fact (1 ≤ p)] (f : α →ₛ E) (hf : mem_ℒp f p μ) : - ∥to_Lp f hf∥ = ennreal.to_real (snorm f p μ) := + ‖to_Lp f hf‖ = ennreal.to_real (snorm f p μ) := norm_to_Lp f hf end to_Lp @@ -588,7 +588,7 @@ begin end lemma norm_to_simple_func [fact (1 ≤ p)] (f : Lp.simple_func E p μ) : - ∥f∥ = ennreal.to_real (snorm (to_simple_func f) p μ) := + ‖f‖ = ennreal.to_real (snorm (to_simple_func f) p μ) := by simpa [to_Lp_to_simple_func] using norm_to_Lp (to_simple_func f) (simple_func.mem_ℒp f) end to_simple_func diff --git a/src/measure_theory/function/strongly_measurable/basic.lean b/src/measure_theory/function/strongly_measurable/basic.lean index 8c563d88b43ee..789d5830408d7 100644 --- a/src/measure_theory/function/strongly_measurable/basic.lean +++ b/src/measure_theory/function/strongly_measurable/basic.lean @@ -180,30 +180,30 @@ protected lemma tendsto_approx {m : measurable_space α} (hf : strongly_measurab hf.some_spec /-- Similar to `strongly_measurable.approx`, but enforces that the norm of every function in the -sequence is less than `c` everywhere. If `∥f x∥ ≤ c` this sequence of simple functions verifies +sequence is less than `c` everywhere. If `‖f x‖ ≤ c` this sequence of simple functions verifies `tendsto (λ n, hf.approx_bounded n x) at_top (𝓝 (f x))`. -/ noncomputable def approx_bounded {m : measurable_space α} [has_norm β] [has_smul ℝ β] (hf : strongly_measurable f) (c : ℝ) : ℕ → simple_func α β := -λ n, (hf.approx n).map (λ x, (min 1 (c / ∥x∥)) • x) +λ n, (hf.approx n).map (λ x, (min 1 (c / ‖x‖)) • x) lemma tendsto_approx_bounded_of_norm_le {β} {f : α → β} [normed_add_comm_group β] [normed_space ℝ β] - {m : measurable_space α} (hf : strongly_measurable[m] f) {c : ℝ} {x : α} (hfx : ∥f x∥ ≤ c) : + {m : measurable_space α} (hf : strongly_measurable[m] f) {c : ℝ} {x : α} (hfx : ‖f x‖ ≤ c) : tendsto (λ n, hf.approx_bounded c n x) at_top (𝓝 (f x)) := begin have h_tendsto := hf.tendsto_approx x, simp only [strongly_measurable.approx_bounded, simple_func.coe_map, function.comp_app], - by_cases hfx0 : ∥f x∥ = 0, + by_cases hfx0 : ‖f x‖ = 0, { rw norm_eq_zero at hfx0, rw hfx0 at h_tendsto ⊢, - have h_tendsto_norm : tendsto (λ n, ∥hf.approx n x∥) at_top (𝓝 0), + have h_tendsto_norm : tendsto (λ n, ‖hf.approx n x‖) at_top (𝓝 0), { convert h_tendsto.norm, rw norm_zero, }, refine squeeze_zero_norm (λ n, _) h_tendsto_norm, - calc ∥min 1 (c / ∥hf.approx n x∥) • hf.approx n x∥ - = ∥min 1 (c / ∥hf.approx n x∥)∥ * ∥hf.approx n x∥ : norm_smul _ _ - ... ≤ ∥(1 : ℝ)∥ * ∥hf.approx n x∥ : + calc ‖min 1 (c / ‖hf.approx n x‖) • hf.approx n x‖ + = ‖min 1 (c / ‖hf.approx n x‖)‖ * ‖hf.approx n x‖ : norm_smul _ _ + ... ≤ ‖(1 : ℝ)‖ * ‖hf.approx n x‖ : begin refine mul_le_mul_of_nonneg_right _ (norm_nonneg _), rw [norm_one, real.norm_of_nonneg], @@ -211,10 +211,10 @@ begin { exact le_min zero_le_one (div_nonneg ((norm_nonneg _).trans hfx) (norm_nonneg _)), }, end - ... = ∥hf.approx n x∥ : by rw [norm_one, one_mul], }, + ... = ‖hf.approx n x‖ : by rw [norm_one, one_mul], }, rw ← one_smul ℝ (f x), refine tendsto.smul _ h_tendsto, - have : min 1 (c / ∥f x∥) = 1, + have : min 1 (c / ‖f x‖) = 1, { rw [min_eq_left_iff, one_le_div (lt_of_le_of_ne (norm_nonneg _) (ne.symm hfx0))], exact hfx, }, nth_rewrite 0 this.symm, @@ -225,20 +225,20 @@ end lemma tendsto_approx_bounded_ae {β} {f : α → β} [normed_add_comm_group β] [normed_space ℝ β] {m m0 : measurable_space α} {μ : measure α} (hf : strongly_measurable[m] f) {c : ℝ} - (hf_bound : ∀ᵐ x ∂μ, ∥f x∥ ≤ c) : + (hf_bound : ∀ᵐ x ∂μ, ‖f x‖ ≤ c) : ∀ᵐ x ∂μ, tendsto (λ n, hf.approx_bounded c n x) at_top (𝓝 (f x)) := by filter_upwards [hf_bound] with x hfx using tendsto_approx_bounded_of_norm_le hf hfx lemma norm_approx_bounded_le {β} {f : α → β} [seminormed_add_comm_group β] [normed_space ℝ β] {m : measurable_space α} {c : ℝ} (hf : strongly_measurable[m] f) (hc : 0 ≤ c) (n : ℕ) (x : α) : - ∥hf.approx_bounded c n x∥ ≤ c := + ‖hf.approx_bounded c n x‖ ≤ c := begin simp only [strongly_measurable.approx_bounded, simple_func.coe_map, function.comp_app], refine (norm_smul _ _).le.trans _, - by_cases h0 : ∥hf.approx n x∥ = 0, + by_cases h0 : ‖hf.approx n x‖ = 0, { simp only [h0, div_zero, min_eq_right, zero_le_one, norm_zero, mul_zero], exact hc, }, - cases le_total (∥hf.approx n x∥) c, + cases le_total (‖hf.approx n x‖) c, { rw min_eq_left _, { simpa only [norm_one, one_mul] using h, }, { rwa one_le_div (lt_of_le_of_ne (norm_nonneg _) (ne.symm h0)), }, }, @@ -781,17 +781,17 @@ continuous_dist.comp_strongly_measurable (hf.prod_mk hg) protected lemma norm {m : measurable_space α} {β : Type*} [seminormed_add_comm_group β] {f : α → β} (hf : strongly_measurable f) : - strongly_measurable (λ x, ∥f x∥) := + strongly_measurable (λ x, ‖f x‖) := continuous_norm.comp_strongly_measurable hf protected lemma nnnorm {m : measurable_space α} {β : Type*} [seminormed_add_comm_group β] {f : α → β} (hf : strongly_measurable f) : - strongly_measurable (λ x, ∥f x∥₊) := + strongly_measurable (λ x, ‖f x‖₊) := continuous_nnnorm.comp_strongly_measurable hf protected lemma ennnorm {m : measurable_space α} {β : Type*} [seminormed_add_comm_group β] {f : α → β} (hf : strongly_measurable f) : - measurable (λ a, (∥f a∥₊ : ℝ≥0∞)) := + measurable (λ a, (‖f a‖₊ : ℝ≥0∞)) := (ennreal.continuous_coe.comp_strongly_measurable hf.nnnorm).measurable protected lemma real_to_nnreal {m : measurable_space α} {f : α → ℝ} @@ -916,14 +916,14 @@ norm. In particular, `f` is integrable on each of those sets. -/ lemma exists_spanning_measurable_set_norm_le [seminormed_add_comm_group β] {m m0 : measurable_space α} (hm : m ≤ m0) (hf : strongly_measurable[m] f) (μ : measure α) [sigma_finite (μ.trim hm)] : - ∃ s : ℕ → set α, (∀ n, measurable_set[m] (s n) ∧ μ (s n) < ∞ ∧ ∀ x ∈ s n, ∥f x∥ ≤ n) + ∃ s : ℕ → set α, (∀ n, measurable_set[m] (s n) ∧ μ (s n) < ∞ ∧ ∀ x ∈ s n, ‖f x‖ ≤ n) ∧ (⋃ i, s i) = set.univ := begin let sigma_finite_sets := spanning_sets (μ.trim hm), - let norm_sets := λ (n : ℕ), {x | ∥f x∥ ≤ n}, + let norm_sets := λ (n : ℕ), {x | ‖f x‖ ≤ n}, have norm_sets_spanning : (⋃ n, norm_sets n) = set.univ, { ext1 x, simp only [set.mem_Union, set.mem_set_of_eq, set.mem_univ, iff_true], - exact ⟨⌈∥f x∥⌉₊, nat.le_ceil (∥f x∥)⟩, }, + exact ⟨⌈‖f x‖⌉₊, nat.le_ceil (‖f x‖)⟩, }, let sets := λ n, sigma_finite_sets n ∩ norm_sets n, have h_meas : ∀ n, measurable_set[m] (sets n), { refine λ n, measurable_set.inter _ _, @@ -1381,17 +1381,17 @@ continuous_dist.comp_ae_strongly_measurable (hf.prod_mk hg) protected lemma norm {β : Type*} [seminormed_add_comm_group β] {f : α → β} (hf : ae_strongly_measurable f μ) : - ae_strongly_measurable (λ x, ∥f x∥) μ := + ae_strongly_measurable (λ x, ‖f x‖) μ := continuous_norm.comp_ae_strongly_measurable hf protected lemma nnnorm {β : Type*} [seminormed_add_comm_group β] {f : α → β} (hf : ae_strongly_measurable f μ) : - ae_strongly_measurable (λ x, ∥f x∥₊) μ := + ae_strongly_measurable (λ x, ‖f x‖₊) μ := continuous_nnnorm.comp_ae_strongly_measurable hf protected lemma ennnorm {β : Type*} [seminormed_add_comm_group β] {f : α → β} (hf : ae_strongly_measurable f μ) : - ae_measurable (λ a, (∥f a∥₊ : ℝ≥0∞)) μ := + ae_measurable (λ a, (‖f a‖₊ : ℝ≥0∞)) μ := (ennreal.continuous_coe.comp_ae_strongly_measurable hf.nnnorm).ae_measurable protected lemma edist {β : Type*} [seminormed_add_comm_group β] {f g : α → β} diff --git a/src/measure_theory/function/uniform_integrable.lean b/src/measure_theory/function/uniform_integrable.lean index 6c9e3178bac49..789d0c95688e2 100644 --- a/src/measure_theory/function/uniform_integrable.lean +++ b/src/measure_theory/function/uniform_integrable.lean @@ -144,9 +144,9 @@ lemma unif_integrable_congr_ae {p : ℝ≥0∞} {f g : ι → α → β} (hfg : ⟨λ hf, hf.ae_eq hfg, λ hg, hg.ae_eq (λ n, (hfg n).symm)⟩ lemma tendsto_indicator_ge (f : α → β) (x : α): - tendsto (λ M : ℕ, {x | (M : ℝ) ≤ ∥f x∥₊}.indicator f x) at_top (𝓝 0) := + tendsto (λ M : ℕ, {x | (M : ℝ) ≤ ‖f x‖₊}.indicator f x) at_top (𝓝 0) := begin - refine @tendsto_at_top_of_eventually_const _ _ _ _ _ _ _ (nat.ceil (∥f x∥₊ : ℝ) + 1) (λ n hn, _), + refine @tendsto_at_top_of_eventually_const _ _ _ _ _ _ _ (nat.ceil (‖f x‖₊ : ℝ) + 1) (λ n hn, _), rw indicator_of_not_mem, simp only [not_le, mem_set_of_eq], refine lt_of_le_of_lt (nat.le_ceil _) _, @@ -165,16 +165,16 @@ variables {f : α → β} as the latter provides `0 ≤ M` and does not require the measurability of `f`. -/ lemma mem_ℒp.integral_indicator_norm_ge_le (hf : mem_ℒp f 1 μ) (hmeas : strongly_measurable f) {ε : ℝ} (hε : 0 < ε) : - ∃ M : ℝ, ∫⁻ x, ∥{x | M ≤ ∥f x∥₊}.indicator f x∥₊ ∂μ ≤ ennreal.of_real ε := + ∃ M : ℝ, ∫⁻ x, ‖{x | M ≤ ‖f x‖₊}.indicator f x‖₊ ∂μ ≤ ennreal.of_real ε := begin - have htendsto : ∀ᵐ x ∂μ, tendsto (λ M : ℕ, {x | (M : ℝ) ≤ ∥f x∥₊}.indicator f x) at_top (𝓝 0) := + have htendsto : ∀ᵐ x ∂μ, tendsto (λ M : ℕ, {x | (M : ℝ) ≤ ‖f x‖₊}.indicator f x) at_top (𝓝 0) := univ_mem' (id $ λ x, tendsto_indicator_ge f x), - have hmeas : ∀ M : ℕ, ae_strongly_measurable ({x | (M : ℝ) ≤ ∥f x∥₊}.indicator f) μ, + have hmeas : ∀ M : ℕ, ae_strongly_measurable ({x | (M : ℝ) ≤ ‖f x‖₊}.indicator f) μ, { assume M, apply hf.1.indicator, apply strongly_measurable.measurable_set_le strongly_measurable_const hmeas.nnnorm.measurable.coe_nnreal_real.strongly_measurable }, - have hbound : has_finite_integral (λ x, ∥f x∥) μ, + have hbound : has_finite_integral (λ x, ‖f x‖) μ, { rw mem_ℒp_one_iff_integrable at hf, exact hf.norm.2 }, have := tendsto_lintegral_norm_of_dominated_convergence hmeas hbound _ htendsto, @@ -188,7 +188,7 @@ begin simp only [coe_nnnorm, ennreal.of_real_eq_coe_nnreal (norm_nonneg _)], refl }, { refine λ n, univ_mem' (id $ λ x, _), - by_cases hx : (n : ℝ) ≤ ∥f x∥, + by_cases hx : (n : ℝ) ≤ ‖f x‖, { dsimp, rwa indicator_of_mem }, { dsimp, @@ -201,12 +201,12 @@ end which does not require measurability. -/ lemma mem_ℒp.integral_indicator_norm_ge_nonneg_le_of_meas (hf : mem_ℒp f 1 μ) (hmeas : strongly_measurable f) {ε : ℝ} (hε : 0 < ε) : - ∃ M : ℝ, 0 ≤ M ∧ ∫⁻ x, ∥{x | M ≤ ∥f x∥₊}.indicator f x∥₊ ∂μ ≤ ennreal.of_real ε := + ∃ M : ℝ, 0 ≤ M ∧ ∫⁻ x, ‖{x | M ≤ ‖f x‖₊}.indicator f x‖₊ ∂μ ≤ ennreal.of_real ε := let ⟨M, hM⟩ := hf.integral_indicator_norm_ge_le μ hmeas hε in ⟨max M 0, le_max_right _ _, by simpa⟩ lemma mem_ℒp.integral_indicator_norm_ge_nonneg_le (hf : mem_ℒp f 1 μ) {ε : ℝ} (hε : 0 < ε) : - ∃ M : ℝ, 0 ≤ M ∧ ∫⁻ x, ∥{x | M ≤ ∥f x∥₊}.indicator f x∥₊ ∂μ ≤ ennreal.of_real ε := + ∃ M : ℝ, 0 ≤ M ∧ ∫⁻ x, ‖{x | M ≤ ‖f x‖₊}.indicator f x‖₊ ∂μ ≤ ennreal.of_real ε := begin have hf_mk : mem_ℒp (hf.1.mk f) 1 μ := (mem_ℒp_congr_ae hf.1.ae_eq_mk).mp hf, obtain ⟨M, hM_pos, hfM⟩ := hf_mk.integral_indicator_norm_ge_nonneg_le_of_meas μ @@ -219,15 +219,15 @@ end lemma mem_ℒp.snorm_ess_sup_indicator_norm_ge_eq_zero (hf : mem_ℒp f ∞ μ) (hmeas : strongly_measurable f) : - ∃ M : ℝ, snorm_ess_sup ({x | M ≤ ∥f x∥₊}.indicator f) μ = 0 := + ∃ M : ℝ, snorm_ess_sup ({x | M ≤ ‖f x‖₊}.indicator f) μ = 0 := begin have hbdd : snorm_ess_sup f μ < ∞ := hf.snorm_lt_top, refine ⟨(snorm f ∞ μ + 1).to_real, _⟩, rw snorm_ess_sup_indicator_eq_snorm_ess_sup_restrict, - have : μ.restrict {x : α | (snorm f ⊤ μ + 1).to_real ≤ ∥f x∥₊} = 0, + have : μ.restrict {x : α | (snorm f ⊤ μ + 1).to_real ≤ ‖f x‖₊} = 0, { simp only [coe_nnnorm, snorm_exponent_top, measure.restrict_eq_zero], - have : {x : α | (snorm_ess_sup f μ + 1).to_real ≤ ∥f x∥} ⊆ - {x : α | snorm_ess_sup f μ < ∥f x∥₊}, + have : {x : α | (snorm_ess_sup f μ + 1).to_real ≤ ‖f x‖} ⊆ + {x : α | snorm_ess_sup f μ < ‖f x‖₊}, { intros x hx, rw [mem_set_of_eq, ← ennreal.to_real_lt_to_real hbdd.ne ennreal.coe_lt_top.ne, ennreal.coe_to_real, coe_nnnorm], @@ -248,7 +248,7 @@ end latter provides `0 < M`. -/ lemma mem_ℒp.snorm_indicator_norm_ge_le (hf : mem_ℒp f p μ) (hmeas : strongly_measurable f) {ε : ℝ} (hε : 0 < ε) : - ∃ M : ℝ, snorm ({x | M ≤ ∥f x∥₊}.indicator f) p μ ≤ ennreal.of_real ε := + ∃ M : ℝ, snorm ({x | M ≤ ‖f x‖₊}.indicator f) p μ ≤ ennreal.of_real ε := begin by_cases hp_ne_zero : p = 0, { refine ⟨1, hp_ne_zero.symm ▸ _⟩, @@ -259,7 +259,7 @@ begin refine ⟨M, _⟩, simp only [snorm_exponent_top, hM, zero_le] }, obtain ⟨M, hM', hM⟩ := @mem_ℒp.integral_indicator_norm_ge_nonneg_le _ _ _ μ _ - (λ x, ∥f x∥^p.to_real) (hf.norm_rpow hp_ne_zero hp_ne_top) _ + (λ x, ‖f x‖^p.to_real) (hf.norm_rpow hp_ne_zero hp_ne_top) _ (real.rpow_pos_of_pos hε p.to_real), refine ⟨M ^(1 / p.to_real), _⟩, rw [snorm_eq_lintegral_rpow_nnnorm hp_ne_zero hp_ne_top, @@ -272,13 +272,13 @@ begin ext1 x, rw [ennreal.coe_rpow_of_nonneg _ ennreal.to_real_nonneg, nnnorm_indicator_eq_indicator_nnnorm, nnnorm_indicator_eq_indicator_nnnorm], - have hiff : M ^ (1 / p.to_real) ≤ ∥f x∥₊ ↔ M ≤ ∥∥f x∥ ^ p.to_real∥₊, + have hiff : M ^ (1 / p.to_real) ≤ ‖f x‖₊ ↔ M ≤ ‖‖f x‖ ^ p.to_real‖₊, { rw [coe_nnnorm, coe_nnnorm, real.norm_rpow_of_nonneg (norm_nonneg _), norm_norm, ← real.rpow_le_rpow_iff hM' (real.rpow_nonneg_of_nonneg (norm_nonneg _) _) (one_div_pos.2 $ ennreal.to_real_pos hp_ne_zero hp_ne_top), ← real.rpow_mul (norm_nonneg _), mul_one_div_cancel (ennreal.to_real_pos hp_ne_zero hp_ne_top).ne.symm, real.rpow_one] }, - by_cases hx : x ∈ {x : α | M ^ (1 / p.to_real) ≤ ∥f x∥₊}, + by_cases hx : x ∈ {x : α | M ^ (1 / p.to_real) ≤ ‖f x‖₊}, { rw [set.indicator_of_mem hx,set.indicator_of_mem, real.nnnorm_of_nonneg], refl, change _ ≤ _, rwa ← hiff }, @@ -291,7 +291,7 @@ end /-- This lemma implies that a single function is uniformly integrable (in the probability sense). -/ lemma mem_ℒp.snorm_indicator_norm_ge_pos_le (hf : mem_ℒp f p μ) (hmeas : strongly_measurable f) {ε : ℝ} (hε : 0 < ε) : - ∃ M : ℝ, 0 < M ∧ snorm ({x | M ≤ ∥f x∥₊}.indicator f) p μ ≤ ennreal.of_real ε := + ∃ M : ℝ, 0 < M ∧ snorm ({x | M ≤ ‖f x‖₊}.indicator f) p μ ≤ ennreal.of_real ε := begin obtain ⟨M, hM⟩ := hf.snorm_indicator_norm_ge_le μ hmeas hε, refine ⟨max M 1, lt_of_lt_of_le zero_lt_one (le_max_right _ _), @@ -305,7 +305,7 @@ end end lemma snorm_indicator_le_of_bound {f : α → β} (hp_top : p ≠ ∞) - {ε : ℝ} (hε : 0 < ε) {M : ℝ} (hf : ∀ x, ∥f x∥ < M) : + {ε : ℝ} (hε : 0 < ε) {M : ℝ} (hf : ∀ x, ‖f x‖ < M) : ∃ (δ : ℝ) (hδ : 0 < δ), ∀ s, measurable_set s → μ s ≤ ennreal.of_real δ → snorm (s.indicator f) p μ ≤ ennreal.of_real ε := begin @@ -321,7 +321,7 @@ begin by_cases hp : p = 0, { simp [hp] }, rw snorm_indicator_eq_snorm_restrict hs, - have haebdd : ∀ᵐ x ∂μ.restrict s, ∥f x∥ ≤ M, + have haebdd : ∀ᵐ x ∂μ.restrict s, ‖f x‖ ≤ M, { filter_upwards, exact (λ x, (hf x).le) }, refine le_trans (snorm_le_of_ae_bound haebdd) _, @@ -347,9 +347,9 @@ lemma mem_ℒp.snorm_indicator_le' (hp_one : 1 ≤ p) (hp_top : p ≠ ∞) begin obtain ⟨M, hMpos, hM⟩ := hf.snorm_indicator_norm_ge_pos_le μ hmeas hε, obtain ⟨δ, hδpos, hδ⟩ := @snorm_indicator_le_of_bound _ _ _ μ _ _ - ({x | ∥f x∥ < M}.indicator f) hp_top _ hε M _, + ({x | ‖f x‖ < M}.indicator f) hp_top _ hε M _, { refine ⟨δ, hδpos, λ s hs hμs, _⟩, - rw (_ : f = {x : α | M ≤ ∥f x∥₊}.indicator f + {x : α | ∥f x∥ < M}.indicator f), + rw (_ : f = {x : α | M ≤ ‖f x‖₊}.indicator f + {x : α | ‖f x‖ < M}.indicator f), { rw snorm_indicator_eq_snorm_restrict hs, refine le_trans (snorm_add_le _ _ hp_one) _, { exact strongly_measurable.ae_strongly_measurable (hmeas.indicator @@ -361,7 +361,7 @@ begin rw ← snorm_indicator_eq_snorm_restrict hs, exact hδ s hs hμs } }, { ext x, - by_cases hx : M ≤ ∥f x∥, + by_cases hx : M ≤ ‖f x‖, { rw [pi.add_apply, indicator_of_mem, indicator_of_not_mem, add_zero]; simpa }, { rw [pi.add_apply, indicator_of_not_mem, indicator_of_mem, zero_add]; @@ -475,7 +475,7 @@ lemma snorm_sub_le_of_dist_bdd begin by_cases hp : p = 0, { simp [hp], }, - have : ∀ x, ∥s.indicator (f - g) x∥ ≤ ∥s.indicator (λ x, c) x∥, + have : ∀ x, ‖s.indicator (f - g) x‖ ≤ ‖s.indicator (λ x, c) x‖, { intro x, by_cases hx : x ∈ s, { rw [indicator_of_mem hx, indicator_of_mem hx, pi.sub_apply, ← dist_eq_norm, @@ -646,7 +646,7 @@ lemma tendsto_in_measure_iff_tendsto_Lp [is_finite_measure μ] lemma unif_integrable_of' (hp : 1 ≤ p) (hp' : p ≠ ∞) {f : ι → α → β} (hf : ∀ i, strongly_measurable (f i)) (h : ∀ ε : ℝ, 0 < ε → ∃ C : ℝ≥0, 0 < C ∧ - ∀ i, snorm ({x | C ≤ ∥f i x∥₊}.indicator (f i)) p μ ≤ ennreal.of_real ε) : + ∀ i, snorm ({x | C ≤ ‖f i x‖₊}.indicator (f i)) p μ ≤ ennreal.of_real ε) : unif_integrable f p μ := begin have hpzero := (lt_of_lt_of_le ennreal.zero_lt_one hp).ne.symm, @@ -661,32 +661,32 @@ begin { rw (snorm_eq_zero_iff ((hf i).indicator hs).ae_strongly_measurable hpzero).2 (indicator_meas_zero hμs'), norm_num }, - calc snorm (indicator s (f i)) p μ ≤ snorm (indicator (s ∩ {x | C ≤ ∥f i x∥₊}) (f i)) p μ + - snorm (indicator (s ∩ {x | ∥f i x∥₊ < C}) (f i)) p μ : + calc snorm (indicator s (f i)) p μ ≤ snorm (indicator (s ∩ {x | C ≤ ‖f i x‖₊}) (f i)) p μ + + snorm (indicator (s ∩ {x | ‖f i x‖₊ < C}) (f i)) p μ : begin refine le_trans (eq.le _) (snorm_add_le (strongly_measurable.ae_strongly_measurable ((hf i).indicator (hs.inter (strongly_measurable_const.measurable_set_le (hf i).nnnorm)))) (strongly_measurable.ae_strongly_measurable ((hf i).indicator (hs.inter ((hf i).nnnorm.measurable_set_lt strongly_measurable_const)))) hp), congr, - change _ = λ x, (s ∩ {x : α | C ≤ ∥f i x∥₊}).indicator (f i) x + - (s ∩ {x : α | ∥f i x∥₊ < C}).indicator (f i) x, + change _ = λ x, (s ∩ {x : α | C ≤ ‖f i x‖₊}).indicator (f i) x + + (s ∩ {x : α | ‖f i x‖₊ < C}).indicator (f i) x, rw ← set.indicator_union_of_disjoint, { congr, rw [← inter_union_distrib_left, (by { ext, simp [le_or_lt] } : - {x : α | C ≤ ∥f i x∥₊} ∪ {x : α | ∥f i x∥₊ < C} = set.univ), inter_univ] }, + {x : α | C ≤ ‖f i x‖₊} ∪ {x : α | ‖f i x‖₊ < C} = set.univ), inter_univ] }, { refine (disjoint.inf_right' _ _).inf_left' _, rw disjoint_iff_inf_le, rintro x ⟨hx₁ : _ ≤ _, hx₂ : _ < _⟩, exact false.elim (hx₂.ne (eq_of_le_of_not_lt hx₁ (not_lt.2 hx₂.le)).symm) } end - ... ≤ snorm (indicator ({x | C ≤ ∥f i x∥₊}) (f i)) p μ + C * μ s ^ (1 / ennreal.to_real p) : + ... ≤ snorm (indicator ({x | C ≤ ‖f i x‖₊}) (f i)) p μ + C * μ s ^ (1 / ennreal.to_real p) : begin refine add_le_add (snorm_mono $ λ x, norm_indicator_le_of_subset (inter_subset_right _ _) _ _) _, rw ← indicator_indicator, rw snorm_indicator_eq_snorm_restrict, - have : ∀ᵐ x ∂(μ.restrict s), ∥({x : α | ∥f i x∥₊ < C}).indicator (f i) x∥ ≤ C, + have : ∀ᵐ x ∂(μ.restrict s), ‖({x : α | ‖f i x‖₊ < C}).indicator (f i) x‖ ≤ C, { refine ae_of_all _ _, simp_rw norm_indicator_eq_indicator_norm, exact indicator_le' (λ x (hx : _ < _), hx.le) (λ _ _, nnreal.coe_nonneg _) }, @@ -718,18 +718,18 @@ end lemma unif_integrable_of (hp : 1 ≤ p) (hp' : p ≠ ∞) {f : ι → α → β} (hf : ∀ i, ae_strongly_measurable (f i) μ) (h : ∀ ε : ℝ, 0 < ε → ∃ C : ℝ≥0, - ∀ i, snorm ({x | C ≤ ∥f i x∥₊}.indicator (f i)) p μ ≤ ennreal.of_real ε) : + ∀ i, snorm ({x | C ≤ ‖f i x‖₊}.indicator (f i)) p μ ≤ ennreal.of_real ε) : unif_integrable f p μ := begin set g : ι → α → β := λ i, (hf i).some, refine (unif_integrable_of' μ hp hp' (λ i, (Exists.some_spec $hf i).1) (λ ε hε, _)).ae_eq (λ i, (Exists.some_spec $ hf i).2.symm), obtain ⟨C, hC⟩ := h ε hε, - have hCg : ∀ i, snorm ({x | C ≤ ∥g i x∥₊}.indicator (g i)) p μ ≤ ennreal.of_real ε, + have hCg : ∀ i, snorm ({x | C ≤ ‖g i x‖₊}.indicator (g i)) p μ ≤ ennreal.of_real ε, { intro i, refine le_trans (le_of_eq $ snorm_congr_ae _) (hC i), filter_upwards [(Exists.some_spec $ hf i).2] with x hx, - by_cases hfx : x ∈ {x | C ≤ ∥f i x∥₊}, + by_cases hfx : x ∈ {x | C ≤ ‖f i x‖₊}, { rw [indicator_of_mem hfx, indicator_of_mem, hx], rwa [mem_set_of, hx] at hfx }, { rw [indicator_of_not_mem hfx, indicator_of_not_mem], @@ -815,15 +815,15 @@ lemma uniform_integrable_const {g : α → β} (hp : 1 ≤ p) (hp_ne_top : p ≠ lemma uniform_integrable_of' [is_finite_measure μ] (hp : 1 ≤ p) (hp' : p ≠ ∞) (hf : ∀ i, strongly_measurable (f i)) (h : ∀ ε : ℝ, 0 < ε → ∃ C : ℝ≥0, - ∀ i, snorm ({x | C ≤ ∥f i x∥₊}.indicator (f i)) p μ ≤ ennreal.of_real ε) : + ∀ i, snorm ({x | C ≤ ‖f i x‖₊}.indicator (f i)) p μ ≤ ennreal.of_real ε) : uniform_integrable f p μ := begin refine ⟨λ i, (hf i).ae_strongly_measurable, unif_integrable_of μ hp hp' (λ i, (hf i).ae_strongly_measurable) h, _⟩, obtain ⟨C, hC⟩ := h 1 one_pos, refine ⟨(C * (μ univ ^ (p.to_real⁻¹)) + 1 : ℝ≥0∞).to_nnreal, λ i, _⟩, - calc snorm (f i) p μ ≤ snorm ({x : α | ∥f i x∥₊ < C}.indicator (f i)) p μ + - snorm ({x : α | C ≤ ∥f i x∥₊}.indicator (f i)) p μ : + calc snorm (f i) p μ ≤ snorm ({x : α | ‖f i x‖₊ < C}.indicator (f i)) p μ + + snorm ({x : α | C ≤ ‖f i x‖₊}.indicator (f i)) p μ : begin refine le_trans (snorm_mono (λ x, _)) (snorm_add_le (strongly_measurable.ae_strongly_measurable ((hf i).indicator @@ -839,7 +839,7 @@ begin end ... ≤ C * μ univ ^ (p.to_real⁻¹) + 1 : begin - have : ∀ᵐ x ∂μ, ∥{x : α | ∥f i x∥₊ < C}.indicator (f i) x∥₊ ≤ C, + have : ∀ᵐ x ∂μ, ‖{x : α | ‖f i x‖₊ < C}.indicator (f i) x‖₊ ≤ C, { refine eventually_of_forall _, simp_rw nnnorm_indicator_eq_indicator_nnnorm, exact indicator_le (λ x (hx : _ < _), hx.le) }, @@ -861,7 +861,7 @@ end lemma uniform_integrable_of [is_finite_measure μ] (hp : 1 ≤ p) (hp' : p ≠ ∞) (hf : ∀ i, ae_strongly_measurable (f i) μ) (h : ∀ ε : ℝ, 0 < ε → ∃ C : ℝ≥0, - ∀ i, snorm ({x | C ≤ ∥f i x∥₊}.indicator (f i)) p μ ≤ ennreal.of_real ε) : + ∀ i, snorm ({x | C ≤ ‖f i x‖₊}.indicator (f i)) p μ ≤ ennreal.of_real ε) : uniform_integrable f p μ := begin set g : ι → α → β := λ i, (hf i).some, @@ -871,7 +871,7 @@ begin obtain ⟨C, hC⟩ := h ε hε, refine ⟨C, λ i, le_trans (le_of_eq $ snorm_congr_ae _) (hC i)⟩, filter_upwards [(Exists.some_spec $ hf i).2] with x hx, - by_cases hfx : x ∈ {x | C ≤ ∥f i x∥₊}, + by_cases hfx : x ∈ {x | C ≤ ‖f i x‖₊}, { rw [indicator_of_mem hfx, indicator_of_mem, hx], rwa [mem_set_of, hx] at hfx }, { rw [indicator_of_not_mem hfx, indicator_of_not_mem], @@ -882,24 +882,24 @@ end lemma uniform_integrable.spec' (hp : p ≠ 0) (hp' : p ≠ ∞) (hf : ∀ i, strongly_measurable (f i)) (hfu : uniform_integrable f p μ) {ε : ℝ} (hε : 0 < ε) : - ∃ C : ℝ≥0, ∀ i, snorm ({x | C ≤ ∥f i x∥₊}.indicator (f i)) p μ ≤ ennreal.of_real ε := + ∃ C : ℝ≥0, ∀ i, snorm ({x | C ≤ ‖f i x‖₊}.indicator (f i)) p μ ≤ ennreal.of_real ε := begin obtain ⟨-, hfu, M, hM⟩ := hfu, obtain ⟨δ, hδpos, hδ⟩ := hfu hε, - obtain ⟨C, hC⟩ : ∃ C : ℝ≥0, ∀ i, μ {x | C ≤ ∥f i x∥₊} ≤ ennreal.of_real δ, + obtain ⟨C, hC⟩ : ∃ C : ℝ≥0, ∀ i, μ {x | C ≤ ‖f i x‖₊} ≤ ennreal.of_real δ, { by_contra hcon, push_neg at hcon, choose ℐ hℐ using hcon, lift δ to ℝ≥0 using hδpos.le, have : ∀ C : ℝ≥0, C • (δ : ℝ≥0∞) ^ (1 / p.to_real) ≤ snorm (f (ℐ C)) p μ, { intros C, - calc C • (δ : ℝ≥0∞) ^ (1 / p.to_real) ≤ C • μ {x | C ≤ ∥f (ℐ C) x∥₊} ^ (1 / p.to_real): + calc C • (δ : ℝ≥0∞) ^ (1 / p.to_real) ≤ C • μ {x | C ≤ ‖f (ℐ C) x‖₊} ^ (1 / p.to_real): begin rw [ennreal.smul_def, ennreal.smul_def, smul_eq_mul, smul_eq_mul], simp_rw ennreal.of_real_coe_nnreal at hℐ, refine ennreal.mul_le_mul le_rfl (ennreal.rpow_le_rpow (hℐ C).le (one_div_nonneg.2 ennreal.to_real_nonneg)), end - ... ≤ snorm ({x | C ≤ ∥f (ℐ C) x∥₊}.indicator (f (ℐ C))) p μ : + ... ≤ snorm ({x | C ≤ ‖f (ℐ C) x‖₊}.indicator (f (ℐ C))) p μ : begin refine snorm_indicator_ge_of_bdd_below hp hp' _ (measurable_set_le measurable_const (hf _).nnnorm.measurable) @@ -921,7 +921,7 @@ end lemma uniform_integrable.spec (hp : p ≠ 0) (hp' : p ≠ ∞) (hfu : uniform_integrable f p μ) {ε : ℝ} (hε : 0 < ε) : - ∃ C : ℝ≥0, ∀ i, snorm ({x | C ≤ ∥f i x∥₊}.indicator (f i)) p μ ≤ ennreal.of_real ε := + ∃ C : ℝ≥0, ∀ i, snorm ({x | C ≤ ‖f i x‖₊}.indicator (f i)) p μ ≤ ennreal.of_real ε := begin set g : ι → α → β := λ i, (hfu.1 i).some, have hgmeas : ∀ i, strongly_measurable (g i) := λ i, (Exists.some_spec $ hfu.1 i).1, @@ -929,7 +929,7 @@ begin obtain ⟨C, hC⟩ := hgunif.spec' hp hp' hgmeas hε, refine ⟨C, λ i, le_trans (le_of_eq $ snorm_congr_ae _) (hC i)⟩, filter_upwards [(Exists.some_spec $ hfu.1 i).2] with x hx, - by_cases hfx : x ∈ {x | C ≤ ∥f i x∥₊}, + by_cases hfx : x ∈ {x | C ≤ ‖f i x‖₊}, { rw [indicator_of_mem hfx, indicator_of_mem, hx], rwa [mem_set_of, hx] at hfx }, { rw [indicator_of_not_mem hfx, indicator_of_not_mem], @@ -941,7 +941,7 @@ found in literature. -/ lemma uniform_integrable_iff [is_finite_measure μ] (hp : 1 ≤ p) (hp' : p ≠ ∞) : uniform_integrable f p μ ↔ (∀ i, ae_strongly_measurable (f i) μ) ∧ ∀ ε : ℝ, 0 < ε → ∃ C : ℝ≥0, - ∀ i, snorm ({x | C ≤ ∥f i x∥₊}.indicator (f i)) p μ ≤ ennreal.of_real ε := + ∀ i, snorm ({x | C ≤ ‖f i x‖₊}.indicator (f i)) p μ ≤ ennreal.of_real ε := ⟨λ h, ⟨h.1, λ ε, h.spec (lt_of_lt_of_le ennreal.zero_lt_one hp).ne.symm hp'⟩, λ h, uniform_integrable_of hp hp' h.1 h.2⟩ diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index 66a4d15b5934e..30fb799a56aac 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -51,7 +51,7 @@ file `set_to_L1`). * `integral_sub` : `∫ x, f x - g x ∂μ = ∫ x, f x ∂μ - ∫ x, g x ∂μ` * `integral_smul` : `∫ x, r • f x ∂μ = r • ∫ x, f x ∂μ` * `integral_congr_ae` : `f =ᵐ[μ] g → ∫ x, f x ∂μ = ∫ x, g x ∂μ` - * `norm_integral_le_integral_norm` : `∥∫ x, f x ∂μ∥ ≤ ∫ x, ∥f x∥ ∂μ` + * `norm_integral_le_integral_norm` : `‖∫ x, f x ∂μ‖ ≤ ∫ x, ‖f x‖ ∂μ` 2. Basic properties of the Bochner integral on functions of type `α → ℝ`, where `α` is a measure space. @@ -99,7 +99,7 @@ functions : 1. First go to the `L¹` space. - For example, if you see `ennreal.to_real (∫⁻ a, ennreal.of_real $ ∥f a∥)`, that is the norm of + For example, if you see `ennreal.to_real (∫⁻ a, ennreal.of_real $ ‖f a‖)`, that is the norm of `f` in `L¹` space. Rewrite using `L1.norm_of_fun_eq_lintegral_norm`. 2. Show that the set `{f ∈ L¹ | ∫ f = ∫⁻ f⁺ - ∫⁻ f⁻}` is closed in `L¹` using `is_closed_eq`. @@ -219,10 +219,10 @@ lemma weighted_smul_smul [normed_field 𝕜] [normed_space 𝕜 F] [smul_comm_cl weighted_smul μ s (c • x) = c • weighted_smul μ s x := by { simp_rw [weighted_smul_apply, smul_comm], } -lemma norm_weighted_smul_le (s : set α) : ∥(weighted_smul μ s : F →L[ℝ] F)∥ ≤ (μ s).to_real := -calc ∥(weighted_smul μ s : F →L[ℝ] F)∥ = ∥(μ s).to_real∥ * ∥continuous_linear_map.id ℝ F∥ : +lemma norm_weighted_smul_le (s : set α) : ‖(weighted_smul μ s : F →L[ℝ] F)‖ ≤ (μ s).to_real := +calc ‖(weighted_smul μ s : F →L[ℝ] F)‖ = ‖(μ s).to_real‖ * ‖continuous_linear_map.id ℝ F‖ : norm_smul _ _ -... ≤ ∥(μ s).to_real∥ : (mul_le_mul_of_nonneg_left norm_id_le (norm_nonneg _)).trans (mul_one _).le +... ≤ ‖(μ s).to_real‖ : (mul_le_mul_of_nonneg_left norm_id_le (norm_nonneg _)).trans (mul_one _).le ... = abs (μ s).to_real : real.norm_eq_abs _ ... = (μ s).to_real : abs_eq_self.mpr ennreal.to_real_nonneg @@ -389,16 +389,16 @@ lemma integral_smul (c : 𝕜) {f : α →ₛ E} (hf : integrable f μ) : set_to_simple_func_smul _ weighted_smul_union weighted_smul_smul c hf lemma norm_set_to_simple_func_le_integral_norm (T : set α → E →L[ℝ] F) {C : ℝ} - (hT_norm : ∀ s, measurable_set s → μ s < ∞ → ∥T s∥ ≤ C * (μ s).to_real) {f : α →ₛ E} + (hT_norm : ∀ s, measurable_set s → μ s < ∞ → ‖T s‖ ≤ C * (μ s).to_real) {f : α →ₛ E} (hf : integrable f μ) : - ∥f.set_to_simple_func T∥ ≤ C * (f.map norm).integral μ := -calc ∥f.set_to_simple_func T∥ - ≤ C * ∑ x in f.range, ennreal.to_real (μ (f ⁻¹' {x})) * ∥x∥ : + ‖f.set_to_simple_func T‖ ≤ C * (f.map norm).integral μ := +calc ‖f.set_to_simple_func T‖ + ≤ C * ∑ x in f.range, ennreal.to_real (μ (f ⁻¹' {x})) * ‖x‖ : norm_set_to_simple_func_le_sum_mul_norm_of_integrable T hT_norm f hf ... = C * (f.map norm).integral μ : by { rw map_integral f norm hf norm_zero, simp_rw smul_eq_mul, } lemma norm_integral_le_integral_norm (f : α →ₛ E) (hf : integrable f μ) : - ∥f.integral μ∥ ≤ (f.map norm).integral μ := + ‖f.integral μ‖ ≤ (f.map norm).integral μ := begin refine (norm_set_to_simple_func_le_integral_norm _ (λ s _ _, _) hf).trans (one_mul _).le, exact (norm_weighted_smul_le s).trans (one_mul _).symm.le, @@ -429,7 +429,7 @@ variables {α E μ} namespace simple_func -lemma norm_eq_integral (f : α →₁ₛ[μ] E) : ∥f∥ = ((to_simple_func f).map norm).integral μ := +lemma norm_eq_integral (f : α →₁ₛ[μ] E) : ‖f‖ = ((to_simple_func f).map norm).integral μ := begin rw [norm_eq_sum_mul f, (to_simple_func f).map_integral norm (simple_func.integrable f) norm_zero], simp_rw smul_eq_mul, @@ -491,7 +491,7 @@ lemma integral_smul (c : 𝕜) (f : α →₁ₛ[μ] E) : integral (c • f) = c • integral f := set_to_L1s_smul _ (λ _ _, weighted_smul_null) weighted_smul_union weighted_smul_smul c f -lemma norm_integral_le_norm (f : α →₁ₛ[μ] E) : ∥integral f∥ ≤ ∥f∥ := +lemma norm_integral_le_norm (f : α →₁ₛ[μ] E) : ‖integral f‖ ≤ ‖f‖ := begin rw [integral, norm_eq_integral], exact (to_simple_func f).norm_integral_le_integral_norm (simple_func.integrable f) @@ -515,7 +515,7 @@ local notation (name := simple_func.integral_clm) `Integral` := integral_clm α open continuous_linear_map -lemma norm_Integral_le_one : ∥Integral∥ ≤ 1 := +lemma norm_Integral_le_one : ‖Integral‖ ≤ 1 := linear_map.mk_continuous_norm_le _ (zero_le_one) _ section pos_part @@ -545,7 +545,7 @@ begin end lemma integral_eq_norm_pos_part_sub (f : α →₁ₛ[μ] ℝ) : - integral f = ∥pos_part f∥ - ∥neg_part f∥ := + integral f = ‖pos_part f‖ - ‖neg_part f‖ := begin -- Convert things in `L¹` to their `simple_func` counterpart have ae_eq₁ : (to_simple_func f).pos_part =ᵐ[μ] (to_simple_func (pos_part f)).map norm, @@ -641,14 +641,14 @@ local notation (name := integral_clm) `Integral` := @integral_clm α E _ _ μ _ local notation (name := simple_func.integral_clm') `sIntegral` := @simple_func.integral_clm α E _ _ μ _ -lemma norm_Integral_le_one : ∥Integral∥ ≤ 1 := +lemma norm_Integral_le_one : ‖Integral‖ ≤ 1 := norm_set_to_L1_le (dominated_fin_meas_additive_weighted_smul μ) zero_le_one -lemma norm_integral_le (f : α →₁[μ] E) : ∥integral f∥ ≤ ∥f∥ := -calc ∥integral f∥ = ∥Integral f∥ : rfl - ... ≤ ∥Integral∥ * ∥f∥ : le_op_norm _ _ - ... ≤ 1 * ∥f∥ : mul_le_mul_of_nonneg_right norm_Integral_le_one $ norm_nonneg _ - ... = ∥f∥ : one_mul _ +lemma norm_integral_le (f : α →₁[μ] E) : ‖integral f‖ ≤ ‖f‖ := +calc ‖integral f‖ = ‖Integral f‖ : rfl + ... ≤ ‖Integral‖ * ‖f‖ : le_op_norm _ _ + ... ≤ 1 * ‖f‖ : mul_le_mul_of_nonneg_right norm_Integral_le_one $ norm_nonneg _ + ... = ‖f‖ : one_mul _ @[continuity] lemma continuous_integral : continuous (λ (f : α →₁[μ] E), integral f) := @@ -657,11 +657,11 @@ L1.integral_clm.continuous section pos_part lemma integral_eq_norm_pos_part_sub (f : α →₁[μ] ℝ) : - integral f = ∥Lp.pos_part f∥ - ∥Lp.neg_part f∥ := + integral f = ‖Lp.pos_part f‖ - ‖Lp.neg_part f‖ := begin -- Use `is_closed_property` and `is_closed_eq` refine @is_closed_property _ _ _ (coe : (α →₁ₛ[μ] ℝ) → (α →₁[μ] ℝ)) - (λ f : α →₁[μ] ℝ, integral f = ∥Lp.pos_part f∥ - ∥Lp.neg_part f∥) + (λ f : α →₁[μ] ℝ, integral f = ‖Lp.pos_part f‖ - ‖Lp.neg_part f‖) (simple_func.dense_range one_ne_top) (is_closed_eq _ _) _ f, { exact cont _ }, { refine continuous.sub (continuous_norm.comp Lp.continuous_pos_part) @@ -699,7 +699,7 @@ if hf : integrable f μ then L1.integral (hf.to_L1 f) else 0 end -/-! In the notation for integrals, an expression like `∫ x, g ∥x∥ ∂μ` will not be parsed correctly, +/-! In the notation for integrals, an expression like `∫ x, g ‖x‖ ∂μ` will not be parsed correctly, and needs parentheses. We do not set the binding power of `r` to `0`, because then `∫ x, f x = 0` will be parsed incorrectly. -/ notation `∫` binders `, ` r:(scoped:60 f, f) ` ∂` μ:70 := integral μ r @@ -791,7 +791,7 @@ lemma continuous_integral : continuous (λ (f : α →₁[μ] E), ∫ a, f a ∂ continuous_set_to_fun (dominated_fin_meas_additive_weighted_smul μ) lemma norm_integral_le_lintegral_norm (f : α → E) : - ∥∫ a, f a ∂μ∥ ≤ ennreal.to_real (∫⁻ a, (ennreal.of_real ∥f a∥) ∂μ) := + ‖∫ a, f a ∂μ‖ ≤ ennreal.to_real (∫⁻ a, (ennreal.of_real ‖f a‖) ∂μ) := begin by_cases hf : integrable f μ, { rw [integral_eq f hf, ← integrable.norm_to_L1_eq_lintegral_norm f hf], @@ -800,7 +800,7 @@ begin end lemma ennnorm_integral_le_lintegral_ennnorm (f : α → E) : - (∥∫ a, f a ∂μ∥₊ : ℝ≥0∞) ≤ ∫⁻ a, ∥f a∥₊ ∂μ := + (‖∫ a, f a ∂μ‖₊ : ℝ≥0∞) ≤ ∫⁻ a, ‖f a‖₊ ∂μ := by { simp_rw [← of_real_norm_eq_coe_nnnorm], apply ennreal.of_real_le_of_le_to_real, exact norm_integral_le_lintegral_norm f } @@ -831,7 +831,7 @@ hf.2.tendsto_set_integral_nhds_zero hs /-- If `F i → f` in `L1`, then `∫ x, F i x ∂μ → ∫ x, f x ∂μ`. -/ lemma tendsto_integral_of_L1 {ι} (f : α → E) (hfi : integrable f μ) {F : ι → α → E} {l : filter ι} (hFi : ∀ᶠ i in l, integrable (F i) μ) - (hF : tendsto (λ i, ∫⁻ x, ∥F i x - f x∥₊ ∂μ) l (𝓝 0)) : + (hF : tendsto (λ i, ∫⁻ x, ‖F i x - f x‖₊ ∂μ) l (𝓝 0)) : tendsto (λ i, ∫ x, F i x ∂μ) l (𝓝 $ ∫ x, f x ∂μ) := tendsto_set_to_fun_of_L1 (dominated_fin_meas_additive_weighted_smul μ) f hfi hFi hF @@ -843,7 +843,7 @@ tendsto_set_to_fun_of_L1 (dominated_fin_meas_additive_weighted_smul μ) f hfi hF theorem tendsto_integral_of_dominated_convergence {F : ℕ → α → E} {f : α → E} (bound : α → ℝ) (F_measurable : ∀ n, ae_strongly_measurable (F n) μ) (bound_integrable : integrable bound μ) - (h_bound : ∀ n, ∀ᵐ a ∂μ, ∥F n a∥ ≤ bound a) + (h_bound : ∀ n, ∀ᵐ a ∂μ, ‖F n a‖ ≤ bound a) (h_lim : ∀ᵐ a ∂μ, tendsto (λ n, F n a) at_top (𝓝 (f a))) : tendsto (λn, ∫ a, F n a ∂μ) at_top (𝓝 $ ∫ a, f a ∂μ) := tendsto_set_to_fun_of_dominated_convergence (dominated_fin_meas_additive_weighted_smul μ) bound @@ -854,7 +854,7 @@ lemma tendsto_integral_filter_of_dominated_convergence {ι} {l : filter ι} [l.is_countably_generated] {F : ι → α → E} {f : α → E} (bound : α → ℝ) (hF_meas : ∀ᶠ n in l, ae_strongly_measurable (F n) μ) - (h_bound : ∀ᶠ n in l, ∀ᵐ a ∂μ, ∥F n a∥ ≤ bound a) + (h_bound : ∀ᶠ n in l, ∀ᵐ a ∂μ, ‖F n a‖ ≤ bound a) (bound_integrable : integrable bound μ) (h_lim : ∀ᵐ a ∂μ, tendsto (λ n, F n a) l (𝓝 (f a))) : tendsto (λn, ∫ a, F n a ∂μ) l (𝓝 $ ∫ a, f a ∂μ) := @@ -865,7 +865,7 @@ tendsto_set_to_fun_filter_of_dominated_convergence (dominated_fin_meas_additive_ lemma has_sum_integral_of_dominated_convergence {ι} [countable ι] {F : ι → α → E} {f : α → E} (bound : ι → α → ℝ) (hF_meas : ∀ n, ae_strongly_measurable (F n) μ) - (h_bound : ∀ n, ∀ᵐ a ∂μ, ∥F n a∥ ≤ bound n a) + (h_bound : ∀ n, ∀ᵐ a ∂μ, ‖F n a‖ ≤ bound n a) (bound_summable : ∀ᵐ a ∂μ, summable (λ n, bound n a)) (bound_integrable : integrable (λ a, ∑' n, bound n a) μ) (h_lim : ∀ᵐ a ∂μ, has_sum (λ n, F n a) (f a)) : @@ -887,7 +887,7 @@ begin { refine eventually_of_forall (λ s, _), filter_upwards [eventually_countable_forall.2 h_bound, hb_nonneg, bound_summable] with a hFa ha0 has, - calc ∥∑ n in s, F n a∥ ≤ ∑ n in s, bound n a : norm_sum_le_of_le _ (λ n hn, hFa n) + calc ‖∑ n in s, F n a‖ ≤ ∑ n in s, bound n a : norm_sum_le_of_le _ (λ n hn, hFa n) ... ≤ ∑' n, bound n a : sum_le_tsum _ (λ n hn, ha0 n) has }, end @@ -895,14 +895,14 @@ variables {X : Type*} [topological_space X] [first_countable_topology X] lemma continuous_at_of_dominated {F : X → α → E} {x₀ : X} {bound : α → ℝ} (hF_meas : ∀ᶠ x in 𝓝 x₀, ae_strongly_measurable (F x) μ) - (h_bound : ∀ᶠ x in 𝓝 x₀, ∀ᵐ a ∂μ, ∥F x a∥ ≤ bound a) + (h_bound : ∀ᶠ x in 𝓝 x₀, ∀ᵐ a ∂μ, ‖F x a‖ ≤ bound a) (bound_integrable : integrable bound μ) (h_cont : ∀ᵐ a ∂μ, continuous_at (λ x, F x a) x₀) : continuous_at (λ x, ∫ a, F x a ∂μ) x₀ := continuous_at_set_to_fun_of_dominated (dominated_fin_meas_additive_weighted_smul μ) hF_meas h_bound bound_integrable h_cont lemma continuous_of_dominated {F : X → α → E} {bound : α → ℝ} - (hF_meas : ∀ x, ae_strongly_measurable (F x) μ) (h_bound : ∀ x, ∀ᵐ a ∂μ, ∥F x a∥ ≤ bound a) + (hF_meas : ∀ x, ae_strongly_measurable (F x) μ) (h_bound : ∀ x, ∀ᵐ a ∂μ, ‖F x a‖ ≤ bound a) (bound_integrable : integrable bound μ) (h_cont : ∀ᵐ a ∂μ, continuous (λ x, F x a)) : continuous (λ x, ∫ a, F x a ∂μ) := continuous_set_to_fun_of_dominated (dominated_fin_meas_additive_weighted_smul μ) hF_meas h_bound @@ -916,7 +916,7 @@ lemma integral_eq_lintegral_pos_part_sub_lintegral_neg_part {f : α → ℝ} (hf ennreal.to_real (∫⁻ a, (ennreal.of_real $ - f a) ∂μ) := let f₁ := hf.to_L1 f in -- Go to the `L¹` space -have eq₁ : ennreal.to_real (∫⁻ a, (ennreal.of_real $ f a) ∂μ) = ∥Lp.pos_part f₁∥ := +have eq₁ : ennreal.to_real (∫⁻ a, (ennreal.of_real $ f a) ∂μ) = ‖Lp.pos_part f₁‖ := begin rw L1.norm_def, congr' 1, @@ -929,7 +929,7 @@ begin simp only [real.coe_to_nnreal', subtype.coe_mk], end, -- Go to the `L¹` space -have eq₂ : ennreal.to_real (∫⁻ a, (ennreal.of_real $ - f a) ∂μ) = ∥Lp.neg_part f₁∥ := +have eq₂ : ennreal.to_real (∫⁻ a, (ennreal.of_real $ - f a) ∂μ) = ‖Lp.neg_part f₁‖ := begin rw L1.norm_def, congr' 1, @@ -963,7 +963,7 @@ begin { rw integral_undef hfi, simp_rw [integrable, hfm, has_finite_integral_iff_norm, lt_top_iff_ne_top, ne.def, true_and, not_not] at hfi, - have : ∫⁻ (a : α), ennreal.of_real (f a) ∂μ = ∫⁻ a, (ennreal.of_real ∥f a∥) ∂μ, + have : ∫⁻ (a : α), ennreal.of_real (f a) ∂μ = ∫⁻ a, (ennreal.of_real ‖f a‖) ∂μ, { refine lintegral_congr_ae (hf.mono $ assume a h, _), rw [real.norm_eq_abs, abs_of_nonneg h] }, rw [this, hfi], refl } @@ -971,7 +971,7 @@ end lemma integral_norm_eq_lintegral_nnnorm {G} [normed_add_comm_group G] {f : α → G} (hf : ae_strongly_measurable f μ) : - ∫ x, ∥f x∥ ∂μ = ennreal.to_real ∫⁻ x, ∥f x∥₊ ∂μ := + ∫ x, ‖f x‖ ∂μ = ennreal.to_real ∫⁻ x, ‖f x‖₊ ∂μ := begin rw integral_eq_lintegral_of_nonneg_ae _ hf.norm, { simp_rw [of_real_norm_eq_coe_nnnorm], }, @@ -980,7 +980,7 @@ end lemma of_real_integral_norm_eq_lintegral_nnnorm {G} [normed_add_comm_group G] {f : α → G} (hf : integrable f μ) : - ennreal.of_real ∫ x, ∥f x∥ ∂μ = ∫⁻ x, ∥f x∥₊ ∂μ := + ennreal.of_real ∫ x, ‖f x‖ ∂μ = ∫⁻ x, ‖f x‖₊ ∂μ := by rw [integral_norm_eq_lintegral_nnnorm hf.ae_strongly_measurable, ennreal.of_real_to_real (lt_top_iff_ne_top.mp hf.2)] @@ -1007,7 +1007,7 @@ end lemma of_real_integral_eq_lintegral_of_real {f : α → ℝ} (hfi : integrable f μ) (f_nn : 0 ≤ᵐ[μ] f) : ennreal.of_real (∫ x, f x ∂μ) = ∫⁻ x, ennreal.of_real (f x) ∂μ := begin - simp_rw [integral_congr_ae (show f =ᵐ[μ] λ x, ∥f x∥, + simp_rw [integral_congr_ae (show f =ᵐ[μ] λ x, ‖f x‖, by { filter_upwards [f_nn] with x hx, rw [real.norm_eq_abs, abs_eq_self.mpr hx], }), of_real_integral_norm_eq_lintegral_nnnorm hfi, ←of_real_norm_eq_coe_nnnorm], @@ -1077,7 +1077,7 @@ integral_pos_iff_support_of_nonneg_ae (eventually_of_forall hf) hfi section normed_add_comm_group variables {H : Type*} [normed_add_comm_group H] -lemma L1.norm_eq_integral_norm (f : α →₁[μ] H) : ∥f∥ = ∫ a, ∥f a∥ ∂μ := +lemma L1.norm_eq_integral_norm (f : α →₁[μ] H) : ‖f‖ = ∫ a, ‖f a‖ ∂μ := begin simp only [snorm, snorm', ennreal.one_to_real, ennreal.rpow_one, Lp.norm_def, if_false, ennreal.one_ne_top, one_ne_zero, _root_.div_one], @@ -1087,7 +1087,7 @@ begin end lemma L1.norm_of_fun_eq_integral_norm {f : α → H} (hf : integrable f μ) : - ∥hf.to_L1 f∥ = ∫ a, ∥f a∥ ∂μ := + ‖hf.to_L1 f‖ = ∫ a, ‖f a‖ ∂μ := begin rw L1.norm_eq_integral_norm, refine integral_congr_ae _, @@ -1098,9 +1098,9 @@ end lemma mem_ℒp.snorm_eq_integral_rpow_norm {f : α → H} {p : ℝ≥0∞} (hp1 : p ≠ 0) (hp2 : p ≠ ∞) (hf : mem_ℒp f p μ) : - snorm f p μ = ennreal.of_real ((∫ a, ∥f a∥ ^ p.to_real ∂μ) ^ (p.to_real ⁻¹)) := + snorm f p μ = ennreal.of_real ((∫ a, ‖f a‖ ^ p.to_real ∂μ) ^ (p.to_real ⁻¹)) := begin - have A : ∫⁻ (a : α), ennreal.of_real (∥f a∥ ^ p.to_real) ∂μ = ∫⁻ (a : α), ∥f a∥₊ ^ p.to_real ∂μ, + have A : ∫⁻ (a : α), ennreal.of_real (‖f a‖ ^ p.to_real) ∂μ = ∫⁻ (a : α), ‖f a‖₊ ^ p.to_real ∂μ, { apply lintegral_congr (λ x, _), rw [← of_real_rpow_of_nonneg (norm_nonneg _) to_real_nonneg, of_real_norm_eq_coe_nnnorm] }, simp only [snorm_eq_lintegral_rpow_nnnorm hp1 hp2, one_div], @@ -1144,13 +1144,13 @@ begin ((has_finite_integral_iff_of_real hf).1 hfi.2).ne] end -lemma norm_integral_le_integral_norm (f : α → E) : ∥(∫ a, f a ∂μ)∥ ≤ ∫ a, ∥f a∥ ∂μ := -have le_ae : ∀ᵐ a ∂μ, 0 ≤ ∥f a∥ := eventually_of_forall (λa, norm_nonneg _), +lemma norm_integral_le_integral_norm (f : α → E) : ‖(∫ a, f a ∂μ)‖ ≤ ∫ a, ‖f a‖ ∂μ := +have le_ae : ∀ᵐ a ∂μ, 0 ≤ ‖f a‖ := eventually_of_forall (λa, norm_nonneg _), classical.by_cases ( λh : ae_strongly_measurable f μ, - calc ∥∫ a, f a ∂μ∥ ≤ ennreal.to_real (∫⁻ a, (ennreal.of_real ∥f a∥) ∂μ) : + calc ‖∫ a, f a ∂μ‖ ≤ ennreal.to_real (∫⁻ a, (ennreal.of_real ‖f a‖) ∂μ) : norm_integral_le_lintegral_norm _ - ... = ∫ a, ∥f a∥ ∂μ : (integral_eq_lintegral_of_nonneg_ae le_ae $ h.norm).symm ) + ... = ∫ a, ‖f a‖ ∂μ : (integral_eq_lintegral_of_nonneg_ae le_ae $ h.norm).symm ) ( λh : ¬ae_strongly_measurable f μ, begin rw [integral_non_ae_strongly_measurable h, norm_zero], @@ -1158,8 +1158,8 @@ classical.by_cases end ) lemma norm_integral_le_of_norm_le {f : α → E} {g : α → ℝ} (hg : integrable g μ) - (h : ∀ᵐ x ∂μ, ∥f x∥ ≤ g x) : ∥∫ x, f x ∂μ∥ ≤ ∫ x, g x ∂μ := -calc ∥∫ x, f x ∂μ∥ ≤ ∫ x, ∥f x∥ ∂μ : norm_integral_le_integral_norm f + (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ g x) : ‖∫ x, f x ∂μ‖ ≤ ∫ x, g x ∂μ := +calc ‖∫ x, f x ∂μ‖ ≤ ∫ x, ‖f x‖ ∂μ : norm_integral_le_integral_norm f ... ≤ ∫ x, g x ∂μ : integral_mono_of_nonneg (eventually_of_forall $ λ x, norm_nonneg _) hg h @@ -1189,9 +1189,9 @@ begin end lemma norm_integral_le_of_norm_le_const [is_finite_measure μ] {f : α → E} {C : ℝ} - (h : ∀ᵐ x ∂μ, ∥f x∥ ≤ C) : - ∥∫ x, f x ∂μ∥ ≤ C * (μ univ).to_real := -calc ∥∫ x, f x ∂μ∥ ≤ ∫ x, C ∂μ : norm_integral_le_of_norm_le (integrable_const C) h + (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : + ‖∫ x, f x ∂μ‖ ≤ C * (μ univ).to_real := +calc ‖∫ x, f x ∂μ‖ ≤ ∫ x, C ∂μ : norm_integral_le_of_norm_le (integrable_const C) h ... = C * (μ univ).to_real : by rw [integral_const, smul_eq_mul, mul_comm] lemma tendsto_integral_approx_on_of_measurable @@ -1258,7 +1258,7 @@ begin end lemma nndist_integral_add_measure_le_lintegral (h₁ : integrable f μ) (h₂ : integrable f ν) : - (nndist (∫ x, f x ∂μ) (∫ x, f x ∂(μ + ν)) : ℝ≥0∞) ≤ ∫⁻ x, ∥f x∥₊ ∂ν := + (nndist (∫ x, f x ∂μ) (∫ x, f x ∂(μ + ν)) : ℝ≥0∞) ≤ ∫⁻ x, ‖f x‖₊ ∂ν := begin rw [integral_add_measure h₁ h₂, nndist_comm, nndist_eq_nnnorm, add_sub_cancel'], exact ennnorm_integral_le_lintegral_ennnorm _ @@ -1272,11 +1272,11 @@ begin simp only [has_sum, ← integral_finset_sum_measure (λ i _, hfi i)], refine metric.nhds_basis_ball.tendsto_right_iff.mpr (λ ε ε0, _), lift ε to ℝ≥0 using ε0.le, - have hf_lt : ∫⁻ x, ∥f x∥₊ ∂(measure.sum μ) < ∞ := hf.2, - have hmem : ∀ᶠ y in 𝓝 ∫⁻ x, ∥f x∥₊ ∂(measure.sum μ), ∫⁻ x, ∥f x∥₊ ∂(measure.sum μ) < y + ε, + have hf_lt : ∫⁻ x, ‖f x‖₊ ∂(measure.sum μ) < ∞ := hf.2, + have hmem : ∀ᶠ y in 𝓝 ∫⁻ x, ‖f x‖₊ ∂(measure.sum μ), ∫⁻ x, ‖f x‖₊ ∂(measure.sum μ) < y + ε, { refine tendsto_id.add tendsto_const_nhds (lt_mem_nhds $ ennreal.lt_add_right _ _), exacts [hf_lt.ne, ennreal.coe_ne_zero.2 (nnreal.coe_ne_zero.1 ε0.ne')] }, - refine ((has_sum_lintegral_measure (λ x, ∥f x∥₊) μ).eventually hmem).mono (λ s hs, _), + refine ((has_sum_lintegral_measure (λ x, ‖f x‖₊) μ).eventually hmem).mono (λ s hs, _), obtain ⟨ν, hν⟩ : ∃ ν, (∑ i in s, μ i) + ν = measure.sum μ, { refine ⟨measure.sum (λ i : ↥(sᶜ : set ι), μ i), _⟩, simpa only [← measure.sum_coe_finset] using measure.sum_add_sum_compl (s : set ι) μ }, @@ -1413,7 +1413,7 @@ begin { simp only [ne.def, with_top.mul_eq_top_iff, ennreal.of_real_eq_zero, not_le, ennreal.of_real_ne_top, false_and, or_false, not_and], exact λ _, measure_ne_top _ _, }, - { have h_lt_top : ∫⁻ a, ∥f a∥₊ ∂μ < ∞ := hf_int.has_finite_integral, + { have h_lt_top : ∫⁻ a, ‖f a‖₊ ∂μ < ∞ := hf_int.has_finite_integral, simp_rw [← of_real_norm_eq_coe_nnnorm, real.norm_eq_abs] at h_lt_top, convert h_lt_top.ne, ext1 x, @@ -1426,7 +1426,7 @@ conjugate exponents. -/ theorem integral_mul_norm_le_Lp_mul_Lq {E} [normed_add_comm_group E] {f g : α → E} {p q : ℝ} (hpq : p.is_conjugate_exponent q) (hf : mem_ℒp f (ennreal.of_real p) μ) (hg : mem_ℒp g (ennreal.of_real q) μ) : - ∫ a, ∥f a∥ * ∥g a∥ ∂μ ≤ (∫ a, ∥f a∥ ^ p ∂μ) ^ (1/p) * (∫ a, ∥g a∥ ^ q ∂μ) ^ (1/q) := + ∫ a, ‖f a‖ * ‖g a‖ ∂μ ≤ (∫ a, ‖f a‖ ^ p ∂μ) ^ (1/p) * (∫ a, ‖g a‖ ^ q ∂μ) ^ (1/q) := begin -- translate the Bochner integrals into Lebesgue integrals. rw [integral_eq_lintegral_of_nonneg_ae, integral_eq_lintegral_of_nonneg_ae, @@ -1440,13 +1440,13 @@ begin { exact hf.1.norm.mul hg.1.norm, }, rw [ennreal.to_real_rpow, ennreal.to_real_rpow, ← ennreal.to_real_mul], -- replace norms by nnnorm - have h_left : ∫⁻ a, ennreal.of_real (∥f a∥ * ∥g a∥) ∂μ - = ∫⁻ a, ((λ x, (∥f x∥₊ : ℝ≥0∞)) * (λ x, ∥g x∥₊)) a ∂μ, + have h_left : ∫⁻ a, ennreal.of_real (‖f a‖ * ‖g a‖) ∂μ + = ∫⁻ a, ((λ x, (‖f x‖₊ : ℝ≥0∞)) * (λ x, ‖g x‖₊)) a ∂μ, { simp_rw [pi.mul_apply, ← of_real_norm_eq_coe_nnnorm, ennreal.of_real_mul (norm_nonneg _)], }, - have h_right_f : ∫⁻ a, ennreal.of_real (∥f a∥ ^ p) ∂μ = ∫⁻ a, ∥f a∥₊ ^ p ∂μ, + have h_right_f : ∫⁻ a, ennreal.of_real (‖f a‖ ^ p) ∂μ = ∫⁻ a, ‖f a‖₊ ^ p ∂μ, { refine lintegral_congr (λ x, _), rw [← of_real_norm_eq_coe_nnnorm, ennreal.of_real_rpow_of_nonneg (norm_nonneg _) hpq.nonneg], }, - have h_right_g : ∫⁻ a, ennreal.of_real (∥g a∥ ^ q) ∂μ = ∫⁻ a, ∥g a∥₊ ^ q ∂μ, + have h_right_g : ∫⁻ a, ennreal.of_real (‖g a‖ ^ q) ∂μ = ∫⁻ a, ‖g a‖₊ ^ q ∂μ, { refine lintegral_congr (λ x, _), rw [← of_real_norm_eq_coe_nnnorm, ennreal.of_real_rpow_of_nonneg (norm_nonneg _) hpq.symm.nonneg], }, @@ -1478,15 +1478,15 @@ theorem integral_mul_le_Lp_mul_Lq_of_nonneg {p q : ℝ} (hf : mem_ℒp f (ennreal.of_real p) μ) (hg : mem_ℒp g (ennreal.of_real q) μ) : ∫ a, f a * g a ∂μ ≤ (∫ a, (f a) ^ p ∂μ) ^ (1/p) * (∫ a, (g a) ^ q ∂μ) ^ (1/q) := begin - have h_left : ∫ a, f a * g a ∂μ = ∫ a, ∥f a∥ * ∥g a∥ ∂μ, + have h_left : ∫ a, f a * g a ∂μ = ∫ a, ‖f a‖ * ‖g a‖ ∂μ, { refine integral_congr_ae _, filter_upwards [hf_nonneg, hg_nonneg] with x hxf hxg, rw [real.norm_of_nonneg hxf, real.norm_of_nonneg hxg], }, - have h_right_f : ∫ a, (f a) ^ p ∂μ = ∫ a, ∥f a∥ ^ p ∂μ, + have h_right_f : ∫ a, (f a) ^ p ∂μ = ∫ a, ‖f a‖ ^ p ∂μ, { refine integral_congr_ae _, filter_upwards [hf_nonneg] with x hxf, rw real.norm_of_nonneg hxf, }, - have h_right_g : ∫ a, (g a) ^ q ∂μ = ∫ a, ∥g a∥ ^ q ∂μ, + have h_right_g : ∫ a, (g a) ^ q ∂μ = ∫ a, ‖g a‖ ^ q ∂μ, { refine integral_congr_ae _, filter_upwards [hg_nonneg] with x hxg, rw real.norm_of_nonneg hxg, }, diff --git a/src/measure_theory/integral/circle_integral.lean b/src/measure_theory/integral/circle_integral.lean index be24aaf99983c..9cd190c8bdbcd 100644 --- a/src/measure_theory/integral/circle_integral.lean +++ b/src/measure_theory/integral/circle_integral.lean @@ -334,11 +334,11 @@ lemma integral_sub {f g : ℂ → E} {c : ℂ} {R : ℝ} (hf : circle_integrable by simp only [circle_integral, smul_sub, interval_integral.integral_sub hf.out hg.out] lemma norm_integral_le_of_norm_le_const' {f : ℂ → E} {c : ℂ} {R C : ℝ} - (hf : ∀ z ∈ sphere c (|R|), ∥f z∥ ≤ C) : - ∥∮ z in C(c, R), f z∥ ≤ 2 * π * |R| * C := -calc ∥∮ z in C(c, R), f z∥ ≤ |R| * C * |2 * π - 0| : + (hf : ∀ z ∈ sphere c (|R|), ‖f z‖ ≤ C) : + ‖∮ z in C(c, R), f z‖ ≤ 2 * π * |R| * C := +calc ‖∮ z in C(c, R), f z‖ ≤ |R| * C * |2 * π - 0| : interval_integral.norm_integral_le_of_norm_le_const $ λ θ _, - (calc ∥deriv (circle_map c R) θ • f (circle_map c R θ)∥ = |R| * ∥f (circle_map c R θ)∥ : + (calc ‖deriv (circle_map c R) θ • f (circle_map c R θ)‖ = |R| * ‖f (circle_map c R θ)‖ : by simp [norm_smul] ... ≤ |R| * C : mul_le_mul_of_nonneg_left (hf _ $ circle_map_mem_sphere' _ _ _) (_root_.abs_nonneg _)) @@ -346,33 +346,33 @@ calc ∥∮ z in C(c, R), f z∥ ≤ |R| * C * |2 * π - 0| : by { rw [sub_zero, _root_.abs_of_pos real.two_pi_pos], ac_refl } lemma norm_integral_le_of_norm_le_const {f : ℂ → E} {c : ℂ} {R C : ℝ} (hR : 0 ≤ R) - (hf : ∀ z ∈ sphere c R, ∥f z∥ ≤ C) : - ∥∮ z in C(c, R), f z∥ ≤ 2 * π * R * C := + (hf : ∀ z ∈ sphere c R, ‖f z‖ ≤ C) : + ‖∮ z in C(c, R), f z‖ ≤ 2 * π * R * C := have |R| = R, from _root_.abs_of_nonneg hR, -calc ∥∮ z in C(c, R), f z∥ ≤ 2 * π * |R| * C : +calc ‖∮ z in C(c, R), f z‖ ≤ 2 * π * |R| * C : norm_integral_le_of_norm_le_const' $ by rwa this ... = 2 * π * R * C : by rw this lemma norm_two_pi_I_inv_smul_integral_le_of_norm_le_const {f : ℂ → E} {c : ℂ} {R C : ℝ} (hR : 0 ≤ R) - (hf : ∀ z ∈ sphere c R, ∥f z∥ ≤ C) : - ∥(2 * π * I : ℂ)⁻¹ • ∮ z in C(c, R), f z∥ ≤ R * C := + (hf : ∀ z ∈ sphere c R, ‖f z‖ ≤ C) : + ‖(2 * π * I : ℂ)⁻¹ • ∮ z in C(c, R), f z‖ ≤ R * C := begin - have : ∥(2 * π * I : ℂ)⁻¹∥ = (2 * π)⁻¹, by simp [real.pi_pos.le], + have : ‖(2 * π * I : ℂ)⁻¹‖ = (2 * π)⁻¹, by simp [real.pi_pos.le], rw [norm_smul, this, ← div_eq_inv_mul, div_le_iff real.two_pi_pos, mul_comm (R * C), ← mul_assoc], exact norm_integral_le_of_norm_le_const hR hf end -/-- If `f` is continuous on the circle `|z - c| = R`, `R > 0`, the `∥f z∥` is less than or equal to +/-- If `f` is continuous on the circle `|z - c| = R`, `R > 0`, the `‖f z‖` is less than or equal to `C : ℝ` on this circle, and this norm is strictly less than `C` at some point `z` of the circle, -then `∥∮ z in C(c, R), f z∥ < 2 * π * R * C`. -/ +then `‖∮ z in C(c, R), f z‖ < 2 * π * R * C`. -/ lemma norm_integral_lt_of_norm_le_const_of_lt {f : ℂ → E} {c : ℂ} {R C : ℝ} (hR : 0 < R) - (hc : continuous_on f (sphere c R)) (hf : ∀ z ∈ sphere c R, ∥f z∥ ≤ C) - (hlt : ∃ z ∈ sphere c R, ∥f z∥ < C) : - ∥∮ z in C(c, R), f z∥ < 2 * π * R * C := + (hc : continuous_on f (sphere c R)) (hf : ∀ z ∈ sphere c R, ‖f z‖ ≤ C) + (hlt : ∃ z ∈ sphere c R, ‖f z‖ < C) : + ‖∮ z in C(c, R), f z‖ < 2 * π * R * C := begin rw [← _root_.abs_of_pos hR, ← image_circle_map_Ioc] at hlt, rcases hlt with ⟨_, ⟨θ₀, hmem, rfl⟩, hlt⟩, - calc ∥∮ z in C(c, R), f z∥ ≤ ∫ θ in 0..2 * π, ∥deriv (circle_map c R) θ • f (circle_map c R θ)∥ : + calc ‖∮ z in C(c, R), f z‖ ≤ ∫ θ in 0..2 * π, ‖deriv (circle_map c R) θ • f (circle_map c R θ)‖ : interval_integral.norm_integral_le_integral_norm real.two_pi_pos.le ... < ∫ θ in 0..2 * π, R * C : begin @@ -476,18 +476,18 @@ by simp only [cauchy_power_series, continuous_multilinear_map.mk_pi_field_apply, div_eq_mul_inv, mul_pow, mul_smul, circle_integral.integral_smul, ← smul_comm (w ^ n)] lemma norm_cauchy_power_series_le (f : ℂ → E) (c : ℂ) (R : ℝ) (n : ℕ) : - ∥cauchy_power_series f c R n∥ ≤ - (2 * π)⁻¹ * (∫ θ : ℝ in 0..2*π, ∥f (circle_map c R θ)∥) * (|R|⁻¹) ^ n := -calc ∥cauchy_power_series f c R n∥ - = (2 * π)⁻¹ * ∥∮ z in C(c, R), (z - c)⁻¹ ^ n • (z - c)⁻¹ • f z∥ : + ‖cauchy_power_series f c R n‖ ≤ + (2 * π)⁻¹ * (∫ θ : ℝ in 0..2*π, ‖f (circle_map c R θ)‖) * (|R|⁻¹) ^ n := +calc ‖cauchy_power_series f c R n‖ + = (2 * π)⁻¹ * ‖∮ z in C(c, R), (z - c)⁻¹ ^ n • (z - c)⁻¹ • f z‖ : by simp [cauchy_power_series, norm_smul, real.pi_pos.le] -... ≤ (2 * π)⁻¹ * ∫ θ in 0..2*π, ∥deriv (circle_map c R) θ • (circle_map c R θ - c)⁻¹ ^ n • - (circle_map c R θ - c)⁻¹ • f (circle_map c R θ)∥ : +... ≤ (2 * π)⁻¹ * ∫ θ in 0..2*π, ‖deriv (circle_map c R) θ • (circle_map c R θ - c)⁻¹ ^ n • + (circle_map c R θ - c)⁻¹ • f (circle_map c R θ)‖ : mul_le_mul_of_nonneg_left (interval_integral.norm_integral_le_integral_norm real.two_pi_pos.le) (by simp [real.pi_pos.le]) -... = (2 * π)⁻¹ * (|R|⁻¹ ^ n * (|R| * (|R|⁻¹ * ∫ (x : ℝ) in 0..2 * π, ∥f (circle_map c R x)∥))) : +... = (2 * π)⁻¹ * (|R|⁻¹ ^ n * (|R| * (|R|⁻¹ * ∫ (x : ℝ) in 0..2 * π, ‖f (circle_map c R x)‖))) : by simp [norm_smul, mul_left_comm (|R|)] -... ≤ (2 * π)⁻¹ * (∫ θ : ℝ in 0..2*π, ∥f (circle_map c R θ)∥) * |R|⁻¹ ^ n : +... ≤ (2 * π)⁻¹ * (∫ θ : ℝ in 0..2*π, ‖f (circle_map c R θ)‖) * |R|⁻¹ ^ n : begin rcases eq_or_ne R 0 with rfl|hR, { cases n; simp [-mul_inv_rev, real.two_pi_pos] }, @@ -499,7 +499,7 @@ lemma le_radius_cauchy_power_series (f : ℂ → E) (c : ℂ) (R : ℝ≥0) : ↑R ≤ (cauchy_power_series f c R).radius := begin refine (cauchy_power_series f c R).le_radius_of_bound - ((2 * π)⁻¹ * (∫ θ : ℝ in 0..2*π, ∥f (circle_map c R θ)∥)) (λ n, _), + ((2 * π)⁻¹ * (∫ θ : ℝ in 0..2*π, ‖f (circle_map c R θ)‖)) (λ n, _), refine (mul_le_mul_of_nonneg_right (norm_cauchy_power_series_le _ _ _ _) (pow_nonneg R.coe_nonneg _)).trans _, rw [_root_.abs_of_nonneg R.coe_nonneg], @@ -522,18 +522,18 @@ begin have hwR : abs w / R ∈ Ico (0 : ℝ) 1, from ⟨div_nonneg (complex.abs.nonneg w) hR.le, (div_lt_one hR).2 hw⟩, refine interval_integral.has_sum_integral_of_dominated_convergence - (λ n θ, ∥f (circle_map c R θ)∥ * (abs w / R) ^ n) (λ n, _) (λ n, _) _ _ _, + (λ n θ, ‖f (circle_map c R θ)‖ * (abs w / R) ^ n) (λ n, _) (λ n, _) _ _ _, { simp only [deriv_circle_map], apply_rules [ae_strongly_measurable.smul, hf.def.1]; { apply measurable.ae_strongly_measurable, measurability } }, - { simp [norm_smul, abs_of_pos hR, mul_left_comm R, mul_inv_cancel_left₀ hR.ne', mul_comm (∥_∥)] }, + { simp [norm_smul, abs_of_pos hR, mul_left_comm R, mul_inv_cancel_left₀ hR.ne', mul_comm (‖_‖)] }, { exact eventually_of_forall (λ _ _, (summable_geometric_of_lt_1 hwR.1 hwR.2).mul_left _) }, { simpa only [tsum_mul_left, tsum_geometric_of_lt_1 hwR.1 hwR.2] using hf.norm.mul_continuous_on continuous_on_const }, { refine eventually_of_forall (λ θ hθ, has_sum.const_smul _), simp only [smul_smul], refine has_sum.smul_const _, - have : ∥w / (circle_map c R θ - c)∥ < 1, by simpa [abs_of_pos hR] using hwR.2, + have : ‖w / (circle_map c R θ - c)‖ < 1, by simpa [abs_of_pos hR] using hwR.2, convert (has_sum_geometric_of_norm_lt_1 this).mul_right _, simp [← sub_sub, ← mul_inv, sub_mul, div_mul_cancel _ (circle_map_ne_center hR.ne')] } end diff --git a/src/measure_theory/integral/circle_transform.lean b/src/measure_theory/integral/circle_transform.lean index 074c71cb07615..7ca20fb8487a1 100644 --- a/src/measure_theory/integral/circle_transform.lean +++ b/src/measure_theory/integral/circle_transform.lean @@ -140,7 +140,7 @@ end lemma circle_transform_deriv_bound {R : ℝ} (hR : 0 < R) {z x : ℂ} {f : ℂ → ℂ} (hx : x ∈ ball z R) (hf : continuous_on f (sphere z R)) : ∃ (B ε : ℝ), 0 < ε ∧ ball x ε ⊆ ball z R ∧ - (∀ (t : ℝ) (y ∈ ball x ε), ∥circle_transform_deriv R z y f t∥ ≤ B) := + (∀ (t : ℝ) (y ∈ ball x ε), ‖circle_transform_deriv R z y f t‖ ≤ B) := begin obtain ⟨r, hr, hrx⟩ := exists_lt_mem_ball_of_mem_ball hx, obtain ⟨ε', hε', H⟩ := exists_ball_subset_ball hrx, diff --git a/src/measure_theory/integral/divergence_theorem.lean b/src/measure_theory/integral/divergence_theorem.lean index 7c666a5d4f53b..8219c2b6bc01b 100644 --- a/src/measure_theory/integral/divergence_theorem.lean +++ b/src/measure_theory/integral/divergence_theorem.lean @@ -202,7 +202,7 @@ begin from box.le_iff_Icc.1 (box.face_mono (hJ_le _) i), rw [mem_closed_ball_zero_iff, real.norm_eq_abs, abs_of_nonneg dist_nonneg, dist_eq_norm, ← integral_sub (Hid.mono_set Hsub) ((Hic _).mono_set Hsub)], - calc ∥(∫ x in ((J k).face i).Icc, f (i.insert_nth d x) i - f (i.insert_nth (c k) x) i)∥ + calc ‖(∫ x in ((J k).face i).Icc, f (i.insert_nth d x) i - f (i.insert_nth (c k) x) i)‖ ≤ (ε / ∏ j, ((I.face i).upper j - (I.face i).lower j)) * (volume ((J k).face i).Icc).to_real : begin refine norm_set_integral_le_of_norm_le_const' (((J k).face i).measure_Icc_lt_top _) diff --git a/src/measure_theory/integral/integrable_on.lean b/src/measure_theory/integral/integrable_on.lean index e4425cf536eba..0e9560c8f3522 100644 --- a/src/measure_theory/integral/integrable_on.lean +++ b/src/measure_theory/integral/integrable_on.lean @@ -68,7 +68,7 @@ namespace measure_theory section normed_add_comm_group lemma has_finite_integral_restrict_of_bounded [normed_add_comm_group E] {f : α → E} {s : set α} - {μ : measure α} {C} (hs : μ s < ∞) (hf : ∀ᵐ x ∂(μ.restrict s), ∥f x∥ ≤ C) : + {μ : measure α} {C} (hs : μ s < ∞) (hf : ∀ᵐ x ∂(μ.restrict s), ‖f x‖ ≤ C) : has_finite_integral f (μ.restrict s) := by haveI : is_finite_measure (μ.restrict s) := ⟨by rwa [measure.restrict_apply_univ]⟩; exact has_finite_integral_of_bounded hf @@ -249,7 +249,7 @@ end lemma integrable.lintegral_lt_top {f : α → ℝ} (hf : integrable f μ) : ∫⁻ x, ennreal.of_real (f x) ∂μ < ∞ := calc ∫⁻ x, ennreal.of_real (f x) ∂μ - ≤ ∫⁻ x, ↑∥f x∥₊ ∂μ : lintegral_of_real_le_lintegral_nnnorm f + ≤ ∫⁻ x, ↑‖f x‖₊ ∂μ : lintegral_of_real_le_lintegral_nnnorm f ... < ∞ : hf.2 lemma integrable_on.set_lintegral_lt_top {f : α → ℝ} {s : set α} (hf : integrable_on f s μ) : @@ -300,7 +300,7 @@ lemma measure.finite_at_filter.integrable_at_filter {l : filter α} [is_measurab (hf : l.is_bounded_under (≤) (norm ∘ f)) : integrable_at_filter f l μ := begin - obtain ⟨C, hC⟩ : ∃ C, ∀ᶠ s in l.small_sets, ∀ x ∈ s, ∥f x∥ ≤ C, + obtain ⟨C, hC⟩ : ∃ C, ∀ᶠ s in l.small_sets, ∀ x ∈ s, ‖f x‖ ≤ C, from hf.imp (λ C hC, eventually_small_sets.2 ⟨_, hC, λ t, id⟩), rcases (hfm.eventually.and (hμ.eventually.and hC)).exists_measurable_mem_of_small_sets with ⟨s, hsl, hsm, hfm, hμ, hC⟩, diff --git a/src/measure_theory/integral/integral_eq_improper.lean b/src/measure_theory/integral/integral_eq_improper.lean index fe3dcabd1f97e..17d2de0aed040 100644 --- a/src/measure_theory/integral/integral_eq_improper.lean +++ b/src/measure_theory/integral/integral_eq_improper.lean @@ -43,7 +43,7 @@ as an `ae_cover` w.r.t. `μ.restrict (Iic b)`, instead of using `(λ x, Ioc x b) then `∫⁻ x in φ n, f x ∂μ` tends to `∫⁻ x, f x ∂μ` as `n` tends to `l` - `measure_theory.ae_cover.integrable_of_integral_norm_tendsto` : if `φ` is a `ae_cover μ l`, where `l` is a countably generated filter, if `f` is measurable and integrable on each `φ n`, - and if `∫ x in φ n, ∥f x∥ ∂μ` tends to some `I : ℝ` as n tends to `l`, then `f` is integrable + and if `∫ x in φ n, ‖f x‖ ∂μ` tends to some `I : ℝ` as n tends to `l`, then `f` is integrable - `measure_theory.ae_cover.integral_tendsto_of_countably_generated` : if `φ` is a `ae_cover μ l`, where `l` is a countably generated filter, and if `f` is measurable and integrable (globally), then `∫ x in φ n, f x ∂μ` tends to `∫ x, f x ∂μ` as `n` tends to `+∞`. @@ -389,7 +389,7 @@ variables {α ι E : Type*} [measurable_space α] {μ : measure α} {l : filter lemma ae_cover.integrable_of_lintegral_nnnorm_bounded [l.ne_bot] [l.is_countably_generated] {φ : ι → set α} (hφ : ae_cover μ l φ) {f : α → E} (I : ℝ) (hfm : ae_strongly_measurable f μ) - (hbounded : ∀ᶠ i in l, ∫⁻ x in φ i, ∥f x∥₊ ∂μ ≤ ennreal.of_real I) : + (hbounded : ∀ᶠ i in l, ∫⁻ x in φ i, ‖f x‖₊ ∂μ ≤ ennreal.of_real I) : integrable f μ := begin refine ⟨hfm, (le_of_tendsto _ hbounded).trans_lt ennreal.of_real_lt_top⟩, @@ -399,7 +399,7 @@ end lemma ae_cover.integrable_of_lintegral_nnnorm_tendsto [l.ne_bot] [l.is_countably_generated] {φ : ι → set α} (hφ : ae_cover μ l φ) {f : α → E} (I : ℝ) (hfm : ae_strongly_measurable f μ) - (htendsto : tendsto (λ i, ∫⁻ x in φ i, ∥f x∥₊ ∂μ) l (𝓝 $ ennreal.of_real I)) : + (htendsto : tendsto (λ i, ∫⁻ x in φ i, ‖f x‖₊ ∂μ) l (𝓝 $ ennreal.of_real I)) : integrable f μ := begin refine hφ.integrable_of_lintegral_nnnorm_bounded (max 1 (I + 1)) hfm _, @@ -410,7 +410,7 @@ end lemma ae_cover.integrable_of_lintegral_nnnorm_bounded' [l.ne_bot] [l.is_countably_generated] {φ : ι → set α} (hφ : ae_cover μ l φ) {f : α → E} (I : ℝ≥0) (hfm : ae_strongly_measurable f μ) - (hbounded : ∀ᶠ i in l, ∫⁻ x in φ i, ∥f x∥₊ ∂μ ≤ I) : + (hbounded : ∀ᶠ i in l, ∫⁻ x in φ i, ‖f x‖₊ ∂μ ≤ I) : integrable f μ := hφ.integrable_of_lintegral_nnnorm_bounded I hfm (by simpa only [ennreal.of_real_coe_nnreal] using hbounded) @@ -418,7 +418,7 @@ hφ.integrable_of_lintegral_nnnorm_bounded I hfm lemma ae_cover.integrable_of_lintegral_nnnorm_tendsto' [l.ne_bot] [l.is_countably_generated] {φ : ι → set α} (hφ : ae_cover μ l φ) {f : α → E} (I : ℝ≥0) (hfm : ae_strongly_measurable f μ) - (htendsto : tendsto (λ i, ∫⁻ x in φ i, ∥f x∥₊ ∂μ) l (𝓝 I)) : + (htendsto : tendsto (λ i, ∫⁻ x in φ i, ‖f x‖₊ ∂μ) l (𝓝 I)) : integrable f μ := hφ.integrable_of_lintegral_nnnorm_tendsto I hfm (by simpa only [ennreal.of_real_coe_nnreal] using htendsto) @@ -426,7 +426,7 @@ hφ.integrable_of_lintegral_nnnorm_tendsto I hfm lemma ae_cover.integrable_of_integral_norm_bounded [l.ne_bot] [l.is_countably_generated] {φ : ι → set α} (hφ : ae_cover μ l φ) {f : α → E} (I : ℝ) (hfi : ∀ i, integrable_on f (φ i) μ) - (hbounded : ∀ᶠ i in l, ∫ x in φ i, ∥f x∥ ∂μ ≤ I) : + (hbounded : ∀ᶠ i in l, ∫ x in φ i, ‖f x‖ ∂μ ≤ I) : integrable f μ := begin have hfm : ae_strongly_measurable f μ := @@ -444,7 +444,7 @@ end lemma ae_cover.integrable_of_integral_norm_tendsto [l.ne_bot] [l.is_countably_generated] {φ : ι → set α} (hφ : ae_cover μ l φ) {f : α → E} (I : ℝ) (hfi : ∀ i, integrable_on f (φ i) μ) - (htendsto : tendsto (λ i, ∫ x in φ i, ∥f x∥ ∂μ) l (𝓝 I)) : + (htendsto : tendsto (λ i, ∫ x in φ i, ‖f x‖ ∂μ) l (𝓝 I)) : integrable f μ := let ⟨I', hI'⟩ := htendsto.is_bounded_under_le in hφ.integrable_of_integral_norm_bounded I' hfi hI' @@ -476,7 +476,7 @@ lemma ae_cover.integral_tendsto_of_countably_generated [l.is_countably_generated tendsto (λ i, ∫ x in φ i, f x ∂μ) l (𝓝 $ ∫ x, f x ∂μ) := suffices h : tendsto (λ i, ∫ (x : α), (φ i).indicator f x ∂μ) l (𝓝 (∫ (x : α), f x ∂μ)), by { convert h, ext n, rw integral_indicator (hφ.measurable n) }, -tendsto_integral_filter_of_dominated_convergence (λ x, ∥f x∥) +tendsto_integral_filter_of_dominated_convergence (λ x, ‖f x‖) (eventually_of_forall $ λ i, hfi.ae_strongly_measurable.indicator $ hφ.measurable i) (eventually_of_forall $ λ i, ae_of_all _ $ λ x, norm_indicator_le_norm_self _ _) hfi.norm (hφ.ae_tendsto_indicator f) @@ -511,7 +511,7 @@ variables {ι E : Type*} {μ : measure ℝ} lemma integrable_of_interval_integral_norm_bounded (I : ℝ) (hfi : ∀ i, integrable_on f (Ioc (a i) (b i)) μ) (ha : tendsto a l at_bot) (hb : tendsto b l at_top) - (h : ∀ᶠ i in l, ∫ x in a i .. b i, ∥f x∥ ∂μ ≤ I) : + (h : ∀ᶠ i in l, ∫ x in a i .. b i, ‖f x‖ ∂μ ≤ I) : integrable f μ := begin have hφ : ae_cover μ l _ := ae_cover_Ioc ha hb, @@ -523,19 +523,19 @@ end /-- If `f` is integrable on intervals `Ioc (a i) (b i)`, where `a i` tends to -∞ and `b i` tends to ∞, and -`∫ x in a i .. b i, ∥f x∥ ∂μ` converges to `I : ℝ` along a filter `l`, +`∫ x in a i .. b i, ‖f x‖ ∂μ` converges to `I : ℝ` along a filter `l`, then `f` is integrable on the interval (-∞, ∞) -/ lemma integrable_of_interval_integral_norm_tendsto (I : ℝ) (hfi : ∀ i, integrable_on f (Ioc (a i) (b i)) μ) (ha : tendsto a l at_bot) (hb : tendsto b l at_top) - (h : tendsto (λ i, ∫ x in a i .. b i, ∥f x∥ ∂μ) l (𝓝 I)) : + (h : tendsto (λ i, ∫ x in a i .. b i, ‖f x‖ ∂μ) l (𝓝 I)) : integrable f μ := let ⟨I', hI'⟩ := h.is_bounded_under_le in integrable_of_interval_integral_norm_bounded I' hfi ha hb hI' lemma integrable_on_Iic_of_interval_integral_norm_bounded (I b : ℝ) (hfi : ∀ i, integrable_on f (Ioc (a i) b) μ) (ha : tendsto a l at_bot) - (h : ∀ᶠ i in l, (∫ x in a i .. b, ∥f x∥ ∂μ) ≤ I) : + (h : ∀ᶠ i in l, (∫ x in a i .. b, ‖f x‖ ∂μ) ≤ I) : integrable_on f (Iic b) μ := begin have hφ : ae_cover (μ.restrict $ Iic b) l _ := ae_cover_Ioi ha, @@ -551,18 +551,18 @@ end /-- If `f` is integrable on intervals `Ioc (a i) b`, where `a i` tends to -∞, and -`∫ x in a i .. b, ∥f x∥ ∂μ` converges to `I : ℝ` along a filter `l`, +`∫ x in a i .. b, ‖f x‖ ∂μ` converges to `I : ℝ` along a filter `l`, then `f` is integrable on the interval (-∞, b) -/ lemma integrable_on_Iic_of_interval_integral_norm_tendsto (I b : ℝ) (hfi : ∀ i, integrable_on f (Ioc (a i) b) μ) (ha : tendsto a l at_bot) - (h : tendsto (λ i, ∫ x in a i .. b, ∥f x∥ ∂μ) l (𝓝 I)) : + (h : tendsto (λ i, ∫ x in a i .. b, ‖f x‖ ∂μ) l (𝓝 I)) : integrable_on f (Iic b) μ := let ⟨I', hI'⟩ := h.is_bounded_under_le in integrable_on_Iic_of_interval_integral_norm_bounded I' b hfi ha hI' lemma integrable_on_Ioi_of_interval_integral_norm_bounded (I a : ℝ) (hfi : ∀ i, integrable_on f (Ioc a (b i)) μ) (hb : tendsto b l at_top) - (h : ∀ᶠ i in l, (∫ x in a .. b i, ∥f x∥ ∂μ) ≤ I) : + (h : ∀ᶠ i in l, (∫ x in a .. b i, ‖f x‖ ∂μ) ≤ I) : integrable_on f (Ioi a) μ := begin have hφ : ae_cover (μ.restrict $ Ioi a) l _ := ae_cover_Iic hb, @@ -579,11 +579,11 @@ end /-- If `f` is integrable on intervals `Ioc a (b i)`, where `b i` tends to ∞, and -`∫ x in a .. b i, ∥f x∥ ∂μ` converges to `I : ℝ` along a filter `l`, +`∫ x in a .. b i, ‖f x‖ ∂μ` converges to `I : ℝ` along a filter `l`, then `f` is integrable on the interval (a, ∞) -/ lemma integrable_on_Ioi_of_interval_integral_norm_tendsto (I a : ℝ) (hfi : ∀ i, integrable_on f (Ioc a (b i)) μ) (hb : tendsto b l at_top) - (h : tendsto (λ i, ∫ x in a .. b i, ∥f x∥ ∂μ) l (𝓝 $ I)) : + (h : tendsto (λ i, ∫ x in a .. b i, ‖f x‖ ∂μ) l (𝓝 $ I)) : integrable_on f (Ioi a) μ := let ⟨I', hI'⟩ := h.is_bounded_under_le in integrable_on_Ioi_of_interval_integral_norm_bounded I' a hfi hb hI' @@ -591,7 +591,7 @@ let ⟨I', hI'⟩ := h.is_bounded_under_le in lemma integrable_on_Ioc_of_interval_integral_norm_bounded {I a₀ b₀ : ℝ} (hfi : ∀ i, integrable_on f $ Ioc (a i) (b i)) (ha : tendsto a l $ 𝓝 a₀) (hb : tendsto b l $ 𝓝 b₀) - (h : ∀ᶠ i in l, (∫ x in Ioc (a i) (b i), ∥f x∥) ≤ I) : integrable_on f (Ioc a₀ b₀) := + (h : ∀ᶠ i in l, (∫ x in Ioc (a i) (b i), ‖f x‖) ≤ I) : integrable_on f (Ioc a₀ b₀) := begin refine (ae_cover_Ioc_of_Ioc ha hb).integrable_of_integral_norm_bounded I (λ i, (hfi i).restrict measurable_set_Ioc) (eventually.mono h _), @@ -603,12 +603,12 @@ end lemma integrable_on_Ioc_of_interval_integral_norm_bounded_left {I a₀ b : ℝ} (hfi : ∀ i, integrable_on f $ Ioc (a i) b) (ha : tendsto a l $ 𝓝 a₀) - (h : ∀ᶠ i in l, (∫ x in Ioc (a i) b, ∥f x∥ ) ≤ I) : integrable_on f (Ioc a₀ b) := + (h : ∀ᶠ i in l, (∫ x in Ioc (a i) b, ‖f x‖ ) ≤ I) : integrable_on f (Ioc a₀ b) := integrable_on_Ioc_of_interval_integral_norm_bounded hfi ha tendsto_const_nhds h lemma integrable_on_Ioc_of_interval_integral_norm_bounded_right {I a b₀ : ℝ} (hfi : ∀ i, integrable_on f $ Ioc a (b i)) (hb : tendsto b l $ 𝓝 b₀) - (h : ∀ᶠ i in l, (∫ x in Ioc a (b i), ∥f x∥ ) ≤ I) : integrable_on f (Ioc a b₀) := + (h : ∀ᶠ i in l, (∫ x in Ioc a (b i), ‖f x‖ ) ≤ I) : integrable_on f (Ioc a b₀) := integrable_on_Ioc_of_interval_integral_norm_bounded hfi tendsto_const_nhds hb h end integrable_of_interval_integral diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index c0ad8967d15e1..7a2f7396a83ef 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -75,7 +75,7 @@ be an `FTC_filter` pair of filters around `a` (i.e., `FTC_filter a la la'`) and an `FTC_filter` pair of filters around `b`. If `f` has finite limits `ca` and `cb` almost surely at `la'` and `lb'`, respectively, then `∫ x in va..vb, f x ∂μ - ∫ x in ua..ub, f x ∂μ = ∫ x in ub..vb, cb ∂μ - ∫ x in ua..va, ca ∂μ + - o(∥∫ x in ua..va, (1:ℝ) ∂μ∥ + ∥∫ x in ub..vb, (1:ℝ) ∂μ∥)` as `ua` and `va` tend to `la` while + o(‖∫ x in ua..va, (1:ℝ) ∂μ‖ + ‖∫ x in ub..vb, (1:ℝ) ∂μ‖)` as `ua` and `va` tend to `la` while `ub` and `vb` tend to `lb`. ### FTC-2 and corollaries @@ -311,12 +311,12 @@ lemma neg (h : interval_integrable f μ a b) : interval_integrable (-f) μ a b : ⟨h.1.neg, h.2.neg⟩ lemma norm (h : interval_integrable f μ a b) : - interval_integrable (λ x, ∥f x∥) μ a b := + interval_integrable (λ x, ‖f x‖) μ a b := ⟨h.1.norm, h.2.norm⟩ lemma interval_integrable_norm_iff {f : ℝ → E} {μ : measure ℝ} {a b : ℝ} (hf : ae_strongly_measurable f (μ.restrict (Ι a b))) : - interval_integrable (λ t, ∥f t∥) μ a b ↔ interval_integrable f μ a b := + interval_integrable (λ t, ‖f t‖) μ a b ↔ interval_integrable f μ a b := by { simp_rw [interval_integrable_iff, integrable_on], exact integrable_norm_iff hf } lemma abs {f : ℝ → ℝ} (h : interval_integrable f μ a b) : @@ -346,12 +346,12 @@ hf.mono_set_ae $ eventually_of_forall hsub lemma mono_fun [normed_add_comm_group F] {g : ℝ → F} (hf : interval_integrable f μ a b) (hgm : ae_strongly_measurable g (μ.restrict (Ι a b))) - (hle : (λ x, ∥g x∥) ≤ᵐ[μ.restrict (Ι a b)] (λ x, ∥f x∥)) : interval_integrable g μ a b := + (hle : (λ x, ‖g x‖) ≤ᵐ[μ.restrict (Ι a b)] (λ x, ‖f x‖)) : interval_integrable g μ a b := interval_integrable_iff.2 $ hf.def.integrable.mono hgm hle lemma mono_fun' {g : ℝ → ℝ} (hg : interval_integrable g μ a b) (hfm : ae_strongly_measurable f (μ.restrict (Ι a b))) - (hle : (λ x, ∥f x∥) ≤ᵐ[μ.restrict (Ι a b)] g) : interval_integrable f μ a b := + (hle : (λ x, ‖f x‖) ≤ᵐ[μ.restrict (Ι a b)] g) : interval_integrable f μ a b := interval_integrable_iff.2 $ hg.def.integrable.mono' hfm hle protected lemma ae_strongly_measurable (h : interval_integrable f μ a b) : @@ -556,7 +556,7 @@ begin end lemma norm_interval_integral_eq (f : ℝ → E) (a b : ℝ) (μ : measure ℝ) : - ∥∫ x in a..b, f x ∂μ∥ = ∥∫ x in Ι a b, f x ∂μ∥ := + ‖∫ x in a..b, f x ∂μ‖ = ‖∫ x in Ι a b, f x ∂μ‖ := begin simp_rw [interval_integral_eq_integral_interval_oc, norm_smul], split_ifs; simp only [norm_neg, norm_one, one_mul], @@ -593,11 +593,11 @@ lemma integral_non_ae_strongly_measurable_of_le (h : a ≤ b) integral_non_ae_strongly_measurable $ by rwa [interval_oc_of_le h] lemma norm_integral_min_max (f : ℝ → E) : - ∥∫ x in min a b..max a b, f x ∂μ∥ = ∥∫ x in a..b, f x ∂μ∥ := + ‖∫ x in min a b..max a b, f x ∂μ‖ = ‖∫ x in a..b, f x ∂μ‖ := by cases le_total a b; simp [*, integral_symm a b] lemma norm_integral_eq_norm_integral_Ioc (f : ℝ → E) : - ∥∫ x in a..b, f x ∂μ∥ = ∥∫ x in Ι a b, f x ∂μ∥ := + ‖∫ x in a..b, f x ∂μ‖ = ‖∫ x in Ι a b, f x ∂μ‖ := by rw [← norm_integral_min_max, integral_of_le min_le_max, interval_oc] lemma abs_integral_eq_abs_integral_interval_oc (f : ℝ → ℝ) : @@ -605,33 +605,33 @@ lemma abs_integral_eq_abs_integral_interval_oc (f : ℝ → ℝ) : norm_integral_eq_norm_integral_Ioc f lemma norm_integral_le_integral_norm_Ioc : - ∥∫ x in a..b, f x ∂μ∥ ≤ ∫ x in Ι a b, ∥f x∥ ∂μ := -calc ∥∫ x in a..b, f x ∂μ∥ = ∥∫ x in Ι a b, f x ∂μ∥ : + ‖∫ x in a..b, f x ∂μ‖ ≤ ∫ x in Ι a b, ‖f x‖ ∂μ := +calc ‖∫ x in a..b, f x ∂μ‖ = ‖∫ x in Ι a b, f x ∂μ‖ : norm_integral_eq_norm_integral_Ioc f -... ≤ ∫ x in Ι a b, ∥f x∥ ∂μ : +... ≤ ∫ x in Ι a b, ‖f x‖ ∂μ : norm_integral_le_integral_norm f -lemma norm_integral_le_abs_integral_norm : ∥∫ x in a..b, f x ∂μ∥ ≤ |∫ x in a..b, ∥f x∥ ∂μ| := +lemma norm_integral_le_abs_integral_norm : ‖∫ x in a..b, f x ∂μ‖ ≤ |∫ x in a..b, ‖f x‖ ∂μ| := begin simp only [← real.norm_eq_abs, norm_integral_eq_norm_integral_Ioc], exact le_trans (norm_integral_le_integral_norm _) (le_abs_self _) end lemma norm_integral_le_integral_norm (h : a ≤ b) : - ∥∫ x in a..b, f x ∂μ∥ ≤ ∫ x in a..b, ∥f x∥ ∂μ := + ‖∫ x in a..b, f x ∂μ‖ ≤ ∫ x in a..b, ‖f x‖ ∂μ := norm_integral_le_integral_norm_Ioc.trans_eq $ by rw [interval_oc_of_le h, integral_of_le h] lemma norm_integral_le_of_norm_le {g : ℝ → ℝ} - (h : ∀ᵐ t ∂(μ.restrict $ Ι a b), ∥f t∥ ≤ g t) + (h : ∀ᵐ t ∂(μ.restrict $ Ι a b), ‖f t‖ ≤ g t) (hbound : interval_integrable g μ a b) : - ∥∫ t in a..b, f t ∂μ∥ ≤ |∫ t in a..b, g t ∂μ| := + ‖∫ t in a..b, f t ∂μ‖ ≤ |∫ t in a..b, g t ∂μ| := by simp_rw [norm_interval_integral_eq, abs_interval_integral_eq, abs_eq_self.mpr (integral_nonneg_of_ae $ h.mono $ λ t ht, (norm_nonneg _).trans ht), norm_integral_le_of_norm_le hbound.def h] lemma norm_integral_le_of_norm_le_const_ae {a b C : ℝ} {f : ℝ → E} - (h : ∀ᵐ x, x ∈ Ι a b → ∥f x∥ ≤ C) : - ∥∫ x in a..b, f x∥ ≤ C * |b - a| := + (h : ∀ᵐ x, x ∈ Ι a b → ‖f x‖ ≤ C) : + ‖∫ x in a..b, f x‖ ≤ C * |b - a| := begin rw [norm_integral_eq_norm_integral_Ioc], convert norm_set_integral_le_of_norm_le_const_ae'' _ measurable_set_Ioc h, @@ -640,8 +640,8 @@ begin end lemma norm_integral_le_of_norm_le_const {a b C : ℝ} {f : ℝ → E} - (h : ∀ x ∈ Ι a b, ∥f x∥ ≤ C) : - ∥∫ x in a..b, f x∥ ≤ C * |b - a| := + (h : ∀ x ∈ Ι a b, ‖f x‖ ≤ C) : + ‖∫ x in a..b, f x‖ ≤ C * |b - a| := norm_integral_le_of_norm_le_const_ae $ eventually_of_forall h @[simp] lemma integral_add (hf : interval_integrable f μ a b) (hg : interval_integrable g μ a b) : @@ -992,7 +992,7 @@ end lemma tendsto_integral_filter_of_dominated_convergence {ι} {l : filter ι} [l.is_countably_generated] {F : ι → ℝ → E} (bound : ℝ → ℝ) (hF_meas : ∀ᶠ n in l, ae_strongly_measurable (F n) (μ.restrict (Ι a b))) - (h_bound : ∀ᶠ n in l, ∀ᵐ x ∂μ, x ∈ Ι a b → ∥F n x∥ ≤ bound x) + (h_bound : ∀ᶠ n in l, ∀ᵐ x ∂μ, x ∈ Ι a b → ‖F n x‖ ≤ bound x) (bound_integrable : interval_integrable bound μ a b) (h_lim : ∀ᵐ x ∂μ, x ∈ Ι a b → tendsto (λ n, F n x) l (𝓝 (f x))) : tendsto (λn, ∫ x in a..b, F n x ∂μ) l (𝓝 $ ∫ x in a..b, f x ∂μ) := @@ -1007,7 +1007,7 @@ end lemma has_sum_integral_of_dominated_convergence {ι} [countable ι] {F : ι → ℝ → E} (bound : ι → ℝ → ℝ) (hF_meas : ∀ n, ae_strongly_measurable (F n) (μ.restrict (Ι a b))) - (h_bound : ∀ n, ∀ᵐ t ∂μ, t ∈ Ι a b → ∥F n t∥ ≤ bound n t) + (h_bound : ∀ n, ∀ᵐ t ∂μ, t ∈ Ι a b → ‖F n t‖ ≤ bound n t) (bound_summable : ∀ᵐ t ∂μ, t ∈ Ι a b → summable (λ n, bound n t)) (bound_integrable : interval_integrable (λ t, ∑' n, bound n t) μ a b) (h_lim : ∀ᵐ t ∂μ, t ∈ Ι a b → has_sum (λ n, F n t) (f t)) : @@ -1031,7 +1031,7 @@ variables {X : Type*} [topological_space X] [first_countable_topology X] lemma continuous_within_at_of_dominated_interval {F : X → ℝ → E} {x₀ : X} {bound : ℝ → ℝ} {a b : ℝ} {s : set X} (hF_meas : ∀ᶠ x in 𝓝[s] x₀, ae_strongly_measurable (F x) (μ.restrict $ Ι a b)) - (h_bound : ∀ᶠ x in 𝓝[s] x₀, ∀ᵐ t ∂μ, t ∈ Ι a b → ∥F x t∥ ≤ bound t) + (h_bound : ∀ᶠ x in 𝓝[s] x₀, ∀ᵐ t ∂μ, t ∈ Ι a b → ‖F x t‖ ≤ bound t) (bound_integrable : interval_integrable bound μ a b) (h_cont : ∀ᵐ t ∂μ, t ∈ Ι a b → continuous_within_at (λ x, F x t) s x₀) : continuous_within_at (λ x, ∫ t in a..b, F x t ∂μ) s x₀ := @@ -1046,7 +1046,7 @@ tendsto_integral_filter_of_dominated_convergence bound hF_meas h_bound bound_int lemma continuous_at_of_dominated_interval {F : X → ℝ → E} {x₀ : X} {bound : ℝ → ℝ} {a b : ℝ} (hF_meas : ∀ᶠ x in 𝓝 x₀, ae_strongly_measurable (F x) (μ.restrict $ Ι a b)) - (h_bound : ∀ᶠ x in 𝓝 x₀, ∀ᵐ t ∂μ, t ∈ Ι a b → ∥F x t∥ ≤ bound t) + (h_bound : ∀ᶠ x in 𝓝 x₀, ∀ᵐ t ∂μ, t ∈ Ι a b → ‖F x t‖ ≤ bound t) (bound_integrable : interval_integrable bound μ a b) (h_cont : ∀ᵐ t ∂μ, t ∈ Ι a b → continuous_at (λ x, F x t) x₀) : continuous_at (λ x, ∫ t in a..b, F x t ∂μ) x₀ := @@ -1059,7 +1059,7 @@ tendsto_integral_filter_of_dominated_convergence bound hF_meas h_bound bound_int then the same holds for `(λ x, ∫ t in a..b, F x t ∂μ) s x₀`. -/ lemma continuous_of_dominated_interval {F : X → ℝ → E} {bound : ℝ → ℝ} {a b : ℝ} (hF_meas : ∀ x, ae_strongly_measurable (F x) $ μ.restrict $ Ι a b) - (h_bound : ∀ x, ∀ᵐ t ∂μ, t ∈ Ι a b → ∥F x t∥ ≤ bound t) + (h_bound : ∀ x, ∀ᵐ t ∂μ, t ∈ Ι a b → ‖F x t‖ ≤ bound t) (bound_integrable : interval_integrable bound μ a b) (h_cont : ∀ᵐ t ∂μ, t ∈ Ι a b → continuous (λ x, F x t)) : continuous (λ x, ∫ t in a..b, F x t ∂μ) := @@ -1105,7 +1105,7 @@ begin exact λ b b_in, (integral_indicator b_in).symm }, apply continuous_within_at.congr_of_eventually_eq _ this (integral_indicator h₀).symm, - have : interval_integrable (λ x, ∥f x∥) μ b₁ b₂, + have : interval_integrable (λ x, ‖f x‖) μ b₁ b₂, from interval_integrable.norm (h_int' $ right_mem_Icc.mpr h₁₂), refine continuous_within_at_of_dominated_interval _ _ this _ ; clear this, { apply eventually.mono (self_mem_nhds_within), @@ -1391,7 +1391,7 @@ Namely, let `f` be a measurable function integrable on `a..b`. Let `(la, la')` b `FTC_filter`s around `a`; let `(lb, lb')` be a pair of `FTC_filter`s around `b`. Suppose that `f` has finite limits `ca` and `cb` at `la' ⊓ μ.ae` and `lb' ⊓ μ.ae`, respectively. Then `∫ x in va..vb, f x ∂μ - ∫ x in ua..ub, f x ∂μ = ∫ x in ub..vb, cb ∂μ - ∫ x in ua..va, ca ∂μ + - o(∥∫ x in ua..va, (1:ℝ) ∂μ∥ + ∥∫ x in ub..vb, (1:ℝ) ∂μ∥)` + o(‖∫ x in ua..va, (1:ℝ) ∂μ‖ + ‖∫ x in ub..vb, (1:ℝ) ∂μ‖)` as `ua` and `va` tend to `la` while `ub` and `vb` tend to `lb`. This theorem is formulated with integral of constants instead of measures in the right hand sides @@ -1598,7 +1598,7 @@ around `a`; let `(lb, lb')` be a pair of `FTC_filter`s around `b`. Suppose that limits `ca` and `cb` at `la' ⊓ μ.ae` and `lb' ⊓ μ.ae`, respectively. Then `∫ x in va..vb, f x ∂μ - ∫ x in ua..ub, f x ∂μ = ∫ x in ub..vb, cb ∂μ - ∫ x in ua..va, ca ∂μ + - o(∥∫ x in ua..va, (1:ℝ) ∂μ∥ + ∥∫ x in ub..vb, (1:ℝ) ∂μ∥)` + o(‖∫ x in ua..va, (1:ℝ) ∂μ‖ + ‖∫ x in ub..vb, (1:ℝ) ∂μ‖)` as `ua` and `va` tend to `la` while `ub` and `vb` tend to `lb`. -/ lemma measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae @@ -1610,7 +1610,7 @@ lemma measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae (hub : tendsto ub lt lb) (hvb : tendsto vb lt lb) : (λ t, (∫ x in va t..vb t, f x ∂μ) - (∫ x in ua t..ub t, f x ∂μ) - (∫ x in ub t..vb t, cb ∂μ - ∫ x in ua t..va t, ca ∂μ)) =o[lt] - (λ t, ∥∫ x in ua t..va t, (1:ℝ) ∂μ∥ + ∥∫ x in ub t..vb t, (1:ℝ) ∂μ∥) := + (λ t, ‖∫ x in ua t..va t, (1:ℝ) ∂μ‖ + ‖∫ x in ub t..vb t, (1:ℝ) ∂μ‖) := begin refine ((measure_integral_sub_linear_is_o_of_tendsto_ae hmeas_a ha_lim hua hva).neg_left.add_add @@ -1703,7 +1703,7 @@ If `f` is a measurable function integrable on `a..b`, `(la, la')` is an `FTC_fil `a`, and `(lb, lb')` is an `FTC_filter` pair around `b`, and `f` has finite limits `ca` and `cb` almost surely at `la'` and `lb'`, respectively, then `(∫ x in va..vb, f x) - ∫ x in ua..ub, f x = (vb - ub) • cb - (va - ua) • ca + - o(∥va - ua∥ + ∥vb - ub∥)` as `ua` and `va` tend to `la` while `ub` and `vb` tend to `lb`. + o(‖va - ua‖ + ‖vb - ub‖)` as `ua` and `va` tend to `la` while `ub` and `vb` tend to `lb`. This lemma could've been formulated using `has_strict_fderiv_at_filter` if we had this definition. -/ @@ -1714,7 +1714,7 @@ lemma integral_sub_integral_sub_linear_is_o_of_tendsto_ae (hua : tendsto ua lt la) (hva : tendsto va lt la) (hub : tendsto ub lt lb) (hvb : tendsto vb lt lb) : (λ t, (∫ x in va t..vb t, f x) - (∫ x in ua t..ub t, f x) - - ((vb t - ub t) • cb - (va t - ua t) • ca)) =o[lt] (λ t, ∥va t - ua t∥ + ∥vb t - ub t∥) := + ((vb t - ub t) • cb - (va t - ua t) • ca)) =o[lt] (λ t, ‖va t - ua t‖ + ‖vb t - ub t‖) := by simpa [integral_const] using measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae hab hmeas_a hmeas_b ha_lim hb_lim hua hva hub hvb @@ -1722,7 +1722,7 @@ by simpa [integral_const] /-- Fundamental theorem of calculus-1, strict differentiability at filter in both endpoints. If `f` is a measurable function integrable on `a..b`, `(lb, lb')` is an `FTC_filter` pair around `b`, and `f` has a finite limit `c` almost surely at `lb'`, then -`(∫ x in a..v, f x) - ∫ x in a..u, f x = (v - u) • c + o(∥v - u∥)` as `u` and `v` tend to `lb`. +`(∫ x in a..v, f x) - ∫ x in a..u, f x = (v - u) • c + o(‖v - u‖)` as `u` and `v` tend to `lb`. This lemma could've been formulated using `has_strict_deriv_at_filter` if we had this definition. -/ lemma integral_sub_integral_sub_linear_is_o_of_tendsto_ae_right @@ -1735,7 +1735,7 @@ by simpa only [integral_const, smul_eq_mul, mul_one] using /-- Fundamental theorem of calculus-1, strict differentiability at filter in both endpoints. If `f` is a measurable function integrable on `a..b`, `(la, la')` is an `FTC_filter` pair around `a`, and `f` has a finite limit `c` almost surely at `la'`, then -`(∫ x in v..b, f x) - ∫ x in u..b, f x = -(v - u) • c + o(∥v - u∥)` as `u` and `v` tend to `la`. +`(∫ x in v..b, f x) - ∫ x in u..b, f x = -(v - u) • c + o(‖v - u‖)` as `u` and `v` tend to `la`. This lemma could've been formulated using `has_strict_deriv_at_filter` if we had this definition. -/ lemma integral_sub_integral_sub_linear_is_o_of_tendsto_ae_left @@ -2430,11 +2430,11 @@ begin { apply (ae_measurable_deriv_within_Ioi g _).congr, refine (ae_restrict_mem measurable_set_Ioo).mono (λ x hx, _), exact (hderiv x hx).deriv_within (unique_diff_within_at_Ioi _) }, - suffices H : ∫⁻ x in Ioo a b, ∥g' x∥₊ ≤ ennreal.of_real (g b - g a), + suffices H : ∫⁻ x in Ioo a b, ‖g' x‖₊ ≤ ennreal.of_real (g b - g a), from ⟨meas_g'.ae_strongly_measurable, H.trans_lt ennreal.of_real_lt_top⟩, by_contra' H, obtain ⟨f, fle, fint, hf⟩ : - ∃ (f : simple_func ℝ ℝ≥0), (∀ x, f x ≤ ∥g' x∥₊) ∧ ∫⁻ (x : ℝ) in Ioo a b, f x < ∞ + ∃ (f : simple_func ℝ ℝ≥0), (∀ x, f x ≤ ‖g' x‖₊) ∧ ∫⁻ (x : ℝ) in Ioo a b, f x < ∞ ∧ ennreal.of_real (g b - g a) < ∫⁻ (x : ℝ) in Ioo a b, f x := exists_lt_lintegral_simple_func_of_lt_lintegral H, let F : ℝ → ℝ := coe ∘ f, diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index ba0777986d50b..bbe6fda82fe77 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -1007,7 +1007,7 @@ variables {m : measurable_space α} {μ ν : measure α} @[irreducible] def lintegral {m : measurable_space α} (μ : measure α) (f : α → ℝ≥0∞) : ℝ≥0∞ := ⨆ (g : α →ₛ ℝ≥0∞) (hf : ⇑g ≤ f), g.lintegral μ -/-! In the notation for integrals, an expression like `∫⁻ x, g ∥x∥ ∂μ` will not be parsed correctly, +/-! In the notation for integrals, an expression like `∫⁻ x, g ‖x‖ ∂μ` will not be parsed correctly, and needs parentheses. We do not set the binding power of `r` to `0`, because then `∫⁻ x, f x = 0` will be parsed incorrectly. -/ notation `∫⁻` binders `, ` r:(scoped:60 f, f) ` ∂` μ:70 := lintegral μ r @@ -1229,7 +1229,7 @@ lemma set_lintegral_congr_fun {f g : α → ℝ≥0∞} {s : set α} (hs : measu by { rw lintegral_congr_ae, rw eventually_eq, rwa ae_restrict_iff' hs, } lemma lintegral_of_real_le_lintegral_nnnorm (f : α → ℝ) : - ∫⁻ x, ennreal.of_real (f x) ∂μ ≤ ∫⁻ x, ∥f x∥₊ ∂μ := + ∫⁻ x, ennreal.of_real (f x) ∂μ ≤ ∫⁻ x, ‖f x‖₊ ∂μ := begin simp_rw ← of_real_norm_eq_coe_nnnorm, refine lintegral_mono (λ x, ennreal.of_real_le_of_real _), @@ -1238,7 +1238,7 @@ begin end lemma lintegral_nnnorm_eq_of_ae_nonneg {f : α → ℝ} (h_nonneg : 0 ≤ᵐ[μ] f) : - ∫⁻ x, ∥f x∥₊ ∂μ = ∫⁻ x, ennreal.of_real (f x) ∂μ := + ∫⁻ x, ‖f x‖₊ ∂μ = ∫⁻ x, ennreal.of_real (f x) ∂μ := begin apply lintegral_congr_ae, filter_upwards [h_nonneg] with x hx, @@ -1246,7 +1246,7 @@ begin end lemma lintegral_nnnorm_eq_of_nonneg {f : α → ℝ} (h_nonneg : 0 ≤ f) : - ∫⁻ x, ∥f x∥₊ ∂μ = ∫⁻ x, ennreal.of_real (f x) ∂μ := + ∫⁻ x, ‖f x‖₊ ∂μ = ∫⁻ x, ennreal.of_real (f x) ∂μ := lintegral_nnnorm_eq_of_ae_nonneg (filter.eventually_of_forall h_nonneg) /-- Monotone convergence theorem -- sometimes called Beppo-Levi convergence. diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index 31ca3cc04c9a0..2f0621ddbeeb3 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -32,7 +32,7 @@ for set integral, see `filter.tendsto.integral_sub_linear_is_o_ae` and its corol Namely, consider a measurably generated filter `l`, a measure `μ` finite at this filter, and a function `f` that has a finite limit `c` at `l ⊓ μ.ae`. Then `∫ x in s, f x ∂μ = μ s • c + o(μ s)` as `s` tends to `l.small_sets`, i.e. for any `ε>0` there exists `t ∈ l` such that -`∥∫ x in s, f x ∂μ - μ s • c∥ ≤ ε * μ s` whenever `s ⊆ t`. We also formulate a version of this +`‖∫ x in s, f x ∂μ - μ s • c‖ ≤ ε * μ s` whenever `s ⊆ t`. We also formulate a version of this theorem for a locally finite measure `μ` and a function `f` continuous at a point `a`. ## Notation @@ -151,7 +151,7 @@ lemma of_real_set_integral_one_of_measure_ne_top {α : Type*} {m : measurable_sp ennreal.of_real (∫ x in s, (1 : ℝ) ∂μ) = μ s := calc ennreal.of_real (∫ x in s, (1 : ℝ) ∂μ) - = ennreal.of_real (∫ x in s, ∥(1 : ℝ)∥ ∂μ) : by simp only [norm_one] + = ennreal.of_real (∫ x in s, ‖(1 : ℝ)‖ ∂μ) : by simp only [norm_one] ... = ∫⁻ x in s, 1 ∂μ : begin rw of_real_integral_norm_eq_lintegral_nnnorm (integrable_on_const.2 (or.inr hs.lt_top)), @@ -176,12 +176,12 @@ lemma tendsto_set_integral_of_monotone {ι : Type*} [countable ι] [semilattice_ (h_mono : monotone s) (hfi : integrable_on f (⋃ n, s n) μ) : tendsto (λ i, ∫ a in s i, f a ∂μ) at_top (𝓝 (∫ a in (⋃ n, s n), f a ∂μ)) := begin - have hfi' : ∫⁻ x in ⋃ n, s n, ∥f x∥₊ ∂μ < ∞ := hfi.2, + have hfi' : ∫⁻ x in ⋃ n, s n, ‖f x‖₊ ∂μ < ∞ := hfi.2, set S := ⋃ i, s i, have hSm : measurable_set S := measurable_set.Union hsm, have hsub : ∀ {i}, s i ⊆ S, from subset_Union s, rw [← with_density_apply _ hSm] at hfi', - set ν := μ.with_density (λ x, ∥f x∥₊) with hν, + set ν := μ.with_density (λ x, ‖f x‖₊) with hν, refine metric.nhds_basis_closed_ball.tendsto_right_iff.2 (λ ε ε0, _), lift ε to ℝ≥0 using ε0.le, have : ∀ᶠ i in at_top, ν (s i) ∈ Icc (ν S - ε) (ν S + ε), @@ -255,11 +255,11 @@ end lemma integral_norm_eq_pos_sub_neg {f : α → ℝ} (hf : strongly_measurable f) (hfi : integrable f μ) : - ∫ x, ∥f x∥ ∂μ = ∫ x in {x | 0 ≤ f x}, f x ∂μ - ∫ x in {x | f x ≤ 0}, f x ∂μ := + ∫ x, ‖f x‖ ∂μ = ∫ x in {x | 0 ≤ f x}, f x ∂μ - ∫ x in {x | f x ≤ 0}, f x ∂μ := have h_meas : measurable_set {x | 0 ≤ f x}, from strongly_measurable_const.measurable_set_le hf, -calc ∫ x, ∥f x∥ ∂μ = ∫ x in {x | 0 ≤ f x}, ∥f x∥ ∂μ + ∫ x in {x | 0 ≤ f x}ᶜ, ∥f x∥ ∂μ : +calc ∫ x, ‖f x‖ ∂μ = ∫ x in {x | 0 ≤ f x}, ‖f x‖ ∂μ + ∫ x in {x | 0 ≤ f x}ᶜ, ‖f x‖ ∂μ : by rw ← integral_add_compl h_meas hfi.norm -... = ∫ x in {x | 0 ≤ f x}, f x ∂μ + ∫ x in {x | 0 ≤ f x}ᶜ, ∥f x∥ ∂μ : +... = ∫ x in {x | 0 ≤ f x}, f x ∂μ + ∫ x in {x | 0 ≤ f x}ᶜ, ‖f x‖ ∂μ : begin congr' 1, refine set_integral_congr h_meas (λ x hx, _), @@ -344,8 +344,8 @@ lemma set_integral_map_equiv {β} [measurable_space β] (e : α ≃ᵐ β) (f : e.measurable_embedding.set_integral_map f s lemma norm_set_integral_le_of_norm_le_const_ae {C : ℝ} (hs : μ s < ∞) - (hC : ∀ᵐ x ∂μ.restrict s, ∥f x∥ ≤ C) : - ∥∫ x in s, f x ∂μ∥ ≤ C * (μ s).to_real := + (hC : ∀ᵐ x ∂μ.restrict s, ‖f x‖ ≤ C) : + ‖∫ x in s, f x ∂μ‖ ≤ C * (μ s).to_real := begin rw ← measure.restrict_apply_univ at *, haveI : is_finite_measure (μ.restrict s) := ⟨‹_›⟩, @@ -353,33 +353,33 @@ begin end lemma norm_set_integral_le_of_norm_le_const_ae' {C : ℝ} (hs : μ s < ∞) - (hC : ∀ᵐ x ∂μ, x ∈ s → ∥f x∥ ≤ C) (hfm : ae_strongly_measurable f (μ.restrict s)) : - ∥∫ x in s, f x ∂μ∥ ≤ C * (μ s).to_real := + (hC : ∀ᵐ x ∂μ, x ∈ s → ‖f x‖ ≤ C) (hfm : ae_strongly_measurable f (μ.restrict s)) : + ‖∫ x in s, f x ∂μ‖ ≤ C * (μ s).to_real := begin apply norm_set_integral_le_of_norm_le_const_ae hs, - have A : ∀ᵐ (x : α) ∂μ, x ∈ s → ∥ae_strongly_measurable.mk f hfm x∥ ≤ C, + have A : ∀ᵐ (x : α) ∂μ, x ∈ s → ‖ae_strongly_measurable.mk f hfm x‖ ≤ C, { filter_upwards [hC, hfm.ae_mem_imp_eq_mk] with _ h1 h2 h3, rw [← h2 h3], exact h1 h3 }, - have B : measurable_set {x | ∥(hfm.mk f) x∥ ≤ C} := + have B : measurable_set {x | ‖(hfm.mk f) x‖ ≤ C} := hfm.strongly_measurable_mk.norm.measurable measurable_set_Iic, filter_upwards [hfm.ae_eq_mk, (ae_restrict_iff B).2 A] with _ h1 _, rwa h1, end lemma norm_set_integral_le_of_norm_le_const_ae'' {C : ℝ} (hs : μ s < ∞) (hsm : measurable_set s) - (hC : ∀ᵐ x ∂μ, x ∈ s → ∥f x∥ ≤ C) : - ∥∫ x in s, f x ∂μ∥ ≤ C * (μ s).to_real := + (hC : ∀ᵐ x ∂μ, x ∈ s → ‖f x‖ ≤ C) : + ‖∫ x in s, f x ∂μ‖ ≤ C * (μ s).to_real := norm_set_integral_le_of_norm_le_const_ae hs $ by rwa [ae_restrict_eq hsm, eventually_inf_principal] lemma norm_set_integral_le_of_norm_le_const {C : ℝ} (hs : μ s < ∞) - (hC : ∀ x ∈ s, ∥f x∥ ≤ C) (hfm : ae_strongly_measurable f (μ.restrict s)) : - ∥∫ x in s, f x ∂μ∥ ≤ C * (μ s).to_real := + (hC : ∀ x ∈ s, ‖f x‖ ≤ C) (hfm : ae_strongly_measurable f (μ.restrict s)) : + ‖∫ x in s, f x ∂μ‖ ≤ C * (μ s).to_real := norm_set_integral_le_of_norm_le_const_ae' hs (eventually_of_forall hC) hfm lemma norm_set_integral_le_of_norm_le_const' {C : ℝ} (hs : μ s < ∞) (hsm : measurable_set s) - (hC : ∀ x ∈ s, ∥f x∥ ≤ C) : - ∥∫ x in s, f x ∂μ∥ ≤ C * (μ s).to_real := + (hC : ∀ x ∈ s, ‖f x‖ ≤ C) : + ‖∫ x in s, f x ∂μ‖ ≤ C * (μ s).to_real := norm_set_integral_le_of_norm_le_const_ae'' hs hsm $ eventually_of_forall hC lemma set_integral_eq_zero_iff_of_nonneg_ae {f : α → ℝ} (hf : 0 ≤ᵐ[μ.restrict s] f) @@ -552,7 +552,7 @@ lemma _root_.antitone.tendsto_set_integral (hsm : ∀ i, measurable_set (s i)) (h_anti : antitone s) (hfi : integrable_on f (s 0) μ) : tendsto (λi, ∫ a in s i, f a ∂μ) at_top (𝓝 (∫ a in (⋂ n, s n), f a ∂μ)) := begin - let bound : α → ℝ := indicator (s 0) (λ a, ∥f a∥), + let bound : α → ℝ := indicator (s 0) (λ a, ‖f a‖), have h_int_eq : (λ i, ∫ a in s i, f a ∂μ) = (λ i, ∫ a, (s i).indicator f a ∂μ), from funext (λ i, (integral_indicator (hsm i)).symm), rw h_int_eq, @@ -612,7 +612,7 @@ end /-- For `f : Lp E p μ`, we can define an element of `Lp E p (μ.restrict s)` by `(Lp.mem_ℒp f).restrict s).to_Lp f`. This map is non-expansive. -/ lemma norm_Lp_to_Lp_restrict_le (s : set α) (f : Lp E p μ) : - ∥((Lp.mem_ℒp f).restrict s).to_Lp f∥ ≤ ∥f∥ := + ‖((Lp.mem_ℒp f).restrict s).to_Lp f‖ ≤ ‖f‖ := begin rw [Lp.norm_def, Lp.norm_def, ennreal.to_real_le_to_real (Lp.snorm_ne_top _) (Lp.snorm_ne_top _)], refine (le_of_eq _).trans (snorm_mono_measure _ measure.restrict_le_self), diff --git a/src/measure_theory/integral/set_to_l1.lean b/src/measure_theory/integral/set_to_l1.lean index 30d3fb6be1626..2ef6c317acb9a 100644 --- a/src/measure_theory/integral/set_to_l1.lean +++ b/src/measure_theory/integral/set_to_l1.lean @@ -24,7 +24,7 @@ expectation of an integrable function in `measure_theory.function.conditional_ex - `fin_meas_additive μ T`: the property that `T` is additive on measurable sets with finite measure. For two such sets, `s ∩ t = ∅ → T (s ∪ t) = T s + T t`. -- `dominated_fin_meas_additive μ T C`: `fin_meas_additive μ T ∧ ∀ s, ∥T s∥ ≤ C * (μ s).to_real`. +- `dominated_fin_meas_additive μ T C`: `fin_meas_additive μ T ∧ ∀ s, ‖T s‖ ≤ C * (μ s).to_real`. This is the property needed to perform the extension from indicators to L1. - `set_to_L1 (hT : dominated_fin_meas_additive μ T C) : (α →₁[μ] E) →L[ℝ] F`: the extension of `T` from indicators to L1. @@ -182,7 +182,7 @@ end fin_meas_additive set (up to a multiplicative constant). -/ def dominated_fin_meas_additive {β} [seminormed_add_comm_group β] {m : measurable_space α} (μ : measure α) (T : set α → β) (C : ℝ) : Prop := -fin_meas_additive μ T ∧ ∀ s, measurable_set s → μ s < ∞ → ∥T s∥ ≤ C * (μ s).to_real +fin_meas_additive μ T ∧ ∀ s, measurable_set s → μ s < ∞ → ‖T s‖ ≤ C * (μ s).to_real namespace dominated_fin_meas_additive @@ -222,7 +222,7 @@ end lemma smul [normed_field 𝕜] [normed_space 𝕜 β] (hT : dominated_fin_meas_additive μ T C) (c : 𝕜) : - dominated_fin_meas_additive μ (λ s, c • (T s)) (∥c∥ * C) := + dominated_fin_meas_additive μ (λ s, c • (T s)) (‖c‖ * C) := begin refine ⟨hT.1.smul c, λ s hs hμs, _⟩, dsimp only, @@ -588,29 +588,29 @@ end order lemma norm_set_to_simple_func_le_sum_op_norm {m : measurable_space α} (T : set α → F' →L[ℝ] F) (f : α →ₛ F') : - ∥f.set_to_simple_func T∥ ≤ ∑ x in f.range, ∥T (f ⁻¹' {x})∥ * ∥x∥ := -calc ∥∑ x in f.range, T (f ⁻¹' {x}) x∥ - ≤ ∑ x in f.range, ∥T (f ⁻¹' {x}) x∥ : norm_sum_le _ _ -... ≤ ∑ x in f.range, ∥T (f ⁻¹' {x})∥ * ∥x∥ : + ‖f.set_to_simple_func T‖ ≤ ∑ x in f.range, ‖T (f ⁻¹' {x})‖ * ‖x‖ := +calc ‖∑ x in f.range, T (f ⁻¹' {x}) x‖ + ≤ ∑ x in f.range, ‖T (f ⁻¹' {x}) x‖ : norm_sum_le _ _ +... ≤ ∑ x in f.range, ‖T (f ⁻¹' {x})‖ * ‖x‖ : by { refine finset.sum_le_sum (λb hb, _), simp_rw continuous_linear_map.le_op_norm, } lemma norm_set_to_simple_func_le_sum_mul_norm (T : set α → F →L[ℝ] F') {C : ℝ} - (hT_norm : ∀ s, measurable_set s → ∥T s∥ ≤ C * (μ s).to_real) (f : α →ₛ F) : - ∥f.set_to_simple_func T∥ ≤ C * ∑ x in f.range, (μ (f ⁻¹' {x})).to_real * ∥x∥ := -calc ∥f.set_to_simple_func T∥ - ≤ ∑ x in f.range, ∥T (f ⁻¹' {x})∥ * ∥x∥ : norm_set_to_simple_func_le_sum_op_norm T f -... ≤ ∑ x in f.range, C * (μ (f ⁻¹' {x})).to_real * ∥x∥ + (hT_norm : ∀ s, measurable_set s → ‖T s‖ ≤ C * (μ s).to_real) (f : α →ₛ F) : + ‖f.set_to_simple_func T‖ ≤ C * ∑ x in f.range, (μ (f ⁻¹' {x})).to_real * ‖x‖ := +calc ‖f.set_to_simple_func T‖ + ≤ ∑ x in f.range, ‖T (f ⁻¹' {x})‖ * ‖x‖ : norm_set_to_simple_func_le_sum_op_norm T f +... ≤ ∑ x in f.range, C * (μ (f ⁻¹' {x})).to_real * ‖x‖ : sum_le_sum $ λ b hb, mul_le_mul_of_nonneg_right (hT_norm _ $ simple_func.measurable_set_fiber _ _) $ norm_nonneg _ -... ≤ C * ∑ x in f.range, (μ (f ⁻¹' {x})).to_real * ∥x∥ : by simp_rw [mul_sum, ← mul_assoc] +... ≤ C * ∑ x in f.range, (μ (f ⁻¹' {x})).to_real * ‖x‖ : by simp_rw [mul_sum, ← mul_assoc] lemma norm_set_to_simple_func_le_sum_mul_norm_of_integrable (T : set α → E →L[ℝ] F') {C : ℝ} - (hT_norm : ∀ s, measurable_set s → μ s < ∞ → ∥T s∥ ≤ C * (μ s).to_real) (f : α →ₛ E) + (hT_norm : ∀ s, measurable_set s → μ s < ∞ → ‖T s‖ ≤ C * (μ s).to_real) (f : α →ₛ E) (hf : integrable f μ) : - ∥f.set_to_simple_func T∥ ≤ C * ∑ x in f.range, (μ (f ⁻¹' {x})).to_real * ∥x∥ := -calc ∥f.set_to_simple_func T∥ - ≤ ∑ x in f.range, ∥T (f ⁻¹' {x})∥ * ∥x∥ : norm_set_to_simple_func_le_sum_op_norm T f -... ≤ ∑ x in f.range, C * (μ (f ⁻¹' {x})).to_real * ∥x∥ : + ‖f.set_to_simple_func T‖ ≤ C * ∑ x in f.range, (μ (f ⁻¹' {x})).to_real * ‖x‖ := +calc ‖f.set_to_simple_func T‖ + ≤ ∑ x in f.range, ‖T (f ⁻¹' {x})‖ * ‖x‖ : norm_set_to_simple_func_le_sum_op_norm T f +... ≤ ∑ x in f.range, C * (μ (f ⁻¹' {x})).to_real * ‖x‖ : begin refine finset.sum_le_sum (λ b hb, _), obtain rfl | hb := eq_or_ne b 0, @@ -618,7 +618,7 @@ calc ∥f.set_to_simple_func T∥ exact mul_le_mul_of_nonneg_right (hT_norm _ (simple_func.measurable_set_fiber _ _) $ simple_func.measure_preimage_lt_top_of_integrable _ hf hb) (norm_nonneg _), end -... ≤ C * ∑ x in f.range, (μ (f ⁻¹' {x})).to_real * ∥x∥ : by simp_rw [mul_sum, ← mul_assoc] +... ≤ C * ∑ x in f.range, (μ (f ⁻¹' {x})).to_real * ‖x‖ : by simp_rw [mul_sum, ← mul_assoc] lemma set_to_simple_func_indicator (T : set α → F →L[ℝ] F') (hT_empty : T ∅ = 0) {m : measurable_space α} {s : set α} (hs : measurable_set s) (x : F) : @@ -682,10 +682,10 @@ variables {α E μ} namespace simple_func lemma norm_eq_sum_mul (f : α →₁ₛ[μ] G) : - ∥f∥ = ∑ x in (to_simple_func f).range, (μ ((to_simple_func f) ⁻¹' {x})).to_real * ∥x∥ := + ‖f‖ = ∑ x in (to_simple_func f).range, (μ ((to_simple_func f) ⁻¹' {x})).to_real * ‖x‖ := begin rw [norm_to_simple_func, snorm_one_eq_lintegral_nnnorm], - have h_eq := simple_func.map_apply (λ x, (∥x∥₊ : ℝ≥0∞)) (to_simple_func f), + have h_eq := simple_func.map_apply (λ x, (‖x‖₊ : ℝ≥0∞)) (to_simple_func f), dsimp only at h_eq, simp_rw ← h_eq, rw [simple_func.lintegral_eq_lintegral, simple_func.map_lintegral, ennreal.to_real_sum], @@ -823,8 +823,8 @@ begin end lemma norm_set_to_L1s_le (T : set α → E →L[ℝ] F) {C : ℝ} - (hT_norm : ∀ s, measurable_set s → μ s < ∞ → ∥T s∥ ≤ C * (μ s).to_real) (f : α →₁ₛ[μ] E) : - ∥set_to_L1s T f∥ ≤ C * ∥f∥ := + (hT_norm : ∀ s, measurable_set s → μ s < ∞ → ‖T s‖ ≤ C * (μ s).to_real) (f : α →₁ₛ[μ] E) : + ‖set_to_L1s T f‖ ≤ C * ‖f‖ := begin rw [set_to_L1s, norm_eq_sum_mul f], exact simple_func.norm_set_to_simple_func_le_sum_mul_norm_of_integrable T hT_norm _ @@ -965,12 +965,12 @@ set_to_L1s_smul_left' T T' c h_smul f lemma norm_set_to_L1s_clm_le {T : set α → E →L[ℝ] F} {C : ℝ} (hT : dominated_fin_meas_additive μ T C) (hC : 0 ≤ C) : - ∥set_to_L1s_clm α E μ hT∥ ≤ C := + ‖set_to_L1s_clm α E μ hT‖ ≤ C := linear_map.mk_continuous_norm_le _ hC _ lemma norm_set_to_L1s_clm_le' {T : set α → E →L[ℝ] F} {C : ℝ} (hT : dominated_fin_meas_additive μ T C) : - ∥set_to_L1s_clm α E μ hT∥ ≤ max C 0 := + ‖set_to_L1s_clm α E μ hT‖ ≤ max C 0 := linear_map.mk_continuous_norm_le' _ _ lemma set_to_L1s_clm_const [is_finite_measure μ] {T : set α → E →L[ℝ] F} {C : ℝ} @@ -1246,38 +1246,38 @@ end end order lemma norm_set_to_L1_le_norm_set_to_L1s_clm (hT : dominated_fin_meas_additive μ T C) : - ∥set_to_L1 hT∥ ≤ ∥set_to_L1s_clm α E μ hT∥ := -calc ∥set_to_L1 hT∥ - ≤ (1 : ℝ≥0) * ∥set_to_L1s_clm α E μ hT∥ : begin + ‖set_to_L1 hT‖ ≤ ‖set_to_L1s_clm α E μ hT‖ := +calc ‖set_to_L1 hT‖ + ≤ (1 : ℝ≥0) * ‖set_to_L1s_clm α E μ hT‖ : begin refine continuous_linear_map.op_norm_extend_le (set_to_L1s_clm α E μ hT) (coe_to_Lp α E ℝ) (simple_func.dense_range one_ne_top) (λ x, le_of_eq _), rw [nnreal.coe_one, one_mul], refl, end -... = ∥set_to_L1s_clm α E μ hT∥ : by rw [nnreal.coe_one, one_mul] +... = ‖set_to_L1s_clm α E μ hT‖ : by rw [nnreal.coe_one, one_mul] lemma norm_set_to_L1_le_mul_norm (hT : dominated_fin_meas_additive μ T C) (hC : 0 ≤ C) (f : α →₁[μ] E) : - ∥set_to_L1 hT f∥ ≤ C * ∥f∥ := -calc ∥set_to_L1 hT f∥ - ≤ ∥set_to_L1s_clm α E μ hT∥ * ∥f∥ : + ‖set_to_L1 hT f‖ ≤ C * ‖f‖ := +calc ‖set_to_L1 hT f‖ + ≤ ‖set_to_L1s_clm α E μ hT‖ * ‖f‖ : continuous_linear_map.le_of_op_norm_le _ (norm_set_to_L1_le_norm_set_to_L1s_clm hT) _ -... ≤ C * ∥f∥ : mul_le_mul (norm_set_to_L1s_clm_le hT hC) le_rfl (norm_nonneg _) hC +... ≤ C * ‖f‖ : mul_le_mul (norm_set_to_L1s_clm_le hT hC) le_rfl (norm_nonneg _) hC lemma norm_set_to_L1_le_mul_norm' (hT : dominated_fin_meas_additive μ T C) (f : α →₁[μ] E) : - ∥set_to_L1 hT f∥ ≤ max C 0 * ∥f∥ := -calc ∥set_to_L1 hT f∥ - ≤ ∥set_to_L1s_clm α E μ hT∥ * ∥f∥ : + ‖set_to_L1 hT f‖ ≤ max C 0 * ‖f‖ := +calc ‖set_to_L1 hT f‖ + ≤ ‖set_to_L1s_clm α E μ hT‖ * ‖f‖ : continuous_linear_map.le_of_op_norm_le _ (norm_set_to_L1_le_norm_set_to_L1s_clm hT) _ -... ≤ max C 0 * ∥f∥ : +... ≤ max C 0 * ‖f‖ : mul_le_mul (norm_set_to_L1s_clm_le' hT) le_rfl (norm_nonneg _) (le_max_right _ _) lemma norm_set_to_L1_le (hT : dominated_fin_meas_additive μ T C) (hC : 0 ≤ C) : - ∥set_to_L1 hT∥ ≤ C := + ‖set_to_L1 hT‖ ≤ C := continuous_linear_map.op_norm_le_bound _ hC (norm_set_to_L1_le_mul_norm hT hC) lemma norm_set_to_L1_le' (hT : dominated_fin_meas_additive μ T C) : - ∥set_to_L1 hT∥ ≤ max C 0 := + ‖set_to_L1 hT‖ ≤ max C 0 := continuous_linear_map.op_norm_le_bound _ (le_max_right _ _) (norm_set_to_L1_le_mul_norm' hT) lemma set_to_L1_lipschitz (hT : dominated_fin_meas_additive μ T C) : @@ -1565,7 +1565,7 @@ by { simp_rw L1.set_to_fun_eq_set_to_L1 hT, exact continuous_linear_map.continuo lemma tendsto_set_to_fun_of_L1 (hT : dominated_fin_meas_additive μ T C) {ι} (f : α → E) (hfi : integrable f μ) {fs : ι → α → E} {l : filter ι} (hfsi : ∀ᶠ i in l, integrable (fs i) μ) - (hfs : tendsto (λ i, ∫⁻ x, ∥fs i x - f x∥₊ ∂μ) l (𝓝 0)) : + (hfs : tendsto (λ i, ∫⁻ x, ‖fs i x - f x‖₊ ∂μ) l (𝓝 0)) : tendsto (λ i, set_to_fun μ T hT (fs i)) l (𝓝 $ set_to_fun μ T hT f) := begin classical, @@ -1752,20 +1752,20 @@ end lemma norm_set_to_fun_le_mul_norm (hT : dominated_fin_meas_additive μ T C) (f : α →₁[μ] E) (hC : 0 ≤ C) : - ∥set_to_fun μ T hT f∥ ≤ C * ∥f∥ := + ‖set_to_fun μ T hT f‖ ≤ C * ‖f‖ := by { rw L1.set_to_fun_eq_set_to_L1, exact L1.norm_set_to_L1_le_mul_norm hT hC f, } lemma norm_set_to_fun_le_mul_norm' (hT : dominated_fin_meas_additive μ T C) (f : α →₁[μ] E) : - ∥set_to_fun μ T hT f∥ ≤ max C 0 * ∥f∥ := + ‖set_to_fun μ T hT f‖ ≤ max C 0 * ‖f‖ := by { rw L1.set_to_fun_eq_set_to_L1, exact L1.norm_set_to_L1_le_mul_norm' hT f, } lemma norm_set_to_fun_le (hT : dominated_fin_meas_additive μ T C) (hf : integrable f μ) (hC : 0 ≤ C) : - ∥set_to_fun μ T hT f∥ ≤ C * ∥hf.to_L1 f∥ := + ‖set_to_fun μ T hT f‖ ≤ C * ‖hf.to_L1 f‖ := by { rw set_to_fun_eq hT hf, exact L1.norm_set_to_L1_le_mul_norm hT hC _, } lemma norm_set_to_fun_le' (hT : dominated_fin_meas_additive μ T C) (hf : integrable f μ) : - ∥set_to_fun μ T hT f∥ ≤ max C 0 * ∥hf.to_L1 f∥ := + ‖set_to_fun μ T hT f‖ ≤ max C 0 * ‖hf.to_L1 f‖ := by { rw set_to_fun_eq hT hf, exact L1.norm_set_to_L1_le_mul_norm' hT _, } /-- Lebesgue dominated convergence theorem provides sufficient conditions under which almost @@ -1777,7 +1777,7 @@ by { rw set_to_fun_eq hT hf, exact L1.norm_set_to_L1_le_mul_norm' hT _, } theorem tendsto_set_to_fun_of_dominated_convergence (hT : dominated_fin_meas_additive μ T C) {fs : ℕ → α → E} {f : α → E} (bound : α → ℝ) (fs_measurable : ∀ n, ae_strongly_measurable (fs n) μ) - (bound_integrable : integrable bound μ) (h_bound : ∀ n, ∀ᵐ a ∂μ, ∥fs n a∥ ≤ bound a) + (bound_integrable : integrable bound μ) (h_bound : ∀ n, ∀ᵐ a ∂μ, ‖fs n a‖ ≤ bound a) (h_lim : ∀ᵐ a ∂μ, tendsto (λ n, fs n a) at_top (𝓝 (f a))) : tendsto (λ n, set_to_fun μ T hT (fs n)) at_top (𝓝 $ set_to_fun μ T hT f) := begin @@ -1801,7 +1801,7 @@ begin /- up to some rewriting, what we need to prove is `h_lim` -/ rw tendsto_iff_norm_tendsto_zero, have lintegral_norm_tendsto_zero : - tendsto (λn, ennreal.to_real $ ∫⁻ a, (ennreal.of_real ∥fs n a - f a∥) ∂μ) at_top (𝓝 0) := + tendsto (λn, ennreal.to_real $ ∫⁻ a, (ennreal.of_real ‖fs n a - f a‖) ∂μ) at_top (𝓝 0) := (tendsto_to_real zero_ne_top).comp (tendsto_lintegral_norm_of_dominated_convergence fs_measurable bound_integrable.has_finite_integral h_bound h_lim), @@ -1821,7 +1821,7 @@ lemma tendsto_set_to_fun_filter_of_dominated_convergence (hT : dominated_fin_mea {ι} {l : _root_.filter ι} [l.is_countably_generated] {fs : ι → α → E} {f : α → E} (bound : α → ℝ) (hfs_meas : ∀ᶠ n in l, ae_strongly_measurable (fs n) μ) - (h_bound : ∀ᶠ n in l, ∀ᵐ a ∂μ, ∥fs n a∥ ≤ bound a) + (h_bound : ∀ᶠ n in l, ∀ᵐ a ∂μ, ‖fs n a‖ ≤ bound a) (bound_integrable : integrable bound μ) (h_lim : ∀ᵐ a ∂μ, tendsto (λ n, fs n a) l (𝓝 (f a))) : tendsto (λ n, set_to_fun μ T hT (fs n)) l (𝓝 $ set_to_fun μ T hT f) := @@ -1830,7 +1830,7 @@ begin intros x xl, have hxl : ∀ s ∈ l, ∃ a, ∀ b ≥ a, x b ∈ s, by { rwa tendsto_at_top' at xl, }, have h : {x : ι | (λ n, ae_strongly_measurable (fs n) μ) x} - ∩ {x : ι | (λ n, ∀ᵐ a ∂μ, ∥fs n a∥ ≤ bound a) x} ∈ l, + ∩ {x : ι | (λ n, ∀ᵐ a ∂μ, ‖fs n a‖ ≤ bound a) x} ∈ l, from inter_mem hfs_meas h_bound, obtain ⟨k, h⟩ := hxl _ h, rw ← tendsto_add_at_top_iff_nat k, @@ -1848,14 +1848,14 @@ variables {X : Type*} [topological_space X] [first_countable_topology X] lemma continuous_at_set_to_fun_of_dominated (hT : dominated_fin_meas_additive μ T C) {fs : X → α → E} {x₀ : X} {bound : α → ℝ} (hfs_meas : ∀ᶠ x in 𝓝 x₀, ae_strongly_measurable (fs x) μ) - (h_bound : ∀ᶠ x in 𝓝 x₀, ∀ᵐ a ∂μ, ∥fs x a∥ ≤ bound a) + (h_bound : ∀ᶠ x in 𝓝 x₀, ∀ᵐ a ∂μ, ‖fs x a‖ ≤ bound a) (bound_integrable : integrable bound μ) (h_cont : ∀ᵐ a ∂μ, continuous_at (λ x, fs x a) x₀) : continuous_at (λ x, set_to_fun μ T hT (fs x)) x₀ := tendsto_set_to_fun_filter_of_dominated_convergence hT bound ‹_› ‹_› ‹_› ‹_› lemma continuous_set_to_fun_of_dominated (hT : dominated_fin_meas_additive μ T C) {fs : X → α → E} {bound : α → ℝ} - (hfs_meas : ∀ x, ae_strongly_measurable (fs x) μ) (h_bound : ∀ x, ∀ᵐ a ∂μ, ∥fs x a∥ ≤ bound a) + (hfs_meas : ∀ x, ae_strongly_measurable (fs x) μ) (h_bound : ∀ x, ∀ᵐ a ∂μ, ‖fs x a‖ ≤ bound a) (bound_integrable : integrable bound μ) (h_cont : ∀ᵐ a ∂μ, continuous (λ x, fs x a)) : continuous (λ x, set_to_fun μ T hT (fs x)) := continuous_iff_continuous_at.mpr (λ x₀, continuous_at_set_to_fun_of_dominated hT diff --git a/src/measure_theory/integral/torus_integral.lean b/src/measure_theory/integral/torus_integral.lean index d13836f451806..2d839e67cbd38 100644 --- a/src/measure_theory/integral/torus_integral.lean +++ b/src/measure_theory/integral/torus_integral.lean @@ -178,14 +178,14 @@ lemma torus_integral_const_mul (a : ℂ) (f : ℂⁿ → ℂ) (c : ℂⁿ) (R : ∯ x in T(c, R), a * f x = a * ∯ x in T(c, R), f x := torus_integral_smul a f c R -/--If for all `θ : ℝⁿ`, `∥f (torus_map c R θ)∥` is less than or equal to a constant `C : ℝ`, then -`∥∯ x in T(c, R), f x∥` is less than or equal to `(2 * π)^n * (∏ i, |R i|) * C`-/ -lemma norm_torus_integral_le_of_norm_le_const {C : ℝ} (hf : ∀ θ, ∥f (torus_map c R θ)∥ ≤ C) : - ∥∯ x in T(c, R), f x∥ ≤ (2 * π)^(n: ℕ) * (∏ i, |R i|) * C := -calc ∥∯ x in T(c, R), f x∥ ≤ (∏ i, |R i|) * C * (volume (Icc (0 : ℝⁿ) (λ _, 2 * π))).to_real : +/--If for all `θ : ℝⁿ`, `‖f (torus_map c R θ)‖` is less than or equal to a constant `C : ℝ`, then +`‖∯ x in T(c, R), f x‖` is less than or equal to `(2 * π)^n * (∏ i, |R i|) * C`-/ +lemma norm_torus_integral_le_of_norm_le_const {C : ℝ} (hf : ∀ θ, ‖f (torus_map c R θ)‖ ≤ C) : + ‖∯ x in T(c, R), f x‖ ≤ (2 * π)^(n: ℕ) * (∏ i, |R i|) * C := +calc ‖∯ x in T(c, R), f x‖ ≤ (∏ i, |R i|) * C * (volume (Icc (0 : ℝⁿ) (λ _, 2 * π))).to_real : norm_set_integral_le_of_norm_le_const' measure_Icc_lt_top measurable_set_Icc $ λ θ hθ, - ( calc ∥(∏ i : fin n, R i * exp (θ i * I) * I : ℂ) • f (torus_map c R θ)∥ - = (∏ i : fin n, |R i|) * ∥f (torus_map c R θ)∥ : by simp [norm_smul] + ( calc ‖(∏ i : fin n, R i * exp (θ i * I) * I : ℂ) • f (torus_map c R θ)‖ + = (∏ i : fin n, |R i|) * ‖f (torus_map c R θ)‖ : by simp [norm_smul] ... ≤ (∏ i : fin n, |R i|) * C : mul_le_mul_of_nonneg_left (hf _) (finset.prod_nonneg $ λ _ _, abs_nonneg _) ) ... = (2 * π)^(n: ℕ) * (∏ i, |R i|) * C : diff --git a/src/measure_theory/measure/finite_measure.lean b/src/measure_theory/measure/finite_measure.lean index 762b1f3cbd378..6a758fb86a18d 100644 --- a/src/measure_theory/measure/finite_measure.lean +++ b/src/measure_theory/measure/finite_measure.lean @@ -551,7 +551,7 @@ lemma integrable_of_bounded_continuous_to_real integrable ⇑f μ := begin refine ⟨f.continuous.measurable.ae_strongly_measurable, _⟩, - have aux : (coe : ℝ≥0 → ℝ) ∘ ⇑f.nnnorm = (λ x, ∥f x∥), + have aux : (coe : ℝ≥0 → ℝ) ∘ ⇑f.nnnorm = (λ x, ‖f x‖), { ext ω, simp only [function.comp_app, bounded_continuous_function.nnnorm_coe_fun_eq, coe_nnnorm], }, apply (has_finite_integral_iff_norm ⇑f).mpr, diff --git a/src/number_theory/number_field/embeddings.lean b/src/number_theory/number_field/embeddings.lean index 0261b284955ef..21eb1f6a2ca08 100644 --- a/src/number_theory/number_field/embeddings.lean +++ b/src/number_theory/number_field/embeddings.lean @@ -70,8 +70,8 @@ open finite_dimensional polynomial set variables {K : Type*} [field K] [number_field K] variables {A : Type*} [normed_field A] [is_alg_closed A] [normed_algebra ℚ A] -lemma coeff_bdd_of_norm_le {B : ℝ} {x : K} (h : ∀ φ : K →+* A, ∥φ x∥ ≤ B) (i : ℕ) : - ∥(minpoly ℚ x).coeff i∥ ≤ (max B 1) ^ (finrank ℚ K) * (finrank ℚ K).choose ((finrank ℚ K) / 2) := +lemma coeff_bdd_of_norm_le {B : ℝ} {x : K} (h : ∀ φ : K →+* A, ‖φ x‖ ≤ B) (i : ℕ) : + ‖(minpoly ℚ x).coeff i‖ ≤ (max B 1) ^ (finrank ℚ K) * (finrank ℚ K).choose ((finrank ℚ K) / 2) := begin have hx := is_separable.is_integral ℚ x, rw [← norm_algebra_map' A, ← coeff_map (algebra_map ℚ A)], @@ -87,7 +87,7 @@ variables (K A) /-- Let `B` be a real number. The set of algebraic integers in `K` whose conjugates are all smaller in norm than `B` is finite. -/ lemma finite_of_norm_le (B : ℝ) : - {x : K | is_integral ℤ x ∧ ∀ φ : K →+* A, ∥φ x∥ ≤ B}.finite := + {x : K | is_integral ℤ x ∧ ∀ φ : K →+* A, ‖φ x‖ ≤ B}.finite := begin let C := nat.ceil ((max B 1) ^ (finrank ℚ K) * (finrank ℚ K).choose ((finrank ℚ K) / 2)), have := bUnion_roots_finite (algebra_map ℤ K) (finrank ℚ K) (finite_Icc (-C : ℤ) C), @@ -103,7 +103,7 @@ end /-- An algebraic integer whose conjugates are all of norm one is a root of unity. -/ lemma pow_eq_one_of_norm_eq_one {x : K} - (hxi : is_integral ℤ x) (hx : ∀ φ : K →+* A, ∥φ x∥ = 1) : + (hxi : is_integral ℤ x) (hx : ∀ φ : K →+* A, ‖φ x‖ = 1) : ∃ (n : ℕ) (hn : 0 < n), x ^ n = 1 := begin obtain ⟨a, -, b, -, habne, h⟩ := @set.infinite.exists_ne_map_eq_of_maps_to _ _ _ _ diff --git a/src/number_theory/padics/hensel.lean b/src/number_theory/padics/hensel.lean index 2651cfcbe2baa..88d82eda6ebd2 100644 --- a/src/number_theory/padics/hensel.lean +++ b/src/number_theory/padics/hensel.lean @@ -38,11 +38,11 @@ open_locale classical topological_space -- We begin with some general lemmas that are used below in the computation. lemma padic_polynomial_dist {p : ℕ} [fact p.prime] (F : polynomial ℤ_[p]) (x y : ℤ_[p]) : - ∥F.eval x - F.eval y∥ ≤ ∥x - y∥ := + ‖F.eval x - F.eval y‖ ≤ ‖x - y‖ := let ⟨z, hz⟩ := F.eval_sub_factor x y in calc - ∥F.eval x - F.eval y∥ = ∥z∥ * ∥x - y∥ : by simp [hz] - ... ≤ 1 * ∥x - y∥ : mul_le_mul_of_nonneg_right (padic_int.norm_le_one _) (norm_nonneg _) - ... = ∥x - y∥ : by simp + ‖F.eval x - F.eval y‖ = ‖z‖ * ‖x - y‖ : by simp [hz] + ... ≤ 1 * ‖x - y‖ : mul_le_mul_of_nonneg_right (padic_int.norm_le_one _) (norm_nonneg _) + ... = ‖x - y‖ : by simp open filter metric @@ -53,25 +53,25 @@ F.continuous_at.tendsto.comp ncs.tendsto_limit section parameters {p : ℕ} [fact p.prime] {ncs : cau_seq ℤ_[p] norm} {F : polynomial ℤ_[p]} {a : ℤ_[p]} - (ncs_der_val : ∀ n, ∥F.derivative.eval (ncs n)∥ = ∥F.derivative.eval a∥) + (ncs_der_val : ∀ n, ‖F.derivative.eval (ncs n)‖ = ‖F.derivative.eval a‖) include ncs_der_val private lemma ncs_tendsto_const : - tendsto (λ i, ∥F.derivative.eval (ncs i)∥) at_top (𝓝 ∥F.derivative.eval a∥) := + tendsto (λ i, ‖F.derivative.eval (ncs i)‖) at_top (𝓝 ‖F.derivative.eval a‖) := by convert tendsto_const_nhds; ext; rw ncs_der_val private lemma ncs_tendsto_lim : - tendsto (λ i, ∥F.derivative.eval (ncs i)∥) at_top (𝓝 (∥F.derivative.eval ncs.lim∥)) := + tendsto (λ i, ‖F.derivative.eval (ncs i)‖) at_top (𝓝 (‖F.derivative.eval ncs.lim‖)) := tendsto.comp (continuous_iff_continuous_at.1 continuous_norm _) (comp_tendsto_lim _) -private lemma norm_deriv_eq : ∥F.derivative.eval ncs.lim∥ = ∥F.derivative.eval a∥ := +private lemma norm_deriv_eq : ‖F.derivative.eval ncs.lim‖ = ‖F.derivative.eval a‖ := tendsto_nhds_unique ncs_tendsto_lim ncs_tendsto_const end section parameters {p : ℕ} [fact p.prime] {ncs : cau_seq ℤ_[p] norm} {F : polynomial ℤ_[p]} - (hnorm : tendsto (λ i, ∥F.eval (ncs i)∥) at_top (𝓝 0)) + (hnorm : tendsto (λ i, ‖F.eval (ncs i)‖) at_top (𝓝 0)) include hnorm private lemma tendsto_zero_of_norm_tendsto_zero : tendsto (λ i, F.eval (ncs i)) at_top (𝓝 0) := @@ -86,26 +86,26 @@ section hensel open nat parameters {p : ℕ} [fact p.prime] {F : polynomial ℤ_[p]} {a : ℤ_[p]} - (hnorm : ∥F.eval a∥ < ∥F.derivative.eval a∥^2) (hnsol : F.eval a ≠ 0) + (hnorm : ‖F.eval a‖ < ‖F.derivative.eval a‖^2) (hnsol : F.eval a ≠ 0) include hnorm /-- `T` is an auxiliary value that is used to control the behavior of the polynomial `F`. -/ -private def T : ℝ := ∥(F.eval a / (F.derivative.eval a)^2 : ℚ_[p])∥ +private def T : ℝ := ‖(F.eval a / (F.derivative.eval a)^2 : ℚ_[p])‖ -private lemma deriv_sq_norm_pos : 0 < ∥F.derivative.eval a∥ ^ 2 := +private lemma deriv_sq_norm_pos : 0 < ‖F.derivative.eval a‖ ^ 2 := lt_of_le_of_lt (norm_nonneg _) hnorm -private lemma deriv_sq_norm_ne_zero : ∥F.derivative.eval a∥^2 ≠ 0 := ne_of_gt deriv_sq_norm_pos +private lemma deriv_sq_norm_ne_zero : ‖F.derivative.eval a‖^2 ≠ 0 := ne_of_gt deriv_sq_norm_pos -private lemma deriv_norm_ne_zero : ∥F.derivative.eval a∥ ≠ 0 := +private lemma deriv_norm_ne_zero : ‖F.derivative.eval a‖ ≠ 0 := λ h, deriv_sq_norm_ne_zero (by simp [*, sq]) -private lemma deriv_norm_pos : 0 < ∥F.derivative.eval a∥ := +private lemma deriv_norm_pos : 0 < ‖F.derivative.eval a‖ := lt_of_le_of_ne (norm_nonneg _) (ne.symm deriv_norm_ne_zero) private lemma deriv_ne_zero : F.derivative.eval a ≠ 0 := mt norm_eq_zero.2 deriv_norm_ne_zero -private lemma T_def : T = ∥F.eval a∥ / ∥F.derivative.eval a∥^2 := +private lemma T_def : T = ‖F.eval a‖ / ‖F.derivative.eval a‖^2 := by simp [T, ← padic_int.norm_def] private lemma T_lt_one : T < 1 := @@ -122,44 +122,44 @@ private lemma T_pow' (n : ℕ) : T ^ (2 ^ n) < 1 := T_pow (pow_ne_zero _ two_ne_ /-- We will construct a sequence of elements of ℤ_p satisfying successive values of `ih`. -/ private def ih (n : ℕ) (z : ℤ_[p]) : Prop := -∥F.derivative.eval z∥ = ∥F.derivative.eval a∥ ∧ ∥F.eval z∥ ≤ ∥F.derivative.eval a∥^2 * T ^ (2^n) +‖F.derivative.eval z‖ = ‖F.derivative.eval a‖ ∧ ‖F.eval z‖ ≤ ‖F.derivative.eval a‖^2 * T ^ (2^n) private lemma ih_0 : ih 0 a := ⟨ rfl, by simp [T_def, mul_div_cancel' _ (ne_of_gt (deriv_sq_norm_pos hnorm))] ⟩ private lemma calc_norm_le_one {n : ℕ} {z : ℤ_[p]} (hz : ih n z) : - ∥(↑(F.eval z) : ℚ_[p]) / ↑(F.derivative.eval z)∥ ≤ 1 := -calc ∥(↑(F.eval z) : ℚ_[p]) / ↑(F.derivative.eval z)∥ - = ∥(↑(F.eval z) : ℚ_[p])∥ / ∥(↑(F.derivative.eval z) : ℚ_[p])∥ : norm_div _ _ -... = ∥F.eval z∥ / ∥F.derivative.eval a∥ : by simp [hz.1] -... ≤ ∥F.derivative.eval a∥^2 * T^(2^n) / ∥F.derivative.eval a∥ : + ‖(↑(F.eval z) : ℚ_[p]) / ↑(F.derivative.eval z)‖ ≤ 1 := +calc ‖(↑(F.eval z) : ℚ_[p]) / ↑(F.derivative.eval z)‖ + = ‖(↑(F.eval z) : ℚ_[p])‖ / ‖(↑(F.derivative.eval z) : ℚ_[p])‖ : norm_div _ _ +... = ‖F.eval z‖ / ‖F.derivative.eval a‖ : by simp [hz.1] +... ≤ ‖F.derivative.eval a‖^2 * T^(2^n) / ‖F.derivative.eval a‖ : (div_le_div_right deriv_norm_pos).2 hz.2 -... = ∥F.derivative.eval a∥ * T^(2^n) : div_sq_cancel _ _ +... = ‖F.derivative.eval a‖ * T^(2^n) : div_sq_cancel _ _ ... ≤ 1 : mul_le_one (padic_int.norm_le_one _) (T_pow_nonneg _) (le_of_lt (T_pow' _)) private lemma calc_deriv_dist {z z' z1 : ℤ_[p]} (hz' : z' = z - z1) - (hz1 : ∥z1∥ = ∥F.eval z∥ / ∥F.derivative.eval a∥) {n} (hz : ih n z) : - ∥F.derivative.eval z' - F.derivative.eval z∥ < ∥F.derivative.eval a∥ := + (hz1 : ‖z1‖ = ‖F.eval z‖ / ‖F.derivative.eval a‖) {n} (hz : ih n z) : + ‖F.derivative.eval z' - F.derivative.eval z‖ < ‖F.derivative.eval a‖ := calc - ∥F.derivative.eval z' - F.derivative.eval z∥ - ≤ ∥z' - z∥ : padic_polynomial_dist _ _ _ -... = ∥z1∥ : by simp only [sub_eq_add_neg, add_assoc, hz', add_add_neg_cancel'_right, norm_neg] -... = ∥F.eval z∥ / ∥F.derivative.eval a∥ : hz1 -... ≤ ∥F.derivative.eval a∥^2 * T^(2^n) / ∥F.derivative.eval a∥ : + ‖F.derivative.eval z' - F.derivative.eval z‖ + ≤ ‖z' - z‖ : padic_polynomial_dist _ _ _ +... = ‖z1‖ : by simp only [sub_eq_add_neg, add_assoc, hz', add_add_neg_cancel'_right, norm_neg] +... = ‖F.eval z‖ / ‖F.derivative.eval a‖ : hz1 +... ≤ ‖F.derivative.eval a‖^2 * T^(2^n) / ‖F.derivative.eval a‖ : (div_le_div_right deriv_norm_pos).2 hz.2 -... = ∥F.derivative.eval a∥ * T^(2^n) : div_sq_cancel _ _ -... < ∥F.derivative.eval a∥ : +... = ‖F.derivative.eval a‖ * T^(2^n) : div_sq_cancel _ _ +... < ‖F.derivative.eval a‖ : (mul_lt_iff_lt_one_right deriv_norm_pos).2 (T_pow' _) private def calc_eval_z' {z z' z1 : ℤ_[p]} (hz' : z' = z - z1) {n} (hz : ih n z) - (h1 : ∥(↑(F.eval z) : ℚ_[p]) / ↑(F.derivative.eval z)∥ ≤ 1) (hzeq : z1 = ⟨_, h1⟩) : + (h1 : ‖(↑(F.eval z) : ℚ_[p]) / ↑(F.derivative.eval z)‖ ≤ 1) (hzeq : z1 = ⟨_, h1⟩) : {q : ℤ_[p] // F.eval z' = q * z1^2} := begin have hdzne : F.derivative.eval z ≠ 0 := mt norm_eq_zero.2 (by rw hz.1; apply deriv_norm_ne_zero; assumption), have hdzne' : (↑(F.derivative.eval z) : ℚ_[p]) ≠ 0 := λ h, hdzne (subtype.ext_iff_val.2 h), obtain ⟨q, hq⟩ := F.binom_expansion z (-z1), - have : ∥(↑(F.derivative.eval z) * (↑(F.eval z) / ↑(F.derivative.eval z)) : ℚ_[p])∥ ≤ 1, + have : ‖(↑(F.derivative.eval z) * (↑(F.eval z) / ↑(F.derivative.eval z)) : ℚ_[p])‖ ≤ 1, { rw padic_norm_e.mul, exact mul_le_one (padic_int.norm_le_one _) (norm_nonneg _) h1 }, have : F.derivative.eval z * (-z1) = -F.eval z, { calc F.derivative.eval z * (-z1) @@ -172,39 +172,39 @@ begin end private def calc_eval_z'_norm {z z' z1 : ℤ_[p]} {n} (hz : ih n z) {q} - (heq : F.eval z' = q * z1^2) (h1 : ∥(↑(F.eval z) : ℚ_[p]) / ↑(F.derivative.eval z)∥ ≤ 1) - (hzeq : z1 = ⟨_, h1⟩) : ∥F.eval z'∥ ≤ ∥F.derivative.eval a∥^2 * T^(2^(n+1)) := -calc ∥F.eval z'∥ - = ∥q∥ * ∥z1∥^2 : by simp [heq] -... ≤ 1 * ∥z1∥^2 : + (heq : F.eval z' = q * z1^2) (h1 : ‖(↑(F.eval z) : ℚ_[p]) / ↑(F.derivative.eval z)‖ ≤ 1) + (hzeq : z1 = ⟨_, h1⟩) : ‖F.eval z'‖ ≤ ‖F.derivative.eval a‖^2 * T^(2^(n+1)) := +calc ‖F.eval z'‖ + = ‖q‖ * ‖z1‖^2 : by simp [heq] +... ≤ 1 * ‖z1‖^2 : mul_le_mul_of_nonneg_right (padic_int.norm_le_one _) (pow_nonneg (norm_nonneg _) _) -... = ∥F.eval z∥^2 / ∥F.derivative.eval a∥^2 : +... = ‖F.eval z‖^2 / ‖F.derivative.eval a‖^2 : by simp [hzeq, hz.1, div_pow] -... ≤ (∥F.derivative.eval a∥^2 * T^(2^n))^2 / ∥F.derivative.eval a∥^2 : +... ≤ (‖F.derivative.eval a‖^2 * T^(2^n))^2 / ‖F.derivative.eval a‖^2 : (div_le_div_right deriv_sq_norm_pos).2 (pow_le_pow_of_le_left (norm_nonneg _) hz.2 _) -... = (∥F.derivative.eval a∥^2)^2 * (T^(2^n))^2 / ∥F.derivative.eval a∥^2 : by simp only [mul_pow] -... = ∥F.derivative.eval a∥^2 * (T^(2^n))^2 : div_sq_cancel _ _ -... = ∥F.derivative.eval a∥^2 * T^(2^(n + 1)) : by rw [←pow_mul, pow_succ' 2] +... = (‖F.derivative.eval a‖^2)^2 * (T^(2^n))^2 / ‖F.derivative.eval a‖^2 : by simp only [mul_pow] +... = ‖F.derivative.eval a‖^2 * (T^(2^n))^2 : div_sq_cancel _ _ +... = ‖F.derivative.eval a‖^2 * T^(2^(n + 1)) : by rw [←pow_mul, pow_succ' 2] set_option eqn_compiler.zeta true /-- Given `z : ℤ_[p]` satisfying `ih n z`, construct `z' : ℤ_[p]` satisfying `ih (n+1) z'`. We need the hypothesis `ih n z`, since otherwise `z'` is not necessarily an integer. -/ private def ih_n {n : ℕ} {z : ℤ_[p]} (hz : ih n z) : {z' : ℤ_[p] // ih (n+1) z'} := -have h1 : ∥(↑(F.eval z) : ℚ_[p]) / ↑(F.derivative.eval z)∥ ≤ 1, from calc_norm_le_one hz, +have h1 : ‖(↑(F.eval z) : ℚ_[p]) / ↑(F.derivative.eval z)‖ ≤ 1, from calc_norm_le_one hz, let z1 : ℤ_[p] := ⟨_, h1⟩, z' : ℤ_[p] := z - z1 in ⟨ z', - have hdist : ∥F.derivative.eval z' - F.derivative.eval z∥ < ∥F.derivative.eval a∥, + have hdist : ‖F.derivative.eval z' - F.derivative.eval z‖ < ‖F.derivative.eval a‖, from calc_deriv_dist rfl (by simp [z1, hz.1]) hz, - have hfeq : ∥F.derivative.eval z'∥ = ∥F.derivative.eval a∥, + have hfeq : ‖F.derivative.eval z'‖ = ‖F.derivative.eval a‖, begin rw [sub_eq_add_neg, ← hz.1, ←norm_neg (F.derivative.eval z)] at hdist, have := padic_int.norm_eq_of_norm_add_lt_right hdist, rwa [norm_neg, hz.1] at this end, let ⟨q, heq⟩ := calc_eval_z' rfl hz h1 rfl in - have hnle : ∥F.eval z'∥ ≤ ∥F.derivative.eval a∥^2 * T^(2^(n+1)), + have hnle : ‖F.eval z'‖ ≤ ‖F.derivative.eval a‖^2 * T^(2^(n+1)), from calc_eval_z'_norm hz heq h1 rfl, ⟨hfeq, hnle⟩⟩ @@ -218,26 +218,26 @@ private noncomputable def newton_seq_aux : Π n : ℕ, {z : ℤ_[p] // ih n z} private def newton_seq (n : ℕ) : ℤ_[p] := (newton_seq_aux n).1 private lemma newton_seq_deriv_norm (n : ℕ) : - ∥F.derivative.eval (newton_seq n)∥ = ∥F.derivative.eval a∥ := + ‖F.derivative.eval (newton_seq n)‖ = ‖F.derivative.eval a‖ := (newton_seq_aux n).2.1 private lemma newton_seq_norm_le (n : ℕ) : - ∥F.eval (newton_seq n)∥ ≤ ∥F.derivative.eval a∥^2 * T ^ (2^n) := + ‖F.eval (newton_seq n)‖ ≤ ‖F.derivative.eval a‖^2 * T ^ (2^n) := (newton_seq_aux n).2.2 private lemma newton_seq_norm_eq (n : ℕ) : - ∥newton_seq (n+1) - newton_seq n∥ = - ∥F.eval (newton_seq n)∥ / ∥F.derivative.eval (newton_seq n)∥ := + ‖newton_seq (n+1) - newton_seq n‖ = + ‖F.eval (newton_seq n)‖ / ‖F.derivative.eval (newton_seq n)‖ := by simp [newton_seq, newton_seq_aux, ih_n, sub_eq_add_neg, add_comm] private lemma newton_seq_succ_dist (n : ℕ) : - ∥newton_seq (n+1) - newton_seq n∥ ≤ ∥F.derivative.eval a∥ * T^(2^n) := -calc ∥newton_seq (n+1) - newton_seq n∥ - = ∥F.eval (newton_seq n)∥ / ∥F.derivative.eval (newton_seq n)∥ : newton_seq_norm_eq _ -... = ∥F.eval (newton_seq n)∥ / ∥F.derivative.eval a∥ : by rw newton_seq_deriv_norm -... ≤ ∥F.derivative.eval a∥^2 * T ^ (2^n) / ∥F.derivative.eval a∥ : + ‖newton_seq (n+1) - newton_seq n‖ ≤ ‖F.derivative.eval a‖ * T^(2^n) := +calc ‖newton_seq (n+1) - newton_seq n‖ + = ‖F.eval (newton_seq n)‖ / ‖F.derivative.eval (newton_seq n)‖ : newton_seq_norm_eq _ +... = ‖F.eval (newton_seq n)‖ / ‖F.derivative.eval a‖ : by rw newton_seq_deriv_norm +... ≤ ‖F.derivative.eval a‖^2 * T ^ (2^n) / ‖F.derivative.eval a‖ : (div_le_div_right deriv_norm_pos).2 (newton_seq_norm_le _) -... = ∥F.derivative.eval a∥ * T^(2^n) : div_sq_cancel _ _ +... = ‖F.derivative.eval a‖ * T^(2^n) : div_sq_cancel _ _ include hnsol private lemma T_pos : T > 0 := @@ -247,18 +247,18 @@ begin end private lemma newton_seq_succ_dist_weak (n : ℕ) : - ∥newton_seq (n+2) - newton_seq (n+1)∥ < ∥F.eval a∥ / ∥F.derivative.eval a∥ := + ‖newton_seq (n+2) - newton_seq (n+1)‖ < ‖F.eval a‖ / ‖F.derivative.eval a‖ := have 2 ≤ 2^(n+1), from have _, from pow_le_pow (by norm_num : 1 ≤ 2) (nat.le_add_left _ _ : 1 ≤ n + 1), by simpa using this, -calc ∥newton_seq (n+2) - newton_seq (n+1)∥ - ≤ ∥F.derivative.eval a∥ * T^(2^(n+1)) : newton_seq_succ_dist _ -... ≤ ∥F.derivative.eval a∥ * T^2 : +calc ‖newton_seq (n+2) - newton_seq (n+1)‖ + ≤ ‖F.derivative.eval a‖ * T^(2^(n+1)) : newton_seq_succ_dist _ +... ≤ ‖F.derivative.eval a‖ * T^2 : mul_le_mul_of_nonneg_left (pow_le_pow_of_le_one (norm_nonneg _) (le_of_lt T_lt_one) this) (norm_nonneg _) -... < ∥F.derivative.eval a∥ * T^1 : +... < ‖F.derivative.eval a‖ * T^1 : mul_lt_mul_of_pos_left (pow_lt_pow_of_lt_one T_pos T_lt_one (by norm_num)) deriv_norm_pos -... = ∥F.eval a∥ / ∥F.derivative.eval a∥ : +... = ‖F.eval a‖ / ‖F.derivative.eval a‖ : begin rw [T, sq, pow_one, norm_div, ←mul_div_assoc, padic_norm_e.mul], apply mul_div_mul_left, @@ -266,55 +266,55 @@ calc ∥newton_seq (n+2) - newton_seq (n+1)∥ end private lemma newton_seq_dist_aux (n : ℕ) : - ∀ k : ℕ, ∥newton_seq (n + k) - newton_seq n∥ ≤ ∥F.derivative.eval a∥ * T^(2^n) + ∀ k : ℕ, ‖newton_seq (n + k) - newton_seq n‖ ≤ ‖F.derivative.eval a‖ * T^(2^n) | 0 := by simp [T_pow_nonneg hnorm, mul_nonneg] | (k+1) := have 2^n ≤ 2^(n+k), by {apply pow_le_pow, norm_num, apply nat.le_add_right}, calc - ∥newton_seq (n + (k + 1)) - newton_seq n∥ - = ∥newton_seq ((n + k) + 1) - newton_seq n∥ : by rw add_assoc -... = ∥(newton_seq ((n + k) + 1) - newton_seq (n+k)) + (newton_seq (n+k) - newton_seq n)∥ : + ‖newton_seq (n + (k + 1)) - newton_seq n‖ + = ‖newton_seq ((n + k) + 1) - newton_seq n‖ : by rw add_assoc +... = ‖(newton_seq ((n + k) + 1) - newton_seq (n+k)) + (newton_seq (n+k) - newton_seq n)‖ : by rw ←sub_add_sub_cancel -... ≤ max (∥newton_seq ((n + k) + 1) - newton_seq (n+k)∥) (∥newton_seq (n+k) - newton_seq n∥) : +... ≤ max (‖newton_seq ((n + k) + 1) - newton_seq (n+k)‖) (‖newton_seq (n+k) - newton_seq n‖) : padic_int.nonarchimedean _ _ -... ≤ max (∥F.derivative.eval a∥ * T^(2^((n + k)))) (∥F.derivative.eval a∥ * T^(2^n)) : +... ≤ max (‖F.derivative.eval a‖ * T^(2^((n + k)))) (‖F.derivative.eval a‖ * T^(2^n)) : max_le_max (newton_seq_succ_dist _) (newton_seq_dist_aux _) -... = ∥F.derivative.eval a∥ * T^(2^n) : +... = ‖F.derivative.eval a‖ * T^(2^n) : max_eq_right $ mul_le_mul_of_nonneg_left (pow_le_pow_of_le_one (norm_nonneg _) (le_of_lt T_lt_one) this) (norm_nonneg _) private lemma newton_seq_dist {n k : ℕ} (hnk : n ≤ k) : - ∥newton_seq k - newton_seq n∥ ≤ ∥F.derivative.eval a∥ * T^(2^n) := + ‖newton_seq k - newton_seq n‖ ≤ ‖F.derivative.eval a‖ * T^(2^n) := have hex : ∃ m, k = n + m, from exists_eq_add_of_le hnk, let ⟨_, hex'⟩ := hex in by rw hex'; apply newton_seq_dist_aux; assumption private lemma newton_seq_dist_to_a : - ∀ n : ℕ, 0 < n → ∥newton_seq n - a∥ = ∥F.eval a∥ / ∥F.derivative.eval a∥ + ∀ n : ℕ, 0 < n → ‖newton_seq n - a‖ = ‖F.eval a‖ / ‖F.derivative.eval a‖ | 1 h := by simp [sub_eq_add_neg, add_assoc, newton_seq, newton_seq_aux, ih_n] | (k+2) h := - have hlt : ∥newton_seq (k+2) - newton_seq (k+1)∥ < ∥newton_seq (k+1) - a∥, + have hlt : ‖newton_seq (k+2) - newton_seq (k+1)‖ < ‖newton_seq (k+1) - a‖, by rw newton_seq_dist_to_a (k+1) (succ_pos _); apply newton_seq_succ_dist_weak; assumption, - have hne' : ∥newton_seq (k + 2) - newton_seq (k+1)∥ ≠ ∥newton_seq (k+1) - a∥, from ne_of_lt hlt, - calc ∥newton_seq (k + 2) - a∥ - = ∥(newton_seq (k + 2) - newton_seq (k+1)) + (newton_seq (k+1) - a)∥ : by rw ←sub_add_sub_cancel -... = max (∥newton_seq (k + 2) - newton_seq (k+1)∥) (∥newton_seq (k+1) - a∥) : + have hne' : ‖newton_seq (k + 2) - newton_seq (k+1)‖ ≠ ‖newton_seq (k+1) - a‖, from ne_of_lt hlt, + calc ‖newton_seq (k + 2) - a‖ + = ‖(newton_seq (k + 2) - newton_seq (k+1)) + (newton_seq (k+1) - a)‖ : by rw ←sub_add_sub_cancel +... = max (‖newton_seq (k + 2) - newton_seq (k+1)‖) (‖newton_seq (k+1) - a‖) : padic_int.norm_add_eq_max_of_ne hne' -... = ∥newton_seq (k+1) - a∥ : max_eq_right_of_lt hlt -... = ∥polynomial.eval a F∥ / ∥polynomial.eval a (polynomial.derivative F)∥ : +... = ‖newton_seq (k+1) - a‖ : max_eq_right_of_lt hlt +... = ‖polynomial.eval a F‖ / ‖polynomial.eval a (polynomial.derivative F)‖ : newton_seq_dist_to_a (k+1) (succ_pos _) -private lemma bound' : tendsto (λ n : ℕ, ∥F.derivative.eval a∥ * T^(2^n)) at_top (𝓝 0) := +private lemma bound' : tendsto (λ n : ℕ, ‖F.derivative.eval a‖ * T^(2^n)) at_top (𝓝 0) := begin - rw ←mul_zero (∥F.derivative.eval a∥), + rw ←mul_zero (‖F.derivative.eval a‖), exact tendsto_const_nhds.mul (tendsto.comp (tendsto_pow_at_top_nhds_0_of_lt_1 (norm_nonneg _) (T_lt_one hnorm)) (nat.tendsto_pow_at_top_at_top_of_one_lt (by norm_num))) end -private lemma bound : ∀ {ε}, ε > 0 → ∃ N : ℕ, ∀ {n}, n ≥ N → ∥F.derivative.eval a∥ * T^(2^n) < ε := +private lemma bound : ∀ {ε}, ε > 0 → ∃ N : ℕ, ∀ {n}, n ≥ N → ‖F.derivative.eval a‖ * T^(2^n) < ε := begin have := bound' hnorm hnsol, simp [tendsto, nhds] at this, @@ -324,9 +324,9 @@ begin simpa [abs_of_nonneg (T_nonneg _)] using hN _ hn end -private lemma bound'_sq : tendsto (λ n : ℕ, ∥F.derivative.eval a∥^2 * T^(2^n)) at_top (𝓝 0) := +private lemma bound'_sq : tendsto (λ n : ℕ, ‖F.derivative.eval a‖^2 * T^(2^n)) at_top (𝓝 0) := begin - rw [←mul_zero (∥F.derivative.eval a∥), sq], + rw [←mul_zero (‖F.derivative.eval a‖), sq], simp only [mul_assoc], apply tendsto.mul, { apply tendsto_const_nhds }, @@ -349,28 +349,28 @@ private def newton_cau_seq : cau_seq ℤ_[p] norm := ⟨_, newton_seq_is_cauchy private def soln : ℤ_[p] := newton_cau_seq.lim private lemma soln_spec {ε : ℝ} (hε : ε > 0) : - ∃ (N : ℕ), ∀ {i : ℕ}, i ≥ N → ∥soln - newton_cau_seq i∥ < ε := + ∃ (N : ℕ), ∀ {i : ℕ}, i ≥ N → ‖soln - newton_cau_seq i‖ < ε := setoid.symm (cau_seq.equiv_lim newton_cau_seq) _ hε -private lemma soln_deriv_norm : ∥F.derivative.eval soln∥ = ∥F.derivative.eval a∥ := +private lemma soln_deriv_norm : ‖F.derivative.eval soln‖ = ‖F.derivative.eval a‖ := norm_deriv_eq newton_seq_deriv_norm private lemma newton_seq_norm_tendsto_zero : - tendsto (λ i, ∥F.eval (newton_cau_seq i)∥) at_top (𝓝 0) := + tendsto (λ i, ‖F.eval (newton_cau_seq i)‖) at_top (𝓝 0) := squeeze_zero (λ _, norm_nonneg _) newton_seq_norm_le bound'_sq private lemma newton_seq_dist_tendsto : - tendsto (λ n, ∥newton_cau_seq n - a∥) at_top (𝓝 (∥F.eval a∥ / ∥F.derivative.eval a∥)) := + tendsto (λ n, ‖newton_cau_seq n - a‖) at_top (𝓝 (‖F.eval a‖ / ‖F.derivative.eval a‖)) := tendsto_const_nhds.congr' $ eventually_at_top.2 ⟨1, λ _ hx, (newton_seq_dist_to_a _ hx).symm⟩ private lemma newton_seq_dist_tendsto' : - tendsto (λ n, ∥newton_cau_seq n - a∥) at_top (𝓝 ∥soln - a∥) := + tendsto (λ n, ‖newton_cau_seq n - a‖) at_top (𝓝 ‖soln - a‖) := (continuous_norm.tendsto _).comp (newton_cau_seq.tendsto_limit.sub tendsto_const_nhds) -private lemma soln_dist_to_a : ∥soln - a∥ = ∥F.eval a∥ / ∥F.derivative.eval a∥ := +private lemma soln_dist_to_a : ‖soln - a‖ = ‖F.eval a‖ / ‖F.derivative.eval a‖ := tendsto_nhds_unique newton_seq_dist_tendsto' newton_seq_dist_tendsto -private lemma soln_dist_to_a_lt_deriv : ∥soln - a∥ < ∥F.derivative.eval a∥ := +private lemma soln_dist_to_a_lt_deriv : ‖soln - a‖ < ‖F.derivative.eval a‖ := begin rw [soln_dist_to_a, div_lt_iff], { rwa sq at hnorm }, @@ -381,12 +381,12 @@ private lemma eval_soln : F.eval soln = 0 := limit_zero_of_norm_tendsto_zero newton_seq_norm_tendsto_zero private lemma soln_unique (z : ℤ_[p]) (hev : F.eval z = 0) - (hnlt : ∥z - a∥ < ∥F.derivative.eval a∥) : + (hnlt : ‖z - a‖ < ‖F.derivative.eval a‖) : z = soln := -have soln_dist : ∥z - soln∥ < ∥F.derivative.eval a∥, from calc - ∥z - soln∥ = ∥(z - a) + (a - soln)∥ : by rw sub_add_sub_cancel - ... ≤ max (∥z - a∥) (∥a - soln∥) : padic_int.nonarchimedean _ _ - ... < ∥F.derivative.eval a∥ : max_lt hnlt (norm_sub_rev soln a ▸ soln_dist_to_a_lt_deriv), +have soln_dist : ‖z - soln‖ < ‖F.derivative.eval a‖, from calc + ‖z - soln‖ = ‖(z - a) + (a - soln)‖ : by rw sub_add_sub_cancel + ... ≤ max (‖z - a‖) (‖a - soln‖) : padic_int.nonarchimedean _ _ + ... < ‖F.derivative.eval a‖ : max_lt hnlt (norm_sub_rev soln a ▸ soln_dist_to_a_lt_deriv), let h := z - soln, ⟨q, hq⟩ := F.binom_expansion soln h in have (F.derivative.eval soln + q * h) * h = 0, from eq.symm (calc @@ -397,13 +397,13 @@ have h = 0, from by_contradiction $ λ hne, have F.derivative.eval soln + q * h = 0, from (eq_zero_or_eq_zero_of_mul_eq_zero this).resolve_right hne, have F.derivative.eval soln = (-q) * h, by simpa using eq_neg_of_add_eq_zero_left this, - lt_irrefl ∥F.derivative.eval soln∥ (calc - ∥F.derivative.eval soln∥ = ∥(-q) * h∥ : by rw this -... ≤ 1 * ∥h∥ : + lt_irrefl ‖F.derivative.eval soln‖ (calc + ‖F.derivative.eval soln‖ = ‖(-q) * h‖ : by rw this +... ≤ 1 * ‖h‖ : by { rw padic_int.norm_mul, exact mul_le_mul_of_nonneg_right (padic_int.norm_le_one _) (norm_nonneg _) } -... = ∥z - soln∥ : by simp [h] -... < ∥F.derivative.eval soln∥ : by rw soln_deriv_norm; apply soln_dist), +... = ‖z - soln‖ : by simp [h] +... < ‖F.derivative.eval soln‖ : by rw soln_deriv_norm; apply soln_dist), eq_of_sub_eq_zero (by rw ←this; refl) end hensel @@ -411,7 +411,7 @@ end hensel variables {p : ℕ} [fact p.prime] {F : polynomial ℤ_[p]} {a : ℤ_[p]} private lemma a_soln_is_unique (ha : F.eval a = 0) (z' : ℤ_[p]) (hz' : F.eval z' = 0) - (hnormz' : ∥z' - a∥ < ∥F.derivative.eval a∥) : z' = a := + (hnormz' : ‖z' - a‖ < ‖F.derivative.eval a‖) : z' = a := let h := z' - a, ⟨q, hq⟩ := F.binom_expansion a h in have (F.derivative.eval a + q * h) * h = 0, from eq.symm (calc @@ -422,23 +422,23 @@ have h = 0, from by_contradiction $ λ hne, have F.derivative.eval a + q * h = 0, from (eq_zero_or_eq_zero_of_mul_eq_zero this).resolve_right hne, have F.derivative.eval a = (-q) * h, by simpa using eq_neg_of_add_eq_zero_left this, - lt_irrefl ∥F.derivative.eval a∥ (calc - ∥F.derivative.eval a∥ = ∥q∥*∥h∥ : by simp [this] - ... ≤ 1*∥h∥ : mul_le_mul_of_nonneg_right (padic_int.norm_le_one _) (norm_nonneg _) - ... < ∥F.derivative.eval a∥ : by simpa [h]), + lt_irrefl ‖F.derivative.eval a‖ (calc + ‖F.derivative.eval a‖ = ‖q‖*‖h‖ : by simp [this] + ... ≤ 1*‖h‖ : mul_le_mul_of_nonneg_right (padic_int.norm_le_one _) (norm_nonneg _) + ... < ‖F.derivative.eval a‖ : by simpa [h]), eq_of_sub_eq_zero (by rw ←this; refl) -variable (hnorm : ∥F.eval a∥ < ∥F.derivative.eval a∥^2) +variable (hnorm : ‖F.eval a‖ < ‖F.derivative.eval a‖^2) include hnorm private lemma a_is_soln (ha : F.eval a = 0) : - F.eval a = 0 ∧ ∥a - a∥ < ∥F.derivative.eval a∥ ∧ ∥F.derivative.eval a∥ = ∥F.derivative.eval a∥ ∧ - ∀ z', F.eval z' = 0 → ∥z' - a∥ < ∥F.derivative.eval a∥ → z' = a := + F.eval a = 0 ∧ ‖a - a‖ < ‖F.derivative.eval a‖ ∧ ‖F.derivative.eval a‖ = ‖F.derivative.eval a‖ ∧ + ∀ z', F.eval z' = 0 → ‖z' - a‖ < ‖F.derivative.eval a‖ → z' = a := ⟨ha, by simp [deriv_ne_zero hnorm], rfl, a_soln_is_unique ha⟩ -lemma hensels_lemma : ∃ z : ℤ_[p], F.eval z = 0 ∧ ∥z - a∥ < ∥F.derivative.eval a∥ ∧ - ∥F.derivative.eval z∥ = ∥F.derivative.eval a∥ ∧ - ∀ z', F.eval z' = 0 → ∥z' - a∥ < ∥F.derivative.eval a∥ → z' = z := +lemma hensels_lemma : ∃ z : ℤ_[p], F.eval z = 0 ∧ ‖z - a‖ < ‖F.derivative.eval a‖ ∧ + ‖F.derivative.eval z‖ = ‖F.derivative.eval a‖ ∧ + ∀ z', F.eval z' = 0 → ‖z' - a‖ < ‖F.derivative.eval a‖ → z' = z := if ha : F.eval a = 0 then ⟨a, a_is_soln hnorm ha⟩ else by refine ⟨soln _ _, eval_soln _ _, soln_dist_to_a_lt_deriv _ _, soln_deriv_norm _ _, soln_unique _ _⟩; assumption diff --git a/src/number_theory/padics/padic_integers.lean b/src/number_theory/padics/padic_integers.lean index e46d74d97428c..cf801b2f0b892 100644 --- a/src/number_theory/padics/padic_integers.lean +++ b/src/number_theory/padics/padic_integers.lean @@ -51,7 +51,7 @@ noncomputable theory open_locale classical /-- The `p`-adic integers `ℤ_[p]` are the `p`-adic numbers with norm `≤ 1`. -/ -def padic_int (p : ℕ) [fact p.prime] := {x : ℚ_[p] // ∥x∥ ≤ 1} +def padic_int (p : ℕ) [fact p.prime] := {x : ℚ_[p] // ‖x‖ ≤ 1} notation `ℤ_[`p`]` := padic_int p @@ -69,14 +69,14 @@ variables (p) /-- The `p`-adic integers as a subring of `ℚ_[p]`. -/ def subring : subring (ℚ_[p]) := -{ carrier := {x : ℚ_[p] | ∥x∥ ≤ 1}, +{ carrier := {x : ℚ_[p] | ‖x‖ ≤ 1}, zero_mem' := by norm_num, one_mem' := by norm_num, add_mem' := λ x y hx hy, (padic_norm_e.nonarchimedean _ _).trans $ max_le_iff.2 ⟨hx, hy⟩, mul_mem' := λ x y hx hy, (padic_norm_e.mul _ _).trans_le $ mul_le_one hx (norm_nonneg _) hy, neg_mem' := λ x hx, (norm_neg _).trans_le hx } -@[simp] lemma mem_subring_iff {x : ℚ_[p]} : x ∈ subring p ↔ ∥x∥ ≤ 1 := iff.rfl +@[simp] lemma mem_subring_iff {x : ℚ_[p]} : x ∈ subring p ↔ ‖x‖ ≤ 1 := iff.rfl variables {p} @@ -133,7 +133,7 @@ def coe.ring_hom : ℤ_[p] →+* ℚ_[p] := (subring p).subtype /-- The inverse of a `p`-adic integer with norm equal to `1` is also a `p`-adic integer. Otherwise, the inverse is defined to be `0`. -/ def inv : ℤ_[p] → ℤ_[p] -| ⟨k, _⟩ := if h : ∥k∥ = 1 then ⟨k⁻¹, by simp [h]⟩ else 0 +| ⟨k, _⟩ := if h : ‖k‖ = 1 then ⟨k⁻¹, by simp [h]⟩ else 0 instance : char_zero ℤ_[p] := { cast_injective := @@ -172,14 +172,14 @@ variables (p : ℕ) [fact p.prime] instance : metric_space ℤ_[p] := subtype.metric_space instance complete_space : complete_space ℤ_[p] := -have is_closed {x : ℚ_[p] | ∥x∥ ≤ 1}, from is_closed_le continuous_norm continuous_const, +have is_closed {x : ℚ_[p] | ‖x‖ ≤ 1}, from is_closed_le continuous_norm continuous_const, this.complete_space_coe -instance : has_norm ℤ_[p] := ⟨λ z, ∥(z : ℚ_[p])∥⟩ +instance : has_norm ℤ_[p] := ⟨λ z, ‖(z : ℚ_[p])‖⟩ variables {p} -lemma norm_def {z : ℤ_[p]} : ∥z∥ = ∥(z : ℚ_[p])∥ := rfl +lemma norm_def {z : ℤ_[p]} : ‖z‖ = ‖(z : ℚ_[p])‖ := rfl variables (p) @@ -190,7 +190,7 @@ instance : normed_comm_ring ℤ_[p] := instance : norm_one_class ℤ_[p] := ⟨norm_def.trans norm_one⟩ -instance is_absolute_value : is_absolute_value (λ z : ℤ_[p], ∥z∥) := +instance is_absolute_value : is_absolute_value (λ z : ℤ_[p], ‖z‖) := { abv_nonneg := norm_nonneg, abv_eq_zero := λ ⟨_, _⟩, by simp [norm_eq_zero], abv_add := λ ⟨_,_⟩ ⟨_, _⟩, norm_add_le _ _, @@ -208,43 +208,43 @@ namespace padic_int variables {p : ℕ} [fact p.prime] -lemma norm_le_one (z : ℤ_[p]) : ∥z∥ ≤ 1 := z.2 +lemma norm_le_one (z : ℤ_[p]) : ‖z‖ ≤ 1 := z.2 -@[simp] lemma norm_mul (z1 z2 : ℤ_[p]) : ∥z1 * z2∥ = ∥z1∥ * ∥z2∥ := by simp [norm_def] +@[simp] lemma norm_mul (z1 z2 : ℤ_[p]) : ‖z1 * z2‖ = ‖z1‖ * ‖z2‖ := by simp [norm_def] -@[simp] lemma norm_pow (z : ℤ_[p]) : ∀ n : ℕ, ∥z ^ n∥ = ∥z∥ ^ n +@[simp] lemma norm_pow (z : ℤ_[p]) : ∀ n : ℕ, ‖z ^ n‖ = ‖z‖ ^ n | 0 := by simp | (k + 1) := by { rw [pow_succ, pow_succ, norm_mul], congr, apply norm_pow } -theorem nonarchimedean (q r : ℤ_[p]) : ∥q + r∥ ≤ max (∥q∥) (∥r∥) := padic_norm_e.nonarchimedean _ _ +theorem nonarchimedean (q r : ℤ_[p]) : ‖q + r‖ ≤ max (‖q‖) (‖r‖) := padic_norm_e.nonarchimedean _ _ -theorem norm_add_eq_max_of_ne {q r : ℤ_[p]} : ∥q∥ ≠ ∥r∥ → ∥q+r∥ = max (∥q∥) (∥r∥) := +theorem norm_add_eq_max_of_ne {q r : ℤ_[p]} : ‖q‖ ≠ ‖r‖ → ‖q+r‖ = max (‖q‖) (‖r‖) := padic_norm_e.add_eq_max_of_ne -lemma norm_eq_of_norm_add_lt_right {z1 z2 : ℤ_[p]} (h : ∥z1 + z2∥ < ∥z2∥) : ∥z1∥ = ∥z2∥ := +lemma norm_eq_of_norm_add_lt_right {z1 z2 : ℤ_[p]} (h : ‖z1 + z2‖ < ‖z2‖) : ‖z1‖ = ‖z2‖ := by_contradiction $ λ hne, not_lt_of_ge (by rw norm_add_eq_max_of_ne hne; apply le_max_right) h -lemma norm_eq_of_norm_add_lt_left {z1 z2 : ℤ_[p]} (h : ∥z1 + z2∥ < ∥z1∥) : ∥z1∥ = ∥z2∥ := +lemma norm_eq_of_norm_add_lt_left {z1 z2 : ℤ_[p]} (h : ‖z1 + z2‖ < ‖z1‖) : ‖z1‖ = ‖z2‖ := by_contradiction $ λ hne, not_lt_of_ge (by rw norm_add_eq_max_of_ne hne; apply le_max_left) h -@[simp] lemma padic_norm_e_of_padic_int (z : ℤ_[p]) : ∥(z : ℚ_[p])∥ = ∥z∥ := by simp [norm_def] +@[simp] lemma padic_norm_e_of_padic_int (z : ℤ_[p]) : ‖(z : ℚ_[p])‖ = ‖z‖ := by simp [norm_def] -lemma norm_int_cast_eq_padic_norm (z : ℤ) : ∥(z : ℤ_[p])∥ = ∥(z : ℚ_[p])∥ := by simp [norm_def] +lemma norm_int_cast_eq_padic_norm (z : ℤ) : ‖(z : ℤ_[p])‖ = ‖(z : ℚ_[p])‖ := by simp [norm_def] -@[simp] lemma norm_eq_padic_norm {q : ℚ_[p]} (hq : ∥q∥ ≤ 1) : @norm ℤ_[p] _ ⟨q, hq⟩ = ∥q∥ := rfl +@[simp] lemma norm_eq_padic_norm {q : ℚ_[p]} (hq : ‖q‖ ≤ 1) : @norm ℤ_[p] _ ⟨q, hq⟩ = ‖q‖ := rfl -@[simp] lemma norm_p : ∥(p : ℤ_[p])∥ = p⁻¹ := padic_norm_e.norm_p +@[simp] lemma norm_p : ‖(p : ℤ_[p])‖ = p⁻¹ := padic_norm_e.norm_p -@[simp] lemma norm_p_pow (n : ℕ) : ∥(p : ℤ_[p])^n∥ = p^(-n:ℤ) := padic_norm_e.norm_p_pow n +@[simp] lemma norm_p_pow (n : ℕ) : ‖(p : ℤ_[p])^n‖ = p^(-n:ℤ) := padic_norm_e.norm_p_pow n -private def cau_seq_to_rat_cau_seq (f : cau_seq ℤ_[p] norm) : cau_seq ℚ_[p] (λ a, ∥a∥) := +private def cau_seq_to_rat_cau_seq (f : cau_seq ℤ_[p] norm) : cau_seq ℚ_[p] (λ a, ‖a‖) := ⟨ λ n, f n, λ _ hε, by simpa [norm, norm_def] using f.cauchy hε ⟩ variables (p) instance complete : cau_seq.is_complete ℤ_[p] norm := ⟨ λ f, - have hqn : ∥cau_seq.lim (cau_seq_to_rat_cau_seq f)∥ ≤ 1, + have hqn : ‖cau_seq.lim (cau_seq_to_rat_cau_seq f)‖ ≤ 1, from padic_norm_e_lim_le zero_lt_one (λ _, norm_le_one _), ⟨⟨_, hqn⟩, λ ε, by simpa [norm, norm_def] using cau_seq.equiv_lim (cau_seq_to_rat_cau_seq f) ε⟩⟩ @@ -280,12 +280,12 @@ end variable {p} -lemma norm_int_lt_one_iff_dvd (k : ℤ) : ∥(k : ℤ_[p])∥ < 1 ↔ (p : ℤ) ∣ k := -suffices ∥(k : ℚ_[p])∥ < 1 ↔ ↑p ∣ k, by rwa norm_int_cast_eq_padic_norm, +lemma norm_int_lt_one_iff_dvd (k : ℤ) : ‖(k : ℤ_[p])‖ < 1 ↔ (p : ℤ) ∣ k := +suffices ‖(k : ℚ_[p])‖ < 1 ↔ ↑p ∣ k, by rwa norm_int_cast_eq_padic_norm, padic_norm_e.norm_int_lt_one_iff_dvd k -lemma norm_int_le_pow_iff_dvd {k : ℤ} {n : ℕ} : ∥(k : ℤ_[p])∥ ≤ p ^ (-n : ℤ) ↔ (p ^ n : ℤ) ∣ k := -suffices ∥(k : ℚ_[p])∥ ≤ p ^ (-n : ℤ) ↔ ↑(p ^ n) ∣ k, +lemma norm_int_le_pow_iff_dvd {k : ℤ} {n : ℕ} : ‖(k : ℤ_[p])‖ ≤ p ^ (-n : ℤ) ↔ (p ^ n : ℤ) ∣ k := +suffices ‖(k : ℚ_[p])‖ ≤ p ^ (-n : ℤ) ↔ ↑(p ^ n) ∣ k, by simpa [norm_int_cast_eq_padic_norm], padic_norm_e.norm_int_le_pow_iff_dvd _ _ /-! ### Valuation on `ℤ_[p]` -/ @@ -293,7 +293,7 @@ by simpa [norm_int_cast_eq_padic_norm], padic_norm_e.norm_int_le_pow_iff_dvd _ _ /-- `padic_int.valuation` lifts the `p`-adic valuation on `ℚ` to `ℤ_[p]`. -/ def valuation (x : ℤ_[p]) := padic.valuation (x : ℚ_[p]) -lemma norm_eq_pow_val {x : ℤ_[p]} (hx : x ≠ 0) : ∥x∥ = (p : ℝ) ^ -x.valuation := +lemma norm_eq_pow_val {x : ℤ_[p]} (hx : x ≠ 0) : ‖x‖ = (p : ℝ) ^ -x.valuation := begin convert padic.norm_eq_pow_val _, contrapose! hx, @@ -320,7 +320,7 @@ end @[simp] lemma valuation_p_pow_mul (n : ℕ) (c : ℤ_[p]) (hc : c ≠ 0) : (↑p ^ n * c).valuation = n + c.valuation := begin - have : ∥(↑p ^ n * c)∥ = ∥(p ^ n : ℤ_[p])∥ * ∥c∥, + have : ‖(↑p ^ n * c)‖ = ‖(p ^ n : ℤ_[p])‖ * ‖c‖, { exact norm_mul _ _ }, have aux : (↑p ^ n * c) ≠ 0, { contrapose! hc, rw mul_eq_zero at hc, cases hc, @@ -340,7 +340,7 @@ section units local attribute [reducible] padic_int -lemma mul_inv : ∀ {z : ℤ_[p]}, ∥z∥ = 1 → z * z.inv = 1 +lemma mul_inv : ∀ {z : ℤ_[p]}, ‖z‖ = 1 → z * z.inv = 1 | ⟨k, _⟩ h := begin have hk : k ≠ 0, from λ h', zero_ne_one' ℚ_[p] (by simpa [h'] using h), @@ -351,9 +351,9 @@ lemma mul_inv : ∀ {z : ℤ_[p]}, ∥z∥ = 1 → z * z.inv = 1 simp [mul_inv_cancel hk] end -lemma inv_mul {z : ℤ_[p]} (hz : ∥z∥ = 1) : z.inv * z = 1 := by rw [mul_comm, mul_inv hz] +lemma inv_mul {z : ℤ_[p]} (hz : ‖z‖ = 1) : z.inv * z = 1 := by rw [mul_comm, mul_inv hz] -lemma is_unit_iff {z : ℤ_[p]} : is_unit z ↔ ∥z∥ = 1 := +lemma is_unit_iff {z : ℤ_[p]} : is_unit z ↔ ‖z‖ = 1 := ⟨λ h, begin rcases is_unit_iff_dvd_one.1 h with ⟨w, eq⟩, refine le_antisymm (norm_le_one _) _, @@ -361,29 +361,29 @@ lemma is_unit_iff {z : ℤ_[p]} : is_unit z ↔ ∥z∥ = 1 := rwa [mul_one, ← norm_mul, ← eq, norm_one] at this end, λ h, ⟨⟨z, z.inv, mul_inv h, inv_mul h⟩, rfl⟩⟩ -lemma norm_lt_one_add {z1 z2 : ℤ_[p]} (hz1 : ∥z1∥ < 1) (hz2 : ∥z2∥ < 1) : ∥z1 + z2∥ < 1 := +lemma norm_lt_one_add {z1 z2 : ℤ_[p]} (hz1 : ‖z1‖ < 1) (hz2 : ‖z2‖ < 1) : ‖z1 + z2‖ < 1 := lt_of_le_of_lt (nonarchimedean _ _) (max_lt hz1 hz2) -lemma norm_lt_one_mul {z1 z2 : ℤ_[p]} (hz2 : ∥z2∥ < 1) : ∥z1 * z2∥ < 1 := -calc ∥z1 * z2∥ = ∥z1∥ * ∥z2∥ : by simp +lemma norm_lt_one_mul {z1 z2 : ℤ_[p]} (hz2 : ‖z2‖ < 1) : ‖z1 * z2‖ < 1 := +calc ‖z1 * z2‖ = ‖z1‖ * ‖z2‖ : by simp ... < 1 : mul_lt_one_of_nonneg_of_lt_one_right (norm_le_one _) (norm_nonneg _) hz2 -@[simp] lemma mem_nonunits {z : ℤ_[p]} : z ∈ nonunits ℤ_[p] ↔ ∥z∥ < 1 := +@[simp] lemma mem_nonunits {z : ℤ_[p]} : z ∈ nonunits ℤ_[p] ↔ ‖z‖ < 1 := by rw lt_iff_le_and_ne; simp [norm_le_one z, nonunits, is_unit_iff] -/-- A `p`-adic number `u` with `∥u∥ = 1` is a unit of `ℤ_[p]`. -/ -def mk_units {u : ℚ_[p]} (h : ∥u∥ = 1) : ℤ_[p]ˣ := +/-- A `p`-adic number `u` with `‖u‖ = 1` is a unit of `ℤ_[p]`. -/ +def mk_units {u : ℚ_[p]} (h : ‖u‖ = 1) : ℤ_[p]ˣ := let z : ℤ_[p] := ⟨u, le_of_eq h⟩ in ⟨z, z.inv, mul_inv h, inv_mul h⟩ -@[simp] lemma mk_units_eq {u : ℚ_[p]} (h : ∥u∥ = 1) : ((mk_units h : ℤ_[p]) : ℚ_[p]) = u := rfl +@[simp] lemma mk_units_eq {u : ℚ_[p]} (h : ‖u‖ = 1) : ((mk_units h : ℤ_[p]) : ℚ_[p]) = u := rfl -@[simp] lemma norm_units (u : ℤ_[p]ˣ) : ∥(u : ℤ_[p])∥ = 1 := is_unit_iff.mp $ by simp +@[simp] lemma norm_units (u : ℤ_[p]ˣ) : ‖(u : ℤ_[p])‖ = 1 := is_unit_iff.mp $ by simp /-- `unit_coeff hx` is the unit `u` in the unique representation `x = u * p ^ n`. See `unit_coeff_spec`. -/ def unit_coeff {x : ℤ_[p]} (hx : x ≠ 0) : ℤ_[p]ˣ := let u : ℚ_[p] := x * p ^ -x.valuation in -have hu : ∥u∥ = 1, +have hu : ‖u‖ = 1, by simp [hx, nat.zpow_ne_zero_of_pos (by exact_mod_cast hp.1.pos) x.valuation, norm_eq_pow_val, zpow_neg, inv_mul_cancel], mk_units hu @@ -411,7 +411,7 @@ section norm_le_iff /-! ### Various characterizations of open unit balls -/ lemma norm_le_pow_iff_le_valuation (x : ℤ_[p]) (hx : x ≠ 0) (n : ℕ) : - ∥x∥ ≤ p ^ (-n : ℤ) ↔ ↑n ≤ x.valuation := + ‖x‖ ≤ p ^ (-n : ℤ) ↔ ↑n ≤ x.valuation := begin rw norm_eq_pow_val hx, lift x.valuation to ℕ using x.valuation_nonneg with k hk, @@ -442,7 +442,7 @@ begin end lemma norm_le_pow_iff_mem_span_pow (x : ℤ_[p]) (n : ℕ) : - ∥x∥ ≤ p ^ (-n : ℤ) ↔ x ∈ (ideal.span {p ^ n} : ideal ℤ_[p]) := + ‖x‖ ≤ p ^ (-n : ℤ) ↔ x ∈ (ideal.span {p ^ n} : ideal ℤ_[p]) := begin by_cases hx : x = 0, { subst hx, @@ -451,15 +451,15 @@ begin rw [norm_le_pow_iff_le_valuation x hx, mem_span_pow_iff_le_valuation x hx] end -lemma norm_le_pow_iff_norm_lt_pow_add_one (x : ℤ_[p]) (n : ℤ) : ∥x∥ ≤ p ^ n ↔ ∥x∥ < p ^ (n + 1) := +lemma norm_le_pow_iff_norm_lt_pow_add_one (x : ℤ_[p]) (n : ℤ) : ‖x‖ ≤ p ^ n ↔ ‖x‖ < p ^ (n + 1) := begin rw norm_def, exact padic.norm_le_pow_iff_norm_lt_pow_add_one _ _, end -lemma norm_lt_pow_iff_norm_le_pow_sub_one (x : ℤ_[p]) (n : ℤ) : ∥x∥ < p ^ n ↔ ∥x∥ ≤ p ^ (n - 1) := +lemma norm_lt_pow_iff_norm_le_pow_sub_one (x : ℤ_[p]) (n : ℤ) : ‖x‖ < p ^ n ↔ ‖x‖ ≤ p ^ (n - 1) := by rw [norm_le_pow_iff_norm_lt_pow_add_one, sub_add_cancel] -lemma norm_lt_one_iff_dvd (x : ℤ_[p]) : ∥x∥ < 1 ↔ ↑p ∣ x := +lemma norm_lt_one_iff_dvd (x : ℤ_[p]) : ‖x‖ < 1 ↔ ↑p ∣ x := begin have := norm_le_pow_iff_mem_span_pow x 1, rw [ideal.mem_span_singleton, pow_one] at this, @@ -547,7 +547,7 @@ instance is_fraction_ring : is_fraction_ring ℤ_[p] ℚ_[p] := end, surj := λ x, begin - by_cases hx : ∥ x ∥ ≤ 1, + by_cases hx : ‖ x ‖ ≤ 1, { use (⟨x, hx⟩, 1), rw [submonoid.coe_one, map_one, mul_one], refl, }, @@ -558,7 +558,7 @@ instance is_fraction_ring : is_fraction_ring ℤ_[p] ℚ_[p] := rw padic.norm_le_one_iff_val_nonneg at hx, exact le_of_lt (not_le.mp hx) }, set a := x * p^n with ha, - have ha_norm : ∥ a ∥ = 1, + have ha_norm : ‖ a ‖ = 1, { have hx : x ≠ 0, { intro h0, rw [h0, norm_zero] at hx, diff --git a/src/number_theory/padics/padic_numbers.lean b/src/number_theory/padics/padic_numbers.lean index 468d7b870ea78..e1b87d64add6b 100644 --- a/src/number_theory/padics/padic_numbers.lean +++ b/src/number_theory/padics/padic_numbers.lean @@ -39,11 +39,11 @@ indices in the proof that the norm extends. `padic_norm_e` is the rational-valued `p`-adic norm on `ℚ_[p]`. To instantiate `ℚ_[p]` as a normed field, we must cast this into a `ℝ`-valued norm. -The `ℝ`-valued norm, using notation `∥ ∥` from normed spaces, +The `ℝ`-valued norm, using notation `‖ ‖` from normed spaces, is the canonical representation of this norm. `simp` prefers `padic_norm` to `padic_norm_e` when possible. -Since `padic_norm_e` and `∥ ∥` have different types, `simp` does not rewrite one to the other. +Since `padic_norm_e` and `‖ ‖` have different types, `simp` does not rewrite one to the other. Coercions from `ℚ` to `ℚ_[p]` are set up to work with the `norm_cast` tactic. @@ -470,7 +470,7 @@ end completion end padic /-- The rational-valued `p`-adic norm on `ℚ_[p]` is lifted from the norm on Cauchy sequences. The -canonical form of this function is the normed space instance, with notation `∥ ∥`. -/ +canonical form of this function is the normed space instance, with notation `‖ ‖`. -/ def padic_norm_e {p : ℕ} [hp : fact p.prime] : absolute_value ℚ_[p] ℚ := { to_fun := quotient.lift padic_seq.norm $ @padic_seq.norm_equiv _ _, map_mul' := λ q r, quotient.induction_on₂ q r $ padic_seq.norm_mul, @@ -509,13 +509,13 @@ begin end /-- Theorems about `padic_norm_e` are named with a `'` so the names do not conflict with the -equivalent theorems about `norm` (`∥ ∥`). -/ +equivalent theorems about `norm` (`‖ ‖`). -/ theorem nonarchimedean' (q r : ℚ_[p]) : padic_norm_e (q + r) ≤ max (padic_norm_e q) (padic_norm_e r) := quotient.induction_on₂ q r $ norm_nonarchimedean /-- Theorems about `padic_norm_e` are named with a `'` so the names do not conflict with the -equivalent theorems about `norm` (`∥ ∥`). -/ +equivalent theorems about `norm` (`‖ ‖`). -/ theorem add_eq_max_of_ne' {q r : ℚ_[p]} : padic_norm_e q ≠ padic_norm_e r → padic_norm_e (q + r) = max (padic_norm_e q) (padic_norm_e r) := quotient.induction_on₂ q r $ λ _ _, padic_seq.add_eq_max_of_ne @@ -669,13 +669,13 @@ instance : normed_field ℚ_[p] := norm_mul' := by simp [has_norm.norm, map_mul], norm := norm, .. padic.field, .. padic.metric_space p } -instance is_absolute_value : is_absolute_value (λ a : ℚ_[p], ∥a∥) := +instance is_absolute_value : is_absolute_value (λ a : ℚ_[p], ‖a‖) := { abv_nonneg := norm_nonneg, abv_eq_zero := λ _, norm_eq_zero, abv_add := norm_add_le, abv_mul := by simp [has_norm.norm, map_mul] } -theorem rat_dense (q : ℚ_[p]) {ε : ℝ} (hε : 0 < ε) : ∃ r : ℚ, ∥q - r∥ < ε := +theorem rat_dense (q : ℚ_[p]) {ε : ℝ} (hε : 0 < ε) : ∃ r : ℚ, ‖q - r‖ < ε := let ⟨ε', hε'l, hε'r⟩ := exists_rat_btwn hε, ⟨r, hr⟩ := rat_dense' q (by simpa using hε'l) in ⟨r, lt_trans (by simpa [has_norm.norm] using hr) hε'r⟩ @@ -688,18 +688,18 @@ section normed_space variables {p : ℕ} [hp : fact p.prime] include hp -@[simp] protected lemma mul (q r : ℚ_[p]) : ∥q * r∥ = ∥q∥ * ∥r∥ := +@[simp] protected lemma mul (q r : ℚ_[p]) : ‖q * r‖ = ‖q‖ * ‖r‖ := by simp [has_norm.norm, map_mul] -protected lemma is_norm (q : ℚ_[p]) : ↑(padic_norm_e q) = ∥q∥ := rfl +protected lemma is_norm (q : ℚ_[p]) : ↑(padic_norm_e q) = ‖q‖ := rfl -theorem nonarchimedean (q r : ℚ_[p]) : ∥q + r∥ ≤ max (∥q∥) (∥r∥) := +theorem nonarchimedean (q r : ℚ_[p]) : ‖q + r‖ ≤ max (‖q‖) (‖r‖) := begin unfold has_norm.norm, exact_mod_cast nonarchimedean' _ _ end -theorem add_eq_max_of_ne {q r : ℚ_[p]} (h : ∥q∥ ≠ ∥r∥) : ∥q + r∥ = max (∥q∥) (∥r∥) := +theorem add_eq_max_of_ne {q r : ℚ_[p]} (h : ‖q‖ ≠ ‖r‖) : ‖q + r‖ = max (‖q‖) (‖r‖) := begin unfold has_norm.norm, apply_mod_cast add_eq_max_of_ne', @@ -709,13 +709,13 @@ begin exact_mod_cast h' end -@[simp] lemma eq_padic_norm (q : ℚ) : ∥(q : ℚ_[p])∥ = padic_norm p q := +@[simp] lemma eq_padic_norm (q : ℚ) : ‖(q : ℚ_[p])‖ = padic_norm p q := begin unfold has_norm.norm, rw [← padic_norm_e.eq_padic_norm'] end -@[simp] lemma norm_p : ∥(p : ℚ_[p])∥ = p⁻¹ := +@[simp] lemma norm_p : ‖(p : ℚ_[p])‖ = p⁻¹ := begin have p₀ : p ≠ 0 := hp.1.ne_zero, have p₁ : p ≠ 1 := hp.1.ne_one, @@ -726,14 +726,14 @@ begin simp [p₀, p₁, padic_val_int, zpow_neg] end -lemma norm_p_lt_one : ∥(p : ℚ_[p])∥ < 1 := +lemma norm_p_lt_one : ‖(p : ℚ_[p])‖ < 1 := begin rw norm_p, apply inv_lt_one, exact_mod_cast hp.1.one_lt end -@[simp] lemma norm_p_pow (n : ℤ) : ∥(p ^ n : ℚ_[p])∥ = p ^ -n := +@[simp] lemma norm_p_pow (n : ℤ) : ‖(p ^ n : ℚ_[p])‖ = p ^ -n := by rw [norm_zpow, norm_p]; field_simp instance : nontrivially_normed_field ℚ_[p] := @@ -743,23 +743,23 @@ instance : nontrivially_normed_field ℚ_[p] := end⟩, .. padic.normed_field p } -protected theorem image {q : ℚ_[p]} : q ≠ 0 → ∃ n : ℤ, ∥q∥ = ↑((p : ℚ) ^ -n) := +protected theorem image {q : ℚ_[p]} : q ≠ 0 → ∃ n : ℤ, ‖q‖ = ↑((p : ℚ) ^ -n) := quotient.induction_on q $ λ f hf, have ¬ f ≈ 0, from (padic_seq.ne_zero_iff_nequiv_zero f).1 hf, let ⟨n, hn⟩ := padic_seq.norm_values_discrete f this in ⟨n, congr_arg coe hn⟩ -protected lemma is_rat (q : ℚ_[p]) : ∃ q' : ℚ, ∥q∥ = q' := +protected lemma is_rat (q : ℚ_[p]) : ∃ q' : ℚ, ‖q‖ = q' := if h : q = 0 then ⟨0, by simp [h]⟩ else let ⟨n, hn⟩ := padic_norm_e.image h in ⟨_, hn⟩ /--`rat_norm q`, for a `p`-adic number `q` is the `p`-adic norm of `q`, as rational number. -The lemma `padic_norm_e.eq_rat_norm` asserts `∥q∥ = rat_norm q`. -/ +The lemma `padic_norm_e.eq_rat_norm` asserts `‖q‖ = rat_norm q`. -/ def rat_norm (q : ℚ_[p]) : ℚ := classical.some (padic_norm_e.is_rat q) -lemma eq_rat_norm (q : ℚ_[p]) : ∥q∥ = rat_norm q := classical.some_spec (padic_norm_e.is_rat q) +lemma eq_rat_norm (q : ℚ_[p]) : ‖q‖ = rat_norm q := classical.some_spec (padic_norm_e.is_rat q) -theorem norm_rat_le_one : ∀ {q : ℚ} (hq : ¬ p ∣ q.denom), ∥(q : ℚ_[p])∥ ≤ 1 +theorem norm_rat_le_one : ∀ {q : ℚ} (hq : ¬ p ∣ q.denom), ‖(q : ℚ_[p])‖ ≤ 1 | ⟨n, d, hn, hd⟩ := λ hq : ¬ p ∣ d, if hnz : n = 0 then have (⟨n, d, hn, hd⟩ : ℚ) = 0, @@ -781,17 +781,17 @@ theorem norm_rat_le_one : ∀ {q : ℚ} (hq : ¬ p ∣ q.denom), ∥(q : ℚ_[p] exact hp.1.pos } end -theorem norm_int_le_one (z : ℤ) : ∥(z : ℚ_[p])∥ ≤ 1 := -suffices ∥((z : ℚ) : ℚ_[p])∥ ≤ 1, by simpa, norm_rat_le_one $ by simp [hp.1.ne_one] +theorem norm_int_le_one (z : ℤ) : ‖(z : ℚ_[p])‖ ≤ 1 := +suffices ‖((z : ℚ) : ℚ_[p])‖ ≤ 1, by simpa, norm_rat_le_one $ by simp [hp.1.ne_one] -lemma norm_int_lt_one_iff_dvd (k : ℤ) : ∥(k : ℚ_[p])∥ < 1 ↔ ↑p ∣ k := +lemma norm_int_lt_one_iff_dvd (k : ℤ) : ‖(k : ℚ_[p])‖ < 1 ↔ ↑p ∣ k := begin split, { intro h, contrapose! h, apply le_of_eq, rw eq_comm, - calc ∥(k : ℚ_[p])∥ = ∥((k : ℚ) : ℚ_[p])∥ : by { norm_cast } + calc ‖(k : ℚ_[p])‖ = ‖((k : ℚ) : ℚ_[p])‖ : by { norm_cast } ... = padic_norm p k : padic_norm_e.eq_padic_norm _ ... = 1 : _, rw padic_norm, @@ -809,7 +809,7 @@ begin { rintro ⟨x, rfl⟩, push_cast, rw padic_norm_e.mul, - calc _ ≤ ∥(p : ℚ_[p])∥ * 1 : mul_le_mul le_rfl (by simpa using norm_int_le_one _) + calc _ ≤ ‖(p : ℚ_[p])‖ * 1 : mul_le_mul le_rfl (by simpa using norm_int_le_one _) (norm_nonneg _) (norm_nonneg _) ... < 1 : _, { rw [mul_one, padic_norm_e.norm_p], @@ -817,7 +817,7 @@ begin exact_mod_cast hp.1.one_lt } } end -lemma norm_int_le_pow_iff_dvd (k : ℤ) (n : ℕ) : ∥(k : ℚ_[p])∥ ≤ ↑p ^ (-n : ℤ) ↔ ↑(p ^ n) ∣ k := +lemma norm_int_le_pow_iff_dvd (k : ℤ) (n : ℕ) : ‖(k : ℚ_[p])‖ ≤ ↑p ^ (-n : ℤ) ↔ ↑(p ^ n) ∣ k := begin have : (p : ℝ) ^ (-n : ℤ) = ↑(p ^ (-n : ℤ) : ℚ), {simp}, rw [show (k : ℚ_[p]) = ((k : ℚ) : ℚ_[p]), by norm_cast, eq_padic_norm, this], @@ -825,11 +825,11 @@ begin rw ← padic_norm.dvd_iff_norm_le end -lemma eq_of_norm_add_lt_right {z1 z2 : ℚ_[p]} (h : ∥z1 + z2∥ < ∥z2∥) : ∥z1∥ = ∥z2∥ := +lemma eq_of_norm_add_lt_right {z1 z2 : ℚ_[p]} (h : ‖z1 + z2‖ < ‖z2‖) : ‖z1‖ = ‖z2‖ := by_contradiction $ λ hne, not_lt_of_ge (by rw padic_norm_e.add_eq_max_of_ne hne; apply le_max_right) h -lemma eq_of_norm_add_lt_left {z1 z2 : ℚ_[p]} (h : ∥z1 + z2∥ < ∥z1∥) : ∥z1∥ = ∥z2∥ := +lemma eq_of_norm_add_lt_left {z1 z2 : ℚ_[p]} (h : ‖z1 + z2‖ < ‖z1‖) : ‖z1‖ = ‖z2‖ := by_contradiction $ λ hne, not_lt_of_ge (by rw padic_norm_e.add_eq_max_of_ne hne; apply le_max_left) h @@ -862,11 +862,11 @@ begin exact_mod_cast hN i hi end -lemma padic_norm_e_lim_le {f : cau_seq ℚ_[p] norm} {a : ℝ} (ha : 0 < a) (hf : ∀ i, ∥f i∥ ≤ a) : - ∥f.lim∥ ≤ a := +lemma padic_norm_e_lim_le {f : cau_seq ℚ_[p] norm} {a : ℝ} (ha : 0 < a) (hf : ∀ i, ‖f i‖ ≤ a) : + ‖f.lim‖ ≤ a := let ⟨N, hN⟩ := setoid.symm (cau_seq.equiv_lim f) _ ha in -calc ∥f.lim∥ = ∥f.lim - f N + f N∥ : by simp - ... ≤ max (∥f.lim - f N∥) (∥f N∥) : padic_norm_e.nonarchimedean _ _ +calc ‖f.lim‖ = ‖f.lim - f N + f N‖ : by simp + ... ≤ max (‖f.lim - f N‖) (‖f N‖) : padic_norm_e.nonarchimedean _ _ ... ≤ a : max_le (le_of_lt (hN _ le_rfl)) (hf _) open filter set @@ -907,7 +907,7 @@ begin simp end -lemma norm_eq_pow_val {x : ℚ_[p]} : x ≠ 0 → ∥x∥ = p ^ -x.valuation := +lemma norm_eq_pow_val {x : ℚ_[p]} : x ≠ 0 → ‖x‖ = p ^ -x.valuation := begin apply quotient.induction_on' x, clear x, intros f hf, @@ -938,7 +938,7 @@ begin { by_cases hy : y = 0, { rw [hy, add_zero], exact min_le_left _ _ }, - { have h_norm : ∥x + y∥ ≤ (max ∥x∥ ∥y∥) := padic_norm_e.nonarchimedean x y, + { have h_norm : ‖x + y‖ ≤ (max ‖x‖ ‖y‖) := padic_norm_e.nonarchimedean x y, have hp_one : (1 : ℝ) < p, { rw [← nat.cast_one, nat.cast_lt], exact nat.prime.one_lt hp.elim }, @@ -949,7 +949,7 @@ end @[simp] lemma valuation_map_mul {x y : ℚ_[p]} (hx : x ≠ 0) (hy : y ≠ 0) : valuation (x * y) = valuation x + valuation y := begin - have h_norm : ∥x * y∥ = ∥x∥ * ∥y∥ := norm_mul x y, + have h_norm : ‖x * y‖ = ‖x‖ * ‖y‖ := norm_mul x y, have hp_ne_one : (p : ℝ) ≠ 1, { rw [← nat.cast_one, ne.def, nat.cast_inj], exact nat.prime.ne_one hp.elim }, @@ -1009,7 +1009,7 @@ section norm_le_iff /-! ### Various characterizations of open unit balls -/ -lemma norm_le_pow_iff_norm_lt_pow_add_one (x : ℚ_[p]) (n : ℤ) : ∥x∥ ≤ p ^ n ↔ ∥x∥ < p ^ (n + 1) := +lemma norm_le_pow_iff_norm_lt_pow_add_one (x : ℚ_[p]) (n : ℤ) : ‖x‖ ≤ p ^ n ↔ ‖x‖ < p ^ (n + 1) := begin have aux : ∀ n : ℤ, 0 < (p ^ n : ℝ), { apply nat.zpow_pos_of_pos, exact hp.1.pos }, @@ -1020,10 +1020,10 @@ begin rw [H.le_iff_le, H.lt_iff_lt, int.lt_add_one_iff] end -lemma norm_lt_pow_iff_norm_le_pow_sub_one (x : ℚ_[p]) (n : ℤ) : ∥x∥ < p ^ n ↔ ∥x∥ ≤ p ^ (n - 1) := +lemma norm_lt_pow_iff_norm_le_pow_sub_one (x : ℚ_[p]) (n : ℤ) : ‖x‖ < p ^ n ↔ ‖x‖ ≤ p ^ (n - 1) := by rw [norm_le_pow_iff_norm_lt_pow_add_one, sub_add_cancel] -lemma norm_le_one_iff_val_nonneg (x : ℚ_[p]) : ∥ x ∥ ≤ 1 ↔ 0 ≤ x.valuation := +lemma norm_le_one_iff_val_nonneg (x : ℚ_[p]) : ‖ x ‖ ≤ 1 ↔ 0 ≤ x.valuation := begin by_cases hx : x = 0, { simp only [hx, norm_zero, valuation_zero, zero_le_one, le_refl], }, diff --git a/src/number_theory/padics/ring_homs.lean b/src/number_theory/padics/ring_homs.lean index d699e84a98519..e8153eb6910c8 100644 --- a/src/number_theory/padics/ring_homs.lean +++ b/src/number_theory/padics/ring_homs.lean @@ -54,7 +54,7 @@ variables (p) (r : ℚ) omit hp_prime /-- `mod_part p r` is an integer that satisfies -`∥(r - mod_part p r : ℚ_[p])∥ < 1` when `∥(r : ℚ_[p])∥ ≤ 1`, +`‖(r - mod_part p r : ℚ_[p])‖ < 1` when `‖(r : ℚ_[p])‖ ≤ 1`, see `padic_int.norm_sub_mod_part`. It is the unique non-negative integer that is `< p` with this property. @@ -77,16 +77,16 @@ end lemma mod_part_nonneg : 0 ≤ mod_part p r := int.mod_nonneg _ $ by exact_mod_cast hp_prime.1.ne_zero -lemma is_unit_denom (r : ℚ) (h : ∥(r : ℚ_[p])∥ ≤ 1) : is_unit (r.denom : ℤ_[p]) := +lemma is_unit_denom (r : ℚ) (h : ‖(r : ℚ_[p])‖ ≤ 1) : is_unit (r.denom : ℤ_[p]) := begin rw is_unit_iff, apply le_antisymm (r.denom : ℤ_[p]).2, rw [← not_lt, val_eq_coe, coe_nat_cast], intro norm_denom_lt, - have hr : ∥(r * r.denom : ℚ_[p])∥ = ∥(r.num : ℚ_[p])∥, + have hr : ‖(r * r.denom : ℚ_[p])‖ = ‖(r.num : ℚ_[p])‖, { rw_mod_cast @rat.mul_denom_eq_num r, refl, }, rw padic_norm_e.mul at hr, - have key : ∥(r.num : ℚ_[p])∥ < 1, + have key : ‖(r.num : ℚ_[p])‖ < 1, { calc _ = _ : hr.symm ... < 1 * 1 : mul_lt_mul' h norm_denom_lt (norm_nonneg _) zero_lt_one ... = 1 : mul_one 1 }, @@ -97,7 +97,7 @@ begin rwa [← r.cop.gcd_eq_one, nat.dvd_gcd_iff, ← int.coe_nat_dvd_left, ← int.coe_nat_dvd], end -lemma norm_sub_mod_part_aux (r : ℚ) (h : ∥(r : ℚ_[p])∥ ≤ 1) : +lemma norm_sub_mod_part_aux (r : ℚ) (h : ‖(r : ℚ_[p])‖ ≤ 1) : ↑p ∣ r.num - r.num * r.denom.gcd_a p % p * ↑(r.denom) := begin rw ← zmod.int_coe_zmod_eq_zero_iff_dvd, @@ -117,7 +117,7 @@ begin exact is_unit_denom r h, end -lemma norm_sub_mod_part (h : ∥(r : ℚ_[p])∥ ≤ 1) : ∥(⟨r,h⟩ - mod_part p r : ℤ_[p])∥ < 1 := +lemma norm_sub_mod_part (h : ‖(r : ℚ_[p])‖ ≤ 1) : ‖(⟨r,h⟩ - mod_part p r : ℤ_[p])‖ < 1 := begin let n := mod_part p r, rw [norm_lt_one_iff_dvd, ← (is_unit_denom r h).dvd_mul_right], @@ -131,8 +131,8 @@ begin exact norm_sub_mod_part_aux r h end -lemma exists_mem_range_of_norm_rat_le_one (h : ∥(r : ℚ_[p])∥ ≤ 1) : - ∃ n : ℤ, 0 ≤ n ∧ n < p ∧ ∥(⟨r,h⟩ - n : ℤ_[p])∥ < 1 := +lemma exists_mem_range_of_norm_rat_le_one (h : ‖(r : ℚ_[p])‖ ≤ 1) : + ∃ n : ℤ, 0 ≤ n ∧ n < p ∧ ‖(⟨r,h⟩ - n : ℤ_[p])‖ < 1 := ⟨mod_part p r, mod_part_nonneg _, mod_part_lt_p _, norm_sub_mod_part _ h⟩ lemma zmod_congr_of_sub_mem_span_aux (n : ℕ) (x : ℤ_[p]) (a b : ℤ) @@ -173,9 +173,9 @@ lemma exists_mem_range : ∃ n : ℕ, n < p ∧ (x - n ∈ maximal_ideal ℤ_[p] begin simp only [maximal_ideal_eq_span_p, ideal.mem_span_singleton, ← norm_lt_one_iff_dvd], obtain ⟨r, hr⟩ := rat_dense p (x : ℚ_[p]) zero_lt_one, - have H : ∥(r : ℚ_[p])∥ ≤ 1, + have H : ‖(r : ℚ_[p])‖ ≤ 1, { rw norm_sub_rev at hr, - calc _ = ∥(r : ℚ_[p]) - x + x∥ : by ring_nf + calc _ = ‖(r : ℚ_[p]) - x + x‖ : by ring_nf ... ≤ _ : padic_norm_e.nonarchimedean _ _ ... ≤ _ : max_le (le_of_lt hr) x.2 }, obtain ⟨n, hzn, hnp, hn⟩ := exists_mem_range_of_norm_rat_le_one r H, @@ -191,7 +191,7 @@ end /-- `zmod_repr x` is the unique natural number smaller than `p` -satisfying `∥(x - zmod_repr x : ℤ_[p])∥ < 1`. +satisfying `‖(x - zmod_repr x : ℤ_[p])‖ < 1`. -/ def zmod_repr : ℕ := classical.some (exists_mem_range x) @@ -572,7 +572,7 @@ def lim_nth_hom (r : R) : ℤ_[p] := of_int_seq (nth_hom f r) (is_cau_seq_nth_hom f_compat r) lemma lim_nth_hom_spec (r : R) : - ∀ ε : ℝ, 0 < ε → ∃ N : ℕ, ∀ n ≥ N, ∥lim_nth_hom f_compat r - nth_hom f r n∥ < ε := + ∀ ε : ℝ, 0 < ε → ∃ N : ℕ, ∀ n ≥ N, ‖lim_nth_hom f_compat r - nth_hom f r n‖ < ε := begin intros ε hε, obtain ⟨ε', hε'0, hε'⟩ : ∃ v : ℚ, (0 : ℝ) < v ∧ ↑v < ε := exists_rat_btwn hε, diff --git a/src/probability/density.lean b/src/probability/density.lean index 7dcbe20223560..5545c5925759d 100644 --- a/src/probability/density.lean +++ b/src/probability/density.lean @@ -192,7 +192,7 @@ begin { refine ⟨hf.ae_strongly_measurable, _⟩, rw [has_finite_integral, lintegral_with_density_eq_lintegral_mul _ (measurable_pdf _ _ _) hf.nnnorm.coe_nnreal_ennreal], - have : (λ x, (pdf X ℙ μ * λ x, ↑∥f x∥₊) x) =ᵐ[μ] (λ x, ∥f x * (pdf X ℙ μ x).to_real∥₊), + have : (λ x, (pdf X ℙ μ * λ x, ↑‖f x‖₊) x) =ᵐ[μ] (λ x, ‖f x * (pdf X ℙ μ x).to_real‖₊), { simp_rw [← smul_eq_mul, nnnorm_smul, ennreal.coe_mul], rw [smul_eq_mul, mul_comm], refine filter.eventually_eq.mul (ae_eq_refl _) (ae_eq_trans of_real_to_real_ae_eq.symm _), @@ -295,12 +295,12 @@ lemma integral_mul_eq_integral [has_pdf X ℙ] : integral_fun_mul_eq_integral measurable_id lemma has_finite_integral_mul {f : ℝ → ℝ} {g : ℝ → ℝ≥0∞} - (hg : pdf X ℙ =ᵐ[volume] g) (hgi : ∫⁻ x, ∥f x∥₊ * g x ≠ ∞) : + (hg : pdf X ℙ =ᵐ[volume] g) (hgi : ∫⁻ x, ‖f x‖₊ * g x ≠ ∞) : has_finite_integral (λ x, f x * (pdf X ℙ volume x).to_real) := begin rw has_finite_integral, - have : (λ x, ↑∥f x∥₊ * g x) =ᵐ[volume] (λ x, ∥f x * (pdf X ℙ volume x).to_real∥₊), - { refine ae_eq_trans (filter.eventually_eq.mul (ae_eq_refl (λ x, ∥f x∥₊)) + have : (λ x, ↑‖f x‖₊ * g x) =ᵐ[volume] (λ x, ‖f x * (pdf X ℙ volume x).to_real‖₊), + { refine ae_eq_trans (filter.eventually_eq.mul (ae_eq_refl (λ x, ‖f x‖₊)) (ae_eq_trans hg.symm of_real_to_real_ae_eq.symm)) _, simp_rw [← smul_eq_mul, nnnorm_smul, ennreal.coe_mul, smul_eq_mul], refine filter.eventually_eq.mul (ae_eq_refl _) _, @@ -390,8 +390,8 @@ begin (measurable_pdf X ℙ).ae_measurable.ennreal_to_real.ae_strongly_measurable, _⟩, refine has_finite_integral_mul huX _, set ind := (volume s)⁻¹ • (1 : ℝ → ℝ≥0∞) with hind, - have : ∀ x, ↑∥x∥₊ * s.indicator ind x = s.indicator (λ x, ∥x∥₊ * ind x) x := - λ x, (s.indicator_mul_right (λ x, ↑∥x∥₊) ind).symm, + have : ∀ x, ↑‖x‖₊ * s.indicator ind x = s.indicator (λ x, ‖x‖₊ * ind x) x := + λ x, (s.indicator_mul_right (λ x, ↑‖x‖₊) ind).symm, simp only [this, lintegral_indicator _ hms, hind, mul_one, algebra.id.smul_eq_mul, pi.one_apply, pi.smul_apply], rw lintegral_mul_const _ measurable_nnnorm.coe_nnreal_ennreal, diff --git a/src/probability/ident_distrib.lean b/src/probability/ident_distrib.lean index 89b563ddf294d..8a0ca1a2a7348 100644 --- a/src/probability/ident_distrib.lean +++ b/src/probability/ident_distrib.lean @@ -44,7 +44,7 @@ instance: We also register several dot notation shortcuts for convenience. For instance, if `h : ident_distrib f g μ ν`, then `h.sq` states that `f^2` and `g^2` are -identically distributed, and `h.norm` states that `∥f∥` and `∥g∥` are identically distributed, and +identically distributed, and `h.norm` states that `‖f‖` and `‖g‖` are identically distributed, and so on. -/ @@ -241,11 +241,11 @@ lemma integrable_iff [normed_add_comm_group γ] [borel_space γ] (h : ident_dist ⟨λ hf, h.integrable_snd hf, λ hg, h.symm.integrable_snd hg⟩ protected lemma norm [normed_add_comm_group γ] [borel_space γ] (h : ident_distrib f g μ ν) : - ident_distrib (λ x, ∥f x∥) (λ x, ∥g x∥) μ ν := + ident_distrib (λ x, ‖f x‖) (λ x, ‖g x‖) μ ν := h.comp measurable_norm protected lemma nnnorm [normed_add_comm_group γ] [borel_space γ] (h : ident_distrib f g μ ν) : - ident_distrib (λ x, ∥f x∥₊) (λ x, ∥g x∥₊) μ ν := + ident_distrib (λ x, ‖f x‖₊) (λ x, ‖g x‖₊) μ ν := h.comp measurable_nnnorm protected lemma pow [has_pow γ ℕ] [has_measurable_pow γ ℕ] (h : ident_distrib f g μ ν) {n : ℕ} : @@ -313,11 +313,11 @@ begin by_cases hι : nonempty ι, swap, { exact ⟨0, λ i, false.elim (hι $ nonempty.intro i)⟩ }, obtain ⟨C, hC₁, hC₂⟩ := hℒp.snorm_indicator_norm_ge_pos_le μ (hfmeas _) hε, - have hmeas : ∀ i, measurable_set {x | (⟨C, hC₁.le⟩ : ℝ≥0) ≤ ∥f i x∥₊} := + have hmeas : ∀ i, measurable_set {x | (⟨C, hC₁.le⟩ : ℝ≥0) ≤ ‖f i x‖₊} := λ i, measurable_set_le measurable_const (hfmeas _).measurable.nnnorm, refine ⟨⟨C, hC₁.le⟩, λ i, le_trans (le_of_eq _) hC₂⟩, - have : {x : α | (⟨C, hC₁.le⟩ : ℝ≥0) ≤ ∥f i x∥₊}.indicator (f i) = - (λ x : E, if (⟨C, hC₁.le⟩ : ℝ≥0) ≤ ∥x∥₊ then x else 0) ∘ (f i), + have : {x : α | (⟨C, hC₁.le⟩ : ℝ≥0) ≤ ‖f i x‖₊}.indicator (f i) = + (λ x : E, if (⟨C, hC₁.le⟩ : ℝ≥0) ≤ ‖x‖₊ then x else 0) ∘ (f i), { ext x, simp only [set.indicator, set.mem_set_of_eq] }, simp_rw [coe_nnnorm, this], diff --git a/src/probability/integration.lean b/src/probability/integration.lean index 1612be45178b8..20776ca4089ee 100644 --- a/src/probability/integration.lean +++ b/src/probability/integration.lean @@ -140,10 +140,10 @@ lemma indep_fun.integrable_mul {β : Type*} [measurable_space β] {X Y : Ω → (hXY : indep_fun X Y μ) (hX : integrable X μ) (hY : integrable Y μ) : integrable (X * Y) μ := begin - let nX : Ω → ennreal := λ a, ∥X a∥₊, - let nY : Ω → ennreal := λ a, ∥Y a∥₊, + let nX : Ω → ennreal := λ a, ‖X a‖₊, + let nY : Ω → ennreal := λ a, ‖Y a‖₊, - have hXY' : indep_fun (λ a, ∥X a∥₊) (λ a, ∥Y a∥₊) μ := + have hXY' : indep_fun (λ a, ‖X a‖₊) (λ a, ‖Y a‖₊) μ := hXY.comp measurable_nnnorm measurable_nnnorm, have hXY'' : indep_fun nX nY μ := hXY'.comp measurable_coe_nnreal_ennreal measurable_coe_nnreal_ennreal, @@ -168,17 +168,17 @@ lemma indep_fun.integrable_left_of_integrable_mul {β : Type*} [measurable_space integrable X μ := begin refine ⟨hX, _⟩, - have I : ∫⁻ ω, ∥Y ω∥₊ ∂μ ≠ 0, + have I : ∫⁻ ω, ‖Y ω‖₊ ∂μ ≠ 0, { assume H, - have I : (λ ω, ↑∥Y ω∥₊) =ᵐ[μ] 0 := (lintegral_eq_zero_iff' hY.ennnorm).1 H, + have I : (λ ω, ↑‖Y ω‖₊) =ᵐ[μ] 0 := (lintegral_eq_zero_iff' hY.ennnorm).1 H, apply h'Y, filter_upwards [I] with ω hω, simpa using hω }, apply lt_top_iff_ne_top.2 (λ H, _), - have J : indep_fun (λ ω, ↑∥X ω∥₊) (λ ω, ↑∥Y ω∥₊) μ, - { have M : measurable (λ (x : β), (∥x∥₊ : ℝ≥0∞)) := measurable_nnnorm.coe_nnreal_ennreal, + have J : indep_fun (λ ω, ↑‖X ω‖₊) (λ ω, ↑‖Y ω‖₊) μ, + { have M : measurable (λ (x : β), (‖x‖₊ : ℝ≥0∞)) := measurable_nnnorm.coe_nnreal_ennreal, apply indep_fun.comp hXY M M }, - have A : ∫⁻ ω, ∥X ω * Y ω∥₊ ∂μ < ∞ := h'XY.2, + have A : ∫⁻ ω, ‖X ω * Y ω‖₊ ∂μ < ∞ := h'XY.2, simp only [nnnorm_mul, ennreal.coe_mul] at A, rw [lintegral_mul_eq_lintegral_mul_lintegral_of_indep_fun'' hX.ennnorm hY.ennnorm J, H] at A, simpa [ennreal.top_mul, I] using A, @@ -193,17 +193,17 @@ lemma indep_fun.integrable_right_of_integrable_mul {β : Type*} [measurable_spac integrable Y μ := begin refine ⟨hY, _⟩, - have I : ∫⁻ ω, ∥X ω∥₊ ∂μ ≠ 0, + have I : ∫⁻ ω, ‖X ω‖₊ ∂μ ≠ 0, { assume H, - have I : (λ ω, ↑∥X ω∥₊) =ᵐ[μ] 0 := (lintegral_eq_zero_iff' hX.ennnorm).1 H, + have I : (λ ω, ↑‖X ω‖₊) =ᵐ[μ] 0 := (lintegral_eq_zero_iff' hX.ennnorm).1 H, apply h'X, filter_upwards [I] with ω hω, simpa using hω }, apply lt_top_iff_ne_top.2 (λ H, _), - have J : indep_fun (λ ω, ↑∥X ω∥₊) (λ ω, ↑∥Y ω∥₊) μ, - { have M : measurable (λ (x : β), (∥x∥₊ : ℝ≥0∞)) := measurable_nnnorm.coe_nnreal_ennreal, + have J : indep_fun (λ ω, ↑‖X ω‖₊) (λ ω, ↑‖Y ω‖₊) μ, + { have M : measurable (λ (x : β), (‖x‖₊ : ℝ≥0∞)) := measurable_nnnorm.coe_nnreal_ennreal, apply indep_fun.comp hXY M M }, - have A : ∫⁻ ω, ∥X ω * Y ω∥₊ ∂μ < ∞ := h'XY.2, + have A : ∫⁻ ω, ‖X ω * Y ω‖₊ ∂μ < ∞ := h'XY.2, simp only [nnnorm_mul, ennreal.coe_mul] at A, rw [lintegral_mul_eq_lintegral_mul_lintegral_of_indep_fun'' hX.ennnorm hY.ennnorm J, H] at A, simpa [ennreal.top_mul, I] using A, diff --git a/src/probability/martingale/convergence.lean b/src/probability/martingale/convergence.lean index 012be9ecce518..69fca0645ba62 100644 --- a/src/probability/martingale/convergence.lean +++ b/src/probability/martingale/convergence.lean @@ -136,7 +136,7 @@ convergent. We use the spelling `< ∞` instead of the standard `≠ ∞` in the assumptions since it is not as easy to change `<` to `≠` under binders. -/ lemma tendsto_of_uncrossing_lt_top - (hf₁ : liminf (λ n, (∥f n ω∥₊ : ℝ≥0∞)) at_top < ∞) + (hf₁ : liminf (λ n, (‖f n ω‖₊ : ℝ≥0∞)) at_top < ∞) (hf₂ : ∀ a b : ℚ, a < b → upcrossings a b f ω < ∞) : ∃ c, tendsto (λ n, f n ω) at_top (𝓝 c) := begin @@ -160,16 +160,16 @@ begin have := hf.mul_lintegral_upcrossings_le_lintegral_pos_part a b, rw [mul_comm, ← ennreal.le_div_iff_mul_le] at this, { refine (lt_of_le_of_lt this (ennreal.div_lt_top _ _)).ne, - { have hR' : ∀ n, ∫⁻ ω, ∥f n ω - a∥₊ ∂μ ≤ R + ∥a∥₊ * μ set.univ, + { have hR' : ∀ n, ∫⁻ ω, ‖f n ω - a‖₊ ∂μ ≤ R + ‖a‖₊ * μ set.univ, { simp_rw snorm_one_eq_lintegral_nnnorm at hbdd, intro n, - refine (lintegral_mono _ : ∫⁻ ω, ∥f n ω - a∥₊ ∂μ ≤ ∫⁻ ω, ∥f n ω∥₊ + ∥a∥₊ ∂μ).trans _, + refine (lintegral_mono _ : ∫⁻ ω, ‖f n ω - a‖₊ ∂μ ≤ ∫⁻ ω, ‖f n ω‖₊ + ‖a‖₊ ∂μ).trans _, { intro ω, simp_rw [sub_eq_add_neg, ← nnnorm_neg a, ← ennreal.coe_add, ennreal.coe_le_coe], exact nnnorm_add_le _ _ }, { simp_rw [ lintegral_add_right _ measurable_const, lintegral_const], exact add_le_add (hbdd _) le_rfl } }, - refine ne_of_lt (supr_lt_iff.2 ⟨R + ∥a∥₊ * μ set.univ, ennreal.add_lt_top.2 + refine ne_of_lt (supr_lt_iff.2 ⟨R + ‖a‖₊ * μ set.univ, ennreal.add_lt_top.2 ⟨ennreal.coe_lt_top, ennreal.mul_lt_top ennreal.coe_lt_top.ne (measure_ne_top _ _)⟩, λ n, le_trans _ (hR' n)⟩), refine lintegral_mono (λ ω, _), diff --git a/src/probability/martingale/upcrossing.lean b/src/probability/martingale/upcrossing.lean index 5642ee3ae3642..a333845c99b71 100644 --- a/src/probability/martingale/upcrossing.lean +++ b/src/probability/martingale/upcrossing.lean @@ -841,7 +841,7 @@ lemma adapted.integrable_upcrossings_before [is_finite_measure μ] (hf : adapted ℱ f) (hab : a < b) : integrable (λ ω, (upcrossings_before a b f N ω : ℝ)) μ := begin - have : ∀ᵐ ω ∂μ, ∥(upcrossings_before a b f N ω : ℝ)∥ ≤ N, + have : ∀ᵐ ω ∂μ, ‖(upcrossings_before a b f N ω : ℝ)‖ ≤ N, { refine eventually_of_forall (λ ω, _), rw [real.norm_eq_abs, nat.abs_cast, nat.cast_le], refine upcrossings_before_le _ _ hab }, diff --git a/src/probability/variance.lean b/src/probability/variance.lean index 189e780273c24..8b67daebd8f56 100644 --- a/src/probability/variance.lean +++ b/src/probability/variance.lean @@ -42,7 +42,7 @@ namespace probability_theory /-- The `ℝ≥0∞`-valued variance of a real-valued random variable defined as the Lebesgue integral of `(X - 𝔼[X])^2`. -/ def evariance {Ω : Type*} {m : measurable_space Ω} (X : Ω → ℝ) (μ : measure Ω) : ℝ≥0∞ := -∫⁻ ω, ∥X ω - μ[X]∥₊^2 ∂μ +∫⁻ ω, ‖X ω - μ[X]‖₊^2 ∂μ /-- The `ℝ`-valued variance of a real-valued random variable defined by applying `ennreal.to_real` to `evariance`. -/ @@ -252,7 +252,7 @@ end lemma evariance_def' [is_probability_measure (ℙ : measure Ω)] {X : Ω → ℝ} (hX : ae_strongly_measurable X ℙ) : - eVar[X] = (∫⁻ ω, ∥X ω∥₊^2) - ennreal.of_real (𝔼[X]^2) := + eVar[X] = (∫⁻ ω, ‖X ω‖₊^2) - ennreal.of_real (𝔼[X]^2) := begin by_cases hℒ : mem_ℒp X 2, { rw [← hℒ.of_real_variance_eq, variance_def' hℒ, ennreal.of_real_sub _ (sq_nonneg _)], diff --git a/src/ring_theory/polynomial/cyclotomic/eval.lean b/src/ring_theory/polynomial/cyclotomic/eval.lean index 18995dd7b46e1..2866bc9cd568a 100644 --- a/src/ring_theory/polynomial/cyclotomic/eval.lean +++ b/src/ring_theory/polynomial/cyclotomic/eval.lean @@ -184,7 +184,7 @@ lemma sub_one_pow_totient_lt_cyclotomic_eval {n : ℕ} {q : ℝ} (hn' : 2 ≤ n) begin have hn : 0 < n := pos_of_gt hn', have hq := zero_lt_one.trans hq', - have hfor : ∀ ζ' ∈ primitive_roots n ℂ, q - 1 ≤ ∥↑q - ζ'∥, + have hfor : ∀ ζ' ∈ primitive_roots n ℂ, q - 1 ≤ ‖↑q - ζ'‖, { intros ζ' hζ', rw mem_primitive_roots hn at hζ', convert norm_sub_norm_le (↑q) ζ', @@ -192,7 +192,7 @@ begin { rw [hζ'.norm'_eq_one hn.ne'] } }, let ζ := complex.exp (2 * ↑real.pi * complex.I / ↑n), have hζ : is_primitive_root ζ n := complex.is_primitive_root_exp n hn.ne', - have hex : ∃ ζ' ∈ primitive_roots n ℂ, q - 1 < ∥↑q - ζ'∥, + have hex : ∃ ζ' ∈ primitive_roots n ℂ, q - 1 < ‖↑q - ζ'‖, { refine ⟨ζ, (mem_primitive_roots hn).mpr hζ, _⟩, suffices : ¬ same_ray ℝ (q : ℂ) ζ, { convert lt_norm_sub_of_not_same_ray this; @@ -209,7 +209,7 @@ begin simpa only [complex.coe_algebra_map, complex.of_real_eq_zero] using (cyclotomic_pos' n hq').ne' }, suffices : (units.mk0 (real.to_nnreal (q - 1)) (by simp [hq'])) ^ totient n - < units.mk0 (∥(cyclotomic n ℂ).eval q∥₊) (by simp [this]), + < units.mk0 (‖(cyclotomic n ℂ).eval q‖₊) (by simp [this]), { simp only [←units.coe_lt_coe, units.coe_pow, units.coe_mk0, ← nnreal.coe_lt_coe, hq'.le, real.to_nnreal_lt_to_nnreal_iff_of_nonneg, coe_nnnorm, complex.norm_eq_abs, nnreal.coe_pow, real.coe_to_nnreal', max_eq_left, sub_nonneg] at this, @@ -237,7 +237,7 @@ lemma cyclotomic_eval_lt_sub_one_pow_totient {n : ℕ} {q : ℝ} (hn' : 3 ≤ n) begin have hn : 0 < n := pos_of_gt hn', have hq := zero_lt_one.trans hq', - have hfor : ∀ ζ' ∈ primitive_roots n ℂ, ∥↑q - ζ'∥ ≤ q + 1, + have hfor : ∀ ζ' ∈ primitive_roots n ℂ, ‖↑q - ζ'‖ ≤ q + 1, { intros ζ' hζ', rw mem_primitive_roots hn at hζ', convert norm_sub_le (↑q) ζ', @@ -245,7 +245,7 @@ begin { rw [hζ'.norm'_eq_one hn.ne'] }, }, let ζ := complex.exp (2 * ↑real.pi * complex.I / ↑n), have hζ : is_primitive_root ζ n := complex.is_primitive_root_exp n hn.ne', - have hex : ∃ ζ' ∈ primitive_roots n ℂ, ∥↑q - ζ'∥ < q + 1, + have hex : ∃ ζ' ∈ primitive_roots n ℂ, ‖↑q - ζ'‖ < q + 1, { refine ⟨ζ, (mem_primitive_roots hn).mpr hζ, _⟩, suffices : ¬ same_ray ℝ (q : ℂ) (-ζ), { convert norm_add_lt_of_not_same_ray this; @@ -270,7 +270,7 @@ begin { erw cyclotomic.eval_apply q n (algebra_map ℝ ℂ), simp only [complex.coe_algebra_map, complex.of_real_eq_zero], exact (cyclotomic_pos' n hq').ne.symm, }, - suffices : units.mk0 (∥(cyclotomic n ℂ).eval q∥₊) (by simp [this]) + suffices : units.mk0 (‖(cyclotomic n ℂ).eval q‖₊) (by simp [this]) < (units.mk0 (real.to_nnreal (q + 1)) (by simp; linarith)) ^ totient n, { simp only [←units.coe_lt_coe, units.coe_pow, units.coe_mk0, ← nnreal.coe_lt_coe, hq'.le, real.to_nnreal_lt_to_nnreal_iff_of_nonneg, coe_nnnorm, complex.norm_eq_abs, diff --git a/src/set_theory/game/basic.lean b/src/set_theory/game/basic.lean index a7cd5b1cba3a8..e68c2d903471c 100644 --- a/src/set_theory/game/basic.lean +++ b/src/set_theory/game/basic.lean @@ -94,13 +94,13 @@ theorem _root_.pgame.equiv_iff_game_eq {x y : pgame} : x ≈ y ↔ ⟦x⟧ = ⟦ /-- The fuzzy, confused, or incomparable relation on games. -If `x ∥ 0`, then the first player can always win `x`. -/ +If `x ‖ 0`, then the first player can always win `x`. -/ def fuzzy : game → game → Prop := quotient.lift₂ fuzzy (λ x₁ y₁ x₂ y₂ hx hy, propext (fuzzy_congr hx hy)) -local infix ` ∥ `:50 := fuzzy +local infix ` ‖ `:50 := fuzzy -theorem _root_.pgame.fuzzy_iff_game_fuzzy {x y : pgame} : pgame.fuzzy x y ↔ ⟦x⟧ ∥ ⟦y⟧ := iff.rfl +theorem _root_.pgame.fuzzy_iff_game_fuzzy {x y : pgame} : pgame.fuzzy x y ↔ ⟦x⟧ ‖ ⟦y⟧ := iff.rfl instance covariant_class_add_le : covariant_class game game (+) (≤) := ⟨by { rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ h, exact @add_le_add_left _ _ _ _ b c h a }⟩ diff --git a/src/set_theory/game/impartial.lean b/src/set_theory/game/impartial.lean index d650ca0fb92df..1befe167356f3 100644 --- a/src/set_theory/game/impartial.lean +++ b/src/set_theory/game/impartial.lean @@ -121,7 +121,7 @@ lemma nonneg : ¬ G < 0 := end /-- In an impartial game, either the first player always wins, or the second player always wins. -/ -lemma equiv_or_fuzzy_zero : G ≈ 0 ∨ G ∥ 0 := +lemma equiv_or_fuzzy_zero : G ≈ 0 ∨ G ‖ 0 := begin rcases lt_or_equiv_or_gt_or_fuzzy G 0 with h | h | h | h, { exact ((nonneg G) h).elim }, @@ -130,10 +130,10 @@ begin { exact or.inr h } end -@[simp] lemma not_equiv_zero_iff : ¬ G ≈ 0 ↔ G ∥ 0 := +@[simp] lemma not_equiv_zero_iff : ¬ G ≈ 0 ↔ G ‖ 0 := ⟨(equiv_or_fuzzy_zero G).resolve_left, fuzzy.not_equiv⟩ -@[simp] lemma not_fuzzy_zero_iff : ¬ G ∥ 0 ↔ G ≈ 0 := +@[simp] lemma not_fuzzy_zero_iff : ¬ G ‖ 0 ↔ G ≈ 0 := ⟨(equiv_or_fuzzy_zero G).resolve_right, equiv.not_fuzzy⟩ lemma add_self : G + G ≈ 0 := @@ -156,11 +156,11 @@ lemma lf_zero_iff {G : pgame} [G.impartial] : G ⧏ 0 ↔ 0 ⧏ G := by rw [←zero_lf_neg_iff, lf_congr_right (neg_equiv_self G)] lemma equiv_zero_iff_le: G ≈ 0 ↔ G ≤ 0 := ⟨and.left, λ h, ⟨h, le_zero_iff.1 h⟩⟩ -lemma fuzzy_zero_iff_lf : G ∥ 0 ↔ G ⧏ 0 := ⟨and.left, λ h, ⟨h, lf_zero_iff.1 h⟩⟩ +lemma fuzzy_zero_iff_lf : G ‖ 0 ↔ G ⧏ 0 := ⟨and.left, λ h, ⟨h, lf_zero_iff.1 h⟩⟩ lemma equiv_zero_iff_ge : G ≈ 0 ↔ 0 ≤ G := ⟨and.right, λ h, ⟨le_zero_iff.2 h, h⟩⟩ -lemma fuzzy_zero_iff_gf : G ∥ 0 ↔ 0 ⧏ G := ⟨and.right, λ h, ⟨lf_zero_iff.2 h, h⟩⟩ +lemma fuzzy_zero_iff_gf : G ‖ 0 ↔ 0 ⧏ G := ⟨and.right, λ h, ⟨lf_zero_iff.2 h, h⟩⟩ -lemma forall_left_moves_fuzzy_iff_equiv_zero : (∀ i, G.move_left i ∥ 0) ↔ G ≈ 0 := +lemma forall_left_moves_fuzzy_iff_equiv_zero : (∀ i, G.move_left i ‖ 0) ↔ G ≈ 0 := begin refine ⟨λ hb, _, λ hp i, _⟩, { rw [equiv_zero_iff_le G, le_zero_lf], @@ -169,7 +169,7 @@ begin exact hp.1.move_left_lf i } end -lemma forall_right_moves_fuzzy_iff_equiv_zero : (∀ j, G.move_right j ∥ 0) ↔ G ≈ 0 := +lemma forall_right_moves_fuzzy_iff_equiv_zero : (∀ j, G.move_right j ‖ 0) ↔ G ≈ 0 := begin refine ⟨λ hb, _, λ hp i, _⟩, { rw [equiv_zero_iff_ge G, zero_le_lf], @@ -178,7 +178,7 @@ begin exact hp.2.lf_move_right i } end -lemma exists_left_move_equiv_iff_fuzzy_zero : (∃ i, G.move_left i ≈ 0) ↔ G ∥ 0 := +lemma exists_left_move_equiv_iff_fuzzy_zero : (∃ i, G.move_left i ≈ 0) ↔ G ‖ 0 := begin refine ⟨λ ⟨i, hi⟩, (fuzzy_zero_iff_gf G).2 (lf_of_le_move_left hi.2), λ hn, _⟩, rw [fuzzy_zero_iff_gf G, zero_lf_le] at hn, @@ -186,7 +186,7 @@ begin exact ⟨i, (equiv_zero_iff_ge _).2 hi⟩ end -lemma exists_right_move_equiv_iff_fuzzy_zero : (∃ j, G.move_right j ≈ 0) ↔ G ∥ 0 := +lemma exists_right_move_equiv_iff_fuzzy_zero : (∃ j, G.move_right j ≈ 0) ↔ G ‖ 0 := begin refine ⟨λ ⟨i, hi⟩, (fuzzy_zero_iff_lf G).2 (lf_of_move_right_le hi.1), λ hn, _⟩, rw [fuzzy_zero_iff_lf G, lf_zero_le] at hn, diff --git a/src/set_theory/game/nim.lean b/src/set_theory/game/nim.lean index 6c70b45933468..97f04116f4c21 100644 --- a/src/set_theory/game/nim.lean +++ b/src/set_theory/game/nim.lean @@ -188,7 +188,7 @@ lemma exists_ordinal_move_left_eq {o : ordinal} (i) : ∃ o' < o, (nim o).move_l lemma exists_move_left_eq {o o' : ordinal} (h : o' < o) : ∃ i, (nim o).move_left i = nim o' := ⟨to_left_moves_nim ⟨o', h⟩, by simp⟩ -lemma nim_fuzzy_zero_of_ne_zero {o : ordinal} (ho : o ≠ 0) : nim o ∥ 0 := +lemma nim_fuzzy_zero_of_ne_zero {o : ordinal} (ho : o ≠ 0) : nim o ‖ 0 := begin rw [impartial.fuzzy_zero_iff_lf, nim_def, lf_zero_le], rw ←ordinal.pos_iff_ne_zero at ho, @@ -212,7 +212,7 @@ begin exact impartial.add_self (nim o₁) } end -@[simp] lemma nim_add_fuzzy_zero_iff {o₁ o₂ : ordinal} : nim o₁ + nim o₂ ∥ 0 ↔ o₁ ≠ o₂ := +@[simp] lemma nim_add_fuzzy_zero_iff {o₁ o₂ : ordinal} : nim o₁ + nim o₂ ‖ 0 ↔ o₁ ≠ o₂ := by rw [iff_not_comm, impartial.not_fuzzy_zero_iff, nim_add_equiv_zero_iff] @[simp] lemma nim_equiv_iff_eq {o₁ o₂ : ordinal} : nim o₁ ≈ nim o₂ ↔ o₁ = o₂ := diff --git a/src/set_theory/game/pgame.lean b/src/set_theory/game/pgame.lean index bd4e75ace9ff1..061e61c787aad 100644 --- a/src/set_theory/game/pgame.lean +++ b/src/set_theory/game/pgame.lean @@ -45,8 +45,8 @@ by Left as the second player. It turns out to be quite convenient to define various relations on top of these. We define the "less or fuzzy" relation `x ⧏ y` as `¬ y ≤ x`, the equivalence relation `x ≈ y` as `x ≤ y ∧ y ≤ x`, and -the fuzzy relation `x ∥ y` as `x ⧏ y ∧ y ⧏ x`. If `0 ⧏ x`, then `x` can be won by Left as the -first player. If `x ≈ 0`, then `x` can be won by the second player. If `x ∥ 0`, then `x` can be won +the fuzzy relation `x ‖ y` as `x ⧏ y ∧ y ⧏ x`. If `0 ⧏ x`, then `x` can be won by Left as the +first player. If `x ≈ 0`, then `x` can be won by the second player. If `x ‖ 0`, then `x` can be won by the first player. Statements like `zero_le_lf`, `zero_lf_le`, etc. unfold these definitions. The theorems `le_def` and @@ -605,58 +605,58 @@ end /-- The fuzzy, confused, or incomparable relation on pre-games. -If `x ∥ 0`, then the first player can always win `x`. -/ +If `x ‖ 0`, then the first player can always win `x`. -/ def fuzzy (x y : pgame) : Prop := x ⧏ y ∧ y ⧏ x -localized "infix (name := pgame.fuzzy) ` ∥ `:50 := pgame.fuzzy" in pgame +localized "infix (name := pgame.fuzzy) ` ‖ `:50 := pgame.fuzzy" in pgame -@[symm] theorem fuzzy.swap {x y : pgame} : x ∥ y → y ∥ x := and.swap -instance : is_symm _ (∥) := ⟨λ x y, fuzzy.swap⟩ -theorem fuzzy.swap_iff {x y : pgame} : x ∥ y ↔ y ∥ x := ⟨fuzzy.swap, fuzzy.swap⟩ +@[symm] theorem fuzzy.swap {x y : pgame} : x ‖ y → y ‖ x := and.swap +instance : is_symm _ (‖) := ⟨λ x y, fuzzy.swap⟩ +theorem fuzzy.swap_iff {x y : pgame} : x ‖ y ↔ y ‖ x := ⟨fuzzy.swap, fuzzy.swap⟩ -theorem fuzzy_irrefl (x : pgame) : ¬ x ∥ x := λ h, lf_irrefl x h.1 -instance : is_irrefl _ (∥) := ⟨fuzzy_irrefl⟩ +theorem fuzzy_irrefl (x : pgame) : ¬ x ‖ x := λ h, lf_irrefl x h.1 +instance : is_irrefl _ (‖) := ⟨fuzzy_irrefl⟩ -theorem lf_iff_lt_or_fuzzy {x y : pgame} : x ⧏ y ↔ x < y ∨ x ∥ y := +theorem lf_iff_lt_or_fuzzy {x y : pgame} : x ⧏ y ↔ x < y ∨ x ‖ y := by { simp only [lt_iff_le_and_lf, fuzzy, ←pgame.not_le], tauto! } -theorem lf_of_fuzzy {x y : pgame} (h : x ∥ y) : x ⧏ y := lf_iff_lt_or_fuzzy.2 (or.inr h) +theorem lf_of_fuzzy {x y : pgame} (h : x ‖ y) : x ⧏ y := lf_iff_lt_or_fuzzy.2 (or.inr h) alias lf_of_fuzzy ← fuzzy.lf -theorem lt_or_fuzzy_of_lf {x y : pgame} : x ⧏ y → x < y ∨ x ∥ y := +theorem lt_or_fuzzy_of_lf {x y : pgame} : x ⧏ y → x < y ∨ x ‖ y := lf_iff_lt_or_fuzzy.1 -theorem fuzzy.not_equiv {x y : pgame} (h : x ∥ y) : ¬ x ≈ y := +theorem fuzzy.not_equiv {x y : pgame} (h : x ‖ y) : ¬ x ≈ y := λ h', h'.1.not_gf h.2 -theorem fuzzy.not_equiv' {x y : pgame} (h : x ∥ y) : ¬ y ≈ x := +theorem fuzzy.not_equiv' {x y : pgame} (h : x ‖ y) : ¬ y ≈ x := λ h', h'.2.not_gf h.2 -theorem not_fuzzy_of_le {x y : pgame} (h : x ≤ y) : ¬ x ∥ y := +theorem not_fuzzy_of_le {x y : pgame} (h : x ≤ y) : ¬ x ‖ y := λ h', h'.2.not_ge h -theorem not_fuzzy_of_ge {x y : pgame} (h : y ≤ x) : ¬ x ∥ y := +theorem not_fuzzy_of_ge {x y : pgame} (h : y ≤ x) : ¬ x ‖ y := λ h', h'.1.not_ge h -theorem equiv.not_fuzzy {x y : pgame} (h : x ≈ y) : ¬ x ∥ y := +theorem equiv.not_fuzzy {x y : pgame} (h : x ≈ y) : ¬ x ‖ y := not_fuzzy_of_le h.1 -theorem equiv.not_fuzzy' {x y : pgame} (h : x ≈ y) : ¬ y ∥ x := +theorem equiv.not_fuzzy' {x y : pgame} (h : x ≈ y) : ¬ y ‖ x := not_fuzzy_of_le h.2 -theorem fuzzy_congr {x₁ y₁ x₂ y₂ : pgame} (hx : x₁ ≈ x₂) (hy : y₁ ≈ y₂) : x₁ ∥ y₁ ↔ x₂ ∥ y₂ := +theorem fuzzy_congr {x₁ y₁ x₂ y₂ : pgame} (hx : x₁ ≈ x₂) (hy : y₁ ≈ y₂) : x₁ ‖ y₁ ↔ x₂ ‖ y₂ := show _ ∧ _ ↔ _ ∧ _, by rw [lf_congr hx hy, lf_congr hy hx] -theorem fuzzy_congr_imp {x₁ y₁ x₂ y₂ : pgame} (hx : x₁ ≈ x₂) (hy : y₁ ≈ y₂) : x₁ ∥ y₁ → x₂ ∥ y₂ := +theorem fuzzy_congr_imp {x₁ y₁ x₂ y₂ : pgame} (hx : x₁ ≈ x₂) (hy : y₁ ≈ y₂) : x₁ ‖ y₁ → x₂ ‖ y₂ := (fuzzy_congr hx hy).1 -theorem fuzzy_congr_left {x₁ x₂ y} (hx : x₁ ≈ x₂) : x₁ ∥ y ↔ x₂ ∥ y := +theorem fuzzy_congr_left {x₁ x₂ y} (hx : x₁ ≈ x₂) : x₁ ‖ y ↔ x₂ ‖ y := fuzzy_congr hx equiv_rfl -theorem fuzzy_congr_right {x y₁ y₂} (hy : y₁ ≈ y₂) : x ∥ y₁ ↔ x ∥ y₂ := +theorem fuzzy_congr_right {x y₁ y₂} (hy : y₁ ≈ y₂) : x ‖ y₁ ↔ x ‖ y₂ := fuzzy_congr equiv_rfl hy -@[trans] theorem fuzzy_of_fuzzy_of_equiv {x y z} (h₁ : x ∥ y) (h₂ : y ≈ z) : x ∥ z := +@[trans] theorem fuzzy_of_fuzzy_of_equiv {x y z} (h₁ : x ‖ y) (h₂ : y ≈ z) : x ‖ z := (fuzzy_congr_right h₂).1 h₁ -@[trans] theorem fuzzy_of_equiv_of_fuzzy {x y z} (h₁ : x ≈ y) (h₂ : y ∥ z) : x ∥ z := +@[trans] theorem fuzzy_of_equiv_of_fuzzy {x y z} (h₁ : x ≈ y) (h₂ : y ‖ z) : x ‖ z := (fuzzy_congr_left h₁).2 h₂ /-- Exactly one of the following is true (although we don't prove this here). -/ -theorem lt_or_equiv_or_gt_or_fuzzy (x y : pgame) : x < y ∨ x ≈ y ∨ y < x ∨ x ∥ y := +theorem lt_or_equiv_or_gt_or_fuzzy (x y : pgame) : x < y ∨ x ≈ y ∨ y < x ∨ x ‖ y := begin cases le_or_gf x y with h₁ h₁; cases le_or_gf y x with h₂ h₂, @@ -928,7 +928,7 @@ by rw [lt_iff_le_and_lf, lt_iff_le_and_lf, neg_le_neg_iff, neg_lf_neg_iff] @[simp] theorem neg_equiv_neg_iff {x y : pgame} : -x ≈ -y ↔ x ≈ y := by rw [equiv, equiv, neg_le_neg_iff, neg_le_neg_iff, and.comm] -@[simp] theorem neg_fuzzy_neg_iff {x y : pgame} : -x ∥ -y ↔ x ∥ y := +@[simp] theorem neg_fuzzy_neg_iff {x y : pgame} : -x ‖ -y ↔ x ‖ y := by rw [fuzzy, fuzzy, neg_lf_neg_iff, neg_lf_neg_iff, and.comm] theorem neg_le_iff {x y : pgame} : -y ≤ x ↔ -x ≤ y := @@ -943,7 +943,7 @@ by rw [←neg_neg x, neg_lt_neg_iff, neg_neg] theorem neg_equiv_iff {x y : pgame} : -x ≈ y ↔ x ≈ -y := by rw [←neg_neg y, neg_equiv_neg_iff, neg_neg] -theorem neg_fuzzy_iff {x y : pgame} : -x ∥ y ↔ x ∥ -y := +theorem neg_fuzzy_iff {x y : pgame} : -x ‖ y ↔ x ‖ -y := by rw [←neg_neg y, neg_fuzzy_neg_iff, neg_neg] theorem le_neg_iff {x y : pgame} : y ≤ -x ↔ x ≤ -y := @@ -976,13 +976,13 @@ by rw [lt_neg_iff, neg_zero] @[simp] theorem neg_equiv_zero_iff {x : pgame} : -x ≈ 0 ↔ x ≈ 0 := by rw [neg_equiv_iff, neg_zero] -@[simp] theorem neg_fuzzy_zero_iff {x : pgame} : -x ∥ 0 ↔ x ∥ 0 := +@[simp] theorem neg_fuzzy_zero_iff {x : pgame} : -x ‖ 0 ↔ x ‖ 0 := by rw [neg_fuzzy_iff, neg_zero] @[simp] theorem zero_equiv_neg_iff {x : pgame} : 0 ≈ -x ↔ 0 ≈ x := by rw [←neg_equiv_iff, neg_zero] -@[simp] theorem zero_fuzzy_neg_iff {x : pgame} : 0 ∥ -x ↔ 0 ∥ x := +@[simp] theorem zero_fuzzy_neg_iff {x : pgame} : 0 ‖ -x ↔ 0 ‖ x := by rw [←neg_fuzzy_iff, neg_zero] /-! ### Addition and subtraction -/ @@ -1355,7 +1355,7 @@ def star : pgame.{u} := ⟨punit, punit, λ _, 0, λ _, 0⟩ instance unique_star_left_moves : unique star.left_moves := punit.unique instance unique_star_right_moves : unique star.right_moves := punit.unique -theorem star_fuzzy_zero : star ∥ 0 := +theorem star_fuzzy_zero : star ‖ 0 := ⟨by { rw lf_zero, use default, rintros ⟨⟩ }, by { rw zero_lf, use default, rintros ⟨⟩ }⟩ @[simp] theorem neg_star : -star = star := diff --git a/src/topology/algebra/module/finite_dimension.lean b/src/topology/algebra/module/finite_dimension.lean index 877d9486194a2..3144352d5b467 100644 --- a/src/topology/algebra/module/finite_dimension.lean +++ b/src/topology/algebra/module/finite_dimension.lean @@ -110,7 +110,7 @@ begin refine topological_add_group.ext h₁ infer_instance (le_antisymm _ _), { -- To show `𝓣 ≤ 𝓣₀`, we have to show that closed balls are `𝓣`-neighborhoods of 0. rw metric.nhds_basis_closed_ball.ge_iff, - -- Let `ε > 0`. Since `𝕜` is nontrivially normed, we have `0 < ∥ξ₀∥ < ε` for some `ξ₀ : 𝕜`. + -- Let `ε > 0`. Since `𝕜` is nontrivially normed, we have `0 < ‖ξ₀‖ < ε` for some `ξ₀ : 𝕜`. intros ε hε, rcases normed_field.exists_norm_lt 𝕜 hε with ⟨ξ₀, hξ₀, hξ₀ε⟩, -- Since `ξ₀ ≠ 0` and `𝓣` is T2, we know that `{ξ₀}ᶜ` is a `𝓣`-neighborhood of 0. @@ -120,18 +120,18 @@ begin -- `𝓑`, which will imply that the closed ball is indeed a `𝓣`-neighborhood of 0. have : balanced_core 𝕜 {ξ₀}ᶜ ∈ @nhds 𝕜 t 0 := balanced_core_mem_nhds_zero this, refine mem_of_superset this (λ ξ hξ, _), - -- Let `ξ ∈ 𝓑`. We want to show `∥ξ∥ < ε`. If `ξ = 0`, this is trivial. + -- Let `ξ ∈ 𝓑`. We want to show `‖ξ‖ < ε`. If `ξ = 0`, this is trivial. by_cases hξ0 : ξ = 0, { rw hξ0, exact metric.mem_closed_ball_self hε.le }, { rw [mem_closed_ball_zero_iff], - -- Now suppose `ξ ≠ 0`. By contradiction, let's assume `ε < ∥ξ∥`, and show that + -- Now suppose `ξ ≠ 0`. By contradiction, let's assume `ε < ‖ξ‖`, and show that -- `ξ₀ ∈ 𝓑 ⊆ {ξ₀}ᶜ`, which is a contradiction. by_contra' h, suffices : (ξ₀ * ξ⁻¹) • ξ ∈ balanced_core 𝕜 {ξ₀}ᶜ, { rw [smul_eq_mul 𝕜, mul_assoc, inv_mul_cancel hξ0, mul_one] at this, exact not_mem_compl_iff.mpr (mem_singleton ξ₀) ((balanced_core_subset _) this) }, - -- For that, we use that `𝓑` is balanced : since `∥ξ₀∥ < ε < ∥ξ∥`, we have `∥ξ₀ / ξ∥ ≤ 1`, + -- For that, we use that `𝓑` is balanced : since `‖ξ₀‖ < ε < ‖ξ‖`, we have `‖ξ₀ / ξ‖ ≤ 1`, -- hence `ξ₀ = (ξ₀ / ξ) • ξ ∈ 𝓑` because `ξ ∈ 𝓑`. refine (balanced_core_balanced _).smul_mem _ hξ, rw [norm_mul, norm_inv, mul_inv_le_iff (norm_pos_iff.mpr hξ0), mul_one], diff --git a/src/topology/algebra/polynomial.lean b/src/topology/algebra/polynomial.lean index 28b88b9438475..949efba5a60ef 100644 --- a/src/topology/algebra/polynomial.lean +++ b/src/topology/algebra/polynomial.lean @@ -21,8 +21,8 @@ In this file we prove the following lemmas. * `polynomial.continuous`: `polynomial.eval` defines a continuous functions; we also prove convenience lemmas `polynomial.continuous_at`, `polynomial.continuous_within_at`, `polynomial.continuous_on`. -* `polynomial.tendsto_norm_at_top`: `λ x, ∥polynomial.eval (z x) p∥` tends to infinity provided that - `λ x, ∥z x∥` tends to infinity and `0 < degree p`; +* `polynomial.tendsto_norm_at_top`: `λ x, ‖polynomial.eval (z x) p‖` tends to infinity provided that + `λ x, ‖z x‖` tends to infinity and `0 < degree p`; * `polynomial.tendsto_abv_eval₂_at_top`, `polynomial.tendsto_abv_at_top`, `polynomial.tendsto_abv_aeval_at_top`: a few versions of the previous statement for `is_absolute_value abv` instead of norm. @@ -121,12 +121,12 @@ tendsto_abv_eval₂_at_top _ abv p hd h₀ hz variables {α R : Type*} [normed_ring R] [is_absolute_value (norm : R → ℝ)] lemma tendsto_norm_at_top (p : R[X]) (h : 0 < degree p) {l : filter α} {z : α → R} - (hz : tendsto (λ x, ∥z x∥) l at_top) : - tendsto (λ x, ∥p.eval (z x)∥) l at_top := + (hz : tendsto (λ x, ‖z x‖) l at_top) : + tendsto (λ x, ‖p.eval (z x)‖) l at_top := p.tendsto_abv_at_top norm h hz lemma exists_forall_norm_le [proper_space R] (p : R[X]) : - ∃ x, ∀ y, ∥p.eval x∥ ≤ ∥p.eval y∥ := + ∃ x, ∀ y, ‖p.eval x‖ ≤ ‖p.eval y‖ := if hp0 : 0 < degree p then p.continuous.norm.exists_forall_le $ p.tendsto_norm_at_top hp0 tendsto_norm_cocompact_at_top else ⟨p.coeff 0, by rw [eq_C_of_degree_le_zero (le_of_not_gt hp0)]; simp⟩ @@ -140,7 +140,7 @@ variables {F K : Type*} [comm_ring F] [normed_field K] open multiset lemma eq_one_of_roots_le {p : F[X]} {f : F →+* K} {B : ℝ} (hB : B < 0) - (h1 : p.monic) (h2 : splits f p) (h3 : ∀ z ∈ (map f p).roots, ∥z∥ ≤ B) : + (h1 : p.monic) (h2 : splits f p) (h3 : ∀ z ∈ (map f p).roots, ‖z‖ ≤ B) : p = 1 := h1.nat_degree_eq_zero_iff_eq_one.mp begin contrapose !hB, @@ -150,8 +150,8 @@ h1.nat_degree_eq_zero_iff_eq_one.mp begin end lemma coeff_le_of_roots_le {p : F[X]} {f : F →+* K} {B : ℝ} (i : ℕ) - (h1 : p.monic) (h2 : splits f p) (h3 : ∀ z ∈ (map f p).roots, ∥z∥ ≤ B) : - ∥ (map f p).coeff i ∥ ≤ B^(p.nat_degree - i) * p.nat_degree.choose i := + (h1 : p.monic) (h2 : splits f p) (h3 : ∀ z ∈ (map f p).roots, ‖z‖ ≤ B) : + ‖ (map f p).coeff i ‖ ≤ B^(p.nat_degree - i) * p.nat_degree.choose i := begin obtain hB | hB := lt_or_le B 0, { rw [eq_one_of_roots_le hB h1 h2 h3, polynomial.map_one, @@ -179,8 +179,8 @@ end /-- The coefficients of the monic polynomials of bounded degree with bounded roots are uniformely bounded. -/ lemma coeff_bdd_of_roots_le {B : ℝ} {d : ℕ} (f : F →+* K) {p : F[X]} - (h1 : p.monic) (h2 : splits f p) (h3 : p.nat_degree ≤ d) (h4 : ∀ z ∈ (map f p).roots, ∥z∥ ≤ B) - (i : ℕ) : ∥(map f p).coeff i∥ ≤ (max B 1) ^ d * d.choose (d / 2) := + (h1 : p.monic) (h2 : splits f p) (h3 : p.nat_degree ≤ d) (h4 : ∀ z ∈ (map f p).roots, ‖z‖ ≤ B) + (i : ℕ) : ‖(map f p).coeff i‖ ≤ (max B 1) ^ d * d.choose (d / 2) := begin obtain hB | hB := le_or_lt 0 B, { apply (coeff_le_of_roots_le i h1 h2 h4).trans, diff --git a/src/topology/algebra/uniform_convergence.lean b/src/topology/algebra/uniform_convergence.lean index 767333f59a91c..5f2b06dbaa9e0 100644 --- a/src/topology/algebra/uniform_convergence.lean +++ b/src/topology/algebra/uniform_convergence.lean @@ -194,7 +194,7 @@ begin { rw mem_ball_zero_iff at ha, rw [smul_hom_class.map_smul, pi.smul_apply], have : φ u x ∈ a⁻¹ • V, - { have ha0 : 0<∥a∥ := norm_pos_iff.mpr ha0, + { have ha0 : 0<‖a‖ := norm_pos_iff.mpr ha0, refine (hr a⁻¹ _) (set.mem_image_of_mem (φ u) hx), rw [norm_inv, le_inv hrpos ha0], exact ha.le }, diff --git a/src/topology/continuous_function/bounded.lean b/src/topology/continuous_function/bounded.lean index d5b3c89982987..66f269c3eda5b 100644 --- a/src/topology/continuous_function/bounded.lean +++ b/src/topology/continuous_function/bounded.lean @@ -690,17 +690,17 @@ variables (f g : α →ᵇ β) {x : α} {C : ℝ} instance : has_norm (α →ᵇ β) := ⟨λu, dist u 0⟩ -lemma norm_def : ∥f∥ = dist f 0 := rfl +lemma norm_def : ‖f‖ = dist f 0 := rfl -/-- The norm of a bounded continuous function is the supremum of `∥f x∥`. +/-- The norm of a bounded continuous function is the supremum of `‖f x‖`. We use `Inf` to ensure that the definition works if `α` has no elements. -/ lemma norm_eq (f : α →ᵇ β) : - ∥f∥ = Inf {C : ℝ | 0 ≤ C ∧ ∀ (x : α), ∥f x∥ ≤ C} := + ‖f‖ = Inf {C : ℝ | 0 ≤ C ∧ ∀ (x : α), ‖f x‖ ≤ C} := by simp [norm_def, bounded_continuous_function.dist_eq] -/-- When the domain is non-empty, we do not need the `0 ≤ C` condition in the formula for ∥f∥ as an +/-- When the domain is non-empty, we do not need the `0 ≤ C` condition in the formula for ‖f‖ as an `Inf`. -/ -lemma norm_eq_of_nonempty [h : nonempty α] : ∥f∥ = Inf {C : ℝ | ∀ (x : α), ∥f x∥ ≤ C} := +lemma norm_eq_of_nonempty [h : nonempty α] : ‖f‖ = Inf {C : ℝ | ∀ (x : α), ‖f x‖ ≤ C} := begin unfreezingI { obtain ⟨a⟩ := h, }, rw norm_eq, @@ -710,45 +710,45 @@ begin exact λ h', le_trans (norm_nonneg (f a)) (h' a), end -@[simp] lemma norm_eq_zero_of_empty [h : is_empty α] : ∥f∥ = 0 := +@[simp] lemma norm_eq_zero_of_empty [h : is_empty α] : ‖f‖ = 0 := dist_zero_of_empty -lemma norm_coe_le_norm (x : α) : ∥f x∥ ≤ ∥f∥ := calc - ∥f x∥ = dist (f x) ((0 : α →ᵇ β) x) : by simp [dist_zero_right] - ... ≤ ∥f∥ : dist_coe_le_dist _ +lemma norm_coe_le_norm (x : α) : ‖f x‖ ≤ ‖f‖ := calc + ‖f x‖ = dist (f x) ((0 : α →ᵇ β) x) : by simp [dist_zero_right] + ... ≤ ‖f‖ : dist_coe_le_dist _ -lemma dist_le_two_norm' {f : γ → β} {C : ℝ} (hC : ∀ x, ∥f x∥ ≤ C) (x y : γ) : +lemma dist_le_two_norm' {f : γ → β} {C : ℝ} (hC : ∀ x, ‖f x‖ ≤ C) (x y : γ) : dist (f x) (f y) ≤ 2 * C := -calc dist (f x) (f y) ≤ ∥f x∥ + ∥f y∥ : dist_le_norm_add_norm _ _ +calc dist (f x) (f y) ≤ ‖f x‖ + ‖f y‖ : dist_le_norm_add_norm _ _ ... ≤ C + C : add_le_add (hC x) (hC y) ... = 2 * C : (two_mul _).symm /-- Distance between the images of any two points is at most twice the norm of the function. -/ -lemma dist_le_two_norm (x y : α) : dist (f x) (f y) ≤ 2 * ∥f∥ := +lemma dist_le_two_norm (x y : α) : dist (f x) (f y) ≤ 2 * ‖f‖ := dist_le_two_norm' f.norm_coe_le_norm x y variable {f} /-- The norm of a function is controlled by the supremum of the pointwise norms -/ -lemma norm_le (C0 : (0 : ℝ) ≤ C) : ∥f∥ ≤ C ↔ ∀x:α, ∥f x∥ ≤ C := +lemma norm_le (C0 : (0 : ℝ) ≤ C) : ‖f‖ ≤ C ↔ ∀x:α, ‖f x‖ ≤ C := by simpa using @dist_le _ _ _ _ f 0 _ C0 lemma norm_le_of_nonempty [nonempty α] - {f : α →ᵇ β} {M : ℝ} : ∥f∥ ≤ M ↔ ∀ x, ∥f x∥ ≤ M := + {f : α →ᵇ β} {M : ℝ} : ‖f‖ ≤ M ↔ ∀ x, ‖f x‖ ≤ M := begin simp_rw [norm_def, ←dist_zero_right], exact dist_le_iff_of_nonempty, end lemma norm_lt_iff_of_compact [compact_space α] - {f : α →ᵇ β} {M : ℝ} (M0 : 0 < M) : ∥f∥ < M ↔ ∀ x, ∥f x∥ < M := + {f : α →ᵇ β} {M : ℝ} (M0 : 0 < M) : ‖f‖ < M ↔ ∀ x, ‖f x‖ < M := begin simp_rw [norm_def, ←dist_zero_right], exact dist_lt_iff_of_compact M0, end lemma norm_lt_iff_of_nonempty_compact [nonempty α] [compact_space α] - {f : α →ᵇ β} {M : ℝ} : ∥f∥ < M ↔ ∀ x, ∥f x∥ < M := + {f : α →ᵇ β} {M : ℝ} : ‖f‖ < M ↔ ∀ x, ‖f x‖ < M := begin simp_rw [norm_def, ←dist_zero_right], exact dist_lt_iff_of_nonempty_compact, @@ -756,28 +756,28 @@ end variable (f) -/-- Norm of `const α b` is less than or equal to `∥b∥`. If `α` is nonempty, -then it is equal to `∥b∥`. -/ -lemma norm_const_le (b : β) : ∥const α b∥ ≤ ∥b∥ := +/-- Norm of `const α b` is less than or equal to `‖b‖`. If `α` is nonempty, +then it is equal to `‖b‖`. -/ +lemma norm_const_le (b : β) : ‖const α b‖ ≤ ‖b‖ := (norm_le (norm_nonneg b)).2 $ λ x, le_rfl -@[simp] lemma norm_const_eq [h : nonempty α] (b : β) : ∥const α b∥ = ∥b∥ := +@[simp] lemma norm_const_eq [h : nonempty α] (b : β) : ‖const α b‖ = ‖b‖ := le_antisymm (norm_const_le b) $ h.elim $ λ x, (const α b).norm_coe_le_norm x /-- Constructing a bounded continuous function from a uniformly bounded continuous function taking values in a normed group. -/ def of_normed_add_comm_group {α : Type u} {β : Type v} [topological_space α] - [seminormed_add_comm_group β] (f : α → β) (Hf : continuous f) (C : ℝ) (H : ∀x, ∥f x∥ ≤ C) : + [seminormed_add_comm_group β] (f : α → β) (Hf : continuous f) (C : ℝ) (H : ∀x, ‖f x‖ ≤ C) : α →ᵇ β := ⟨⟨λn, f n, Hf⟩, ⟨_, dist_le_two_norm' H⟩⟩ @[simp] lemma coe_of_normed_add_comm_group {α : Type u} {β : Type v} [topological_space α] [seminormed_add_comm_group β] - (f : α → β) (Hf : continuous f) (C : ℝ) (H : ∀x, ∥f x∥ ≤ C) : + (f : α → β) (Hf : continuous f) (C : ℝ) (H : ∀x, ‖f x‖ ≤ C) : (of_normed_add_comm_group f Hf C H : α → β) = f := rfl lemma norm_of_normed_add_comm_group_le {f : α → β} (hfc : continuous f) {C : ℝ} (hC : 0 ≤ C) - (hfC : ∀ x, ∥f x∥ ≤ C) : ∥of_normed_add_comm_group f hfc C hfC∥ ≤ C := + (hfC : ∀ x, ‖f x‖ ≤ C) : ‖of_normed_add_comm_group f hfc C hfC‖ ≤ C := (norm_le hC).2 hfC /-- Constructing a bounded continuous function from a uniformly bounded @@ -788,7 +788,7 @@ def of_normed_add_comm_group_discrete {α : Type u} {β : Type v} of_normed_add_comm_group f continuous_of_discrete_topology C H @[simp] lemma coe_of_normed_add_comm_group_discrete {α : Type u} {β : Type v} [topological_space α] - [discrete_topology α] [seminormed_add_comm_group β] (f : α → β) (C : ℝ) (H : ∀x, ∥f x∥ ≤ C) : + [discrete_topology α] [seminormed_add_comm_group β] (f : α → β) (C : ℝ) (H : ∀x, ‖f x‖ ≤ C) : (of_normed_add_comm_group_discrete f C H : α → β) = f := rfl /-- Taking the pointwise norm of a bounded continuous function with values in a @@ -798,29 +798,29 @@ f.comp norm lipschitz_with_one_norm @[simp] lemma coe_norm_comp : (f.norm_comp : α → ℝ) = norm ∘ f := rfl -@[simp] lemma norm_norm_comp : ∥f.norm_comp∥ = ∥f∥ := +@[simp] lemma norm_norm_comp : ‖f.norm_comp‖ = ‖f‖ := by simp only [norm_eq, coe_norm_comp, norm_norm] lemma bdd_above_range_norm_comp : bdd_above $ set.range $ norm ∘ f := (real.bounded_iff_bdd_below_bdd_above.mp $ @bounded_range _ _ _ _ f.norm_comp).2 -lemma norm_eq_supr_norm : ∥f∥ = ⨆ x : α, ∥f x∥ := +lemma norm_eq_supr_norm : ‖f‖ = ⨆ x : α, ‖f x‖ := by simp_rw [norm_def, dist_eq_supr, bounded_continuous_function.coe_zero, pi.zero_apply, dist_zero_right] -/-- If `∥(1 : β)∥ = 1`, then `∥(1 : α →ᵇ β)∥ = 1` if `α` is nonempty. -/ +/-- If `‖(1 : β)‖ = 1`, then `‖(1 : α →ᵇ β)‖ = 1` if `α` is nonempty. -/ instance [nonempty α] [has_one β] [norm_one_class β] : norm_one_class (α →ᵇ β) := { norm_one := by simp only [norm_eq_supr_norm, bounded_continuous_function.coe_one, pi.one_apply, norm_one, csupr_const] } /-- The pointwise opposite of a bounded continuous function is again bounded continuous. -/ instance : has_neg (α →ᵇ β) := -⟨λf, of_normed_add_comm_group (-f) f.continuous.neg ∥f∥ $ λ x, +⟨λf, of_normed_add_comm_group (-f) f.continuous.neg ‖f‖ $ λ x, trans_rel_right _ (norm_neg _) (f.norm_coe_le_norm x)⟩ /-- The pointwise difference of two bounded continuous functions is again bounded continuous. -/ instance : has_sub (α →ᵇ β) := -⟨λf g, of_normed_add_comm_group (f - g) (f.continuous.sub g.continuous) (∥f∥ + ∥g∥) $ λ x, +⟨λf g, of_normed_add_comm_group (f - g) (f.continuous.sub g.continuous) (‖f‖ + ‖g‖) $ λ x, by { simp only [sub_eq_add_neg], exact le_trans (norm_add_le _ _) (add_le_add (f.norm_coe_le_norm x) $ trans_rel_right _ (norm_neg _) (g.norm_coe_le_norm x)) }⟩ @@ -862,33 +862,33 @@ instance : seminormed_add_comm_group (α →ᵇ β) := instance {α β} [topological_space α] [normed_add_comm_group β] : normed_add_comm_group (α →ᵇ β) := { ..bounded_continuous_function.seminormed_add_comm_group } -lemma nnnorm_def : ∥f∥₊ = nndist f 0 := rfl +lemma nnnorm_def : ‖f‖₊ = nndist f 0 := rfl -lemma nnnorm_coe_le_nnnorm (x : α) : ∥f x∥₊ ≤ ∥f∥₊ := norm_coe_le_norm _ _ +lemma nnnorm_coe_le_nnnorm (x : α) : ‖f x‖₊ ≤ ‖f‖₊ := norm_coe_le_norm _ _ -lemma nndist_le_two_nnnorm (x y : α) : nndist (f x) (f y) ≤ 2 * ∥f∥₊ := dist_le_two_norm _ _ _ +lemma nndist_le_two_nnnorm (x y : α) : nndist (f x) (f y) ≤ 2 * ‖f‖₊ := dist_le_two_norm _ _ _ /-- The nnnorm of a function is controlled by the supremum of the pointwise nnnorms -/ -lemma nnnorm_le (C : ℝ≥0) : ∥f∥₊ ≤ C ↔ ∀x:α, ∥f x∥₊ ≤ C := +lemma nnnorm_le (C : ℝ≥0) : ‖f‖₊ ≤ C ↔ ∀x:α, ‖f x‖₊ ≤ C := norm_le C.prop -lemma nnnorm_const_le (b : β) : ∥const α b∥₊ ≤ ∥b∥₊ := +lemma nnnorm_const_le (b : β) : ‖const α b‖₊ ≤ ‖b‖₊ := norm_const_le _ -@[simp] lemma nnnorm_const_eq [h : nonempty α] (b : β) : ∥const α b∥₊ = ∥b∥₊ := +@[simp] lemma nnnorm_const_eq [h : nonempty α] (b : β) : ‖const α b‖₊ = ‖b‖₊ := subtype.ext $ norm_const_eq _ -lemma nnnorm_eq_supr_nnnorm : ∥f∥₊ = ⨆ x : α, ∥f x∥₊ := +lemma nnnorm_eq_supr_nnnorm : ‖f‖₊ = ⨆ x : α, ‖f x‖₊ := subtype.ext $ (norm_eq_supr_norm f).trans $ by simp_rw [nnreal.coe_supr, coe_nnnorm] -lemma abs_diff_coe_le_dist : ∥f x - g x∥ ≤ dist f g := +lemma abs_diff_coe_le_dist : ‖f x - g x‖ ≤ dist f g := by { rw dist_eq_norm, exact (f - g).norm_coe_le_norm x } lemma coe_le_coe_add_dist {f g : α →ᵇ ℝ} : f x ≤ g x + dist f g := sub_le_iff_le_add'.1 $ (abs_le.1 $ @dist_coe_le_dist _ _ _ _ f g x).2 lemma norm_comp_continuous_le [topological_space γ] (f : α →ᵇ β) (g : C(γ, α)) : - ∥f.comp_continuous g∥ ≤ ∥f∥ := + ‖f.comp_continuous g‖ ≤ ‖f‖ := ((lipschitz_comp_continuous g).dist_le_mul f 0).trans $ by rw [nnreal.coe_one, one_mul, dist_zero_right] @@ -1022,11 +1022,11 @@ linear_map.mk_continuous { to_fun := λ f, of_normed_add_comm_group (g ∘ f) (g.continuous.comp f.continuous) - (∥g∥ * ∥f∥) + (‖g‖ * ‖f‖) (λ x, (g.le_op_norm_of_le (f.norm_coe_le_norm x))), map_add' := λ f g, by ext; simp, map_smul' := λ c f, by ext; simp } - ∥g∥ + ‖g‖ (λ f, norm_of_normed_add_comm_group_le _ (mul_nonneg (norm_nonneg g) (norm_nonneg f)) _) @[simp] lemma _root_.continuous_linear_map.comp_left_continuous_bounded_apply (g : β →L[𝕜] γ) @@ -1052,7 +1052,7 @@ section semi_normed variables [non_unital_semi_normed_ring R] instance : has_mul (α →ᵇ R) := -{ mul := λ f g, of_normed_add_comm_group (f * g) (f.continuous.mul g.continuous) (∥f∥ * ∥g∥) $ λ x, +{ mul := λ f g, of_normed_add_comm_group (f * g) (f.continuous.mul g.continuous) (‖f‖ * ‖g‖) $ λ x, le_trans (norm_mul_le (f x) (g x)) $ mul_le_mul (f.norm_coe_le_norm x) (g.norm_coe_le_norm x) (norm_nonneg _) (norm_nonneg _) } @@ -1192,9 +1192,9 @@ functions from `α` to `𝕜`. -/ instance has_smul' : has_smul (α →ᵇ 𝕜) (α →ᵇ β) := ⟨λ (f : α →ᵇ 𝕜) (g : α →ᵇ β), of_normed_add_comm_group (λ x, (f x) • (g x)) -(f.continuous.smul g.continuous) (∥f∥ * ∥g∥) (λ x, calc - ∥f x • g x∥ ≤ ∥f x∥ * ∥g x∥ : normed_space.norm_smul_le _ _ - ... ≤ ∥f∥ * ∥g∥ : mul_le_mul (f.norm_coe_le_norm _) (g.norm_coe_le_norm _) (norm_nonneg _) +(f.continuous.smul g.continuous) (‖f‖ * ‖g‖) (λ x, calc + ‖f x • g x‖ ≤ ‖f x‖ * ‖g x‖ : normed_space.norm_smul_le _ _ + ... ≤ ‖f‖ * ‖g‖ : mul_le_mul (f.norm_coe_le_norm _) (g.norm_coe_le_norm _) (norm_nonneg _) (norm_nonneg _)) ⟩ instance module' : module (α →ᵇ 𝕜) (α →ᵇ β) := @@ -1205,7 +1205,7 @@ module.of_core $ mul_smul := λ c₁ c₂ f, ext $ λ x, mul_smul _ _ _, one_smul := λ f, ext $ λ x, one_smul 𝕜 (f x) } -lemma norm_smul_le (f : α →ᵇ 𝕜) (g : α →ᵇ β) : ∥f • g∥ ≤ ∥f∥ * ∥g∥ := +lemma norm_smul_le (f : α →ᵇ 𝕜) (g : α →ᵇ β) : ‖f • g‖ ≤ ‖f‖ * ‖g‖ := norm_of_normed_add_comm_group_le _ (mul_nonneg (norm_nonneg _) (norm_nonneg _)) _ /- TODO: When `normed_module` has been added to `normed_space.basic`, the above facts @@ -1356,7 +1356,7 @@ instance : normed_lattice_add_comm_group (α →ᵇ β) := solid := begin intros f g h, - have i1: ∀ t, ∥f t∥ ≤ ∥g t∥ := λ t, solid (h t), + have i1: ∀ t, ‖f t‖ ≤ ‖g t‖ := λ t, solid (h t), rw norm_le (norm_nonneg _), exact λ t, (i1 t).trans (norm_coe_le_norm g t), end, @@ -1380,7 +1380,7 @@ bounded_continuous_function.comp _ continuous `ℝ≥0`-valued function. -/ def nnnorm (f : α →ᵇ ℝ) : α →ᵇ ℝ≥0 := bounded_continuous_function.comp _ - (show lipschitz_with 1 (λ (x : ℝ), ∥x∥₊), from lipschitz_with_one_norm) f + (show lipschitz_with 1 (λ (x : ℝ), ‖x‖₊), from lipschitz_with_one_norm) f @[simp] lemma nnnorm_coe_fun_eq (f : α →ᵇ ℝ) : ⇑(f.nnnorm) = has_nnnorm.nnnorm ∘ ⇑f := rfl diff --git a/src/topology/continuous_function/compact.lean b/src/topology/continuous_function/compact.lean index d113fe1c3e47d..a9e1834733f46 100644 --- a/src/topology/continuous_function/compact.lean +++ b/src/topology/continuous_function/compact.lean @@ -147,10 +147,10 @@ instance : has_norm C(α, E) := { norm := λ x, dist x 0 } @[simp] lemma _root_.bounded_continuous_function.norm_mk_of_compact (f : C(α, E)) : - ∥mk_of_compact f∥ = ∥f∥ := rfl + ‖mk_of_compact f‖ = ‖f‖ := rfl @[simp] lemma _root_.bounded_continuous_function.norm_to_continuous_map_eq (f : α →ᵇ E) : - ∥f.to_continuous_map∥ = ∥f∥ := + ‖f.to_continuous_map‖ = ‖f‖ := rfl open bounded_continuous_function @@ -168,42 +168,42 @@ variables (f : C(α, E)) -- The corresponding lemmas for `bounded_continuous_function` are stated with `{f}`, -- and so can not be used in dot notation. -lemma norm_coe_le_norm (x : α) : ∥f x∥ ≤ ∥f∥ := +lemma norm_coe_le_norm (x : α) : ‖f x‖ ≤ ‖f‖ := (mk_of_compact f).norm_coe_le_norm x /-- Distance between the images of any two points is at most twice the norm of the function. -/ -lemma dist_le_two_norm (x y : α) : dist (f x) (f y) ≤ 2 * ∥f∥ := +lemma dist_le_two_norm (x y : α) : dist (f x) (f y) ≤ 2 * ‖f‖ := (mk_of_compact f).dist_le_two_norm x y /-- The norm of a function is controlled by the supremum of the pointwise norms -/ -lemma norm_le {C : ℝ} (C0 : (0 : ℝ) ≤ C) : ∥f∥ ≤ C ↔ ∀x:α, ∥f x∥ ≤ C := +lemma norm_le {C : ℝ} (C0 : (0 : ℝ) ≤ C) : ‖f‖ ≤ C ↔ ∀x:α, ‖f x‖ ≤ C := @bounded_continuous_function.norm_le _ _ _ _ (mk_of_compact f) _ C0 -lemma norm_le_of_nonempty [nonempty α] {M : ℝ} : ∥f∥ ≤ M ↔ ∀ x, ∥f x∥ ≤ M := +lemma norm_le_of_nonempty [nonempty α] {M : ℝ} : ‖f‖ ≤ M ↔ ∀ x, ‖f x‖ ≤ M := @bounded_continuous_function.norm_le_of_nonempty _ _ _ _ _ (mk_of_compact f) _ -lemma norm_lt_iff {M : ℝ} (M0 : 0 < M) : ∥f∥ < M ↔ ∀ x, ∥f x∥ < M := +lemma norm_lt_iff {M : ℝ} (M0 : 0 < M) : ‖f‖ < M ↔ ∀ x, ‖f x‖ < M := @bounded_continuous_function.norm_lt_iff_of_compact _ _ _ _ _ (mk_of_compact f) _ M0 -theorem nnnorm_lt_iff {M : ℝ≥0} (M0 : 0 < M) : ∥f∥₊ < M ↔ ∀ (x : α), ∥f x∥₊ < M := +theorem nnnorm_lt_iff {M : ℝ≥0} (M0 : 0 < M) : ‖f‖₊ < M ↔ ∀ (x : α), ‖f x‖₊ < M := f.norm_lt_iff M0 lemma norm_lt_iff_of_nonempty [nonempty α] {M : ℝ} : - ∥f∥ < M ↔ ∀ x, ∥f x∥ < M := + ‖f‖ < M ↔ ∀ x, ‖f x‖ < M := @bounded_continuous_function.norm_lt_iff_of_nonempty_compact _ _ _ _ _ _ (mk_of_compact f) _ lemma nnnorm_lt_iff_of_nonempty [nonempty α] {M : ℝ≥0} : - ∥f∥₊ < M ↔ ∀ x, ∥f x∥₊ < M := + ‖f‖₊ < M ↔ ∀ x, ‖f x‖₊ < M := f.norm_lt_iff_of_nonempty -lemma apply_le_norm (f : C(α, ℝ)) (x : α) : f x ≤ ∥f∥ := +lemma apply_le_norm (f : C(α, ℝ)) (x : α) : f x ≤ ‖f‖ := le_trans (le_abs.mpr (or.inl (le_refl (f x)))) (f.norm_coe_le_norm x) -lemma neg_norm_le_apply (f : C(α, ℝ)) (x : α) : -∥f∥ ≤ f x := +lemma neg_norm_le_apply (f : C(α, ℝ)) (x : α) : -‖f‖ ≤ f x := le_trans (neg_le_neg (f.norm_coe_le_norm x)) (neg_le.mp (neg_le_abs_self (f x))) -lemma norm_eq_supr_norm : ∥f∥ = ⨆ x : α, ∥f x∥ := +lemma norm_eq_supr_norm : ‖f‖ = ⨆ x : α, ‖f x‖ := (mk_of_compact f).norm_eq_supr_norm end @@ -409,7 +409,7 @@ variables {X : Type*} [topological_space X] [t2_space X] [locally_compact_space variables {E : Type*} [normed_add_comm_group E] [complete_space E] lemma summable_of_locally_summable_norm {ι : Type*} {F : ι → C(X, E)} - (hF : ∀ K : compacts X, summable (λ i, ∥(F i).restrict K∥)) : + (hF : ∀ K : compacts X, summable (λ i, ‖(F i).restrict K‖)) : summable F := begin refine (continuous_map.exists_tendsto_compact_open_iff_forall _).2 (λ K hK, _), diff --git a/src/topology/continuous_function/ideals.lean b/src/topology/continuous_function/ideals.lean index 3555557dd5340..53c3dd8fb9a7f 100644 --- a/src/topology/continuous_function/ideals.lean +++ b/src/topology/continuous_function/ideals.lean @@ -188,8 +188,8 @@ begin replace hε := (show (0 : ℝ≥0) < ε, from hε), simp_rw dist_nndist, norm_cast, - -- Let `t := {x : X | ε / 2 ≤ ∥f x∥₊}}` which is closed and disjoint from `set_of_ideal I`. - set t := {x : X | ε / 2 ≤ ∥f x∥₊}, + -- Let `t := {x : X | ε / 2 ≤ ‖f x‖₊}}` which is closed and disjoint from `set_of_ideal I`. + set t := {x : X | ε / 2 ≤ ‖f x‖₊}, have ht : is_closed t := is_closed_le continuous_const (map_continuous f).nnnorm, have htI : disjoint t (set_of_ideal I)ᶜ, { refine set.subset_compl_iff_disjoint_left.mp (λ x hx, _), @@ -197,8 +197,8 @@ begin using (nnnorm_eq_zero.mpr (mem_ideal_of_set.mp hf hx)).trans_lt (half_pos hε), }, /- It suffices to produce `g : C(X, ℝ≥0)` which takes values in `[0,1]` and is constantly `1` on `t` such that when composed with the natural embedding of `ℝ≥0` into `𝕜` lies in the ideal `I`. - Indeed, then `∥f - f * ↑g∥ ≤ ∥f * (1 - ↑g)∥ ≤ ⨆ ∥f * (1 - ↑g) x∥`. When `x ∉ t`, `∥f x∥ < ε / 2` - and `∥(1 - ↑g) x∥ ≤ 1`, and when `x ∈ t`, `(1 - ↑g) x = 0`, and clearly `f * ↑g ∈ I`. -/ + Indeed, then `‖f - f * ↑g‖ ≤ ‖f * (1 - ↑g)‖ ≤ ⨆ ‖f * (1 - ↑g) x‖`. When `x ∉ t`, `‖f x‖ < ε / 2` + and `‖(1 - ↑g) x‖ ≤ 1`, and when `x ∈ t`, `(1 - ↑g) x = 0`, and clearly `f * ↑g ∈ I`. -/ suffices : ∃ g : C(X, ℝ≥0), (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g ∈ I ∧ (∀ x, g x ≤ 1) ∧ t.eq_on g 1, { obtain ⟨g, hgI, hg, hgt⟩ := this, @@ -210,28 +210,28 @@ begin { simpa only [hgt hx, comp_apply, pi.one_apply, continuous_map.coe_coe, algebra_map_clm_apply, map_one, mul_one, sub_self, nnnorm_zero] using hε, }, { refine lt_of_le_of_lt _ (half_lt_self hε), - have := calc ∥((1 - (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g) x : 𝕜)∥₊ - = ∥1 - algebra_map ℝ≥0 𝕜 (g x)∥₊ + have := calc ‖((1 - (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g) x : 𝕜)‖₊ + = ‖1 - algebra_map ℝ≥0 𝕜 (g x)‖₊ : by simp only [continuous_map.coe_sub, continuous_map.coe_one, coe_comp, continuous_map.coe_coe, pi.sub_apply, pi.one_apply, function.comp_app, algebra_map_clm_apply] - ... = ∥algebra_map ℝ≥0 𝕜 (1 - g x)∥₊ + ... = ‖algebra_map ℝ≥0 𝕜 (1 - g x)‖₊ : by simp only [algebra.algebra_map_eq_smul_one, nnreal.smul_def, nnreal.coe_sub (hg x), sub_smul, nonneg.coe_one, one_smul] ... ≤ 1 : (nnnorm_algebra_map_nnreal 𝕜 (1 - g x)).trans_le tsub_le_self, - calc ∥f x - f x * (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g x∥₊ - = ∥f x * (1 - (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g) x∥₊ + calc ‖f x - f x * (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g x‖₊ + = ‖f x * (1 - (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g) x‖₊ : by simp only [mul_sub, continuous_map.coe_sub, continuous_map.coe_one, pi.sub_apply, pi.one_apply, mul_one] - ... ≤ (ε / 2) * ∥(1 - (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g) x∥₊ + ... ≤ (ε / 2) * ‖(1 - (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g) x‖₊ : (nnnorm_mul_le _ _).trans (mul_le_mul_right' - (not_le.mp $ show ¬ ε / 2 ≤ ∥f x∥₊, from hx).le _) + (not_le.mp $ show ¬ ε / 2 ≤ ‖f x‖₊, from hx).le _) ... ≤ ε / 2 : by simpa only [mul_one] using mul_le_mul_left' this _, } }, /- There is some `g' : C(X, ℝ≥0)` which is strictly positive on `t` such that the composition `↑g` with the natural embedding of `ℝ≥0` into `𝕜` lies in `I`. This follows from compactness of `t` and that we can do it in any neighborhood of a point `x ∈ t`. Indeed, since `x ∈ t`, then - `fₓ x ≠ 0` for some `fₓ ∈ I` and so `λ y, ∥(star fₓ * fₓ) y∥₊` is strictly posiive in a - neighborhood of `y`. Moreover, `(∥(star fₓ * fₓ) y∥₊ : 𝕜) = (star fₓ * fₓ) y`, so composition of + `fₓ x ≠ 0` for some `fₓ ∈ I` and so `λ y, ‖(star fₓ * fₓ) y‖₊` is strictly posiive in a + neighborhood of `y`. Moreover, `(‖(star fₓ * fₓ) y‖₊ : 𝕜) = (star fₓ * fₓ) y`, so composition of this map with the natural embedding is just `star fₓ * fₓ ∈ I`. -/ have : ∃ g' : C(X, ℝ≥0), (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g' ∈ I ∧ (∀ x ∈ t, 0 < g' x), { refine @is_compact.induction_on _ _ _ ht.is_compact (λ s, ∃ g' : C(X, ℝ≥0), @@ -257,7 +257,7 @@ begin obtain ⟨g, hI, hgx⟩ := hx, have := (map_continuous g).continuous_at.eventually_ne hgx, refine ⟨{y : X | g y ≠ 0} ∩ t, mem_nhds_within_iff_exists_mem_nhds_inter.mpr - ⟨_, this, set.subset.rfl⟩, ⟨⟨λ x, ∥g x∥₊ ^ 2, (map_continuous g).nnnorm.pow 2⟩, _, + ⟨_, this, set.subset.rfl⟩, ⟨⟨λ x, ‖g x‖₊ ^ 2, (map_continuous g).nnnorm.pow 2⟩, _, λ x hx, pow_pos (norm_pos_iff.mpr hx.1) 2⟩⟩, convert I.mul_mem_left (star g) hI, ext, diff --git a/src/topology/continuous_function/stone_weierstrass.lean b/src/topology/continuous_function/stone_weierstrass.lean index 2ad73af260012..4ebcf29084c25 100644 --- a/src/topology/continuous_function/stone_weierstrass.lean +++ b/src/topology/continuous_function/stone_weierstrass.lean @@ -15,7 +15,7 @@ separates points, then it is dense. We argue as follows. * In any subalgebra `A` of `C(X, ℝ)`, if `f ∈ A`, then `abs f ∈ A.topological_closure`. - This follows from the Weierstrass approximation theorem on `[-∥f∥, ∥f∥]` by + This follows from the Weierstrass approximation theorem on `[-‖f‖, ‖f‖]` by approximating `abs` uniformly thereon by polynomials. * This ensures that `A.topological_closure` is actually a sublattice: if it contains `f` and `g`, then it contains the pointwise supremum `f ⊔ g` @@ -48,16 +48,16 @@ variables {X : Type*} [topological_space X] [compact_space X] open_locale polynomial /-- -Turn a function `f : C(X, ℝ)` into a continuous map into `set.Icc (-∥f∥) (∥f∥)`, +Turn a function `f : C(X, ℝ)` into a continuous map into `set.Icc (-‖f‖) (‖f‖)`, thereby explicitly attaching bounds. -/ -def attach_bound (f : C(X, ℝ)) : C(X, set.Icc (-∥f∥) (∥f∥)) := +def attach_bound (f : C(X, ℝ)) : C(X, set.Icc (-‖f‖) (‖f‖)) := { to_fun := λ x, ⟨f x, ⟨neg_norm_le_apply f x, apply_le_norm f x⟩⟩ } @[simp] lemma attach_bound_apply_coe (f : C(X, ℝ)) (x : X) : ((attach_bound f) x : ℝ) = f x := rfl lemma polynomial_comp_attach_bound (A : subalgebra ℝ C(X, ℝ)) (f : A) (g : ℝ[X]) : - (g.to_continuous_map_on (set.Icc (-∥f∥) ∥f∥)).comp (f : C(X, ℝ)).attach_bound = + (g.to_continuous_map_on (set.Icc (-‖f‖) ‖f‖)).comp (f : C(X, ℝ)).attach_bound = polynomial.aeval f g := begin ext, @@ -74,23 +74,23 @@ Given a continuous function `f` in a subalgebra of `C(X, ℝ)`, postcomposing by gives another function in `A`. This lemma proves something slightly more subtle than this: -we take `f`, and think of it as a function into the restricted target `set.Icc (-∥f∥) ∥f∥)`, +we take `f`, and think of it as a function into the restricted target `set.Icc (-‖f‖) ‖f‖)`, and then postcompose with a polynomial function on that interval. This is in fact the same situation as above, and so also gives a function in `A`. -/ lemma polynomial_comp_attach_bound_mem (A : subalgebra ℝ C(X, ℝ)) (f : A) (g : ℝ[X]) : - (g.to_continuous_map_on (set.Icc (-∥f∥) ∥f∥)).comp (f : C(X, ℝ)).attach_bound ∈ A := + (g.to_continuous_map_on (set.Icc (-‖f‖) ‖f‖)).comp (f : C(X, ℝ)).attach_bound ∈ A := begin rw polynomial_comp_attach_bound, apply set_like.coe_mem, end theorem comp_attach_bound_mem_closure - (A : subalgebra ℝ C(X, ℝ)) (f : A) (p : C(set.Icc (-∥f∥) (∥f∥), ℝ)) : + (A : subalgebra ℝ C(X, ℝ)) (f : A) (p : C(set.Icc (-‖f‖) (‖f‖), ℝ)) : p.comp (attach_bound f) ∈ A.topological_closure := begin -- `p` itself is in the closure of polynomials, by the Weierstrass theorem, - have mem_closure : p ∈ (polynomial_functions (set.Icc (-∥f∥) (∥f∥))).topological_closure := + have mem_closure : p ∈ (polynomial_functions (set.Icc (-‖f‖) (‖f‖))).topological_closure := continuous_map_mem_polynomial_functions_closure _ _ p, -- and so there are polynomials arbitrarily close. have frequently_mem_polynomials := mem_closure_iff_frequently.mp mem_closure, @@ -110,10 +110,10 @@ end theorem abs_mem_subalgebra_closure (A : subalgebra ℝ C(X, ℝ)) (f : A) : (f : C(X, ℝ)).abs ∈ A.topological_closure := begin - let M := ∥f∥, + let M := ‖f‖, let f' := attach_bound (f : C(X, ℝ)), - let abs : C(set.Icc (-∥f∥) (∥f∥), ℝ) := - { to_fun := λ x : set.Icc (-∥f∥) (∥f∥), |(x : ℝ)| }, + let abs : C(set.Icc (-‖f‖) (‖f‖), ℝ) := + { to_fun := λ x : set.Icc (-‖f‖) (‖f‖), |(x : ℝ)| }, change (abs.comp f') ∈ A.topological_closure, apply comp_attach_bound_mem_closure, end @@ -331,7 +331,7 @@ every real-valued continuous function on `X` is within any `ε > 0` of some elem theorem exists_mem_subalgebra_near_continuous_map_of_separates_points (A : subalgebra ℝ C(X, ℝ)) (w : A.separates_points) (f : C(X, ℝ)) (ε : ℝ) (pos : 0 < ε) : - ∃ (g : A), ∥(g : C(X, ℝ)) - f∥ < ε := + ∃ (g : A), ‖(g : C(X, ℝ)) - f‖ < ε := begin have w := mem_closure_iff_frequently.mp (continuous_map_mem_subalgebra_closure_of_separates_points A w f), @@ -351,7 +351,7 @@ every real-valued continuous function on `X` is within any `ε > 0` of some elem theorem exists_mem_subalgebra_near_continuous_of_separates_points (A : subalgebra ℝ C(X, ℝ)) (w : A.separates_points) (f : X → ℝ) (c : continuous f) (ε : ℝ) (pos : 0 < ε) : - ∃ (g : A), ∀ x, ∥g x - f x∥ < ε := + ∃ (g : A), ∀ x, ‖g x - f x‖ < ε := begin obtain ⟨g, b⟩ := exists_mem_subalgebra_near_continuous_map_of_separates_points A w ⟨f, c⟩ ε pos, use g, diff --git a/src/topology/continuous_function/weierstrass.lean b/src/topology/continuous_function/weierstrass.lean index 988124d481403..d60d6f9f63f88 100644 --- a/src/topology/continuous_function/weierstrass.lean +++ b/src/topology/continuous_function/weierstrass.lean @@ -103,7 +103,7 @@ Every real-valued continuous function on `[a,b]` is within any `ε > 0` of some -/ theorem exists_polynomial_near_continuous_map (a b : ℝ) (f : C(set.Icc a b, ℝ)) (ε : ℝ) (pos : 0 < ε) : - ∃ (p : ℝ[X]), ∥p.to_continuous_map_on _ - f∥ < ε := + ∃ (p : ℝ[X]), ‖p.to_continuous_map_on _ - f‖ < ε := begin have w := mem_closure_iff_frequently.mp (continuous_map_mem_polynomial_functions_closure _ _ f), rw metric.nhds_basis_ball.frequently_iff at w, diff --git a/src/topology/continuous_function/zero_at_infty.lean b/src/topology/continuous_function/zero_at_infty.lean index dae214f8e216b..41448a5c1765b 100644 --- a/src/topology/continuous_function/zero_at_infty.lean +++ b/src/topology/continuous_function/zero_at_infty.lean @@ -396,7 +396,7 @@ normed_add_comm_group.induced C₀(α, β) (α →ᵇ β) (⟨to_bcf, rfl, λ x (to_bcf_injective α β) @[simp] -lemma norm_to_bcf_eq_norm {f : C₀(α, β)} : ∥f.to_bcf∥ = ∥f∥ := rfl +lemma norm_to_bcf_eq_norm {f : C₀(α, β)} : ‖f.to_bcf‖ = ‖f‖ := rfl instance : normed_space 𝕜 C₀(α, β) := { norm_smul_le := λ k f, (norm_smul k f.to_bcf).le } diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index 7202015659d06..77b12cd7519c7 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -654,7 +654,7 @@ end topological_space section liminf lemma exists_frequently_lt_of_liminf_ne_top - {ι : Type*} {l : filter ι} {x : ι → ℝ} (hx : liminf (λ n, (∥x n∥₊ : ℝ≥0∞)) l ≠ ∞) : + {ι : Type*} {l : filter ι} {x : ι → ℝ} (hx : liminf (λ n, (‖x n‖₊ : ℝ≥0∞)) l ≠ ∞) : ∃ R, ∃ᶠ n in l, x n < R := begin by_contra h, @@ -665,7 +665,7 @@ begin end lemma exists_frequently_lt_of_liminf_ne_top' - {ι : Type*} {l : filter ι} {x : ι → ℝ} (hx : liminf (λ n, (∥x n∥₊ : ℝ≥0∞)) l ≠ ∞) : + {ι : Type*} {l : filter ι} {x : ι → ℝ} (hx : liminf (λ n, (‖x n‖₊ : ℝ≥0∞)) l ≠ ∞) : ∃ R, ∃ᶠ n in l, R < x n := begin by_contra h, @@ -677,7 +677,7 @@ end lemma exists_upcrossings_of_not_bounded_under {ι : Type*} {l : filter ι} {x : ι → ℝ} - (hf : liminf (λ i, (∥x i∥₊ : ℝ≥0∞)) l ≠ ∞) + (hf : liminf (λ i, (‖x i‖₊ : ℝ≥0∞)) l ≠ ∞) (hbdd : ¬ is_bounded_under (≤) l (λ i, |x i|)) : ∃ a b : ℚ, a < b ∧ (∃ᶠ i in l, x i < a) ∧ (∃ᶠ i in l, ↑b < x i) := begin diff --git a/src/topology/tietze_extension.lean b/src/topology/tietze_extension.lean index db91d89119046..89bd87a13f8d9 100644 --- a/src/topology/tietze_extension.lean +++ b/src/topology/tietze_extension.lean @@ -41,26 +41,26 @@ namespace bounded_continuous_function /-- One step in the proof of the Tietze extension theorem. If `e : C(X, Y)` is a closed embedding of a topological space into a normal topological space and `f : X →ᵇ ℝ` is a bounded continuous -function, then there exists a bounded continuous function `g : Y →ᵇ ℝ` of the norm `∥g∥ ≤ ∥f∥ / 3` -such that the distance between `g ∘ e` and `f` is at most `(2 / 3) * ∥f∥`. -/ +function, then there exists a bounded continuous function `g : Y →ᵇ ℝ` of the norm `‖g‖ ≤ ‖f‖ / 3` +such that the distance between `g ∘ e` and `f` is at most `(2 / 3) * ‖f‖`. -/ lemma tietze_extension_step (f : X →ᵇ ℝ) (e : C(X, Y)) (he : closed_embedding e) : - ∃ g : Y →ᵇ ℝ, ∥g∥ ≤ ∥f∥ / 3 ∧ dist (g.comp_continuous e) f ≤ (2 / 3) * ∥f∥ := + ∃ g : Y →ᵇ ℝ, ‖g‖ ≤ ‖f‖ / 3 ∧ dist (g.comp_continuous e) f ≤ (2 / 3) * ‖f‖ := begin have h3 : (0 : ℝ) < 3 := by norm_num1, have h23 : 0 < (2 / 3 : ℝ) := by norm_num1, -- In the trivial case `f = 0`, we take `g = 0` rcases eq_or_ne f 0 with (rfl|hf), { use 0, simp }, - replace hf : 0 < ∥f∥ := norm_pos_iff.2 hf, - /- Otherwise, the closed sets `e '' (f ⁻¹' (Iic (-∥f∥ / 3)))` and `e '' (f ⁻¹' (Ici (∥f∥ / 3)))` - are disjoint, hence by Urysohn's lemma there exists a function `g` that is equal to `-∥f∥ / 3` - on the former set and is equal to `∥f∥ / 3` on the latter set. This function `g` satisfies the + replace hf : 0 < ‖f‖ := norm_pos_iff.2 hf, + /- Otherwise, the closed sets `e '' (f ⁻¹' (Iic (-‖f‖ / 3)))` and `e '' (f ⁻¹' (Ici (‖f‖ / 3)))` + are disjoint, hence by Urysohn's lemma there exists a function `g` that is equal to `-‖f‖ / 3` + on the former set and is equal to `‖f‖ / 3` on the latter set. This function `g` satisfies the assertions of the lemma. -/ - have hf3 : -∥f∥ / 3 < ∥f∥ / 3, from (div_lt_div_right h3).2 (left.neg_lt_self hf), - have hc₁ : is_closed (e '' (f ⁻¹' (Iic (-∥f∥ / 3)))), + have hf3 : -‖f‖ / 3 < ‖f‖ / 3, from (div_lt_div_right h3).2 (left.neg_lt_self hf), + have hc₁ : is_closed (e '' (f ⁻¹' (Iic (-‖f‖ / 3)))), from he.is_closed_map _ (is_closed_Iic.preimage f.continuous), - have hc₂ : is_closed (e '' (f ⁻¹' (Ici (∥f∥ / 3)))), + have hc₂ : is_closed (e '' (f ⁻¹' (Ici (‖f‖ / 3)))), from he.is_closed_map _ (is_closed_Ici.preimage f.continuous), - have hd : disjoint (e '' (f ⁻¹' (Iic (-∥f∥ / 3)))) (e '' (f ⁻¹' (Ici (∥f∥ / 3)))), + have hd : disjoint (e '' (f ⁻¹' (Iic (-‖f‖ / 3)))) (e '' (f ⁻¹' (Ici (‖f‖ / 3)))), { refine disjoint_image_of_injective he.inj (disjoint.preimage _ _), rwa [Iic_disjoint_Ici, not_le] }, rcases exists_bounded_mem_Icc_of_closed_of_le hc₁ hc₂ hd hf3.le with ⟨g, hg₁, hg₂, hgf⟩, @@ -68,21 +68,21 @@ begin { refine (norm_le $ div_nonneg hf.le h3.le).mpr (λ y, _), simpa [abs_le, neg_div] using hgf y }, { refine (dist_le $ mul_nonneg h23.le hf.le).mpr (λ x, _), - have hfx : -∥f∥ ≤ f x ∧ f x ≤ ∥f∥, + have hfx : -‖f‖ ≤ f x ∧ f x ≤ ‖f‖, by simpa only [real.norm_eq_abs, abs_le] using f.norm_coe_le_norm x, - cases le_total (f x) (-∥f∥ / 3) with hle₁ hle₁, - { calc |g (e x) - f x| = -∥f∥ / 3 - f x: + cases le_total (f x) (-‖f‖ / 3) with hle₁ hle₁, + { calc |g (e x) - f x| = -‖f‖ / 3 - f x: by rw [hg₁ (mem_image_of_mem _ hle₁), abs_of_nonneg (sub_nonneg.2 hle₁)] - ... ≤ (2 / 3) * ∥f∥ : by linarith }, - { cases le_total (f x) (∥f∥ / 3) with hle₂ hle₂, + ... ≤ (2 / 3) * ‖f‖ : by linarith }, + { cases le_total (f x) (‖f‖ / 3) with hle₂ hle₂, { simp only [neg_div] at *, calc dist (g (e x)) (f x) ≤ |g (e x)| + |f x| : dist_le_norm_add_norm _ _ - ... ≤ ∥f∥ / 3 + ∥f∥ / 3 : + ... ≤ ‖f‖ / 3 + ‖f‖ / 3 : add_le_add (abs_le.2 $ hgf _) (abs_le.2 ⟨hle₁, hle₂⟩) - ... = (2 / 3) * ∥f∥ : by linarith }, - { calc |g (e x) - f x| = f x - ∥f∥ / 3 : + ... = (2 / 3) * ‖f‖ : by linarith }, + { calc |g (e x) - f x| = f x - ‖f‖ / 3 : by rw [hg₂ (mem_image_of_mem _ hle₂), abs_sub_comm, abs_of_nonneg (sub_nonneg.2 hle₂)] - ... ≤ (2 / 3) * ∥f∥ : by linarith } } } + ... ≤ (2 / 3) * ‖f‖ : by linarith } } } end /-- **Tietze extension theorem** for real-valued bounded continuous maps, a version with a closed @@ -91,7 +91,7 @@ into a normal topological space and `f : X →ᵇ ℝ` is a bounded continuous f a bounded continuous function `g : Y →ᵇ ℝ` of the same norm such that `g ∘ e = f`. -/ lemma exists_extension_norm_eq_of_closed_embedding' (f : X →ᵇ ℝ) (e : C(X, Y)) (he : closed_embedding e) : - ∃ g : Y →ᵇ ℝ, ∥g∥ = ∥f∥ ∧ g.comp_continuous e = f := + ∃ g : Y →ᵇ ℝ, ‖g‖ = ‖f‖ ∧ g.comp_continuous e = f := begin /- For the proof, we iterate `tietze_extension_step`. Each time we apply it to the difference between the previous approximation and `f`. -/ @@ -100,29 +100,29 @@ begin have g0 : g 0 = 0 := rfl, have g_succ : ∀ n, g (n + 1) = g n + F (f - (g n).comp_continuous e), from λ n, function.iterate_succ_apply' _ _ _, - have hgf : ∀ n, dist ((g n).comp_continuous e) f ≤ (2 / 3) ^ n * ∥f∥, + have hgf : ∀ n, dist ((g n).comp_continuous e) f ≤ (2 / 3) ^ n * ‖f‖, { intro n, induction n with n ihn, { simp [g0] }, { rw [g_succ n, add_comp_continuous, ← dist_sub_right, add_sub_cancel', pow_succ, mul_assoc], refine (hF_dist _).trans (mul_le_mul_of_nonneg_left _ (by norm_num1)), rwa ← dist_eq_norm' } }, - have hg_dist : ∀ n, dist (g n) (g (n + 1)) ≤ 1 / 3 * ∥f∥ * (2 / 3) ^ n, + have hg_dist : ∀ n, dist (g n) (g (n + 1)) ≤ 1 / 3 * ‖f‖ * (2 / 3) ^ n, { intro n, - calc dist (g n) (g (n + 1)) = ∥F (f - (g n).comp_continuous e)∥ : + calc dist (g n) (g (n + 1)) = ‖F (f - (g n).comp_continuous e)‖ : by rw [g_succ, dist_eq_norm', add_sub_cancel'] - ... ≤ ∥f - (g n).comp_continuous e∥ / 3 : hF_norm _ + ... ≤ ‖f - (g n).comp_continuous e‖ / 3 : hF_norm _ ... = (1 / 3) * dist ((g n).comp_continuous e) f : by rw [dist_eq_norm', one_div, div_eq_inv_mul] - ... ≤ (1 / 3) * ((2 / 3) ^ n * ∥f∥) : + ... ≤ (1 / 3) * ((2 / 3) ^ n * ‖f‖) : mul_le_mul_of_nonneg_left (hgf n) (by norm_num1) - ... = 1 / 3 * ∥f∥ * (2 / 3) ^ n : by ac_refl }, + ... = 1 / 3 * ‖f‖ * (2 / 3) ^ n : by ac_refl }, have hg_cau : cauchy_seq g, from cauchy_seq_of_le_geometric _ _ (by norm_num1) hg_dist, have : tendsto (λ n, (g n).comp_continuous e) at_top (𝓝 $ (lim at_top g).comp_continuous e), from ((continuous_comp_continuous e).tendsto _).comp hg_cau.tendsto_lim, have hge : (lim at_top g).comp_continuous e = f, { refine tendsto_nhds_unique this (tendsto_iff_dist_tendsto_zero.2 _), refine squeeze_zero (λ _, dist_nonneg) hgf _, - rw ← zero_mul (∥f∥), + rw ← zero_mul (‖f‖), refine (tendsto_pow_at_top_nhds_0_of_lt_1 _ _).mul tendsto_const_nhds; norm_num1 }, refine ⟨lim at_top g, le_antisymm _ _, hge⟩, { rw [← dist_zero_left, ← g0], @@ -138,7 +138,7 @@ into a normal topological space and `f : X →ᵇ ℝ` is a bounded continuous f a bounded continuous function `g : Y →ᵇ ℝ` of the same norm such that `g ∘ e = f`. -/ lemma exists_extension_norm_eq_of_closed_embedding (f : X →ᵇ ℝ) {e : X → Y} (he : closed_embedding e) : - ∃ g : Y →ᵇ ℝ, ∥g∥ = ∥f∥ ∧ g ∘ e = f := + ∃ g : Y →ᵇ ℝ, ‖g‖ = ‖f‖ ∧ g ∘ e = f := begin rcases exists_extension_norm_eq_of_closed_embedding' f ⟨e, he.continuous⟩ he with ⟨g, hg, rfl⟩, exact ⟨g, hg, rfl⟩ @@ -149,7 +149,7 @@ set. If `f` is a bounded continuous real-valued function defined on a closed set topological space, then it can be extended to a bounded continuous function of the same norm defined on the whole space. -/ lemma exists_norm_eq_restrict_eq_of_closed {s : set Y} (f : s →ᵇ ℝ) (hs : is_closed s) : - ∃ g : Y →ᵇ ℝ, ∥g∥ = ∥f∥ ∧ g.restrict s = f := + ∃ g : Y →ᵇ ℝ, ‖g‖ = ‖f‖ ∧ g.restrict s = f := exists_extension_norm_eq_of_closed_embedding' f ((continuous_map.id _).restrict s) (closed_embedding_subtype_coe hs) @@ -168,7 +168,7 @@ begin rcases exists_extension_norm_eq_of_closed_embedding (f - const X ((a + b) / 2)) he with ⟨g, hgf, hge⟩, refine ⟨const Y ((a + b) / 2) + g, λ y, _, _⟩, - { suffices : ∥f - const X ((a + b) / 2)∥ ≤ (b - a) / 2, + { suffices : ‖f - const X ((a + b) / 2)‖ ≤ (b - a) / 2, by simpa [real.Icc_eq_closed_ball, add_mem_closed_ball_iff_norm] using (norm_coe_le_norm g y).trans (hgf.trans_le this), refine (norm_le $ div_nonneg (sub_nonneg.2 hle) zero_le_two).2 (λ x, _), diff --git a/test/positivity.lean b/test/positivity.lean index 9ec81a8302f9d..57e2c76c2d101 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -221,7 +221,7 @@ example {α : Type*} [fintype α] [nonempty α] : 0 < fintype.card α := by posi example {r : ℝ} : 0 < real.exp r := by positivity -example {V : Type*} [normed_add_comm_group V] (x : V) : 0 ≤ ∥x∥ := by positivity +example {V : Type*} [normed_add_comm_group V] (x : V) : 0 ≤ ‖x‖ := by positivity example [metric_space α] (x y : α) : 0 ≤ dist x y := by positivity example [metric_space α] {s : set α} : 0 ≤ metric.diam s := by positivity