working...

This commit is contained in:
Stefan Kebekus 2024-06-29 17:17:19 +02:00
parent 602296031d
commit 323b133c88
2 changed files with 86 additions and 18 deletions

View File

@ -1,5 +1,34 @@
/-
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 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 open TensorProduct
@ -14,26 +43,28 @@ lemma OrthonormalBasis.sum_repr'
simp_rw [b.repr_apply_apply v] simp_rw [b.repr_apply_apply v]
noncomputable def InnerProductSpace.canonicalTensor noncomputable def InnerProductSpace.canonicalContravariantTensor
(E : Type u_2) [NormedAddCommGroup E] [InnerProductSpace E] [FiniteDimensional E] {E : Type*} [NormedAddCommGroup E] [InnerProductSpace E]
: E ⊗[] E := by : E ⊗[] E →ₗ[] := TensorProduct.lift bilinFormOfRealInner
let v := stdOrthonormalBasis E
exact ∑ i, (v i) ⊗ₜ[] (v i)
theorem InnerProductSpace.InvariantTensor noncomputable def InnerProductSpace.canonicalCovariantTensor
(E : Type*) [NormedAddCommGroup E] [InnerProductSpace E] [FiniteDimensional E] (E : Type*) [NormedAddCommGroup E] [InnerProductSpace E] [FiniteDimensional E]
(v : OrthonormalBasis (Fin (FiniteDimensional.finrank E)) E) : E ⊗[] E := ∑ i, ((stdOrthonormalBasis E) i) ⊗ₜ[] ((stdOrthonormalBasis E) i)
: InnerProductSpace.canonicalTensor E = ∑ i, (v i) ⊗ₜ[] (v i) := by
unfold InnerProductSpace.canonicalTensor
let v₁ := stdOrthonormalBasis E
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 => conv =>
right right
arg 2 arg 2
intro i intro i
rw [v₁.sum_repr' (v i)] rw [w.sum_repr' (v i)]
simp_rw [TensorProduct.sum_tmul, TensorProduct.tmul_sum, TensorProduct.smul_tmul_smul] simp_rw [TensorProduct.sum_tmul, TensorProduct.tmul_sum, TensorProduct.smul_tmul_smul]
conv => conv =>
@ -48,11 +79,11 @@ theorem InnerProductSpace.InvariantTensor
arg 1 arg 1
arg 2 arg 2
intro i intro i
rw [← real_inner_comm (v₁ x)] rw [← real_inner_comm (w x)]
simp_rw [OrthonormalBasis.sum_inner_mul_inner v] simp_rw [OrthonormalBasis.sum_inner_mul_inner v]
have {x : Fin (FiniteDimensional.finrank E)} : ∑ x_1 : Fin (FiniteDimensional.finrank E), ⟪v₁ x, v₁ x_1⟫_ • v₁ x ⊗ₜ[] v₁ x_1 = v₁ x ⊗ₜ[] v₁ x := by have {i} : ∑ j, ⟪w i, w j⟫_ • w i ⊗ₜ[] w j = w i ⊗ₜ[] w i := by
rw [Fintype.sum_eq_single x, orthonormal_iff_ite.1 v₁.orthonormal]; simp rw [Fintype.sum_eq_single i, orthonormal_iff_ite.1 w.orthonormal]; simp
intro r₁ hr₁ intro _ _; rw [orthonormal_iff_ite.1 w.orthonormal]; simp; tauto
rw [orthonormal_iff_ite.1 v₁.orthonormal]; simp; tauto
simp_rw [this] simp_rw [this]
rfl

View File

@ -10,7 +10,6 @@ import Mathlib.LinearAlgebra.Contraction
open BigOperators open BigOperators
open Finset open Finset
lemma OrthonormalBasis.sum_repr' lemma OrthonormalBasis.sum_repr'
{𝕜 : Type*} [RCLike 𝕜] {𝕜 : Type*} [RCLike 𝕜]
{E : Type*} [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] {E : Type*} [NormedAddCommGroup E] [InnerProductSpace 𝕜 E]
@ -21,15 +20,53 @@ lemma OrthonormalBasis.sum_repr'
nth_rw 1 [← (b.sum_repr v)] nth_rw 1 [← (b.sum_repr v)]
simp_rw [b.repr_apply_apply v] simp_rw [b.repr_apply_apply v]
variable {E : Type*} [NormedAddCommGroup E] [InnerProductSpace E] [FiniteDimensional E]
noncomputable def realInnerAsElementOfDualTensorprod noncomputable def realInnerAsElementOfDualTensorprod
{E : Type*} [NormedAddCommGroup E] [InnerProductSpace E] {E : Type*} [NormedAddCommGroup E] [InnerProductSpace E]
: TensorProduct E E →ₗ[] := TensorProduct.lift bilinFormOfRealInner : TensorProduct E E →ₗ[] := TensorProduct.lift bilinFormOfRealInner
instance
{E : Type*} [NormedAddCommGroup E] [InnerProductSpace E] [CompleteSpace E] [FiniteDimensional E]
: NormedAddCommGroup (TensorProduct E E) where
norm := by
sorry
dist_self := by
sorry
sorry
/-
instance
{E : Type*} [NormedAddCommGroup E] [InnerProductSpace E] [CompleteSpace E] [FiniteDimensional E]
: InnerProductSpace (TensorProduct E E) where
smul := _
one_smul := _
mul_smul := _
smul_zero := _
smul_add := _
add_smul := _
zero_smul := _
norm_smul_le := _
inner := _
norm_sq_eq_inner := _
conj_symm := _
add_left := _
smul_left := _
-/
noncomputable def dual' noncomputable def dual'
{E : Type*} [NormedAddCommGroup E] [InnerProductSpace E] [CompleteSpace E] [FiniteDimensional E] {E : Type*} [NormedAddCommGroup E] [InnerProductSpace E] [CompleteSpace E] [FiniteDimensional E]
: (TensorProduct E E →ₗ[] ) ≃ₗ[] TensorProduct E E := by : (TensorProduct E E →ₗ[] ) ≃ₗ[] TensorProduct E E := by
let d := InnerProductSpace.toDual E let d := InnerProductSpace.toDual E
let e := d.toLinearEquiv let e := d.toLinearEquiv
let a := TensorProduct.congr e e let a := TensorProduct.congr e e