Linear maps on inner product spaces #
This file studies linear maps on inner product spaces.
Main results #
- We define
innerSL
as the inner product bundled as a continuous sesquilinear map, andinnerₛₗ
for the non-continuous version. - We prove a general polarization identity for linear maps (
inner_map_polarization
) - We show that a linear map preserving the inner product is an isometry
(
LinearMap.isometryOfInner
) and conversely an isometry preserves the inner product (LinearIsometry.inner_map_map
).
Tags #
inner product space, Hilbert space, norm
A complex polarization identity, with a linear map.
A linear map T
is zero, if and only if the identity ⟪T x, x⟫_ℂ = 0
holds for all x
.
Two linear maps S
and T
are equal, if and only if the identity ⟪S x, x⟫_ℂ = ⟪T x, x⟫_ℂ
holds
for all x
.
A linear isometry preserves the inner product.
A linear isometric equivalence preserves the inner product.
The adjoint of a linear isometric equivalence is its inverse.
A linear map that preserves the inner product is a linear isometry.
Equations
- f.isometryOfInner h = { toLinearMap := f, norm_map' := ⋯ }
Instances For
A linear equivalence that preserves the inner product is a linear isometric equivalence.
Equations
- f.isometryOfInner h = { toLinearEquiv := f, norm_map' := ⋯ }
Instances For
A linear map is an isometry if and it preserves the inner product.
The inner product as a sesquilinear map.
Equations
- innerₛₗ 𝕜 = LinearMap.mk₂'ₛₗ (starRingEnd 𝕜) (RingHom.id 𝕜) (fun (v w : E) => inner v w) ⋯ ⋯ ⋯ ⋯
Instances For
The inner product as a continuous sesquilinear map. Note that toDualMap
(resp. toDual
)
in InnerProductSpace.Dual
is a version of this given as a linear isometry (resp. linear
isometric equivalence).
Instances For
The inner product as a continuous sesquilinear map, with the two arguments flipped.
Equations
- innerSLFlip 𝕜 = (ContinuousLinearMap.flipₗᵢ' E E 𝕜 (RingHom.id 𝕜) (starRingEnd 𝕜)) (innerSL 𝕜)
Instances For
Given f : E →L[𝕜] E'
, construct the continuous sesquilinear form fun x y ↦ ⟪x, A y⟫
, given
as a continuous linear map.
Equations
- One or more equations did not get rendered due to their size.
Instances For
innerSL
is an isometry. Note that the associated LinearIsometry
is defined in
InnerProductSpace.Dual
as toDualMap
.
The inner product on an inner product space of dimension 2 can be evaluated in terms of a complex-number representation of the space.
Extract a real bilinear form from an operator T
,
by taking the pairing fun x ↦ re ⟪T x, x⟫
.