Documentation

Mathlib.Analysis.InnerProductSpace.LinearMap

Linear maps on inner product spaces #

This file studies linear maps on inner product spaces.

Main results #

Tags #

inner product space, Hilbert space, norm

theorem inner_map_polarization {V : Type u_4} [SeminormedAddCommGroup V] [InnerProductSpace V] (T : V →ₗ[] V) (x y : V) :
inner (T y) x = (inner (T (x + y)) (x + y) - inner (T (x - y)) (x - y) + Complex.I * inner (T (x + Complex.I y)) (x + Complex.I y) - Complex.I * inner (T (x - Complex.I y)) (x - Complex.I y)) / 4

A complex polarization identity, with a linear map.

theorem inner_map_polarization' {V : Type u_4} [SeminormedAddCommGroup V] [InnerProductSpace V] (T : V →ₗ[] V) (x y : V) :
inner (T x) y = (inner (T (x + y)) (x + y) - inner (T (x - y)) (x - y) - Complex.I * inner (T (x + Complex.I y)) (x + Complex.I y) + Complex.I * inner (T (x - Complex.I y)) (x - Complex.I y)) / 4
theorem inner_map_self_eq_zero {V : Type u_4} [NormedAddCommGroup V] [InnerProductSpace V] (T : V →ₗ[] V) :
(∀ (x : V), inner (T x) x = 0) T = 0

A linear map T is zero, if and only if the identity ⟪T x, x⟫_ℂ = 0 holds for all x.

theorem ext_inner_map {V : Type u_4} [NormedAddCommGroup V] [InnerProductSpace V] (S T : V →ₗ[] V) :
(∀ (x : V), inner (S x) x = inner (T x) x) S = T

Two linear maps S and T are equal, if and only if the identity ⟪S x, x⟫_ℂ = ⟪T x, x⟫_ℂ holds for all x.

@[simp]
theorem LinearIsometry.inner_map_map {𝕜 : Type u_1} {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] (f : E →ₗᵢ[𝕜] E') (x y : E) :
inner 𝕜 (f x) (f y) = inner 𝕜 x y

A linear isometry preserves the inner product.

@[simp]
theorem LinearIsometryEquiv.inner_map_map {𝕜 : Type u_1} {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] (f : E ≃ₗᵢ[𝕜] E') (x y : E) :
inner 𝕜 (f x) (f y) = inner 𝕜 x y

A linear isometric equivalence preserves the inner product.

theorem LinearIsometryEquiv.inner_map_eq_flip {𝕜 : Type u_1} {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] (f : E ≃ₗᵢ[𝕜] E') (x : E) (y : E') :
inner 𝕜 (f x) y = inner 𝕜 x (f.symm y)

The adjoint of a linear isometric equivalence is its inverse.

def LinearMap.isometryOfInner {𝕜 : Type u_1} {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] (f : E →ₗ[𝕜] E') (h : ∀ (x y : E), inner 𝕜 (f x) (f y) = inner 𝕜 x y) :
E →ₗᵢ[𝕜] E'

A linear map that preserves the inner product is a linear isometry.

Equations
    Instances For
      @[simp]
      theorem LinearMap.coe_isometryOfInner {𝕜 : Type u_1} {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] (f : E →ₗ[𝕜] E') (h : ∀ (x y : E), inner 𝕜 (f x) (f y) = inner 𝕜 x y) :
      (f.isometryOfInner h) = f
      @[simp]
      theorem LinearMap.isometryOfInner_toLinearMap {𝕜 : Type u_1} {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] (f : E →ₗ[𝕜] E') (h : ∀ (x y : E), inner 𝕜 (f x) (f y) = inner 𝕜 x y) :
      def LinearEquiv.isometryOfInner {𝕜 : Type u_1} {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] (f : E ≃ₗ[𝕜] E') (h : ∀ (x y : E), inner 𝕜 (f x) (f y) = inner 𝕜 x y) :
      E ≃ₗᵢ[𝕜] E'

      A linear equivalence that preserves the inner product is a linear isometric equivalence.

      Equations
        Instances For
          @[simp]
          theorem LinearEquiv.coe_isometryOfInner {𝕜 : Type u_1} {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] (f : E ≃ₗ[𝕜] E') (h : ∀ (x y : E), inner 𝕜 (f x) (f y) = inner 𝕜 x y) :
          (f.isometryOfInner h) = f
          @[simp]
          theorem LinearEquiv.isometryOfInner_toLinearEquiv {𝕜 : Type u_1} {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] (f : E ≃ₗ[𝕜] E') (h : ∀ (x y : E), inner 𝕜 (f x) (f y) = inner 𝕜 x y) :
          theorem LinearMap.norm_map_iff_inner_map_map {𝕜 : Type u_1} {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] {F : Type u_9} [FunLike F E E'] [LinearMapClass F 𝕜 E E'] (f : F) :
          (∀ (x : E), f x = x) ∀ (x y : E), inner 𝕜 (f x) (f y) = inner 𝕜 x y

          A linear map is an isometry if and it preserves the inner product.

          def innerₛₗ (𝕜 : Type u_1) {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] :
          E →ₗ⋆[𝕜] E →ₗ[𝕜] 𝕜

          The inner product as a sesquilinear map.

          Equations
            Instances For
              @[simp]
              theorem innerₛₗ_apply_coe (𝕜 : Type u_1) {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] (v : E) :
              ((innerₛₗ 𝕜) v) = fun (w : E) => inner 𝕜 v w
              @[simp]
              theorem innerₛₗ_apply (𝕜 : Type u_1) {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] (v w : E) :
              ((innerₛₗ 𝕜) v) w = inner 𝕜 v w

              The inner product as a bilinear map in the real case.

              Equations
                Instances For
                  @[simp]
                  theorem innerₗ_apply {F : Type u_3} [SeminormedAddCommGroup F] [InnerProductSpace F] (v w : F) :
                  ((innerₗ F) v) w = inner v w
                  def innerSL (𝕜 : Type u_1) {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] :
                  E →L⋆[𝕜] E →L[𝕜] 𝕜

                  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).

                  Equations
                    Instances For
                      @[simp]
                      theorem innerSL_apply_coe (𝕜 : Type u_1) {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] (v : E) :
                      ((innerSL 𝕜) v) = fun (w : E) => inner 𝕜 v w
                      @[simp]
                      theorem innerSL_apply (𝕜 : Type u_1) {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] (v w : E) :
                      ((innerSL 𝕜) v) w = inner 𝕜 v w
                      def innerSLFlip (𝕜 : Type u_1) {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] :
                      E →L[𝕜] E →L⋆[𝕜] 𝕜

                      The inner product as a continuous sesquilinear map, with the two arguments flipped.

                      Equations
                        Instances For
                          @[simp]
                          theorem innerSLFlip_apply (𝕜 : Type u_1) {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] (x y : E) :
                          ((innerSLFlip 𝕜) x) y = inner 𝕜 y x
                          noncomputable def ContinuousLinearMap.toSesqForm {𝕜 : Type u_1} {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] {E' : Type u_4} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] :
                          (E →L[𝕜] E') →L[𝕜] E' →L⋆[𝕜] E →L[𝕜] 𝕜

                          Given f : E →L[𝕜] E', construct the continuous sesquilinear form fun x y ↦ ⟪x, A y⟫, given as a continuous linear map.

                          Equations
                            Instances For
                              @[simp]
                              theorem ContinuousLinearMap.toSesqForm_apply_coe {𝕜 : Type u_1} {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] {E' : Type u_4} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] (f : E →L[𝕜] E') (x : E') :
                              (toSesqForm f) x = ((innerSL 𝕜) x).comp f
                              theorem ContinuousLinearMap.toSesqForm_apply_norm_le {𝕜 : Type u_1} {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] {E' : Type u_4} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] {f : E →L[𝕜] E'} {v : E'} :
                              @[simp]
                              theorem innerSL_apply_norm (𝕜 : Type u_1) {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] (x : E) :

                              innerSL is an isometry. Note that the associated LinearIsometry is defined in InnerProductSpace.Dual as toDualMap.

                              theorem norm_innerSL_le (𝕜 : Type u_1) {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] :
                              theorem inner_map_complex {G : Type u_4} [SeminormedAddCommGroup G] [InnerProductSpace G] (f : G ≃ₗᵢ[] ) (x y : G) :
                              inner x y = (f y * (starRingEnd ) (f x)).re

                              The inner product on an inner product space of dimension 2 can be evaluated in terms of a complex-number representation of the space.

                              def ContinuousLinearMap.reApplyInnerSelf {𝕜 : Type u_1} {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] (T : E →L[𝕜] E) (x : E) :

                              Extract a real bilinear form from an operator T, by taking the pairing fun x ↦ re ⟪T x, x⟫.

                              Equations
                                Instances For
                                  theorem ContinuousLinearMap.reApplyInnerSelf_apply {𝕜 : Type u_1} {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] (T : E →L[𝕜] E) (x : E) :
                                  T.reApplyInnerSelf x = RCLike.re (inner 𝕜 (T x) x)
                                  theorem ContinuousLinearMap.reApplyInnerSelf_smul {𝕜 : Type u_1} {E : Type u_2} [RCLike 𝕜] [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] (T : E →L[𝕜] E) (x : E) {c : 𝕜} :