nevanlinna/Nevanlinna/bilinear.lean

79 lines
2.6 KiB
Plaintext
Raw Permalink Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

/-
Copyright (c) 2024 Stefan Kebekus. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Stefan Kebekus
-/
import Mathlib.Analysis.InnerProductSpace.PiL2
/-!
# Canoncial Elements in Tensor Powers of Real Inner Product Spaces
Given an `InnerProductSpace E`, this file defines two canonical tensors, which
are relevant when for a coordinate-free definition of differential operators
such as the Laplacian.
* `InnerProductSpace.canonicalContravariantTensor E : E ⊗[] E →ₗ[] `. This is
the element corresponding to the inner product.
* If `E` is finite-dimensional, then `E ⊗[] E` is canonically isomorphic to its
dual. Accordingly, there exists an element
`InnerProductSpace.canonicalCovariantTensor E : E ⊗[] E` that corresponds to
`InnerProductSpace.canonicalContravariantTensor E` under this identification.
The theorem `InnerProductSpace.canonicalCovariantTensorRepresentation` shows
that `InnerProductSpace.canonicalCovariantTensor E` can be computed from any
orthonormal basis `v` as `∑ i, (v i) ⊗ₜ[] (v i)`.
-/
open TensorProduct
noncomputable def InnerProductSpace.canonicalContravariantTensor
{E : Type*} [NormedAddCommGroup E] [InnerProductSpace E]
: E ⊗[] E →ₗ[] := TensorProduct.lift bilinFormOfRealInner
noncomputable def InnerProductSpace.canonicalCovariantTensor
(E : Type*) [NormedAddCommGroup E] [InnerProductSpace E] [FiniteDimensional E]
: E ⊗[] E := ∑ i, ((stdOrthonormalBasis E) i) ⊗ₜ[] ((stdOrthonormalBasis E) i)
theorem InnerProductSpace.canonicalCovariantTensorRepresentation
(E : Type*) [NormedAddCommGroup E] [InnerProductSpace E] [FiniteDimensional E]
[Fintype ι]
(v : OrthonormalBasis ι E)
: InnerProductSpace.canonicalCovariantTensor E = ∑ i, (v i) ⊗ₜ[] (v i) := by
let w := stdOrthonormalBasis E
conv =>
right
arg 2
intro i
rw [← w.sum_repr' (v i)]
simp_rw [TensorProduct.sum_tmul, TensorProduct.tmul_sum, TensorProduct.smul_tmul_smul]
conv =>
right
rw [Finset.sum_comm]
arg 2
intro y
rw [Finset.sum_comm]
arg 2
intro x
rw [← Finset.sum_smul]
arg 1
arg 2
intro i
rw [← real_inner_comm (w x)]
simp_rw [OrthonormalBasis.sum_inner_mul_inner v]
have {i} : ∑ j, ⟪w i, w j⟫_ • w i ⊗ₜ[] w j = w i ⊗ₜ[] w i := by
rw [Fintype.sum_eq_single i, orthonormal_iff_ite.1 w.orthonormal]; simp
intro _ _; rw [orthonormal_iff_ite.1 w.orthonormal]; simp; tauto
simp_rw [this]
rfl