Skip to content

Commit

Permalink
feat: port MeasureTheory.Function.L2Space (#4737)
Browse files Browse the repository at this point in the history
Co-authored-by: Johan Commelin <johan@commelin.net>
  • Loading branch information
Ruben-VandeVelde and jcommelin committed Jun 7, 2023
1 parent e7eeaa6 commit 2d45c79
Show file tree
Hide file tree
Showing 5 changed files with 345 additions and 9 deletions.
1 change: 1 addition & 0 deletions Mathlib.lean
Expand Up @@ -2011,6 +2011,7 @@ import Mathlib.MeasureTheory.Function.Egorov
import Mathlib.MeasureTheory.Function.EssSup
import Mathlib.MeasureTheory.Function.Floor
import Mathlib.MeasureTheory.Function.L1Space
import Mathlib.MeasureTheory.Function.L2Space
import Mathlib.MeasureTheory.Function.LocallyIntegrable
import Mathlib.MeasureTheory.Function.LpOrder
import Mathlib.MeasureTheory.Function.LpSeminorm
Expand Down
2 changes: 2 additions & 0 deletions Mathlib/Algebra/GroupPower/Order.lean
Expand Up @@ -680,6 +680,8 @@ theorem sq_pos_iff (a : R) : 0 < a ^ 2 ↔ a ≠ 0 :=

variable {x y : R}

-- Porting note: added `simp` to replace `pow_bit0_abs`
@[simp]
theorem sq_abs (x : R) : |x| ^ 2 = x ^ 2 := by simpa only [sq] using abs_mul_abs_self x
#align sq_abs sq_abs

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Analysis/InnerProductSpace/Adjoint.lean
Expand Up @@ -239,7 +239,7 @@ instance : CstarRing (E →L[𝕜] E) :=
refine' op_norm_le_bound _ (Real.sqrt_nonneg _) fun x => _
have :=
calc
re ⟪(A† * A) x, x⟫ ≤ ‖(A† * A) x‖ * ‖x‖ := re_inner_le_norm 𝕜 _ _
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]
Expand Down
19 changes: 11 additions & 8 deletions Mathlib/Analysis/InnerProductSpace/Basic.lean
Expand Up @@ -1023,7 +1023,8 @@ theorem real_inner_self_eq_norm_sq (x : F) : ⟪x, x⟫_ℝ = ‖x‖ ^ 2 := by
rw [pow_two, real_inner_self_eq_norm_mul_norm]
#align real_inner_self_eq_norm_sq real_inner_self_eq_norm_sq

variable (𝕜)
-- Porting note: this was present in mathlib3 but seemingly didn't do anything.
-- variable (𝕜)

/-- Expand the square -/
theorem norm_add_sq (x y : E) : ‖x + y‖ ^ 2 = ‖x‖ ^ 2 + 2 * re ⟪x, y⟫ + ‖y‖ ^ 2 := by
Expand All @@ -1049,7 +1050,7 @@ alias norm_add_sq_real ← norm_add_pow_two_real
theorem norm_add_mul_self (x y : E) :
‖x + y‖ * ‖x + y‖ = ‖x‖ * ‖x‖ + 2 * re ⟪x, y⟫ + ‖y‖ * ‖y‖ := by
repeat' rw [← sq (M := ℝ)]
exact norm_add_sq _ _ _
exact norm_add_sq _ _
#align norm_add_mul_self norm_add_mul_self

/-- Expand the square -/
Expand Down Expand Up @@ -1080,7 +1081,7 @@ alias norm_sub_sq_real ← norm_sub_pow_two_real
theorem norm_sub_mul_self (x y : E) :
‖x - y‖ * ‖x - y‖ = ‖x‖ * ‖x‖ - 2 * re ⟪x, y⟫ + ‖y‖ * ‖y‖ := by
repeat' rw [← sq (M := ℝ)]
exact norm_sub_sq _ _ _
exact norm_sub_sq _ _
#align norm_sub_mul_self norm_sub_mul_self

/-- Expand the square -/
Expand All @@ -1098,23 +1099,25 @@ theorem norm_inner_le_norm (x y : E) : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ :=
#align norm_inner_le_norm norm_inner_le_norm

theorem nnnorm_inner_le_nnnorm (x y : E) : ‖⟪x, y⟫‖₊ ≤ ‖x‖₊ * ‖y‖₊ :=
norm_inner_le_norm 𝕜 x y
norm_inner_le_norm x y
#align nnnorm_inner_le_nnnorm nnnorm_inner_le_nnnorm

theorem re_inner_le_norm (x y : E) : re ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ :=
le_trans (re_le_norm (inner x y)) (norm_inner_le_norm 𝕜 x y)
le_trans (re_le_norm (inner x y)) (norm_inner_le_norm x y)
#align re_inner_le_norm re_inner_le_norm

/-- Cauchy–Schwarz inequality with norm -/
theorem abs_real_inner_le_norm (x y : F) : |⟪x, y⟫_ℝ| ≤ ‖x‖ * ‖y‖ :=
(Real.norm_eq_abs _).ge.trans (norm_inner_le_norm x y)
(Real.norm_eq_abs _).ge.trans (norm_inner_le_norm x y)
#align abs_real_inner_le_norm abs_real_inner_le_norm

/-- Cauchy–Schwarz inequality with norm -/
theorem real_inner_le_norm (x y : F) : ⟪x, y⟫_ℝ ≤ ‖x‖ * ‖y‖ :=
le_trans (le_abs_self _) (abs_real_inner_le_norm _ _)
#align real_inner_le_norm real_inner_le_norm

variable (𝕜)

theorem parallelogram_law_with_norm (x y : E) :
‖x + y‖ * ‖x + y‖ + ‖x - y‖ * ‖x - y‖ = 2 * (‖x‖ * ‖x‖ + ‖y‖ * ‖y‖) := by
simp only [← @inner_self_eq_norm_mul_norm 𝕜]
Expand Down Expand Up @@ -1785,7 +1788,7 @@ set_option linter.uppercaseLean3 false in
@[simp]
theorem innerSL_apply_norm (x : E) : ‖innerSL 𝕜 x‖ = ‖x‖ := by
refine'
le_antisymm ((innerSL 𝕜 x).op_norm_le_bound (norm_nonneg _) fun y => norm_inner_le_norm _ _ _) _
le_antisymm ((innerSL 𝕜 x).op_norm_le_bound (norm_nonneg _) fun y => norm_inner_le_norm _ _) _
rcases eq_or_ne x 0 with (rfl | h)
· simp
· refine' (mul_le_mul_right (norm_pos_iff.2 h)).mp _
Expand Down Expand Up @@ -1864,7 +1867,7 @@ theorem _root_.isBoundedBilinearMap_inner [NormedSpace ℝ E] :
bound :=
1, zero_lt_one, fun x y => by
rw [one_mul]
exact norm_inner_le_norm _ x y⟩ }
exact norm_inner_le_norm x y⟩ }
#align is_bounded_bilinear_map_inner isBoundedBilinearMap_inner

end Norm
Expand Down

0 comments on commit 2d45c79

Please sign in to comment.