Built with Alectryon, running Lean4 v4.19.0. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use instead of Ctrl.
/-
 - Created in 2024 by Gaëtan Serré
-/

/-
 - https://github.com/gaetanserre/SBS-Proofs
-/

import Mathlib.Analysis.Normed.Field.Instances
import Mathlib.Data.Real.StarOrdered
import Mathlib.Topology.MetricSpace.Polish
import SBSProofs.RKHS.Basic
import SBSProofs.Utils

open Classical MeasureTheory

local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y)

set_option maxHeartbeats 600000

variable {d : ℕ} {Ω : Set (Vector ℝ d)} [MeasureSpace Ω]

def L2 (μ : Measure Ω) [IsFiniteMeasure μ] := {f : Ω | MemLp f 2 μ}

def eigen := {v : ℕ //  i, 0 <= v i}

def f_repr (v : eigen) (e : ℕ  Ω  ℝ) (f : Ω  ℝ) (a : ℕ  ℝ) := (f = λ x  (∑' i, (v.1 i) * (a i) * (e i x)))  ( x, Summable (λ i  (v.1 i) * (a i) * (e i x)))

/-
  We define a set of functions that depends on a finite measure μ. Each function is representable by a infinite sum.
-/

def H (v : eigen) (e : ℕ  Ω  ℝ) (μ : Measure Ω) [IsFiniteMeasure μ] := {f | f  L2 μ   (a : ℕ  ℝ), (f_repr v e f a)  Summable (λ i  (v.1 i) * (a i)^2)}

def set_repr {v : eigen} {e : ℕ  Ω  ℝ} {μ : Measure Ω} [IsFiniteMeasure μ] (f : H v e μ) := {a : ℕ | (f_repr v e f.1 a)  (Summable (λ i  (v.1 i) * (a i)^2))}

Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

(set_repr f).Nonempty
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
left✝: f L2 μ
a:
ha: f_repr v e (f) a Summable fun i => v i * a i ^ 2

intro.intro
(set_repr f).Nonempty
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

(set_repr f).Nonempty
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
left✝: f L2 μ
a:
ha: f_repr v e (f) a Summable fun i => v i * a i ^ 2

h
a set_repr f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

(set_repr f).Nonempty
Goals accomplished!

Goals accomplished! 🐙
/- We assume that the the representative of each function in H is unique (property of v and e). -/ axiom set_repr_unique {v : eigen} {e : ℕ Ω ℝ} {μ : Measure Ω} [IsFiniteMeasure μ] {f : H v e μ} {a b : ℕ ℝ} (ha : a set_repr f) (hb : b set_repr f) : a = b
Goals accomplished!
/- We assume that the product of two representative is summable. -/ axiom product_summable {v : eigen} {e : ℕ Ω ℝ} {μ : Measure Ω} [IsFiniteMeasure μ] (f g : H v e μ) : Summable (λ i (v.1 i) * ((set_repr_ne f).some i) * ((set_repr_ne g).some i)) /- We define the multiplication between a real number and a function in H as the pointwise product. We show that the result lies in H. -/ namespace Ring variable {v : eigen} {e : ℕ Ω ℝ} {μ : Measure Ω} [IsFiniteMeasure μ] (f g : H v e μ)
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

f_repr v e (fun x => a * f x) fun i => a * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω

f_repr v e (fun x => a * f x) fun i => a * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

f_repr v e (fun x => a * f x) fun i => a * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:

f_repr v e (fun x => a * f x) fun i => a * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

f_repr v e (fun x => a * f x) fun i => a * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:

f_repr v e (fun x => a * f x) fun i => a * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

f_repr v e (fun x => a * f x) fun i => a * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f

_root_.f_repr v e (fun x => a * f x) fun i => a * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

f_repr v e (fun x => a * f x) fun i => a * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f

left
(fun x => a * f x) = fun x => ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f
(x : ↑Ω), Summable fun i => v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

f_repr v e (fun x => a * f x) fun i => a * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f

left
(fun x => a * f x) = fun x => ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f

left
(fun x => a * f x) = fun x => ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f
(x : ↑Ω), Summable fun i => v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f
x: ↑Ω

left.h
a * f x = ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f

left
(fun x => a * f x) = fun x => ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x

left.h.intro.intro
a * f x = ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f

left
(fun x => a * f x) = fun x => ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x

left.h.intro.intro
a * f x = ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x

g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x

g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x

g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω

h
g x = ∑' (i : ℕ), v i * a * h i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x

g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω

h
g x = ∑' (i : ℕ), v i * a * h i * e i x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
i:

v i * a * h i * e i x = a * v i * h i * e i x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x

g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x

h
g x = ∑' (i : ℕ), v i * a * h i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x

h
g x = ∑' (i : ℕ), a * v i * h i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x

h
g x = ∑' (i : ℕ), a * v i * h i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x

g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x

h
g x = ∑' (i : ℕ), a * v i * h i * e i x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x

(i : ℕ), a * v i * h i * e i x = a * (v i * h i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x

(i : ℕ), a * v i * h i * e i x = a * (v i * h i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x

(i : ℕ), a * v i * h i * e i x = a * (v i * h i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x
i:

a * v i * h i * e i x = a * (v i * h i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x

(i : ℕ), a * v i * h i * e i x = a * (v i * h i * e i x)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x

g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x
summand_comm: (i : ℕ), a * v i * h i * e i x = a * (v i * h i * e i x)

h
g x = ∑' (i : ℕ), a * v i * h i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x
summand_comm: (i : ℕ), a * v i * h i * e i x = a * (v i * h i * e i x)

h
g x = ∑' (i : ℕ), a * (v i * h i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x
summand_comm: (i : ℕ), a * v i * h i * e i x = a * (v i * h i * e i x)

h
g x = ∑' (i : ℕ), a * (v i * h i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x

g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x
summand_comm: (i : ℕ), a * v i * h i * e i x = a * (v i * h i * e i x)

h
g x = ∑' (i : ℕ), a * (v i * h i * e i x)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x
summand_comm: (i : ℕ), a * v i * h i * e i x = a * (v i * h i * e i x)

h
g x = ∑' (i : ℕ), a * (v i * h i * e i x)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x

g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x
summand_comm: (i : ℕ), a * v i * h i * e i x = a * (v i * h i * e i x)
const_out: ∑' (i : ℕ), a * (fun i => v i * h i * e i x) i = a * ∑' (i : ℕ), (fun i => v i * h i * e i x) i

h
g x = ∑' (i : ℕ), a * (v i * h i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x
summand_comm: (i : ℕ), a * v i * h i * e i x = a * (v i * h i * e i x)
const_out: ∑' (i : ℕ), a * (fun i => v i * h i * e i x) i = a * ∑' (i : ℕ), (fun i => v i * h i * e i x) i

h
g x = a * ∑' (i : ℕ), (fun i => v i * h i * e i x) i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x
summand_comm: (i : ℕ), a * v i * h i * e i x = a * (v i * h i * e i x)
const_out: ∑' (i : ℕ), a * (fun i => v i * h i * e i x) i = a * ∑' (i : ℕ), (fun i => v i * h i * e i x) i

h
g x = ∑' (i : ℕ), a * (v i * h i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x✝: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
x: ↑Ω
comm: (i : ℕ), v i * a * h i * e i x = a * v i * h i * e i x
summand_comm: (i : ℕ), a * v i * h i * e i x = a * (v i * h i * e i x)
const_out: ∑' (i : ℕ), a * (fun i => v i * h i * e i x) i = a * ∑' (i : ℕ), (fun i => v i * h i * e i x) i

h
g x = a * f x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f

left
(fun x => a * f x) = fun x => ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
g_eq_tsum: g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
g_x: g x = ∑' (i : ℕ), v i * a * h i * e i x

left.h.intro.intro
a * f x = ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f

left
(fun x => a * f x) = fun x => ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
g_eq_tsum: g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
g_x: g x = ∑' (i : ℕ), v i * a * h i * e i x

left.h.intro.intro
a * f x = ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
g_eq_tsum: g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
g_x: g x = ∑' (i : ℕ), v i * a * h i * e i x

left.h.intro.intro
a * f x = ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
g_eq_tsum: g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
g_x: g x = ∑' (i : ℕ), v i * a * h i * e i x
i:

v i * a * h i * e i x = v i * (a * h i) * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
g_eq_tsum: g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
g_x: g x = ∑' (i : ℕ), v i * a * h i * e i x
i:

v i * a * h i * e i x = v i * (a * h i) * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
g_eq_tsum: g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
g_x: g x = ∑' (i : ℕ), v i * a * h i * e i x

(i : ℕ), v i * a * h i * e i x = v i * (a * h i) * e i x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
g_eq_tsum: g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
g_x: g x = ∑' (i : ℕ), v i * (a * h i) * e i x

left.h.intro.intro
a * f x = ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f

left
(fun x => a * f x) = fun x => ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
g_eq_tsum: g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
g_x: g x = ∑' (i : ℕ), v i * (a * h i) * e i x

left.h.intro.intro
a * f x = ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
g_eq_tsum: g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
g_x: g x = ∑' (i : ℕ), v i * (a * h i) * e i x

left.h.intro.intro
a * f x = ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
g_eq_tsum: g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
g_x: g x = ∑' (i : ℕ), v i * (a * h i) * e i x
i:

a * h i = g_h i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
g_eq_tsum: g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
g_x: g x = ∑' (i : ℕ), v i * (a * h i) * e i x
i:

a * h i = g_h i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
g_eq_tsum: g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
g_x: g x = ∑' (i : ℕ), v i * (a * h i) * e i x

(i : ℕ), a * h i = g_h i
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
x: ↑Ω
right✝¹: Summable fun i => v i * h i ^ 2
f_repr: f = fun x => ∑' (i : ℕ), v i * h i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * h i * e i x
g_eq_tsum: g = fun x => ∑' (i : ℕ), v i * a * h i * e i x
g_x: g x = ∑' (i : ℕ), v i * g_h i * e i x

left.h.intro.intro
a * f x = ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f

left
(fun x => a * f x) = fun x => ∑' (i : ℕ), v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

f_repr v e (fun x => a * f x) fun i => a * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f
x: ↑Ω

right
Summable fun i => v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

f_repr v e (fun x => a * f x) fun i => a * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f
x: ↑Ω

right
Summable fun i => v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f
x: ↑Ω

(fun i => v i * (fun i => a * h i) i * e i x) = fun i => a * (v i * h i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f
x: ↑Ω

(fun i => v i * (fun i => a * h i) i * e i x) = fun i => a * (v i * h i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f
x: ↑Ω

(fun i => v i * (fun i => a * h i) i * e i x) = fun i => a * (v i * h i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f
x: ↑Ω
i:

h
v i * (fun i => a * h i) i * e i x = a * (v i * h i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f
x: ↑Ω

(fun i => v i * (fun i => a * h i) i * e i x) = fun i => a * (v i * h i * e i x)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

f_repr v e (fun x => a * f x) fun i => a * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f
x: ↑Ω
remove_function: (fun i => v i * (fun i => a * h i) i * e i x) = fun i => a * (v i * h i * e i x)

right
Summable fun i => v i * (fun i => a * ⋯.some i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f
x: ↑Ω
remove_function: (fun i => v i * (fun i => a * h i) i * e i x) = fun i => a * (v i * h i * e i x)

right
Summable fun i => a * (v i * h i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
h:= ⋯.some:
g_h:= fun i => a * h i:
f_repr: h set_repr f
x: ↑Ω
remove_function: (fun i => v i * (fun i => a * h i) i * e i x) = fun i => a * (v i * h i * e i x)

right
Summable fun i => a * (v i * h i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

f_repr v e (fun x => a * f x) fun i => a * ⋯.some i
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

Summable fun i => v i * (fun i => a * ⋯.some i) i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:

Summable fun i => v i * (fun i => a * ⋯.some i) i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

Summable fun i => v i * (fun i => a * ⋯.some i) i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:

Summable fun i => v i * (fun i => a * ⋯.some i) i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

Summable fun i => v i * (fun i => a * ⋯.some i) i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
set_repr: h _root_.set_repr f

Summable fun i => v i * (fun i => a * ⋯.some i) i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

Summable fun i => v i * (fun i => a * ⋯.some i) i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2

intro
Summable fun i => v i * (fun i => a * ⋯.some i) i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

Summable fun i => v i * (fun i => a * ⋯.some i) i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2

intro
Summable fun i => v i * (fun i => a * ⋯.some i) i ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2

(fun i => v i * g_h i ^ 2) = fun i => a ^ 2 * (v i * h i ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2

(fun i => v i * g_h i ^ 2) = fun i => a ^ 2 * (v i * h i ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2

(fun i => v i * g_h i ^ 2) = fun i => a ^ 2 * (v i * h i ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2
i:

h
v i * g_h i ^ 2 = a ^ 2 * (v i * h i ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2

(fun i => v i * g_h i ^ 2) = fun i => a ^ 2 * (v i * h i ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2
i:

h
v i * g_h i ^ 2 = a ^ 2 * (v i * h i ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2
i:

h
v i * g_h i ^ 2 = a ^ 2 * (v i * h i ^ 2)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2
i:

h
v i * (a * h i) ^ 2 = a ^ 2 * (v i * h i ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2

(fun i => v i * g_h i ^ 2) = fun i => a ^ 2 * (v i * h i ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2
i:

h
v i * (a * h i) ^ 2 = a ^ 2 * (v i * h i ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2
i:

h
v i * (a ^ 2 * h i ^ 2) = a ^ 2 * (v i * h i ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2
i:

h
v i * (a ^ 2 * h i ^ 2) = a ^ 2 * (v i * h i ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2

(fun i => v i * g_h i ^ 2) = fun i => a ^ 2 * (v i * h i ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2
i:

h
v i * (a ^ 2 * h i ^ 2) = a ^ 2 * (v i * h i ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2
i:

h
v i * (a ^ 2 * h i ^ 2) = a ^ 2 * (v i * h i ^ 2)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

Summable fun i => v i * (fun i => a * ⋯.some i) i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2
comm_fun: (fun i => v i * g_h i ^ 2) = fun i => a ^ 2 * (v i * h i ^ 2)

intro
Summable fun i => v i * (fun i => a * ⋯.some i) i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2
comm_fun: (fun i => v i * g_h i ^ 2) = fun i => a ^ 2 * (v i * h i ^ 2)

intro
Summable fun i => a ^ 2 * (v i * h i ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
h:= ⋯.some:
g_h:= fun i => a * h i:
left✝: f_repr v e (f) h
h_summable: Summable fun i => v i * h i ^ 2
comm_fun: (fun i => v i * g_h i ^ 2) = fun i => a ^ 2 * (v i * h i ^ 2)

intro
Summable fun i => a ^ 2 * (v i * h i ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

Summable fun i => v i * (fun i => a * ⋯.some i) i ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

(fun x => a * f x) H v e μ
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω

(fun x => a * f x) H v e μ
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

(fun x => a * f x) H v e μ
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
g_in_L2: g L2 μ

(fun x => a * f x) H v e μ
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

(fun x => a * f x) H v e μ
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
g_in_L2: g L2 μ
h:= ⋯.some:

(fun x => a * f x) H v e μ
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

(fun x => a * f x) H v e μ
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:
g:= fun x => a * f x: ↑Ω
g_in_L2: g L2 μ
h:= ⋯.some:
g_h:= fun i => a * h i:

(fun x => a * f x) H v e μ
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:

(fun x => a * f x) H v e μ
Goals accomplished!

Goals accomplished! 🐙
instance : HMul ℝ (H v e μ) (H v e μ) where hMul := λ a f λ x a * f.1 x, mul_in_H f a⟩ instance : HSMul ℝ (H v e μ) (H v e μ) where hSMul := λ r f r * f end Ring /- We define the sum between two functions in H as the pointwise sum. We show that the result lies in H. We also define the 0 function of H. We show several properties on the addition. -/ namespace Group variable {v : eigen} {e : ℕ Ω ℝ} {μ : Measure Ω} [IsFiniteMeasure μ] (f g : H v e μ)
Goals accomplished!
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

Summable fun i => v i * (⋯.some i + ⋯.some i) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:

Summable fun i => v i * (⋯.some i + ⋯.some i) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

Summable fun i => v i * (⋯.some i + ⋯.some i) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:

Summable fun i => v i * (⋯.some i + ⋯.some i) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

Summable fun i => v i * (⋯.some i + ⋯.some i) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:

Summable fun i => v i * (⋯.some i + ⋯.some i) ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:

(fun i => v i * (a_f i + a_g i) ^ 2) = fun i => v i * a_f i * a_f i + 2 * (v i * a_f i * a_g i) + v i * a_g i * a_g i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:

(fun i => v i * (a_f i + a_g i) ^ 2) = fun i => v i * a_f i * a_f i + 2 * (v i * a_f i * a_g i) + v i * a_g i * a_g i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:

(fun i => v i * (a_f i + a_g i) ^ 2) = fun i => v i * a_f i * a_f i + 2 * (v i * a_f i * a_g i) + v i * a_g i * a_g i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
i:

h
v i * (a_f i + a_g i) ^ 2 = v i * a_f i * a_f i + 2 * (v i * a_f i * a_g i) + v i * a_g i * a_g i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:

(fun i => v i * (a_f i + a_g i) ^ 2) = fun i => v i * a_f i * a_f i + 2 * (v i * a_f i * a_g i) + v i * a_g i * a_g i
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

Summable fun i => v i * (⋯.some i + ⋯.some i) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
f_eq: (fun i => v i * (a_f i + a_g i) ^ 2) = fun i => v i * a_f i * a_f i + 2 * (v i * a_f i * a_g i) + v i * a_g i * a_g i

Summable fun i => v i * (⋯.some i + ⋯.some i) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
f_eq: (fun i => v i * (a_f i + a_g i) ^ 2) = fun i => v i * a_f i * a_f i + 2 * (v i * a_f i * a_g i) + v i * a_g i * a_g i

Summable fun i => v i * a_f i * a_f i + 2 * (v i * a_f i * a_g i) + v i * a_g i * a_g i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
f_eq: (fun i => v i * (a_f i + a_g i) ^ 2) = fun i => v i * a_f i * a_f i + 2 * (v i * a_f i * a_g i) + v i * a_g i * a_g i

Summable fun i => v i * a_f i * a_f i + 2 * (v i * a_f i * a_g i) + v i * a_g i * a_g i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

Summable fun i => v i * (⋯.some i + ⋯.some i) ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f_repr v e (fun x => f x + g x) fun i => ⋯.some i + ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:

f_repr v e (fun x => f x + g x) fun i => a_f i + ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f_repr v e (fun x => f x + g x) fun i => ⋯.some i + ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:

f_repr v e (fun x => f x + g x) fun i => a_f i + a_g i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f_repr v e (fun x => f x + g x) fun i => ⋯.some i + ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2

intro
f_repr v e (fun x => f x + g x) fun i => a_f i + a_g i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f_repr v e (fun x => f x + g x) fun i => ⋯.some i + ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2

intro.intro
f_repr v e (fun x => f x + g x) fun i => a_f i + a_g i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f_repr v e (fun x => f x + g x) fun i => ⋯.some i + ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2

intro.intro
f_repr v e (fun x => f x + g x) fun i => a_f i + a_g i
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2

intro.intro
f_repr v e (fun x => f x + g x) fun i => a_f i + a_g i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
x: ↑Ω
i:

v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
x: ↑Ω
i:

v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2

intro.intro
f_repr v e (fun x => f x + g x) fun i => a_f i + a_g i
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f_repr v e (fun x => f x + g x) fun i => ⋯.some i + ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x

intro.intro.left
(fun x => f x + g x) = fun x => ∑' (i : ℕ), v i * (fun i => a_f i + a_g i) i * e i x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
(x : ↑Ω), Summable fun i => v i * (fun i => a_f i + a_g i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f_repr v e (fun x => f x + g x) fun i => ⋯.some i + ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x

intro.intro.left
(fun x => f x + g x) = fun x => ∑' (i : ℕ), v i * (fun i => a_f i + a_g i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x

intro.intro.left
(fun x => f x + g x) = fun x => ∑' (i : ℕ), v i * (fun i => a_f i + a_g i) i * e i x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
(x : ↑Ω), Summable fun i => v i * (fun i => a_f i + a_g i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
x: ↑Ω

intro.intro.left.h
f x + g x = ∑' (i : ℕ), v i * (fun i => a_f i + a_g i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x

intro.intro.left
(fun x => f x + g x) = fun x => ∑' (i : ℕ), v i * (fun i => a_f i + a_g i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
x: ↑Ω

intro.intro.left.h
f x + g x = ∑' (i : ℕ), v i * (a_f i + a_g i) * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
x: ↑Ω

intro.intro.left.h
f x + g x = ∑' (i : ℕ), (v i * a_f i * e i x + v i * a_g i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
x: ↑Ω

intro.intro.left.h
f x + g x = ∑' (i : ℕ), (v i * a_f i * e i x + v i * a_g i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x

intro.intro.left
(fun x => f x + g x) = fun x => ∑' (i : ℕ), v i * (fun i => a_f i + a_g i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
x: ↑Ω

intro.intro.left.h
f x + g x = ∑' (i : ℕ), (v i * a_f i * e i x + v i * a_g i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
x: ↑Ω

intro.intro.left.h
f x + g x = ∑' (b : ℕ), v b * ⋯.some b * e b x + ∑' (b : ℕ), v b * ⋯.some b * e b x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
x: ↑Ω

intro.intro.left.h
f x + g x = ∑' (i : ℕ), (v i * a_f i * e i x + v i * a_g i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
x: ↑Ω

intro.intro.left.h
f x + g x = f x + ∑' (b : ℕ), v b * ⋯.some b * e b x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
x: ↑Ω

intro.intro.left.h
f x + g x = ∑' (i : ℕ), (v i * a_f i * e i x + v i * a_g i * e i x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
x: ↑Ω

intro.intro.left.h
f x + g x = f x + g x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
x: ↑Ω

intro.intro.left.h
f x + g x = ∑' (i : ℕ), (v i * a_f i * e i x + v i * a_g i * e i x)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f_repr v e (fun x => f x + g x) fun i => ⋯.some i + ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
x: ↑Ω

intro.intro.right
Summable fun i => v i * (fun i => a_f i + a_g i) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f_repr v e (fun x => f x + g x) fun i => ⋯.some i + ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
x: ↑Ω

intro.intro.right
Summable fun i => v i * (a_f i + a_g i) * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
x: ↑Ω

intro.intro.right
Summable fun i => v i * a_f i * e i x + v i * a_g i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
af_repr: f_repr v e f ⋯.some
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ag_repr: f_repr v e g ⋯.some
right✝: Summable fun i => v i * ⋯.some i ^ 2
summand_distrib: (x : ↑Ω) (i : ℕ), v i * (a_f i + a_g i) * e i x = v i * a_f i * e i x + v i * a_g i * e i x
x: ↑Ω

intro.intro.right
Summable fun i => v i * a_f i * e i x + v i * a_g i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f_repr v e (fun x => f x + g x) fun i => ⋯.some i + ⋯.some i
Goals accomplished!

Goals accomplished! 🐙
/- We define the 0 of H as pointwise 0 function. We show that it lies in H. -/ def zero : Ω ℝ := λ _ 0 omit [MeasureSpace ↑Ω] in
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω

f_repr v e zero fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:

f_repr v e zero fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω

f_repr v e zero fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:

left
zero = fun x => ∑' (i : ℕ), v i * (fun x => 0) i * e i x
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
(x : ↑Ω), Summable fun i => v i * (fun x => 0) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω

f_repr v e zero fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:

left
zero = fun x => ∑' (i : ℕ), v i * (fun x => 0) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:

left
zero = fun x => ∑' (i : ℕ), v i * (fun x => 0) i * e i x
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
(x : ↑Ω), Summable fun i => v i * (fun x => 0) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
x: ↑Ω

left.h
zero x = ∑' (i : ℕ), v i * (fun x => 0) i * e i x
/- have summand_zero : ∀ i, v.1 i * a i * e i x = 0 := by { intro i rw [show v.1 i * a i * e i x = v.1 i * 0 * e i x by rfl] ring } -/
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
x: ↑Ω

left.h
0 = ∑' (i : ℕ), 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:

left
zero = fun x => ∑' (i : ℕ), v i * (fun x => 0) i * e i x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω

f_repr v e zero fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
x: ↑Ω

right
Summable fun i => v i * (fun x => 0) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω

f_repr v e zero fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
x: ↑Ω

right
Summable fun i => v i * (fun x => 0) i * e i x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
x: ↑Ω

(fun i => v i * a i * e i x) = fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
x: ↑Ω

(fun i => v i * a i * e i x) = fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
x: ↑Ω

(fun i => v i * a i * e i x) = fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
x: ↑Ω
i:

h
v i * a i * e i x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
x: ↑Ω

(fun i => v i * a i * e i x) = fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
x: ↑Ω
i:

h
v i * a i * e i x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
x: ↑Ω
i:

h
v i * a i * e i x = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
x: ↑Ω
i:

h
v i * 0 * e i x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
x: ↑Ω

(fun i => v i * a i * e i x) = fun x => 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω

f_repr v e zero fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
x: ↑Ω
null_function: (fun i => v i * a i * e i x) = fun x => 0

right
Summable fun i => v i * (fun x => 0) i * e i x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
x: ↑Ω
null_function: (fun i => v i * a i * e i x) = fun x => 0

right
Summable fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
a:= fun x => 0:
x: ↑Ω
null_function: (fun i => v i * a i * e i x) = fun x => 0

right
Summable fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω

f_repr v e zero fun x => 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
v: eigen

Summable fun i => v i * 0 ^ 2
Goals accomplished!
v: eigen

Summable fun i => v i * 0 ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
v: eigen

(fun i => v i * 0 ^ 2) = fun x => 0
Goals accomplished!
v: eigen

(fun i => v i * 0 ^ 2) = fun x => 0
Goals accomplished!
v: eigen

(fun i => v i * 0 ^ 2) = fun x => 0
Goals accomplished!
v: eigen
i:

h
v i * 0 ^ 2 = 0
Goals accomplished!
v: eigen

(fun i => v i * 0 ^ 2) = fun x => 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
v: eigen

Summable fun i => v i * 0 ^ 2
Goals accomplished!
v: eigen
zero_fun: (fun i => v i * 0 ^ 2) = fun x => 0

Summable fun i => v i * 0 ^ 2
Goals accomplished!
v: eigen
zero_fun: (fun i => v i * 0 ^ 2) = fun x => 0

Summable fun x => 0
Goals accomplished!
v: eigen
zero_fun: (fun i => v i * 0 ^ 2) = fun x => 0

Summable fun x => 0
Goals accomplished!
v: eigen

Summable fun i => v i * 0 ^ 2
Goals accomplished!
v: eigen
zero_fun: (fun i => v i * 0 ^ 2) = fun x => 0
hf: b , (fun x => 0) b = 0

Summable fun x => 0
Goals accomplished!
v: eigen

Summable fun i => v i * 0 ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
instance : Zero (H v e μ) where zero := ⟨zero, zero_in_H⟩
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

⋯.some = fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:

⋯.some = fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

⋯.some = fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:

⋯.some = fun x => 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:

a set_repr 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:

a set_repr 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:

a set_repr 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:

Summable fun i => v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:

a set_repr 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:

Summable fun i => v i * a i ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:

(fun i => v i * a i ^ 2) = fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:

(fun i => v i * a i ^ 2) = fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:

(fun i => v i * a i ^ 2) = fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:
i:

h
v i * a i ^ 2 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:

(fun i => v i * a i ^ 2) = fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:
i:

h
v i * a i ^ 2 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:
i:

h
v i * a i ^ 2 = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:
i:

h
v i * 0 ^ 2 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:

(fun i => v i * a i ^ 2) = fun x => 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:

a set_repr 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:
null_function: (fun i => v i * a i ^ 2) = fun x => 0

Summable fun i => v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:
null_function: (fun i => v i * a i ^ 2) = fun x => 0

Summable fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:
null_function: (fun i => v i * a i ^ 2) = fun x => 0

Summable fun x => 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:= fun x => 0:

a set_repr 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

⋯.some = fun x => 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

(fun x => f x + g x) H v e μ
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
h:= fun i => ⋯.some i + ⋯.some i:

(fun x => f x + g x) H v e μ
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

(fun x => f x + g x) H v e μ
Goals accomplished!

Goals accomplished! 🐙
instance : HAdd (H v e μ) (H v e μ) (H v e μ) where hAdd := λ f g ⟨(λ x f.1 x + g.1 x), add_in_H f g⟩ instance : HSub (H v e μ) (H v e μ) (H v e μ) where hSub := λ f g f + (-1 : ℝ) * g instance : Neg (H v e μ) where neg := λ f (-1 : ℝ) * f
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a + b + c = a + (b + c)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a + b + c) x = (a + (b + c)) x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a + b + c = a + (b + c)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a + b + c) x = (a + (b + c)) x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a + b + c) x = (a + (b + c)) x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
a x + b x + c x = (a + (b + c)) x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a + b + c = a + (b + c)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
a x + b x + c x = (a + (b + c)) x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
a x + b x + c x = (a + (b + c)) x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
a x + b x + c x = a x + (b x + c x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a + b + c = a + (b + c)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

a + b = b + a
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
x: ↑Ω

a.h
(a + b) x = (b + a) x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

a + b = b + a
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
x: ↑Ω

a.h
(a + b) x = (b + a) x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
x: ↑Ω

a.h
(a + b) x = (b + a) x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
x: ↑Ω

a.h
a x + b x = (b + a) x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

a + b = b + a
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
x: ↑Ω

a.h
a x + b x = (b + a) x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
x: ↑Ω

a.h
a x + b x = (b + a) x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
x: ↑Ω

a.h
a x + b x = b x + a x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

a + b = b + a
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a + 0 = a
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
(a + 0) x = a x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a + 0 = a
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
(a + 0) x = a x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
(a + 0) x = a x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
a x + 0 x = a x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a + 0 = a
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
a x + 0 x = a x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
a x + 0 x = a x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
a x + 0 = a x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a + 0 = a
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

0 + a = a
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

0 + a = a
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a + 0 = a
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a + 0 = a
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

0 + a = a
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

-a + a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

-a + a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

-a + a = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

-1 * a + a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

-a + a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
(-1 * a + a) x = 0 x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

-a + a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
(-1 * a + a) x = 0 x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
(-1 * a + a) x = 0 x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
(-1 * a) x + a x = 0 x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

-a + a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
(-1 * a) x + a x = 0 x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
(-1 * a) x + a x = 0 x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
-1 * a x + a x = 0 x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

-a + a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
-1 * a x + a x = 0 x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
-1 * a x + a x = 0 x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
-1 * a x + a x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

-a + a = 0
Goals accomplished!

Goals accomplished! 🐙
end Group /- We define a function : H × H → ℝ. The purpose of the following is to prove that H endowed with this function is a inner product space. -/ noncomputable def H_inner {v : eigen} {e : ℕ Ω ℝ} {μ : Measure Ω} [IsFiniteMeasure μ] (f g : H v e μ) : ℝ := ∑' i, (v.1 i) * ((set_repr_ne f).some i) * ((set_repr_ne g).some i) /- - We show properties on the inner product of H and the induced norm. -/ namespace Inner open Ring Group variable {v : eigen} {e : ℕ Ω ℝ} {μ : Measure Ω} [IsFiniteMeasure μ] (f g : H v e μ) noncomputable instance : Inner ℝ (H v e μ) where inner := H_inner noncomputable instance : Norm (H v e μ) where norm := λ f Real.sqrt (inner f f) noncomputable instance : Dist (H v e μ) where dist := λ f g norm (f - g)
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:

inner (a * f) g = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:

inner (a * f) g = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:

inner (a * f) g = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:

inner (a * f) g = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:

inner (a * f) g = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)

inner (a * f) g = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:

inner (a * f) g = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)

inner (a * f) g = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)

inner (a * f) g = a * inner f g
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)

H_inner (a * f) g = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:

inner (a * f) g = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:

inner (a * f) g = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)

∑' (i : ℕ), v i * h_af i * ⋯.some i = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)

∑' (i : ℕ), v i * h_af i * ⋯.some i = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:

inner (a * f) g = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)

∑' (i : ℕ), v i * h_af i * ⋯.some i = a * inner f g
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)

(i : ℕ), v i * h_af i * ⋯.some i = a * v i * h i * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)

(i : ℕ), v i * h_af i * ⋯.some i = a * v i * h i * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)

(i : ℕ), v i * h_af i * ⋯.some i = a * v i * h i * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)
i:

v i * h_af i * ⋯.some i = a * v i * h i * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)

(i : ℕ), v i * h_af i * ⋯.some i = a * v i * h i * ⋯.some i
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:

inner (a * f) g = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)
comm_summand: (i : ℕ), v i * h_af i * ⋯.some i = a * v i * h i * ⋯.some i

∑' (i : ℕ), v i * h_af i * ⋯.some i = a * inner f g
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)
comm_summand: (i : ℕ), v i * h_af i * ⋯.some i = a * v i * h i * ⋯.some i

(i : ℕ), a * v i * h i * ⋯.some i = a * (fun i => v i * h i * ⋯.some i) i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)
comm_summand: (i : ℕ), v i * h_af i * ⋯.some i = a * v i * h i * ⋯.some i

(i : ℕ), a * v i * h i * ⋯.some i = a * (fun i => v i * h i * ⋯.some i) i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)
comm_summand: (i : ℕ), v i * h_af i * ⋯.some i = a * v i * h i * ⋯.some i

(i : ℕ), a * v i * h i * ⋯.some i = a * (fun i => v i * h i * ⋯.some i) i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)
comm_summand: (i : ℕ), v i * h_af i * ⋯.some i = a * v i * h i * ⋯.some i
i:

a * v i * h i * ⋯.some i = a * (fun i => v i * h i * ⋯.some i) i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)
comm_summand: (i : ℕ), v i * h_af i * ⋯.some i = a * v i * h i * ⋯.some i

(i : ℕ), a * v i * h i * ⋯.some i = a * (fun i => v i * h i * ⋯.some i) i
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:

inner (a * f) g = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)
comm_summand: (i : ℕ), v i * h_af i * ⋯.some i = a * v i * h i * ⋯.some i
lambda_comm: (i : ℕ), a * v i * h i * ⋯.some i = a * (fun i => v i * h i * ⋯.some i) i

∑' (i : ℕ), v i * h_af i * ⋯.some i = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)
comm_summand: (i : ℕ), v i * h_af i * ⋯.some i = a * v i * h i * ⋯.some i
lambda_comm: (i : ℕ), a * v i * h i * ⋯.some i = a * (fun i => v i * h i * ⋯.some i) i

∑' (i : ℕ), a * v i * h i * ⋯.some i = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)
comm_summand: (i : ℕ), v i * h_af i * ⋯.some i = a * v i * h i * ⋯.some i
lambda_comm: (i : ℕ), a * v i * h i * ⋯.some i = a * (fun i => v i * h i * ⋯.some i) i

∑' (i : ℕ), v i * h_af i * ⋯.some i = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)
comm_summand: (i : ℕ), v i * h_af i * ⋯.some i = a * v i * h i * ⋯.some i
lambda_comm: (i : ℕ), a * v i * h i * ⋯.some i = a * (fun i => v i * h i * ⋯.some i) i

∑' (i : ℕ), a * (v i * h i * ⋯.some i) = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:
h:= ⋯.some:
h_af:= fun i => a * h i:
h_af_in: h_af set_repr (a * f)
comm_summand: (i : ℕ), v i * h_af i * ⋯.some i = a * v i * h i * ⋯.some i
lambda_comm: (i : ℕ), a * v i * h i * ⋯.some i = a * (fun i => v i * h i * ⋯.some i) i

∑' (i : ℕ), a * (v i * h i * ⋯.some i) = a * inner f g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
a:

inner (a * f) g = a * inner f g
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)

inner (f + g) h = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)

inner (f + g) h = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)

inner (f + g) h = inner f h + inner g h
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)

H_inner (f + g) h = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)

inner (f + g) h = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)

inner (f + g) h = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)

inner (f + g) h = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)

inner (f + g) h = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:

∑' (i : ℕ), v i * ⋯.some i * a_h i = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)

inner (f + g) h = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:

∑' (i : ℕ), v i * ⋯.some i * a_h i = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)

inner (f + g) h = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)

∑' (i : ℕ), v i * ⋯.some i * a_h i = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)

inner (f + g) h = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)

∑' (i : ℕ), v i * ⋯.some i * a_h i = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)

∑' (i : ℕ), v i * a_fg i * a_h i = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)

∑' (i : ℕ), v i * a_fg i * a_h i = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)

inner (f + g) h = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)

∑' (i : ℕ), v i * a_fg i * a_h i = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)

∑' (i : ℕ), v i * a_fg i * a_h i = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)
i:

v i * a_fg i * a_h i = v i * (a_f i + a_g i) * a_h i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)
i:

v i * a_fg i * a_h i = v i * (a_f i + a_g i) * a_h i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)

(i : ℕ), v i * a_fg i * a_h i = v i * (a_f i + a_g i) * a_h i
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)

∑' (i : ℕ), v i * (a_f i + a_g i) * a_h i = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)

inner (f + g) h = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)

∑' (i : ℕ), v i * (a_f i + a_g i) * a_h i = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)

∑' (i : ℕ), v i * (a_f i + a_g i) * a_h i = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)
i:

v i * (a_f i + a_g i) * a_h i = v i * a_f i * a_h i + v i * a_g i * a_h i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)
i:

v i * (a_f i + a_g i) * a_h i = v i * a_f i * a_h i + v i * a_g i * a_h i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)

(i : ℕ), v i * (a_f i + a_g i) * a_h i = v i * a_f i * a_h i + v i * a_g i * a_h i
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)

∑' (i : ℕ), (v i * a_f i * a_h i + v i * a_g i * a_h i) = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)

inner (f + g) h = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)

∑' (i : ℕ), (v i * a_f i * a_h i + v i * a_g i * a_h i) = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)

∑' (b : ℕ), v b * ⋯.some b * ⋯.some b + ∑' (b : ℕ), v b * ⋯.some b * ⋯.some b = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_h:= ⋯.some:
a_fg:= fun i => a_f i + a_g i:
a_fg_repr: a_fg set_repr (f + g)

∑' (b : ℕ), v b * ⋯.some b * ⋯.some b + ∑' (b : ℕ), v b * ⋯.some b * ⋯.some b = inner f h + inner g h
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)

inner (f + g) h = inner f h + inner g h
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g = inner g f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g = inner g f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g = inner g f
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

H_inner f g = inner g f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g = inner g f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = inner g f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g = inner g f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = inner g f
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
i:

v i * ⋯.some i * ⋯.some i = v i * ⋯.some i * ⋯.some i
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g = inner g f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
comm: (i : ℕ), v i * ⋯.some i * ⋯.some i = v i * ⋯.some i * ⋯.some i

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = inner g f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
comm: (i : ℕ), v i * ⋯.some i * ⋯.some i = v i * ⋯.some i * ⋯.some i

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = inner g f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
comm: (i : ℕ), v i * ⋯.some i * ⋯.some i = v i * ⋯.some i * ⋯.some i

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = inner g f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g = inner g f
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 inner f f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 inner f f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 inner f f
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 H_inner f f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 inner f f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 inner f f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:

0 ∑' (i : ℕ), v i * a i * a i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 inner f f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:

0 ∑' (i : ℕ), v i * a i * a i
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:

(i : ℕ), v i * a i * a i = v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:

(i : ℕ), v i * a i * a i = v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:

(i : ℕ), v i * a i * a i = v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
i:

v i * a i * a i = v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:

(i : ℕ), v i * a i * a i = v i * a i ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 inner f f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq: (i : ℕ), v i * a i * a i = v i * a i ^ 2

0 ∑' (i : ℕ), v i * a i * a i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq: (i : ℕ), v i * a i * a i = v i * a i ^ 2

0 ∑' (i : ℕ), v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq: (i : ℕ), v i * a i * a i = v i * a i ^ 2

0 ∑' (i : ℕ), v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 inner f f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq: (i : ℕ), v i * a i * a i = v i * a i ^ 2

0 ∑' (i : ℕ), v i * a i ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq: (i : ℕ), v i * a i * a i = v i * a i ^ 2

(i : ℕ), 0 v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq: (i : ℕ), v i * a i * a i = v i * a i ^ 2

(i : ℕ), 0 v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq: (i : ℕ), v i * a i * a i = v i * a i ^ 2

(i : ℕ), 0 v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq: (i : ℕ), v i * a i * a i = v i * a i ^ 2
i:

0 v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq: (i : ℕ), v i * a i * a i = v i * a i ^ 2

(i : ℕ), 0 v i * a i ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 inner f f
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 f
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 (inner f f)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 f
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f 0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f 0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f 0 = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

H_inner f 0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f 0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f 0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

∑' (i : ℕ), v i * ⋯.some i * (fun x => 0) i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

∑' (i : ℕ), v i * ⋯.some i * (fun x => 0) i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f 0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

∑' (i : ℕ), v i * ⋯.some i * (fun x => 0) i = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

(i : ℕ), v i * ⋯.some i * (fun i => 0) i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

(i : ℕ), v i * ⋯.some i * (fun i => 0) i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

(i : ℕ), v i * ⋯.some i * (fun i => 0) i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
i:

v i * ⋯.some i * (fun i => 0) i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

(i : ℕ), v i * ⋯.some i * (fun i => 0) i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
i:

v i * ⋯.some i * (fun i => 0) i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
i:

v i * ⋯.some i * (fun i => 0) i = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
i:

v i * ⋯.some i * 0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

(i : ℕ), v i * ⋯.some i * (fun i => 0) i = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f 0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
summand_eq_zero: (i : ℕ), v i * ⋯.some i * (fun i => 0) i = 0

∑' (i : ℕ), v i * ⋯.some i * 0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
summand_eq_zero: (i : ℕ), v i * ⋯.some i * (fun i => 0) i = 0

∑' (i : ℕ), 0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
summand_eq_zero: (i : ℕ), v i * ⋯.some i * (fun i => 0) i = 0

∑' (i : ℕ), 0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f 0 = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
inner_eq_0: inner f f = 0

f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
inner_eq_0: inner f f = 0

f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
inner_eq_0: inner f f = 0

f = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
inner_eq_0: H_inner f f = 0

f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
inner_eq_0: H_inner f f = 0

f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
inner_eq_0: ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = 0

f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
inner_eq_0: ∑' (i : ℕ), v i * a i * a i = 0

f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
inner_eq_0: ∑' (i : ℕ), v i * a i * a i = 0

f = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
inner_eq_0: ∑' (i : ℕ), v i * a i * a i = 0
i:

v i * a i * a i = v i * a i ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
inner_eq_0: ∑' (i : ℕ), v i * a i * a i = 0
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2

f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0

f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0

f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0

f = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0

(i : ℕ), 0 v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0

(i : ℕ), 0 v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0

(i : ℕ), 0 v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
i:

0 v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0

(i : ℕ), 0 v i * a i ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2

f = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2

Summable fun i => v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2

Summable fun i => v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2

Summable fun i => v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2

Summable fun i => v i * a i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2

Summable fun i => v i * a i * a i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2

Summable fun i => v i * a i * a i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2

Summable fun i => v i * a i ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0

f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0

f = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0

(i : ℕ), v i * a i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0

(i : ℕ), v i * a i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0

(i : ℕ), v i * a i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
i:

v i * a i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0

(i : ℕ), v i * a i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
i:
x✝: v i = 0 a i ^ 2 = 0

v i * a i = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
i:
x✝: v i = 0 a i ^ 2 = 0

v i * a i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
i:
hv: v i = 0

inl
v i * a i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
i:
hv: v i = 0

inl
0 * a i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
i:
hv: v i = 0

inl
0 * a i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
i:
hv: v i = 0

inl
v i * a i = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
i:
x✝: v i = 0 a i ^ 2 = 0

v i * a i = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
i:
x✝: v i = 0 a i ^ 2 = 0

v i * a i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
i:
ha: a i ^ 2 = 0

inr
v i * a i = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
i:
ha: a i ^ 2 = 0

inr
v i * 0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
i:
ha: a i ^ 2 = 0

inr
v i * 0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
i:
ha: a i ^ 2 = 0

inr
v i * a i = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω

a.h
f x = 0 x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω

a.h
f x = 0 x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω

a.h
f x = 0 x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω

a.h
f x = zero x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω

a.h
f x = zero x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω

a.h
f x = zero x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω

a.h
f x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x

a.h.intro.intro
f x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x

a.h.intro.intro
f x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x

a.h.intro.intro
(fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x) x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x

a.h.intro.intro
f x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x

a.h.intro.intro
(fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x) x = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x

a.h.intro.intro
(fun x => ∑' (i : ℕ), v i * a i * e i x) x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x

a.h.intro.intro
(fun x => ∑' (i : ℕ), v i * a i * e i x) x = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x

(i : ℕ), v i * a i * e i x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x

(i : ℕ), v i * a i * e i x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x

(i : ℕ), v i * a i * e i x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x
i:

v i * a i * e i x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x

(i : ℕ), v i * a i * e i x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x
i:

v i * a i * e i x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x
i:

0 * e i x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x
i:

0 * e i x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x

(i : ℕ), v i * a i * e i x = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x
summand_eq_0: (i : ℕ), v i * a i * e i x = 0

a.h.intro.intro
∑' (i : ℕ), v i * a i * e i x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x
summand_eq_0: (i : ℕ), v i * a i * e i x = 0

a.h.intro.intro
∑' (i : ℕ), 0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
a:= ⋯.some:
sq_summand: (i : ℕ), v i * a i * a i = v i * a i ^ 2
inner_eq_0: ∑' (i : ℕ), v i * a i ^ 2 = 0
summand_nonneg: (i : ℕ), 0 v i * a i ^ 2
summand_summable: Summable fun i => v i * a i ^ 2
summand_eq_zero: (i : ℕ), v i * a i ^ 2 = 0
mul_v_a_eq_0: (i : ℕ), v i * a i = 0
x: ↑Ω
right✝¹: Summable fun i => v i * ⋯.some i ^ 2
ha_r: f = fun x => ∑' (i : ℕ), v i * ⋯.some i * e i x
right✝: (x : ↑Ω), Summable fun i => v i * ⋯.some i * e i x
summand_eq_0: (i : ℕ), v i * a i * e i x = 0

a.h.intro.intro
∑' (i : ℕ), 0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = 0 f = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = f ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = f ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = f ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = (inner f f) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = f ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = (inner f f) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = inner f f
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

(inner (f + t * g) (f + t * g)) ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)

(inner (f + t * g) (f + t * g)) ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)

(inner (f + t * g) (f + t * g)) ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)

inner (f + t * g) (f + t * g) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)

inner (f + t * g) (f + t * g) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:

inner (f + t * g) (f + t * g) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:

inner (f + t * g) (f + t * g) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:

inner (f + t * g) (f + t * g) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:

inner (f + t * g) (f + t * g) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:

a_f_tg set_repr (f + t * g)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:

a_f_tg set_repr (f + t * g)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:

a_f_tg set_repr (f + t * g)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
repr_add: (fun i => a_f i + ⋯.some i) set_repr (f + t * g)

a_f_tg set_repr (f + t * g)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:

a_f_tg set_repr (f + t * g)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
repr_add: (fun i => a_f i + ⋯.some i) set_repr (f + t * g)
repr_mul: (fun i => t * a_g i) set_repr (t * g)

a_f_tg set_repr (f + t * g)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:

a_f_tg set_repr (f + t * g)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
repr_add: (fun i => a_f i + ⋯.some i) set_repr (f + t * g)
repr_mul: (fun i => t * a_g i) set_repr (t * g)

a_f_tg set_repr (f + t * g)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
repr_add: (fun i => a_f i + (fun i => t * a_g i) i) set_repr (f + t * g)
repr_mul: (fun i => t * a_g i) set_repr (t * g)

a_f_tg set_repr (f + t * g)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
repr_add: (fun i => a_f i + (fun i => t * a_g i) i) set_repr (f + t * g)
repr_mul: (fun i => t * a_g i) set_repr (t * g)

a_f_tg set_repr (f + t * g)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)

inner (f + t * g) (f + t * g) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)

inner (f + t * g) (f + t * g) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)

H_inner (f + t * g) (f + t * g) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)

∑' (i : ℕ), v i * a_f_tg i * a_f_tg i = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)

∑' (i : ℕ), v i * a_f_tg i * a_f_tg i = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)

∑' (i : ℕ), v i * a_f_tg i * a_f_tg i = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)

∑' (i : ℕ), v i * a_f_tg i * a_f_tg i = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
i:

v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
i:

v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)

∑' (i : ℕ), v i * a_f_tg i * a_f_tg i = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)

∑' (i : ℕ), v i * a_f_tg i * a_f_tg i = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)

∑' (i : ℕ), (v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)

∑' (i : ℕ), (v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)

∑' (i : ℕ), (v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)

∑' (i : ℕ), (v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)

∑' (b : ℕ), (v b * a_f b * a_f b + 2 * t * (v b * a_f b * a_g b)) + ∑' (b : ℕ), t ^ 2 * (v b * ⋯.some b * ⋯.some b) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)

∑' (b : ℕ), (v b * a_f b * a_f b + 2 * t * (v b * a_f b * a_g b)) + ∑' (b : ℕ), t ^ 2 * (v b * ⋯.some b * ⋯.some b) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)

∑' (b : ℕ), (v b * a_f b * a_f b + 2 * t * (v b * a_f b * a_g b)) + ∑' (b : ℕ), t ^ 2 * (v b * ⋯.some b * ⋯.some b) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
h: (H v e μ)

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
h: (H v e μ)

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
h: (H v e μ)

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
h: (H v e μ)

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
h: (H v e μ)

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
h: (H v e μ)

H_inner h h = h ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
h: (H v e μ)

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
h: (H v e μ)

H_inner h h = h ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
h: (H v e μ)

H_inner h h = h ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
h: (H v e μ)

inner h h = h ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
h: (H v e μ)

∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
h: (H v e μ)

inner h h = h ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
h: (H v e μ)

h ^ 2 = h ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2

∑' (b : ℕ), (v b * a_f b * a_f b + 2 * t * (v b * a_f b * a_g b)) + ∑' (b : ℕ), t ^ 2 * (v b * ⋯.some b * ⋯.some b) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2

∑' (b : ℕ), v b * ⋯.some b * ⋯.some b + ∑' (b : ℕ), 2 * t * (v b * ⋯.some b * ⋯.some b) + ∑' (b : ℕ), t ^ 2 * (v b * ⋯.some b * ⋯.some b) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2

∑' (b : ℕ), v b * ⋯.some b * ⋯.some b + ∑' (b : ℕ), 2 * t * (v b * ⋯.some b * ⋯.some b) + ∑' (b : ℕ), t ^ 2 * (v b * ⋯.some b * ⋯.some b) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
const_out: ∑' (i : ℕ), t ^ 2 * (v i * a_g i * a_g i) = t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i

∑' (b : ℕ), v b * ⋯.some b * ⋯.some b + ∑' (b : ℕ), 2 * t * (v b * ⋯.some b * ⋯.some b) + ∑' (b : ℕ), t ^ 2 * (v b * ⋯.some b * ⋯.some b) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
const_out: ∑' (i : ℕ), t ^ 2 * (v i * a_g i * a_g i) = t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i

∑' (b : ℕ), v b * ⋯.some b * ⋯.some b + ∑' (b : ℕ), 2 * t * (v b * ⋯.some b * ⋯.some b) + ∑' (b : ℕ), t ^ 2 * (v b * ⋯.some b * ⋯.some b) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
const_out: ∑' (i : ℕ), t ^ 2 * (v i * a_g i * a_g i) = t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i

∑' (b : ℕ), v b * ⋯.some b * ⋯.some b + ∑' (b : ℕ), 2 * t * (v b * ⋯.some b * ⋯.some b) + t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
const_out: ∑' (i : ℕ), t ^ 2 * (v i * a_g i * a_g i) = t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i

∑' (b : ℕ), v b * ⋯.some b * ⋯.some b + ∑' (b : ℕ), 2 * t * (v b * ⋯.some b * ⋯.some b) + ∑' (b : ℕ), t ^ 2 * (v b * ⋯.some b * ⋯.some b) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
const_out: ∑' (i : ℕ), t ^ 2 * (v i * a_g i * a_g i) = t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i

f ^ 2 + ∑' (b : ℕ), 2 * t * (v b * ⋯.some b * ⋯.some b) + t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
const_out: ∑' (i : ℕ), t ^ 2 * (v i * a_g i * a_g i) = t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i

∑' (b : ℕ), v b * ⋯.some b * ⋯.some b + ∑' (b : ℕ), 2 * t * (v b * ⋯.some b * ⋯.some b) + ∑' (b : ℕ), t ^ 2 * (v b * ⋯.some b * ⋯.some b) = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
const_out: ∑' (i : ℕ), t ^ 2 * (v i * a_g i * a_g i) = t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i

f ^ 2 + ∑' (b : ℕ), 2 * t * (v b * ⋯.some b * ⋯.some b) + t ^ 2 * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
const_out: ∑' (i : ℕ), t ^ 2 * (v i * a_g i * a_g i) = t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i

f ^ 2 + ∑' (b : ℕ), 2 * t * (v b * ⋯.some b * ⋯.some b) + t ^ 2 * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
const_out✝: ∑' (i : ℕ), t ^ 2 * (v i * a_g i * a_g i) = t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i
const_out: ∑' (i : ℕ), 2 * t * (v i * a_f i * a_g i) = 2 * t * ∑' (i : ℕ), v i * a_f i * a_g i

f ^ 2 + ∑' (b : ℕ), 2 * t * (v b * ⋯.some b * ⋯.some b) + t ^ 2 * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
const_out✝: ∑' (i : ℕ), t ^ 2 * (v i * a_g i * a_g i) = t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i
const_out: ∑' (i : ℕ), 2 * t * (v i * a_f i * a_g i) = 2 * t * ∑' (i : ℕ), v i * a_f i * a_g i

f ^ 2 + ∑' (b : ℕ), 2 * t * (v b * ⋯.some b * ⋯.some b) + t ^ 2 * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
const_out✝: ∑' (i : ℕ), t ^ 2 * (v i * a_g i * a_g i) = t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i
const_out: ∑' (i : ℕ), 2 * t * (v i * a_f i * a_g i) = 2 * t * ∑' (i : ℕ), v i * a_f i * a_g i

f ^ 2 + 2 * t * ∑' (i : ℕ), v i * a_f i * a_g i + t ^ 2 * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
const_out✝: ∑' (i : ℕ), t ^ 2 * (v i * a_g i * a_g i) = t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i
const_out: ∑' (i : ℕ), 2 * t * (v i * a_f i * a_g i) = 2 * t * ∑' (i : ℕ), v i * a_f i * a_g i

f ^ 2 + 2 * t * ∑' (i : ℕ), v i * a_f i * a_g i + t ^ 2 * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
const_out✝: ∑' (i : ℕ), t ^ 2 * (v i * a_g i * a_g i) = t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i
const_out: ∑' (i : ℕ), 2 * t * (v i * a_f i * a_g i) = 2 * t * ∑' (i : ℕ), v i * a_f i * a_g i

f ^ 2 + 2 * t * ∑' (i : ℕ), v i * a_f i * a_g i + t ^ 2 * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
const_out✝: ∑' (i : ℕ), t ^ 2 * (v i * a_g i * a_g i) = t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i
const_out: ∑' (i : ℕ), 2 * t * (v i * a_f i * a_g i) = 2 * t * ∑' (i : ℕ), v i * a_f i * a_g i

f ^ 2 + 2 * t * ∑' (i : ℕ), v i * a_f i * a_g i + t ^ 2 * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:

f + t * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
const_out✝: ∑' (i : ℕ), t ^ 2 * (v i * a_g i * a_g i) = t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i
const_out: ∑' (i : ℕ), 2 * t * (v i * a_f i * a_g i) = 2 * t * ∑' (i : ℕ), v i * a_f i * a_g i
tsum_to_inner: ∑' (i : ℕ), v i * a_f i * a_g i = inner f g

f ^ 2 + 2 * t * ∑' (i : ℕ), v i * a_f i * a_g i + t ^ 2 * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
t:
inner_nn: 0 inner (f + t * g) (f + t * g)
a_f:= ⋯.some:
a_g:= ⋯.some:
a_f_tg:= fun i => a_f i + t * a_g i:
a_f_tg_repr: a_f_tg set_repr (f + t * g)
distribute: (i : ℕ), v i * a_f_tg i * a_f_tg i = v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i) + t ^ 2 * (v i * a_g i * a_g i)
add_summable: Summable fun i => v i * a_f i * a_f i + 2 * t * (v i * a_f i * a_g i)
tsum_to_norm: (h : (H v e μ)), ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i = h ^ 2
const_out✝: ∑' (i : ℕ), t ^ 2 * (v i * a_g i * a_g i) = t ^ 2 * ∑' (i : ℕ), v i * a_g i * a_g i
const_out: ∑' (i : ℕ), 2 * t * (v i * a_f i * a_g i) = 2 * t * ∑' (i : ℕ), v i * a_f i * a_g i
tsum_to_inner: ∑' (i : ℕ), v i * a_f i * a_g i = inner f g

f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2 = f ^ 2 + 2 * t * inner f g + t ^ 2 * g ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0

pos
inner f g f * g
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: ¬‖g 0
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0

pos
inner f g f * g
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: ¬‖g 0
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀

pos
inner f g f * g
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀

P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀

P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀

P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀

P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀

P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + inner f g ^ 2 / (g ^ 2 * g ^ 2) * g ^ 2 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀

P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + inner f g ^ 2 / (g ^ 2 * g ^ 2) * g ^ 2 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + inner f g ^ 2 / (g ^ 2 * g ^ 2) * g ^ 2 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + inner f g ^ 2 / g ^ 2 * 1 / g ^ 2 * g ^ 2 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀

P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + inner f g ^ 2 / g ^ 2 * 1 / g ^ 2 * g ^ 2 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + inner f g ^ 2 / g ^ 2 * 1 / g ^ 2 * g ^ 2 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + inner f g ^ 2 / g ^ 2 * (1 / g ^ 2 * g ^ 2) = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀

P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + inner f g ^ 2 / g ^ 2 * (1 / g ^ 2 * g ^ 2) = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + inner f g ^ 2 / g ^ 2 * 1 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + inner f g ^ 2 / g ^ 2 * 1 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀

P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + inner f g ^ 2 / g ^ 2 * 1 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + inner f g ^ 2 / g ^ 2 * 1 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + inner f g ^ 2 / g ^ 2 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀

P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + inner f g ^ 2 / g ^ 2 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + 2 * t₀ * inner f g + inner f g ^ 2 / g ^ 2 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + -2 * inner f g ^ 2 / g ^ 2 + inner f g ^ 2 / g ^ 2 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀

P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + -2 * inner f g ^ 2 / g ^ 2 + inner f g ^ 2 / g ^ 2 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0: P t₀ = f ^ 2 + 2 * t₀ * inner f g + t₀ ^ 2 * g ^ 2

f ^ 2 + -2 * inner f g ^ 2 / g ^ 2 + inner f g ^ 2 / g ^ 2 = f ^ 2 - inner f g ^ 2 / g ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 P t₀
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

pos
inner f g f * g
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

inner f g ^ 2 (f * g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

inner f g ^ 2 (f * g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

inner f g ^ 2 (f * g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

inner f g ^ 2 (f * g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

inner f g ^ 2 (f * g) ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

inner f g ^ 2 f ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

inner f g ^ 2 (f * g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

inner f g ^ 2 f ^ 2 * g ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

inner f g ^ 2 * (g ^ 2)⁻¹ f ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

inner f g ^ 2 * (g ^ 2)⁻¹ f ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

inner f g ^ 2 (f * g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

inner f g ^ 2 * (g ^ 2)⁻¹ f ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

0 f ^ 2 - inner f g ^ 2 * (g ^ 2)⁻¹
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2

0 f ^ 2 - inner f g ^ 2 * (g ^ 2)⁻¹
Goals accomplished!

Goals accomplished! 🐙
--rw [←sq_abs (inner f g)] at sq_ineq
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
sq_ineq: inner f g ^ 2 (f * g) ^ 2

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
sq_ineq: inner f g ^ 2 (f * g) ^ 2

pos
inner f g f * g
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
sq_ineq: inner f g * inner f g (f * g) ^ 2

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
sq_ineq: inner f g * inner f g (f * g) ^ 2

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
sq_ineq: inner f g * inner f g (f * g) ^ 2

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
sq_ineq: inner f g * inner f g (f * g) ^ 2

pos
inner f g f * g
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
sq_ineq: inner f g * inner f g f * g * (f * g)

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
sq_ineq: inner f g * inner f g f * g * (f * g)

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0
hg_sq: g ^ 2 0
P:= fun t => f + t * g ^ 2:
t₀:= -inner f g / g ^ 2:
P_nonneg: 0 f ^ 2 - inner f g ^ 2 / g ^ 2
P_t0_val: P t₀ = f ^ 2 - inner f g ^ 2 / g ^ 2
sq_ineq: inner f g * inner f g f * g * (f * g)
norm_mul_nonneg: 0 f * g

pos
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g 0

pos
inner f g f * g
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g = 0

neg
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g = 0

neg
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: g = 0

neg
inner f g f * g
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: (inner g g) = 0

neg
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: (inner g g) = 0

neg
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: (inner g g) = 0

neg
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0

neg
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0

neg
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0

neg
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0
g_eq_0: g = 0

neg
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0
g_eq_0: g = 0

neg
inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0
g_eq_0: g = 0

neg
inner f 0 f * 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0
g_eq_0: g = 0

neg
inner f 0 f * 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0
g_eq_0: g = 0

neg
inner f 0 f * 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0
g_eq_0: g = 0

neg
inner f 0 f * 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0
g_eq_0: g = 0

neg
inner f 0 f * 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0
g_eq_0: g = 0

neg
inner f 0 f * 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0
g_eq_0: g = 0

neg
0 f * 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0
g_eq_0: g = 0

neg
0 f * 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0
g_eq_0: g = 0

neg
0 f * 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0
g_eq_0: g = 0

neg
0 f * 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0
g_eq_0: g = 0

neg
0 f * (inner 0 0)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g f * g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0
g_eq_0: g = 0

neg
0 f * (inner 0 0)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0
g_eq_0: g = 0

neg
0 f * 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
hg: inner g g = 0
g_eq_0: g = 0

neg
0 f * 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

inner f g f * g
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f + g f + g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

hb
0 f + g
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
f + g * f + g (f + g) * (f + g)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f + g f + g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

hb
0 f + g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

hb
0 f + g
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
f + g * f + g (f + g) * (f + g)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f + g f + g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

h
f + g * f + g (f + g) * (f + g)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

h
f + g * f + g (f + g) * (f + g)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

h
f + g ^ 2 (f + g) * (f + g)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f + g f + g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

h
f + g ^ 2 (f + g) * (f + g)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

h
f + g ^ 2 (f + g) * (f + g)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f + g f + g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

2 * inner f g 2 * (f * g)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

2 * inner f g 2 * (f * g)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

2 * inner f g 2 * (f * g)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
this: inner f g f * g

2 * inner f g 2 * (f * g)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

2 * inner f g 2 * (f * g)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
cauchy_schwarz: 2 * inner f g 2 * (f * g)
ineq: f ^ 2 + 2 * inner f g + g ^ 2 f ^ 2 + 2 * (f * g) + g ^ 2

f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
cauchy_schwarz: 2 * inner f g 2 * (f * g)
ineq: f ^ 2 + 2 * inner f g + g ^ 2 f ^ 2 + 2 * (f * g) + g ^ 2

f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
cauchy_schwarz: 2 * inner f g 2 * (f * g)
ineq: f ^ 2 + 2 * inner f g + g ^ 2 f ^ 2 + 2 * (f * g) + g ^ 2

f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
cauchy_schwarz: 2 * inner f g 2 * (f * g)
ineq: f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2

f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f + g f + g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * 1 * inner f g + 1 ^ 2 * g ^ 2

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f + g f + g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * 1 * inner f g + 1 ^ 2 * g ^ 2

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * 1 * inner f g + 1 ^ 2 * g ^ 2

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + 1 ^ 2 * g ^ 2

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + 1 ^ 2 * g ^ 2

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f + g f + g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + 1 ^ 2 * g ^ 2

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + 1 ^ 2 * g ^ 2

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f + g f + g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f ^ 2 + 2 * inner f g + g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + 1 * g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + 1 * g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + 1 * g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f + g f + g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + 1 * g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + 1 * g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2

1 * g = g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + 1 * g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2

1 * g = g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + 1 * g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2

1 * g = g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + 1 * g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2
x: ↑Ω

a.h
(1 * g) x = g x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + 1 * g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2

1 * g = g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + 1 * g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2
x: ↑Ω

a.h
(1 * g) x = g x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + 1 * g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2
x: ↑Ω

a.h
(1 * g) x = g x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + 1 * g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2
x: ↑Ω

a.h
1 * g x = g x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + 1 * g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2

1 * g = g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + 1 * g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2
x: ↑Ω

a.h
1 * g x = g x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + 1 * g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2
x: ↑Ω

a.h
1 * g x = g x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

f + g f + g
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + 1 * g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2
one_mul_g_eq_g: 1 * g = g

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2
one_mul_g_eq_g: 1 * g = g

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
sq_ineq: f + g ^ 2 (f + g) ^ 2
distrib_norm: f + 1 * g ^ 2 = f ^ 2 + 2 * inner f g + g ^ 2
one_mul_g_eq_g: 1 * g = g

h
f + g ^ 2 (f + g) ^ 2
Goals accomplished!

Goals accomplished! 🐙
end Inner /- We show properties on the distance induced by the inner product of H. -/ namespace Dist open Inner Ring Group variable {v : eigen} {e : ℕ Ω ℝ} {μ : Measure Ω} [IsFiniteMeasure μ] (f g : H v e μ)
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

dist a a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

dist a a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

dist a a = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a - a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

dist a a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a - a = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a - a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a - a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a - a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a - a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a - a = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a + -1 * a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a - a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
(a + -1 * a) x = 0 x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a - a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
(a + -1 * a) x = 0 x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
(a + -1 * a) x = 0 x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
a x + (-1 * a) x = 0 x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a - a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
a x + (-1 * a) x = 0 x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
a x + (-1 * a) x = 0 x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
a x + -1 * a x = 0 x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a - a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
a x + -1 * a x = 0 x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
a x + -1 * a x = 0 x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
x: ↑Ω

a.h
a x + -1 * a x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

a - a = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

dist a a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
a_sub_a_eq_0: a - a = 0

a - a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
a_sub_a_eq_0: a - a = 0

0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
a_sub_a_eq_0: a - a = 0

0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

dist a a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
a_sub_a_eq_0: a - a = 0

0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
a_sub_a_eq_0: a - a = 0

0 = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
a_sub_a_eq_0: a - a = 0

(inner 0 0) = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

dist a a = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
a_sub_a_eq_0: a - a = 0

(inner 0 0) = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
a_sub_a_eq_0: a - a = 0

0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)
a_sub_a_eq_0: a - a = 0

0 = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a: (H v e μ)

dist a a = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

a - b = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

a - b = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

a - b = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

(H_inner (a - b) (a - b)) = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

(∑' (i : ℕ), v i * ⋯.some i * ⋯.some i) = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:

(∑' (i : ℕ), v i * ⋯.some i * ⋯.some i) = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:

(∑' (i : ℕ), v i * ⋯.some i * ⋯.some i) = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:

repr set_repr (a - b)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:

repr set_repr (a - b)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:

repr set_repr (a - b)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_minus_b: (fun i => -1 * ⋯.some i) set_repr (-1 * b)

repr set_repr (a - b)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:

repr set_repr (a - b)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_minus_b: (fun i => -1 * ⋯.some i) set_repr (-1 * b)
repr_a_sub_b: (fun i => ⋯.some i + ⋯.some i) set_repr (a + -1 * b)

repr set_repr (a - b)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:

repr set_repr (a - b)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_minus_b: (fun i => -1 * ⋯.some i) set_repr (-1 * b)
repr_a_sub_b: (fun i => ⋯.some i + ⋯.some i) set_repr (a + -1 * b)

repr set_repr (a - b)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_minus_b: (fun i => -1 * ⋯.some i) set_repr (-1 * b)
repr_a_sub_b: (fun i => ⋯.some i + (fun i => -1 * ⋯.some i) i) set_repr (a + -1 * b)

repr set_repr (a - b)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_minus_b: (fun i => -1 * ⋯.some i) set_repr (-1 * b)
repr_a_sub_b: (fun i => ⋯.some i + (fun i => -1 * ⋯.some i) i) set_repr (a + -1 * b)

repr set_repr (a - b)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_a_minus_b: repr set_repr (a - b)

(∑' (i : ℕ), v i * ⋯.some i * ⋯.some i) = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_a_minus_b: repr set_repr (a - b)

(∑' (i : ℕ), v i * repr i * repr i) = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_a_minus_b: repr set_repr (a - b)

(∑' (i : ℕ), v i * repr i * repr i) = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_a_minus_b: repr set_repr (a - b)
a_r:= ⋯.some:

(∑' (i : ℕ), v i * repr i * repr i) = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_a_minus_b: repr set_repr (a - b)
a_r:= ⋯.some:
b_r:= ⋯.some:

(∑' (i : ℕ), v i * repr i * repr i) = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_a_minus_b: repr set_repr (a - b)
a_r:= ⋯.some:
b_r:= ⋯.some:

(∑' (i : ℕ), v i * repr i * repr i) = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_a_minus_b: repr set_repr (a - b)
a_r:= ⋯.some:
b_r:= ⋯.some:

(∑' (i : ℕ), v i * repr i * repr i) = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_a_minus_b: repr set_repr (a - b)
a_r:= ⋯.some:
b_r:= ⋯.some:
i:

v i * repr i * repr i = v i * repr i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_a_minus_b: repr set_repr (a - b)
a_r:= ⋯.some:
b_r:= ⋯.some:
i:

v i * repr i * repr i = v i * repr i ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_a_minus_b: repr set_repr (a - b)
a_r:= ⋯.some:
b_r:= ⋯.some:

(i : ℕ), v i * repr i * repr i = v i * repr i ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_a_minus_b: repr set_repr (a - b)
a_r:= ⋯.some:
b_r:= ⋯.some:

(∑' (i : ℕ), v i * repr i ^ 2) = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_a_minus_b: repr set_repr (a - b)
a_r:= ⋯.some:
b_r:= ⋯.some:

(∑' (i : ℕ), v i * repr i ^ 2) = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_a_minus_b: repr set_repr (a - b)
a_r:= ⋯.some:
b_r:= ⋯.some:

(∑' (i : ℕ), v i * repr i ^ 2) = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_a_minus_b: repr set_repr (a - b)
a_r:= ⋯.some:
b_r:= ⋯.some:
i:

v i * repr i ^ 2 = v i * (a_r i - b_r i) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_a_minus_b: repr set_repr (a - b)
a_r:= ⋯.some:
b_r:= ⋯.some:
i:

v i * repr i ^ 2 = v i * (a_r i - b_r i) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_a_minus_b: repr set_repr (a - b)
a_r:= ⋯.some:
b_r:= ⋯.some:

(i : ℕ), v i * repr i ^ 2 = v i * (a_r i - b_r i) ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
repr:= fun i => ⋯.some i + -1 * ⋯.some i:
repr_a_minus_b: repr set_repr (a - b)
a_r:= ⋯.some:
b_r:= ⋯.some:

(∑' (i : ℕ), v i * (a_r i - b_r i) ^ 2) = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = dist b a
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = dist b a
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

(∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2) = dist b a
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = dist b a
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

(∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2) = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

(∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2) = (∑' (i : ℕ), v i * (⋯.some i - ⋯.some i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = dist b a
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
a_r:= ⋯.some:

(∑' (i : ℕ), v i * (a_r i - ⋯.some i) ^ 2) = (∑' (i : ℕ), v i * (⋯.some i - a_r i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = dist b a
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
a_r:= ⋯.some:
b_r:= ⋯.some:

(∑' (i : ℕ), v i * (a_r i - b_r i) ^ 2) = (∑' (i : ℕ), v i * (b_r i - a_r i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = dist b a
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
a_r:= ⋯.some:
b_r:= ⋯.some:

(∑' (i : ℕ), v i * (a_r i - b_r i) ^ 2) = (∑' (i : ℕ), v i * (b_r i - a_r i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
a_r:= ⋯.some:
b_r:= ⋯.some:

(∑' (i : ℕ), v i * (a_r i - b_r i) ^ 2) = (∑' (i : ℕ), v i * (b_r i - a_r i) ^ 2)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
a_r:= ⋯.some:
b_r:= ⋯.some:
i:

v i * (a_r i - b_r i) ^ 2 = v i * (b_r i - a_r i) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
a_r:= ⋯.some:
b_r:= ⋯.some:
i:

v i * (a_r i - b_r i) ^ 2 = v i * (b_r i - a_r i) ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
a_r:= ⋯.some:
b_r:= ⋯.some:

(i : ℕ), v i * (a_r i - b_r i) ^ 2 = v i * (b_r i - a_r i) ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

0 dist a b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

0 dist a b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

0 dist a b
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

0 a - b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

0 dist a b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

0 a - b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

0 a - b
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

0 (H_inner (a - b) (a - b))
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

0 dist a b
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = 0 a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: dist a b = 0

a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = 0 a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: dist a b = 0

a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: dist a b = 0

a = b
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: a - b = 0

a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: a - b = 0

a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = 0 a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: a - b = 0

a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: a - b = 0

a = b
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: (inner (a - b) (a - b)) = 0

a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: (inner (a - b) (a - b)) = 0

a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = 0 a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: (inner (a - b) (a - b)) = 0

a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0

a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0

a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0

a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = 0 a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a - b = 0

a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = 0 a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a - b = 0
x: ↑Ω

a.h
a x = b x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = 0 a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a - b = 0
x: ↑Ω

a.h
a x = b x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a - b = 0
x: ↑Ω

a.h
a x = b x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a - b = 0
x: ↑Ω
h: a x - b x = 0

a x = b x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a - b = 0
x: ↑Ω
h: a x - b x = 0

a x = b x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a - b = 0
x: ↑Ω

a.h
a x = b x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = 0 a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a - b = 0
x: ↑Ω
dif_imp_eq: a x - b x = 0 a x = b x

a.h
a x - b x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = 0 a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a - b = 0
x: ↑Ω
dif_imp_eq: a x - b x = 0 a x = b x

a.h
a x - b x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a - b = 0
x: ↑Ω
dif_imp_eq: a x - b x = 0 a x = b x

a.h
a x - b x = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a - b = 0
x: ↑Ω
dif_imp_eq: a x - b x = 0 a x = b x

a.h
a x + -1 * b x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = 0 a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a - b = 0
x: ↑Ω
dif_imp_eq: a x - b x = 0 a x = b x

a.h
a x + -1 * b x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a - b = 0
x: ↑Ω
dif_imp_eq: a x - b x = 0 a x = b x

a.h
a x + -1 * b x = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a + -1 * b = 0
x: ↑Ω
dif_imp_eq: a x - b x = 0 a x = b x

a.h
a x + -1 * b x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a + -1 * b = 0
x: ↑Ω
dif_imp_eq: a x - b x = 0 a x = b x

a.h
a x + -1 * b x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = 0 a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a + -1 * b = 0
x: ↑Ω
dif_imp_eq: a x - b x = 0 a x = b x
a_minus_b_eq_zero_val: (a + -1 * b) = 0

a.h
a x + -1 * b x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = 0 a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a + -1 * b = 0
x: ↑Ω
dif_imp_eq: a x - b x = 0 a x = b x
a_minus_b_eq_zero_val: (a + -1 * b) = 0

a.h
a x + -1 * b x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a + -1 * b = 0
x: ↑Ω
dif_imp_eq: a x - b x = 0 a x = b x
a_minus_b_eq_zero_val: (a + -1 * b) = 0

a.h
a x + -1 * b x = 0
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a + -1 * b = 0
x: ↑Ω
dif_imp_eq: a x - b x = 0 a x = b x
a_minus_b_eq_zero_val: (fun x => a x + -1 * b x) = 0

a.h
a x + -1 * b x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a + -1 * b = 0
x: ↑Ω
dif_imp_eq: a x - b x = 0 a x = b x
a_minus_b_eq_zero_val: (fun x => a x + -1 * b x) = 0

a.h
a x + -1 * b x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = 0 a = b
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a + -1 * b = 0
x: ↑Ω
dif_imp_eq: a x - b x = 0 a x = b x
a_minus_b_eq_zero_val: (fun x => a x + -1 * b x) = 0

a.h
a x + -1 * b x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a + -1 * b = 0
x: ↑Ω
dif_imp_eq: a x - b x = 0 a x = b x
a_minus_b_eq_zero_val: (fun x => a x + -1 * b x) = 0

a.h
0 x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)
zero_dist: inner (a - b) (a - b) = 0
a_minus_b_eq_zero: a + -1 * b = 0
x: ↑Ω
dif_imp_eq: a x - b x = 0 a x = b x
a_minus_b_eq_zero_val: (fun x => a x + -1 * b x) = 0

a.h
0 x = 0
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b: (H v e μ)

dist a b = 0 a = b
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

dist a c dist a b + dist b c
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

dist a c dist a b + dist b c
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

dist a c dist a b + dist b c
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a - c dist a b + dist b c
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

dist a c dist a b + dist b c
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a - c dist a b + dist b c
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a - c dist a b + dist b c
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a - c a - b + dist b c
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

dist a c dist a b + dist b c
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a - c a - b + dist b c
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a - c a - b + dist b c
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a - c a - b + b - c
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

dist a c dist a b + dist b c
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a - c a - b + b - c
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a - c = a - b + (b - c)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a - c = a - b + (b - c)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a - c = a - b + (b - c)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a - c) x = (a - b + (b - c)) x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a - c = a - b + (b - c)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a - c) x = (a - b + (b - c)) x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a - c) x = (a - b + (b - c)) x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a - c) x = (a - b) x + (b - c) x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a - c = a - b + (b - c)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a - c) x = (a - b) x + (b - c) x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a - c) x = (a - b) x + (b - c) x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a - c) x = a x + -1 * b x + (b - c) x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a - c = a - b + (b - c)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a - c) x = a x + -1 * b x + (b - c) x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a - c) x = a x + -1 * b x + (b - c) x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a - c) x = a x + -1 * b x + (b x + -1 * c x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a - c = a - b + (b - c)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a - c) x = a x + -1 * b x + (b x + -1 * c x)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a - c) x = a x + -1 * b x + (b x + -1 * c x)
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a - c) x = a x + -1 * c x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

a - c = a - b + (b - c)
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a - c) x = a x + -1 * c x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
x: ↑Ω

a.h
(a - c) x = a x + -1 * c x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

dist a c dist a b + dist b c
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
split_a_sub_c: a - c = a - b + (b - c)

a - c a - b + b - c
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
split_a_sub_c: a - c = a - b + (b - c)

a - b + (b - c) a - b + b - c
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)
split_a_sub_c: a - c = a - b + (b - c)

a - b + (b - c) a - b + b - c
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a, b, c: (H v e μ)

dist a c dist a b + dist b c
Goals accomplished!

Goals accomplished! 🐙
end Dist /- - We instanciate the `NormedAddCommGroup` and `InnerProductSpace ℝ` typeclasses for H. -/ variable {v : eigen} {e : ℕ Ω ℝ} {μ : Measure Ω} [IsFiniteMeasure μ]
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:

(fun n f => n * f) n f = n * f
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
z:

(fun z f => z * f) z f = z * f
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
n:

(n + 1) = n + 1
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:

(n + 1) * f = n * f + f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:
x: ↑Ω

a.h
((n + 1) * f) x = (n * f + f) x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:

(n + 1) * f = n * f + f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:
x: ↑Ω
n':= n:

a.h
((n + 1) * f) x = (n * f + f) x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:

(n + 1) * f = n * f + f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:
x: ↑Ω
n':= n:

a.h
((n + 1) * f) x = (n * f + f) x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:
x: ↑Ω
n':= n:

a.h
((n + 1) * f) x = (n * f + f) x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:
x: ↑Ω
n':= n:

a.h
((n + 1) * f) x = (n' * f) x + f x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:

(n + 1) * f = n * f + f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:
x: ↑Ω
n':= n:

a.h
((n + 1) * f) x = (n' * f) x + f x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:
x: ↑Ω
n':= n:

a.h
((n + 1) * f) x = (n' * f) x + f x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:
x: ↑Ω
n':= n:

a.h
((n + 1) * f) x = n' * f x + f x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:

(n + 1) * f = n * f + f
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:
x: ↑Ω
n':= n:

a.h
((n + 1) * f) x = n' * f x + f x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:
x: ↑Ω
n':= n:

a.h
((n + 1) * f) x = n' * f x + f x
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:
x: ↑Ω
n':= n:

a.h
(n' + 1) * f x = n' * f x + f x
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
n:

(n + 1) * f = n * f + f
Goals accomplished!

Goals accomplished! 🐙
noncomputable instance : NormedAddCommGroup (H v e μ) where dist := λ f g dist f g edist := λ f g ENNReal.ofReal (dist f g) norm := λ f norm f add := λ f g f + g add_assoc := Group.H_add_assoc zero_add := Group.H_zero_add add_zero := Group.H_add_zero nsmul := λ n f (n : ℝ) * f neg := λ f -f zsmul := λ z f (z : ℝ) * f add_comm := Group.H_add_comm dist_self := Dist.H_dist_self dist_comm := Dist.H_dist_comm dist_triangle := Dist.H_dist_triangle edist_dist := λ f g rfl eq_of_dist_eq_zero := Dist.H_eq_of_dist_eq_zero dist_eq := λ x y rfl neg_add_cancel := λ x Group.H_add_left_neg x nsmul_zero :=

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), 0 * x = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), 0 * x = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), 0 * x = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 * f = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), 0 * x = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 * f = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 * f = 0

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 * f = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), 0 * x = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
(0 * f) x = 0 x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), 0 * x = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
(0 * f) x = 0 x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
(0 * f) x = 0 x

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
0 * f x = 0 x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), 0 * x = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
0 * f x = 0 x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
0 * f x = 0 x

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
0 * f x = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), 0 * x = 0

Goals accomplished! 🐙

Goals accomplished! 🐙
nsmul_succ :=

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (x : (H v e μ)), (n + 1) * x = n * x + x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (x : (H v e μ)), (n + 1) * x = n * x + x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (x : (H v e μ)), (n + 1) * x = n * x + x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(n + 1) * f = n * f + f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (x : (H v e μ)), (n + 1) * x = n * x + x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(n + 1) * f = n * f + f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(n + 1) * f = n * f + f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(n + 1) * f = n * f + f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (x : (H v e μ)), (n + 1) * x = n * x + x

Goals accomplished! 🐙

Goals accomplished! 🐙
zsmul_zero' :=

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : (H v e μ)), 0 * a = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : (H v e μ)), 0 * a = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : (H v e μ)), 0 * a = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 * f = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : (H v e μ)), 0 * a = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 * f = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 * f = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 * f = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : (H v e μ)), 0 * a = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
(0 * f) x = 0 x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : (H v e μ)), 0 * a = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
(0 * f) x = 0 x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
(0 * f) x = 0 x

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
0 * f x = 0 x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : (H v e μ)), 0 * a = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
0 * f x = 0 x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
0 * f x = 0 x

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
0 * f x = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : (H v e μ)), 0 * a = 0

Goals accomplished! 🐙

Goals accomplished! 🐙
zsmul_succ' :=

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), ↑↑n.succ * a = ↑↑n * a + a
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), ↑↑n.succ * a = ↑↑n * a + a
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), ↑↑n.succ * a = ↑↑n * a + a
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

↑↑n.succ * f = ↑↑n * f + f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), ↑↑n.succ * a = ↑↑n * a + a
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

↑↑n.succ * f = ↑↑n * f + f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

↑↑n.succ * f = ↑↑n * f + f

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

↑↑(n + 1) * f = ↑↑n * f + f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), ↑↑n.succ * a = ↑↑n * a + a
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

↑↑(n + 1) * f = ↑↑n * f + f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

↑↑(n + 1) * f = ↑↑n * f + f

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(n + 1) * f = ↑↑n * f + f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), ↑↑n.succ * a = ↑↑n * a + a
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(n + 1) * f = ↑↑n * f + f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(n + 1) * f = ↑↑n * f + f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(n + 1) * f = ↑↑n * f + f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), ↑↑n.succ * a = ↑↑n * a + a

Goals accomplished! 🐙

Goals accomplished! 🐙
zsmul_neg' :=

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), (Int.negSucc n) * a = -(↑↑n.succ * a)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), (Int.negSucc n) * a = -(↑↑n.succ * a)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), (Int.negSucc n) * a = -(↑↑n.succ * a)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(Int.negSucc n) * f = -(↑↑n.succ * f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), (Int.negSucc n) * a = -(↑↑n.succ * a)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(Int.negSucc n) * f = -(↑↑n.succ * f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(Int.negSucc n) * f = -(↑↑n.succ * f)

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -(↑↑n.succ * f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), (Int.negSucc n) * a = -(↑↑n.succ * a)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -(↑↑n.succ * f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -(↑↑n.succ * f)

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -(↑↑(n + 1) * f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), (Int.negSucc n) * a = -(↑↑n.succ * a)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -(↑↑(n + 1) * f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -(↑↑(n + 1) * f)

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -1 * (fun z f => z * f) ((n + 1)) f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), (Int.negSucc n) * a = -(↑↑n.succ * a)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -1 * (fun z f => z * f) ((n + 1)) f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -1 * (fun z f => z * f) ((n + 1)) f

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -1 * (fun z f => z * f) (n + 1) f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), (Int.negSucc n) * a = -(↑↑n.succ * a)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -1 * (fun z f => z * f) (n + 1) f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -1 * ((n + 1) * f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -1 * ((n + 1) * f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), (Int.negSucc n) * a = -(↑↑n.succ * a)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -1 * ((n + 1) * f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -1 * ((n + 1) * f)

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -1 * ((n + 1) * f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), (Int.negSucc n) * a = -(↑↑n.succ * a)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -1 * ((n + 1) * f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

(-(n + 1)) * f = -1 * ((n + 1) * f)

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)

-(n + 1) * f = -1 * ((n + 1) * f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), (Int.negSucc n) * a = -(↑↑n.succ * a)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)
x: ↑Ω

a.h
(-(n + 1) * f) x = (-1 * ((n + 1) * f)) x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), (Int.negSucc n) * a = -(↑↑n.succ * a)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)
x: ↑Ω

a.h
(-(n + 1) * f) x = (-1 * ((n + 1) * f)) x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)
x: ↑Ω

a.h
(-(n + 1) * f) x = (-1 * ((n + 1) * f)) x

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)
x: ↑Ω

a.h
-(n + 1) * f x = (-1 * ((n + 1) * f)) x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), (Int.negSucc n) * a = -(↑↑n.succ * a)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)
x: ↑Ω

a.h
-(n + 1) * f x = (-1 * ((n + 1) * f)) x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)
x: ↑Ω

a.h
-(n + 1) * f x = (-1 * ((n + 1) * f)) x

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)
x: ↑Ω

a.h
-(n + 1) * f x = -1 * ((n + 1) * f) x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), (Int.negSucc n) * a = -(↑↑n.succ * a)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)
x: ↑Ω

a.h
-(n + 1) * f x = -1 * ((n + 1) * f) x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)
x: ↑Ω

a.h
-(n + 1) * f x = -1 * ((n + 1) * f) x

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
n:
f: (H v e μ)
x: ↑Ω

a.h
-(n + 1) * f x = -1 * ((n + 1) * f x)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(n : ℕ) (a : (H v e μ)), (Int.negSucc n) * a = -(↑↑n.succ * a)

Goals accomplished! 🐙

Goals accomplished! 🐙
open Inner noncomputable instance : InnerProductSpace ℝ (H v e μ) where smul := λ a f a * f one_smul :=

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(b : (H v e μ)), 1 b = b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(b : (H v e μ)), 1 b = b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(b : (H v e μ)), 1 b = b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
b: (H v e μ)

1 b = b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(b : (H v e μ)), 1 b = b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
b: (H v e μ)
x: ↑Ω

a.h
(1 b) x = b x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(b : (H v e μ)), 1 b = b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
b: (H v e μ)
x: ↑Ω

a.h
(1 b) x = b x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
b: (H v e μ)
x: ↑Ω

a.h
(1 b) x = b x

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
b: (H v e μ)
x: ↑Ω

a.h
1 * b x = b x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(b : (H v e μ)), 1 b = b

Goals accomplished! 🐙

Goals accomplished! 🐙
mul_smul :=

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : ℝ) (b : (H v e μ)), (x * y) b = x y b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : ℝ) (b : (H v e μ)), (x * y) b = x y b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : ℝ) (b : (H v e μ)), (x * y) b = x y b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
x, y:
b: (H v e μ)

(x * y) b = x y b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : ℝ) (b : (H v e μ)), (x * y) b = x y b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e✝: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
x, y:
b: (H v e μ)
e: ↑Ω

a.h
((x * y) b) e = (x y b) e
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : ℝ) (b : (H v e μ)), (x * y) b = x y b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e✝: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
x, y:
b: (H v e μ)
e: ↑Ω

a.h
((x * y) b) e = (x y b) e
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e✝: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
x, y:
b: (H v e μ)
e: ↑Ω

a.h
((x * y) b) e = (x y b) e

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e✝: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
x, y:
b: (H v e μ)
e: ↑Ω

a.h
x * y * b e = (x y b) e
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : ℝ) (b : (H v e μ)), (x * y) b = x y b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e✝: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
x, y:
b: (H v e μ)
e: ↑Ω

a.h
x * y * b e = (x y b) e
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e✝: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
x, y:
b: (H v e μ)
e: ↑Ω

a.h
x * y * b e = (x y b) e

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e✝: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
x, y:
b: (H v e μ)
e: ↑Ω

a.h
x * y * b e = x * y b e
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : ℝ) (b : (H v e μ)), (x * y) b = x y b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e✝: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
x, y:
b: (H v e μ)
e: ↑Ω

a.h
x * y * b e = x * y b e
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e✝: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
x, y:
b: (H v e μ)
e: ↑Ω

a.h
x * y * b e = x * y b e

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e✝: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
x, y:
b: (H v e μ)
e: ↑Ω

a.h
x * y * b e = x * (y * b e)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : ℝ) (b : (H v e μ)), (x * y) b = x y b

Goals accomplished! 🐙

Goals accomplished! 🐙
smul_zero :=

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ), a 0 = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ), a 0 = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ), a 0 = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:

a 0 = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ), a 0 = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
x: ↑Ω

a.h
(a 0) x = 0 x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ), a 0 = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
x: ↑Ω

a.h
(a 0) x = 0 x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
x: ↑Ω

a.h
(a 0) x = 0 x

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
x: ↑Ω

a.h
a * 0 = 0 x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ), a 0 = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
x: ↑Ω

a.h
a * 0 = 0 x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
x: ↑Ω

a.h
a * 0 = 0 x

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
x: ↑Ω

a.h
a * 0 = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ), a 0 = 0

Goals accomplished! 🐙

Goals accomplished! 🐙
smul_add :=

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (x y : (H v e μ)), a (x + y) = a x + a y
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (x y : (H v e μ)), a (x + y) = a x + a y
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (x y : (H v e μ)), a (x + y) = a x + a y
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
f, g: (H v e μ)

a (f + g) = a f + a g
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (x y : (H v e μ)), a (x + y) = a x + a y
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
f, g: (H v e μ)
x: ↑Ω

a.h
(a (f + g)) x = (a f + a g) x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (x y : (H v e μ)), a (x + y) = a x + a y
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
f, g: (H v e μ)
x: ↑Ω

a.h
(a (f + g)) x = (a f + a g) x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
f, g: (H v e μ)
x: ↑Ω

a.h
(a (f + g)) x = (a f + a g) x

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
f, g: (H v e μ)
x: ↑Ω

a.h
a * (f + g) x = (a f + a g) x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (x y : (H v e μ)), a (x + y) = a x + a y
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
f, g: (H v e μ)
x: ↑Ω

a.h
a * (f + g) x = (a f + a g) x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
f, g: (H v e μ)
x: ↑Ω

a.h
a * (f + g) x = (a f + a g) x

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
f, g: (H v e μ)
x: ↑Ω

a.h
a * (f x + g x) = (a f + a g) x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (x y : (H v e μ)), a (x + y) = a x + a y
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
f, g: (H v e μ)
x: ↑Ω

a.h
a * (f x + g x) = (a f + a g) x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
f, g: (H v e μ)
x: ↑Ω

a.h
a * (f x + g x) = (a f + a g) x

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
a:
f, g: (H v e μ)
x: ↑Ω

a.h
a * (f x + g x) = a * f x + a * g x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (x y : (H v e μ)), a (x + y) = a x + a y

Goals accomplished! 🐙

Goals accomplished! 🐙
add_smul :=

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(r s : ℝ) (x : (H v e μ)), (r + s) x = r x + s x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(r s : ℝ) (x : (H v e μ)), (r + s) x = r x + s x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(r s : ℝ) (x : (H v e μ)), (r + s) x = r x + s x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r, s:
f: (H v e μ)

(r + s) f = r f + s f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(r s : ℝ) (x : (H v e μ)), (r + s) x = r x + s x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r, s:
f: (H v e μ)
x: ↑Ω

a.h
((r + s) f) x = (r f + s f) x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(r s : ℝ) (x : (H v e μ)), (r + s) x = r x + s x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r, s:
f: (H v e μ)
x: ↑Ω

a.h
((r + s) f) x = (r f + s f) x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r, s:
f: (H v e μ)
x: ↑Ω

a.h
((r + s) f) x = (r f + s f) x

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r, s:
f: (H v e μ)
x: ↑Ω

a.h
(r + s) * f x = (r f + s f) x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(r s : ℝ) (x : (H v e μ)), (r + s) x = r x + s x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r, s:
f: (H v e μ)
x: ↑Ω

a.h
(r + s) * f x = (r f + s f) x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r, s:
f: (H v e μ)
x: ↑Ω

a.h
(r + s) * f x = (r f + s f) x

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r, s:
f: (H v e μ)
x: ↑Ω

a.h
(r + s) * f x = r * f x + s * f x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(r s : ℝ) (x : (H v e μ)), (r + s) x = r x + s x

Goals accomplished! 🐙

Goals accomplished! 🐙
zero_smul :=

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), 0 x = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), 0 x = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), 0 x = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

0 f = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), 0 x = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
(0 f) x = 0 x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), 0 x = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
(0 f) x = 0 x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
(0 f) x = 0 x

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
0 * f x = 0 x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), 0 x = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
0 * f x = 0 x
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
0 * f x = 0 x

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)
x: ↑Ω

a.h
0 * f x = 0
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), 0 x = 0

Goals accomplished! 🐙

Goals accomplished! 🐙
norm_smul_le :=

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (b : (H v e μ)), a b a * b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (b : (H v e μ)), a b a * b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (b : (H v e μ)), a b a * b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

r f r * f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (b : (H v e μ)), a b a * b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r f = r * f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (b : (H v e μ)), a b a * b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r f = r * f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r f = r * f

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * f = r * f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (b : (H v e μ)), a b a * b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * f = r * f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * f = r * f

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * f = |r| * f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (b : (H v e μ)), a b a * b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * f = |r| * f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * f * r * f = |r| * f * (|r| * f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * f * r * f = |r| * f * (|r| * f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (b : (H v e μ)), a b a * b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * f * r * f = |r| * f * (|r| * f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * f * r * f = |r| * f * (|r| * f)

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * f ^ 2 = |r| * f * (|r| * f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (b : (H v e μ)), a b a * b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * f ^ 2 = |r| * f * (|r| * f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * f ^ 2 = |r| * f * (|r| * f)

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * f ^ 2 = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (b : (H v e μ)), a b a * b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * f ^ 2 = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
inner (r * f) (r * f) = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * f ^ 2 = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * inner f (r * f) = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * inner f (r * f) = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (b : (H v e μ)), a b a * b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * inner f (r * f) = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * inner (r * f) f = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * inner f (r * f) = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * (r * inner f f) = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * inner f (r * f) = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * (r * f ^ 2) = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * (r * f ^ 2) = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (b : (H v e μ)), a b a * b
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * (r * f ^ 2) = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * (r * f ^ 2) = (|r| * f) ^ 2

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r ^ 2 * f ^ 2 = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
r * (r * f ^ 2) = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
|r| ^ 2 * f ^ 2 = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
r:
f: (H v e μ)

hab
|r| ^ 2 * f ^ 2 = (|r| * f) ^ 2
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(a : ℝ) (b : (H v e μ)), a b a * b

Goals accomplished! 🐙

Goals accomplished! 🐙
inner := H_inner norm_sq_eq_inner :=

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), x ^ 2 = RCLike.re (H_inner x x)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), x ^ 2 = RCLike.re (H_inner x x)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), x ^ 2 = RCLike.re (H_inner x x)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

f ^ 2 = RCLike.re (H_inner f f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), x ^ 2 = RCLike.re (H_inner x x)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

f ^ 2 = RCLike.re (H_inner f f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = RCLike.re (H_inner f f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f: (H v e μ)

inner f f = RCLike.re (H_inner f f)
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x : (H v e μ)), x ^ 2 = RCLike.re (H_inner x x)

Goals accomplished! 🐙

Goals accomplished! 🐙
conj_symm :=

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : (H v e μ)), (starRingEnd ℝ) (H_inner y x) = H_inner x y
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : (H v e μ)), (starRingEnd ℝ) (H_inner y x) = H_inner x y
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : (H v e μ)), (starRingEnd ℝ) (H_inner y x) = H_inner x y
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

(starRingEnd ℝ) (H_inner g f) = H_inner f g
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : (H v e μ)), (starRingEnd ℝ) (H_inner y x) = H_inner x y
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

(starRingEnd ℝ) (H_inner g f) = H_inner f g
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

(starRingEnd ℝ) (H_inner g f) = H_inner f g

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

(starRingEnd ℝ) (H_inner g f) = inner f g
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

(starRingEnd ℝ) (H_inner g f) = H_inner f g
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

(starRingEnd ℝ) (H_inner g f) = inner g f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)

(starRingEnd ℝ) (H_inner g f) = inner g f
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : (H v e μ)), (starRingEnd ℝ) (H_inner y x) = H_inner x y

Goals accomplished! 🐙

Goals accomplished! 🐙
add_left :=

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y z : (H v e μ)), H_inner (x + y) z = H_inner x z + H_inner y z
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y z : (H v e μ)), H_inner (x + y) z = H_inner x z + H_inner y z
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y z : (H v e μ)), H_inner (x + y) z = H_inner x z + H_inner y z
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g, h: (H v e μ)

H_inner (f + g) h = H_inner f h + H_inner g h
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y z : (H v e μ)), H_inner (x + y) z = H_inner x z + H_inner y z

Goals accomplished! 🐙

Goals accomplished! 🐙
smul_left :=

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : (H v e μ)) (r : ℝ), H_inner (r x) y = (starRingEnd ℝ) r * H_inner x y
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : (H v e μ)) (r : ℝ), H_inner (r x) y = (starRingEnd ℝ) r * H_inner x y
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : (H v e μ)) (r : ℝ), H_inner (r x) y = (starRingEnd ℝ) r * H_inner x y
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
r:

H_inner (r f) g = (starRingEnd ℝ) r * H_inner f g
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : (H v e μ)) (r : ℝ), H_inner (r x) y = (starRingEnd ℝ) r * H_inner x y
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
r:

H_inner (r f) g = (starRingEnd ℝ) r * H_inner f g
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
r:

H_inner (r f) g = (starRingEnd ℝ) r * H_inner f g

Goals accomplished! 🐙
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
f, g: (H v e μ)
r:

H_inner (r * f) g = (starRingEnd ℝ) r * H_inner f g
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
v: eigen
e: ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ

(x y : (H v e μ)) (r : ℝ), H_inner (r x) y = (starRingEnd ℝ) r * H_inner x y

Goals accomplished! 🐙

Goals accomplished! 🐙
/- --- MERCER --- -/ /- Given the Mercer's theorem, we prove that H, endowed with k is a RKHS. -/ def mercer (v : eigen) (e : ℕ Ω ℝ) (k : Ω Ω ℝ) := s, (k s = λ t ∑' i, v.1 i * e i s * e i t) ( t, Summable (fun i v.1 i * e i s * e i t)) omit [MeasureSpace ↑Ω] in
Goals accomplished!
omit [MeasureSpace ↑Ω] in
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k

(s : ↑Ω), Summable fun i => v i * e i s ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
s: ↑Ω

Summable fun i => v i * e i s ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k

(s : ↑Ω), Summable fun i => v i * e i s ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
s: ↑Ω

Summable fun i => v i * e i s ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
s: ↑Ω

Summable fun i => v i * e i s ^ 2
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
s: ↑Ω
i:

v i * e i s ^ 2 = v i * e i s * e i s
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
s: ↑Ω
i:

v i * e i s ^ 2 = v i * e i s * e i s
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
s: ↑Ω

(i : ℕ), v i * e i s ^ 2 = v i * e i s * e i s
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
s: ↑Ω

Summable fun i => v i * e i s * e i s
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k

(s : ↑Ω), Summable fun i => v i * e i s ^ 2
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
Goals accomplished!
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ

(f : (H v e μ)) (x : ↑Ω), f x = inner f ⟨k x,
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ
f: (H v e μ)
x: ↑Ω

f x = inner f ⟨k x,
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ

(f : (H v e μ)) (x : ↑Ω), f x = inner f ⟨k x,
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ
f: (H v e μ)
x: ↑Ω
k_H:= ⟨k x, : (H v e μ)

f x = inner f ⟨k x,
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ

(f : (H v e μ)) (x : ↑Ω), f x = inner f ⟨k x,
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ
f: (H v e μ)
x: ↑Ω
k_H:= ⟨k x, : (H v e μ)

f x = inner f ⟨k x,
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ
f: (H v e μ)
x: ↑Ω
k_H:= ⟨k x, : (H v e μ)

f x = inner f ⟨k x,
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ
f: (H v e μ)
x: ↑Ω
k_H:= ⟨k x, : (H v e μ)

f x = H_inner f k_H
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ

(f : (H v e μ)) (x : ↑Ω), f x = inner f ⟨k x,
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ
f: (H v e μ)
x: ↑Ω
k_H:= ⟨k x, : (H v e μ)

f x = ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ

(f : (H v e μ)) (x : ↑Ω), f x = inner f ⟨k x,
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ
f: (H v e μ)
x: ↑Ω
k_H:= ⟨k x, : (H v e μ)
a_f:= ⋯.some:

f x = ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ

(f : (H v e μ)) (x : ↑Ω), f x = inner f ⟨k x,
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ
f: (H v e μ)
x: ↑Ω
k_H:= ⟨k x, : (H v e μ)
a_f:= ⋯.some:
this: ⋯.some = fun i => e i x

f x = ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ

(f : (H v e μ)) (x : ↑Ω), f x = inner f ⟨k x,
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ
f: (H v e μ)
x: ↑Ω
k_H:= ⟨k x, : (H v e μ)
a_f:= ⋯.some:
this: ⋯.some = fun i => e i x

f x = ∑' (i : ℕ), v i * ⋯.some i * ⋯.some i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ
f: (H v e μ)
x: ↑Ω
k_H:= ⟨k x, : (H v e μ)
a_f:= ⋯.some:
this: ⋯.some = fun i => e i x

f x = ∑' (i : ℕ), v i * ⋯.some i * (fun i => e i x) i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ
f: (H v e μ)
x: ↑Ω
k_H:= ⟨k x, : (H v e μ)
a_f:= ⋯.some:
this: ⋯.some = fun i => e i x

f x = ∑' (i : ℕ), v i * ⋯.some i * (fun i => e i x) i
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ

(f : (H v e μ)) (x : ↑Ω), f x = inner f ⟨k x,
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ
f: (H v e μ)
x: ↑Ω
k_H:= ⟨k x, : (H v e μ)
a_f:= ⋯.some:
this: ⋯.some = fun i => e i x

f x = ∑' (i : ℕ), v i * ⋯.some i * (fun i => e i x) i
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ
f: (H v e μ)
x: ↑Ω
k_H:= ⟨k x, : (H v e μ)
a_f:= ⋯.some:
this✝: ⋯.some = fun i => e i x
this: f = fun x => ∑' (i : ℕ), v i * a_f i * e i x

f x = ∑' (i : ℕ), v i * ⋯.some i * (fun i => e i x) i
Goals accomplished!

Goals accomplished! 🐙
Goals accomplished!
d:
Ω: Set (Vector ℝ d)
inst✝¹: MeasureSpace ↑Ω
μ: Measure ↑Ω
inst✝: IsFiniteMeasure μ
v: eigen
e: ↑Ω
k: ↑Ω ↑Ω
h_mercer: mercer v e k
hk_l2: (s : ↑Ω), k s L2 μ

(f : (H v e μ)) (x : ↑Ω), f x = inner f ⟨k x,
Goals accomplished!

Goals accomplished! 🐙
variable (k : Ω Ω ℝ) (v : eigen) (e : ℕ Ω ℝ) (h_mercer : mercer v e k) (hk_l2 : s, k s L2 μ) instance : RKHS (H v e μ) where k := k memb := k_i_H h_mercer hk_l2 repro := k_repro h_mercer hk_l2