Matrices are isomorphic with linear maps between Lp spaces #
This file provides a WithLp version of Matrix.toLin'.
def
Matrix.toLpLin
{m : Type u_1}
{n : Type u_2}
{R : Type u_4}
[Fintype n]
[DecidableEq n]
[CommRing R]
(p q : ENNReal)
:
Matrix.toLin' adapted for PiLp R _.
Equations
- Matrix.toLpLin p q = Matrix.toLin'.trans ((WithLp.linearEquiv p R (n → R)).symm.arrowCongr (WithLp.linearEquiv q R (m → R)).symm)
Instances For
@[simp]
theorem
Matrix.toLpLin_toLp
{m : Type u_1}
{n : Type u_2}
{R : Type u_4}
[Fintype n]
[DecidableEq n]
[CommRing R]
(p q : ENNReal)
(A : Matrix m n R)
(x : n → R)
:
theorem
Matrix.toLpLin_eq_toLin
{m : Type u_1}
{n : Type u_2}
{R : Type u_4}
[Fintype n]
[DecidableEq n]
[CommRing R]
(p q : ENNReal)
[Finite m]
:
@[simp]
theorem
Matrix.toLpLin_one
{n : Type u_2}
{R : Type u_4}
[Fintype n]
[DecidableEq n]
[CommRing R]
(p : ENNReal)
:
theorem
Matrix.toLpLin_mul
{m : Type u_1}
{n : Type u_2}
{o : Type u_3}
{R : Type u_4}
[Fintype n]
[DecidableEq n]
[CommRing R]
(p q r : ENNReal)
[Fintype o]
[DecidableEq o]
(A : Matrix m n R)
(B : Matrix n o R)
:
Note that applying this theorem needs an explicit choice of q.
@[simp]
theorem
Matrix.toLpLin_mul_same
{m : Type u_1}
{n : Type u_2}
{o : Type u_3}
{R : Type u_4}
[Fintype n]
[DecidableEq n]
[CommRing R]
(p : ENNReal)
[Fintype o]
[DecidableEq o]
(A : Matrix m n R)
(B : Matrix n o R)
:
A copy of toLpLin_mul that works for simp, for the common case where the domain and codomain
have the same norm.
@[simp]
theorem
Matrix.toLpLin_symm_id
{n : Type u_2}
{R : Type u_4}
[Fintype n]
[DecidableEq n]
[CommRing R]
(p : ENNReal)
:
theorem
Matrix.toLpLin_symm_comp
{m : Type u_1}
{n : Type u_2}
{o : Type u_3}
{R : Type u_4}
[Fintype n]
[DecidableEq n]
[CommRing R]
(p q r : ENNReal)
[Fintype o]
[DecidableEq o]
(A : WithLp q (n → R) →ₗ[R] WithLp r (m → R))
(B : WithLp p (o → R) →ₗ[R] WithLp q (n → R))
:
Note that applying this theorem needs an explicit choice of q.
def
Matrix.toLpLinAlgEquiv
{n : Type u_2}
{R : Type u_4}
[Fintype n]
[DecidableEq n]
[CommRing R]
(p : ENNReal)
:
Matrix.toLinAlgEquiv' adapted for PiLp R _.
Equations
- Matrix.toLpLinAlgEquiv p = AlgEquiv.ofLinearEquiv (Matrix.toLpLin p p) ⋯ ⋯
Instances For
@[simp]
theorem
Matrix.toLpLinAlgEquiv_apply_apply_ofLp
{n : Type u_2}
{R : Type u_4}
[Fintype n]
[DecidableEq n]
[CommRing R]
(p : ENNReal)
(a : Matrix n n R)
(a✝ : WithLp p (n → R))
(a✝¹ : n)
:
@[simp]
theorem
Matrix.toLpLinAlgEquiv_symm_apply
{n : Type u_2}
{R : Type u_4}
[Fintype n]
[DecidableEq n]
[CommRing R]
(p : ENNReal)
(a : Module.End R (WithLp p (n → R)))
: