PhysLean/HepLean/SpaceTime/LorentzTensor/Basic.lean

517 lines
21 KiB
Text
Raw Normal View History

2024-07-25 16:57:57 -04:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.LinearAlgebra.PiTensorProduct
2024-07-29 08:38:01 -04:00
import Mathlib.LinearAlgebra.DFinsupp
2024-07-25 16:57:57 -04:00
/-!
2024-07-29 08:38:01 -04:00
# Structure of Tensors
2024-07-25 16:57:57 -04:00
2024-07-29 08:38:01 -04:00
This file sets up the structure `TensorStructure` which contains
data of types (or `colors`) of indices, the dual of colors, the associated module,
contraction between color modules, and the unit of contraction.
This structure is extended to `DualizeTensorStructure` which adds a metric to the tensor structure,
allowing a vector to be taken to its dual vector by contraction with a specified metric.
The definition of `DualizeTensorStructure` can be found in
`HepLean.SpaceTime.LorentzTensor.RisingLowering`.
The structure `DualizeTensorStructure` is further extended in
`HepLean.SpaceTime.LorentzTensor.LorentzTensorStruct` to add a group action on the tensor space,
under which contraction and rising and lowering etc, are invariant.
2024-07-25 16:57:57 -04:00
## References
-- For modular operads see: [Raynor][raynor2021graphical]
-/
noncomputable section
open TensorProduct
variable {R : Type} [CommSemiring R]
2024-07-29 16:54:59 -04:00
def contrDualLeftAux {V1 V2 V3 : Type} [AddCommMonoid V1] [AddCommMonoid V2] [AddCommMonoid V3]
[Module R V1] [Module R V2] [Module R V3] (f : V1 ⊗[R] V2 →ₗ[R] R) :
V1 ⊗[R] V2 ⊗[R] V3 →ₗ[R] V3 :=
(TensorProduct.lid R _).toLinearMap ∘ₗ
TensorProduct.map (f) (LinearEquiv.refl R V3).toLinearMap
∘ₗ (TensorProduct.assoc R _ _ _).symm.toLinearMap
def contrDualMidAux {V1 V2 V3 V4 : Type} [AddCommMonoid V1] [AddCommMonoid V2] [AddCommMonoid V3]
[AddCommMonoid V4] [Module R V1] [Module R V2] [Module R V3] [Module R V4] (f : V1 ⊗[R] V2 →ₗ[R] R) :
(V4 ⊗[R] V1) ⊗[R] (V2 ⊗[R] V3) →ₗ[R] V4 ⊗[R] V3 :=
(TensorProduct.map (LinearEquiv.refl R V4).toLinearMap (contrDualLeftAux f)) ∘ₗ
(TensorProduct.assoc R _ _ _).toLinearMap
2024-07-30 07:51:07 -04:00
2024-07-26 14:43:20 -04:00
/-- An initial structure specifying a tensor system (e.g. a system in which you can
2024-07-29 08:38:01 -04:00
define real Lorentz tensors or Einstein notation convention). -/
structure TensorStructure (R : Type) [CommSemiring R] where
2024-07-26 14:43:20 -04:00
/-- The allowed colors of indices.
For example for a real Lorentz tensor these are `{up, down}`. -/
2024-07-25 16:57:57 -04:00
Color : Type
2024-07-26 14:54:09 -04:00
/-- To each color we associate a module. -/
2024-07-25 16:57:57 -04:00
ColorModule : Color → Type
2024-07-26 15:41:26 -04:00
/-- A map taking every color to its dual color. -/
2024-07-25 16:57:57 -04:00
τ : Color → Color
2024-07-26 15:41:26 -04:00
/-- The map `τ` is an involution. -/
2024-07-25 16:57:57 -04:00
τ_involutive : Function.Involutive τ
2024-07-26 15:41:26 -04:00
/-- Each `ColorModule` has the structure of an additive commutative monoid. -/
2024-07-25 16:57:57 -04:00
colorModule_addCommMonoid : ∀ μ, AddCommMonoid (ColorModule μ)
2024-07-26 15:41:26 -04:00
/-- Each `ColorModule` has the structure of a module over `R`. -/
2024-07-25 16:57:57 -04:00
colorModule_module : ∀ μ, Module R (ColorModule μ)
2024-07-26 15:41:26 -04:00
/-- The contraction of a vector with a vector with dual color. -/
2024-07-25 16:57:57 -04:00
contrDual : ∀ μ, ColorModule μ ⊗[R] ColorModule (τ μ) →ₗ[R] R
2024-07-29 08:38:01 -04:00
/-- The contraction is symmetric. -/
contrDual_symm : ∀ μ x y, (contrDual μ) (x ⊗ₜ[R] y) =
(contrDual (τ μ)) (y ⊗ₜ[R] (Equiv.cast (congrArg ColorModule (τ_involutive μ).symm) x))
/-- The unit of the contraction. -/
2024-07-29 16:54:59 -04:00
unit : (μ : Color) → ColorModule μ ⊗[R] ColorModule (τ μ)
2024-07-29 08:38:01 -04:00
/-- The unit is a right identity. -/
2024-07-29 16:54:59 -04:00
unit_lid : ∀ μ (x : ColorModule μ),
TensorProduct.rid R _
2024-07-30 07:51:07 -04:00
(TensorProduct.map (LinearEquiv.refl R (ColorModule μ)).toLinearMap
(contrDual μ ∘ₗ (TensorProduct.comm R _ _).toLinearMap)
((TensorProduct.assoc R _ _ _) (unit μ ⊗ₜ[R] x))) = x
2024-07-29 16:54:59 -04:00
/-- The metric for a given color. -/
metric : (μ : Color) → ColorModule μ ⊗[R] ColorModule μ
/-- The metric contracted with its dual is the unit. -/
metric_dual : ∀ (μ : Color), (contrDualMidAux (contrDual μ)
(metric μ ⊗ₜ[R] metric (τ μ))) = unit μ
2024-07-25 16:57:57 -04:00
2024-07-29 08:38:01 -04:00
namespace TensorStructure
2024-07-25 16:57:57 -04:00
2024-07-29 08:38:01 -04:00
variable (𝓣 : TensorStructure R)
2024-07-25 16:57:57 -04:00
2024-07-26 14:43:20 -04:00
variable {d : } {X Y Y' Z W : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
2024-07-26 14:54:09 -04:00
[Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z] [Fintype W] [DecidableEq W]
2024-07-26 15:41:26 -04:00
{cX cX2 : X → 𝓣.Color} {cY : Y → 𝓣.Color} {cZ : Z → 𝓣.Color}
{cW : W → 𝓣.Color} {cY' : Y' → 𝓣.Color} {μ ν: 𝓣.Color}
2024-07-25 16:57:57 -04:00
instance : AddCommMonoid (𝓣.ColorModule μ) := 𝓣.colorModule_addCommMonoid μ
instance : Module R (𝓣.ColorModule μ) := 𝓣.colorModule_module μ
2024-07-26 15:41:26 -04:00
/-- The type of tensors given a map from an indexing set `X` to the type of colors,
specifying the color of that index. -/
2024-07-26 14:54:09 -04:00
def Tensor (c : X → 𝓣.Color) : Type := ⨂[R] x, 𝓣.ColorModule (c x)
2024-07-25 16:57:57 -04:00
2024-07-26 15:41:26 -04:00
instance : AddCommMonoid (𝓣.Tensor cX) :=
PiTensorProduct.instAddCommMonoid fun i => 𝓣.ColorModule (cX i)
2024-07-25 16:57:57 -04:00
2024-07-26 15:41:26 -04:00
instance : Module R (𝓣.Tensor cX) := PiTensorProduct.instModule
2024-07-25 16:57:57 -04:00
2024-07-26 15:41:26 -04:00
/-- Equivalence of `ColorModule` given an equality of colors. -/
2024-07-25 16:57:57 -04:00
def colorModuleCast (h : μ = ν) : 𝓣.ColorModule μ ≃ₗ[R] 𝓣.ColorModule ν where
2024-07-29 08:38:01 -04:00
toFun := Equiv.cast (congrArg 𝓣.ColorModule h)
invFun := (Equiv.cast (congrArg 𝓣.ColorModule h)).symm
2024-07-25 16:57:57 -04:00
map_add' x y := by
subst h
rfl
map_smul' x y := by
subst h
rfl
2024-07-26 15:41:26 -04:00
left_inv x := Equiv.symm_apply_apply (Equiv.cast (congrArg 𝓣.ColorModule h)) x
right_inv x := Equiv.apply_symm_apply (Equiv.cast (congrArg 𝓣.ColorModule h)) x
2024-07-25 16:57:57 -04:00
2024-07-26 15:41:26 -04:00
lemma tensorProd_piTensorProd_ext {M : Type} [AddCommMonoid M] [Module R M]
{f g : 𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY →ₗ[R] M}
2024-07-26 14:43:20 -04:00
(h : ∀ p q, f (PiTensorProduct.tprod R p ⊗ₜ[R] PiTensorProduct.tprod R q)
= g (PiTensorProduct.tprod R p ⊗ₜ[R] PiTensorProduct.tprod R q)) : f = g := by
apply TensorProduct.ext'
refine fun x ↦
PiTensorProduct.induction_on' x ?_ (by
intro a b hx hy y
simp [map_add, add_tmul, hx, hy])
intro rx fx
refine fun y ↦
PiTensorProduct.induction_on' y ?_ (by
intro a b hx hy
simp at hx hy
simp [map_add, tmul_add, hx, hy])
intro ry fy
2024-07-26 15:41:26 -04:00
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, tmul_smul, LinearMapClass.map_smul]
2024-07-26 14:43:20 -04:00
apply congrArg
2024-07-26 15:41:26 -04:00
simp only [smul_tmul, tmul_smul, LinearMapClass.map_smul]
exact congrArg (HSMul.hSMul rx) (h fx fy)
2024-07-26 14:43:20 -04:00
2024-07-25 16:57:57 -04:00
/-!
## Mapping isomorphisms
-/
2024-07-26 15:41:26 -04:00
/-- An linear equivalence of tensor spaces given a color-preserving equivalence of indexing sets. -/
2024-07-25 16:57:57 -04:00
def mapIso {c : X → 𝓣.Color} {d : Y → 𝓣.Color} (e : X ≃ Y) (h : c = d ∘ e) :
𝓣.Tensor c ≃ₗ[R] 𝓣.Tensor d :=
(PiTensorProduct.reindex R _ e) ≪≫ₗ
(PiTensorProduct.congr (fun y => 𝓣.colorModuleCast (by rw [h]; simp)))
2024-07-26 15:41:26 -04:00
lemma mapIso_trans_cond {e : X ≃ Y} {e' : Y ≃ Z} (h : cX = cY ∘ e) (h' : cY = cZ ∘ e') :
cX = cZ ∘ (e.trans e') := by
2024-07-25 16:57:57 -04:00
funext a
subst h h'
simp
@[simp]
2024-07-26 15:41:26 -04:00
lemma mapIso_trans (e : X ≃ Y) (e' : Y ≃ Z) (h : cX = cY ∘ e) (h' : cY = cZ ∘ e') :
2024-07-25 16:57:57 -04:00
(𝓣.mapIso e h ≪≫ₗ 𝓣.mapIso e' h') = 𝓣.mapIso (e.trans e') (𝓣.mapIso_trans_cond h h') := by
refine LinearEquiv.toLinearMap_inj.mp ?_
apply PiTensorProduct.ext
apply MultilinearMap.ext
intro x
simp only [mapIso, LinearMap.compMultilinearMap_apply, LinearEquiv.coe_coe,
LinearEquiv.trans_apply, PiTensorProduct.reindex_tprod, Equiv.symm_trans_apply]
2024-07-26 15:55:10 -04:00
change (PiTensorProduct.congr fun y => 𝓣.colorModuleCast _)
2024-07-26 15:41:26 -04:00
((PiTensorProduct.reindex R (fun x => 𝓣.ColorModule (cY x)) e')
2024-07-26 15:55:10 -04:00
((PiTensorProduct.congr fun y => 𝓣.colorModuleCast _) _)) =
2024-07-25 16:57:57 -04:00
(PiTensorProduct.congr fun y => 𝓣.colorModuleCast _)
2024-07-26 15:41:26 -04:00
((PiTensorProduct.reindex R (fun x => 𝓣.ColorModule (cX x)) (e.trans e')) _)
2024-07-25 16:57:57 -04:00
rw [PiTensorProduct.congr_tprod, PiTensorProduct.reindex_tprod,
PiTensorProduct.congr_tprod, PiTensorProduct.reindex_tprod, PiTensorProduct.congr]
simp [colorModuleCast]
@[simp]
2024-07-26 15:41:26 -04:00
lemma mapIso_mapIso (e : X ≃ Y) (e' : Y ≃ Z) (h : cX = cY ∘ e) (h' : cY = cZ ∘ e')
(T : 𝓣.Tensor cX) :
2024-07-25 16:57:57 -04:00
(𝓣.mapIso e' h') (𝓣.mapIso e h T) = 𝓣.mapIso (e.trans e') (𝓣.mapIso_trans_cond h h') T := by
rw [← LinearEquiv.trans_apply, mapIso_trans]
@[simp]
2024-07-26 15:41:26 -04:00
lemma mapIso_symm (e : X ≃ Y) (h : cX = cY ∘ e) :
(𝓣.mapIso e h).symm = 𝓣.mapIso e.symm ((Equiv.eq_comp_symm e cY cX).mpr h.symm) := by
2024-07-25 16:57:57 -04:00
refine LinearEquiv.toLinearMap_inj.mp ?_
apply PiTensorProduct.ext
apply MultilinearMap.ext
intro x
2024-07-26 14:54:09 -04:00
simp [mapIso, LinearMap.compMultilinearMap_apply, LinearEquiv.coe_coe,
2024-07-25 16:57:57 -04:00
LinearEquiv.symm_apply_apply, PiTensorProduct.reindex_tprod]
2024-07-26 15:41:26 -04:00
change (PiTensorProduct.reindex R (fun x => 𝓣.ColorModule (cX x)) e).symm
2024-07-25 16:57:57 -04:00
((PiTensorProduct.congr fun y => 𝓣.colorModuleCast _).symm ((PiTensorProduct.tprod R) x)) =
2024-07-26 15:41:26 -04:00
(PiTensorProduct.congr fun y => 𝓣.colorModuleCast _)
((PiTensorProduct.reindex R (fun x => 𝓣.ColorModule (cY x)) e.symm)
((PiTensorProduct.tprod R) x))
2024-07-25 16:57:57 -04:00
rw [PiTensorProduct.reindex_tprod, PiTensorProduct.congr_tprod, PiTensorProduct.congr_symm_tprod,
LinearEquiv.symm_apply_eq, PiTensorProduct.reindex_tprod]
apply congrArg
funext i
simp only [colorModuleCast, Equiv.cast_symm, LinearEquiv.coe_symm_mk,
Equiv.symm_symm_apply, LinearEquiv.coe_mk]
rw [← Equiv.symm_apply_eq]
simp only [Equiv.cast_symm, Equiv.cast_apply, cast_cast]
apply cast_eq_iff_heq.mpr
rw [Equiv.apply_symm_apply]
@[simp]
2024-07-26 15:41:26 -04:00
lemma mapIso_refl : 𝓣.mapIso (Equiv.refl X) (rfl : cX = cX) = LinearEquiv.refl R _ := by
2024-07-25 16:57:57 -04:00
refine LinearEquiv.toLinearMap_inj.mp ?_
apply PiTensorProduct.ext
apply MultilinearMap.ext
intro x
simp only [mapIso, Equiv.refl_symm, Equiv.refl_apply, PiTensorProduct.reindex_refl,
LinearMap.compMultilinearMap_apply, LinearEquiv.coe_coe, LinearEquiv.trans_apply,
LinearEquiv.refl_apply, LinearEquiv.refl_toLinearMap, LinearMap.id, LinearMap.coe_mk,
AddHom.coe_mk, id_eq]
change (PiTensorProduct.congr fun y => 𝓣.colorModuleCast _) ((PiTensorProduct.tprod R) x) = _
rw [PiTensorProduct.congr_tprod]
rfl
@[simp]
2024-07-26 15:41:26 -04:00
lemma mapIso_tprod {c : X → 𝓣.Color} {d : Y → 𝓣.Color} (e : X ≃ Y) (h : c = d ∘ e)
(f : (i : X) → 𝓣.ColorModule (c i)) : (𝓣.mapIso e h) (PiTensorProduct.tprod R f) =
2024-07-25 16:57:57 -04:00
(PiTensorProduct.tprod R (fun i => 𝓣.colorModuleCast (by rw [h]; simp) (f (e.symm i)))) := by
simp [mapIso]
change (PiTensorProduct.congr fun y => 𝓣.colorModuleCast _)
2024-07-26 15:41:26 -04:00
((PiTensorProduct.reindex R _ e) ((PiTensorProduct.tprod R) f)) = _
2024-07-25 16:57:57 -04:00
rw [PiTensorProduct.reindex_tprod]
2024-07-26 15:41:26 -04:00
exact PiTensorProduct.congr_tprod (fun y => 𝓣.colorModuleCast _) fun i => f (e.symm i)
2024-07-25 16:57:57 -04:00
/-!
2024-07-26 14:43:20 -04:00
## Pure tensors
This section is needed since: `PiTensorProduct.tmulEquiv` is not defined for dependent types.
Hence we need to construct a version of it here.
-/
2024-07-26 15:41:26 -04:00
/-- The type of pure tensors, i.e. of the form `v1 ⊗ v2 ⊗ v3 ⊗ ...`. -/
2024-07-26 14:43:20 -04:00
abbrev PureTensor (c : X → 𝓣.Color) := (x : X) → 𝓣.ColorModule (c x)
2024-07-26 15:41:26 -04:00
/-- A pure tensor in `𝓣.PureTensor (Sum.elim cX cY)` constructed from a pure tensor
in `𝓣.PureTensor cX` and a pure tensor in `𝓣.PureTensor cY`. -/
def elimPureTensor (p : 𝓣.PureTensor cX) (q : 𝓣.PureTensor cY) : 𝓣.PureTensor (Sum.elim cX cY) :=
2024-07-26 14:43:20 -04:00
fun x =>
match x with
| Sum.inl x => p x
| Sum.inr x => q x
@[simp]
2024-07-26 15:41:26 -04:00
lemma elimPureTensor_update_right (p : 𝓣.PureTensor cX) (q : 𝓣.PureTensor cY)
(y : Y) (r : 𝓣.ColorModule (cY y)) : 𝓣.elimPureTensor p (Function.update q y r) =
2024-07-26 14:43:20 -04:00
Function.update (𝓣.elimPureTensor p q) (Sum.inr y) r := by
funext x
match x with
| Sum.inl x => rfl
| Sum.inr x =>
change Function.update q y r x = _
simp only [Function.update, Sum.inr.injEq, Sum.elim_inr]
split_ifs
rename_i h
subst h
simp_all only
rfl
@[simp]
2024-07-26 15:41:26 -04:00
lemma elimPureTensor_update_left (p : 𝓣.PureTensor cX) (q : 𝓣.PureTensor cY)
(x : X) (r : 𝓣.ColorModule (cX x)) : 𝓣.elimPureTensor (Function.update p x r) q =
2024-07-26 14:43:20 -04:00
Function.update (𝓣.elimPureTensor p q) (Sum.inl x) r := by
funext y
match y with
| Sum.inl y =>
change (Function.update p x r) y = _
simp only [Function.update, Sum.inl.injEq, Sum.elim_inl]
split_ifs
rename_i h
subst h
simp_all only
rfl
| Sum.inr y => rfl
2024-07-26 15:41:26 -04:00
/-- The projection of a pure tensor in `𝓣.PureTensor (Sum.elim cX cY)` onto a pure tensor in
`𝓣.PureTensor cX`. -/
def inlPureTensor (p : 𝓣.PureTensor (Sum.elim cX cY)) : 𝓣.PureTensor cX := fun x => p (Sum.inl x)
2024-07-26 14:43:20 -04:00
2024-07-26 15:41:26 -04:00
/-- The projection of a pure tensor in `𝓣.PureTensor (Sum.elim cX cY)` onto a pure tensor in
`𝓣.PureTensor cY`. -/
def inrPureTensor (p : 𝓣.PureTensor (Sum.elim cX cY)) : 𝓣.PureTensor cY := fun y => p (Sum.inr y)
2024-07-26 14:43:20 -04:00
@[simp]
2024-07-26 15:41:26 -04:00
lemma inlPureTensor_update_left [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (Sum.elim cX cY)) (x : X)
(v1 : 𝓣.ColorModule (Sum.elim cX cY (Sum.inl x))) :
𝓣.inlPureTensor (Function.update f (Sum.inl x) v1) =
Function.update (𝓣.inlPureTensor f) x v1 := by
2024-07-26 14:43:20 -04:00
funext y
simp [inlPureTensor, Function.update, Sum.inl.injEq, Sum.elim_inl]
split
next h =>
subst h
simp_all only
2024-07-26 15:41:26 -04:00
rfl
2024-07-26 14:43:20 -04:00
@[simp]
2024-07-26 15:41:26 -04:00
lemma inrPureTensor_update_left [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (Sum.elim cX cY)) (x : X)
(v1 : 𝓣.ColorModule (Sum.elim cX cY (Sum.inl x))) :
2024-07-26 14:43:20 -04:00
𝓣.inrPureTensor (Function.update f (Sum.inl x) v1) = (𝓣.inrPureTensor f) := by
funext x
simp [inrPureTensor, Function.update]
@[simp]
2024-07-26 15:41:26 -04:00
lemma inrPureTensor_update_right [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (Sum.elim cX cY)) (y : Y)
(v1 : 𝓣.ColorModule (Sum.elim cX cY (Sum.inr y))) :
𝓣.inrPureTensor (Function.update f (Sum.inr y) v1) =
Function.update (𝓣.inrPureTensor f) y v1 := by
2024-07-26 14:43:20 -04:00
funext y
simp [inrPureTensor, Function.update, Sum.inl.injEq, Sum.elim_inl]
split
next h =>
subst h
simp_all only
2024-07-26 15:41:26 -04:00
rfl
2024-07-26 14:43:20 -04:00
@[simp]
2024-07-26 15:41:26 -04:00
lemma inlPureTensor_update_right [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (Sum.elim cX cY)) (y : Y)
(v1 : 𝓣.ColorModule (Sum.elim cX cY (Sum.inr y))) :
2024-07-26 14:43:20 -04:00
𝓣.inlPureTensor (Function.update f (Sum.inr y) v1) = (𝓣.inlPureTensor f) := by
funext x
simp [inlPureTensor, Function.update]
2024-07-26 15:41:26 -04:00
/-- The multilinear map taking pure tensors a `𝓣.PureTensor cX` and a pure tensor in
`𝓣.PureTensor cY`, and constructing a tensor in `𝓣.Tensor (Sum.elim cX cY))`. -/
def elimPureTensorMulLin : MultilinearMap R (fun i => 𝓣.ColorModule (cX i))
(MultilinearMap R (fun x => 𝓣.ColorModule (cY x)) (𝓣.Tensor (Sum.elim cX cY))) where
2024-07-26 14:43:20 -04:00
toFun p := {
toFun := fun q => PiTensorProduct.tprod R (𝓣.elimPureTensor p q)
map_add' := fun m x v1 v2 => by
simp [Sum.elim_inl, Sum.elim_inr]
map_smul' := fun m x r v => by
simp [Sum.elim_inl, Sum.elim_inr]}
map_add' p x v1 v2 := by
apply MultilinearMap.ext
intro y
simp
map_smul' p x r v := by
apply MultilinearMap.ext
intro y
simp
/-!
## tensorator
-/
/-! TODO: Replace with dependent type version of `MultilinearMap.domCoprod` when in Mathlib. -/
2024-07-26 15:41:26 -04:00
/-- The multi-linear map taking a pure tensor in `𝓣.PureTensor (Sum.elim cX cY)` and constructing
a vector in `𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY`. -/
def domCoprod : MultilinearMap R (fun x => 𝓣.ColorModule (Sum.elim cX cY x))
(𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY) where
2024-07-26 14:43:20 -04:00
toFun f := (PiTensorProduct.tprod R (𝓣.inlPureTensor f)) ⊗ₜ
(PiTensorProduct.tprod R (𝓣.inrPureTensor f))
map_add' f xy v1 v2:= by
match xy with
| Sum.inl x => simp [← TensorProduct.add_tmul]
| Sum.inr y => simp [← TensorProduct.tmul_add]
map_smul' f xy r p := by
match xy with
| Sum.inl x => simp [TensorProduct.tmul_smul, TensorProduct.smul_tmul]
| Sum.inr y => simp [TensorProduct.tmul_smul, TensorProduct.smul_tmul]
2024-07-26 15:41:26 -04:00
/-- The linear map combining two tensors into a single tensor
via the tensor product i.e. `v1 v2 ↦ v1 ⊗ v2`. -/
def tensoratorSymm : 𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY →ₗ[R] 𝓣.Tensor (Sum.elim cX cY) := by
2024-07-26 14:43:20 -04:00
refine TensorProduct.lift {
toFun := fun a ↦
PiTensorProduct.lift <|
PiTensorProduct.lift (𝓣.elimPureTensorMulLin) a,
map_add' := fun a b ↦ by simp
map_smul' := fun r a ↦ by simp}
/-! TODO: Replace with dependent type version of `PiTensorProduct.tmulEquiv` when in Mathlib. -/
2024-07-26 15:41:26 -04:00
/-- Splitting a tensor in `𝓣.Tensor (Sum.elim cX cY)` into the tensor product of two tensors. -/
def tensorator : 𝓣.Tensor (Sum.elim cX cY) →ₗ[R] 𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY :=
2024-07-26 14:43:20 -04:00
PiTensorProduct.lift 𝓣.domCoprod
2024-07-26 15:41:26 -04:00
/-- An equivalence formed by taking the tensor product of tensors. -/
def tensoratorEquiv (c : X → 𝓣.Color) (d : Y → 𝓣.Color) :
𝓣.Tensor c ⊗[R] 𝓣.Tensor d ≃ₗ[R] 𝓣.Tensor (Sum.elim c d) :=
2024-07-26 14:43:20 -04:00
LinearEquiv.ofLinear (𝓣.tensoratorSymm) (𝓣.tensorator)
(by
apply PiTensorProduct.ext
apply MultilinearMap.ext
intro p
simp [tensorator, tensoratorSymm, domCoprod]
2024-07-26 15:55:10 -04:00
change (PiTensorProduct.lift _) ((PiTensorProduct.tprod R) _) =
2024-07-26 14:43:20 -04:00
LinearMap.id ((PiTensorProduct.tprod R) p)
rw [PiTensorProduct.lift.tprod]
simp [elimPureTensorMulLin, elimPureTensor]
2024-07-26 14:54:09 -04:00
change (PiTensorProduct.tprod R) _ = _
2024-07-26 14:43:20 -04:00
apply congrArg
funext x
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl)
(by
apply tensorProd_piTensorProd_ext
intro p q
2024-07-26 14:54:09 -04:00
simp [tensorator, tensoratorSymm]
2024-07-26 15:41:26 -04:00
change (PiTensorProduct.lift 𝓣.domCoprod)
((PiTensorProduct.lift (𝓣.elimPureTensorMulLin p)) ((PiTensorProduct.tprod R) q)) =_
2024-07-26 14:43:20 -04:00
rw [PiTensorProduct.lift.tprod]
simp [elimPureTensorMulLin]
rfl)
@[simp]
2024-07-26 15:41:26 -04:00
lemma tensoratorEquiv_tmul_tprod (p : 𝓣.PureTensor cX) (q : 𝓣.PureTensor cY) :
(𝓣.tensoratorEquiv cX cY) ((PiTensorProduct.tprod R) p ⊗ₜ[R] (PiTensorProduct.tprod R) q) =
2024-07-26 14:43:20 -04:00
(PiTensorProduct.tprod R) (𝓣.elimPureTensor p q) := by
2024-07-26 15:41:26 -04:00
simp only [tensoratorEquiv, tensoratorSymm, tensorator, domCoprod, LinearEquiv.ofLinear_apply,
lift.tmul, LinearMap.coe_mk, AddHom.coe_mk, PiTensorProduct.lift.tprod]
exact PiTensorProduct.lift.tprod q
2024-07-26 14:43:20 -04:00
lemma tensoratorEquiv_mapIso_cond {e : X ≃ Y} {e' : Z ≃ Y} {e'' : W ≃ X}
2024-07-26 15:41:26 -04:00
(h : cX = 𝓣.τ ∘ cY ∘ e) (h' : cZ = cY ∘ e') (h'' : bW = cX ∘ e'') :
Sum.elim bW cZ = Sum.elim cX cY ∘ ⇑(e''.sumCongr e') := by
2024-07-26 14:43:20 -04:00
subst h h' h''
funext x
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl
@[simp]
lemma tensoratorEquiv_mapIso (e : X ≃ Y) (e' : Z ≃ Y) (e'' : W ≃ X)
2024-07-26 15:41:26 -04:00
(h : cX = 𝓣.τ ∘ cY ∘ e) (h' : cZ = cY ∘ e') (h'' : bW = cX ∘ e'') :
(TensorProduct.congr (𝓣.mapIso e'' h'') (𝓣.mapIso e' h')) ≪≫ₗ (𝓣.tensoratorEquiv cX cY)
= (𝓣.tensoratorEquiv bW cZ)
2024-07-26 14:54:09 -04:00
≪≫ₗ (𝓣.mapIso (Equiv.sumCongr e'' e') (𝓣.tensoratorEquiv_mapIso_cond h h' h'')) := by
2024-07-26 14:43:20 -04:00
apply LinearEquiv.toLinearMap_inj.mp
apply tensorProd_piTensorProd_ext
intro p q
2024-07-26 15:41:26 -04:00
simp only [LinearEquiv.coe_coe, LinearEquiv.trans_apply, congr_tmul, mapIso_tprod,
tensoratorEquiv_tmul_tprod, Equiv.sumCongr_symm, Equiv.sumCongr_apply]
2024-07-26 14:43:20 -04:00
apply congrArg
funext x
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl
@[simp]
lemma tensoratorEquiv_mapIso_apply (e : X ≃ Y) (e' : Z ≃ Y) (e'' : W ≃ X)
2024-07-26 15:41:26 -04:00
(h : cX = 𝓣.τ ∘ cY ∘ e) (h' : cZ = cY ∘ e') (h'' : cW = cX ∘ e'')
(x : 𝓣.Tensor cW ⊗[R] 𝓣.Tensor cZ) :
(𝓣.tensoratorEquiv cX cY) ((TensorProduct.congr (𝓣.mapIso e'' h'') (𝓣.mapIso e' h')) x) =
(𝓣.mapIso (Equiv.sumCongr e'' e') (𝓣.tensoratorEquiv_mapIso_cond h h' h''))
((𝓣.tensoratorEquiv cW cZ) x) := by
trans ((TensorProduct.congr (𝓣.mapIso e'' h'') (𝓣.mapIso e' h')) ≪≫ₗ
(𝓣.tensoratorEquiv cX cY)) x
2024-07-26 14:43:20 -04:00
rfl
rw [tensoratorEquiv_mapIso]
rfl
exact e
exact h
lemma tensoratorEquiv_mapIso_tmul (e : X ≃ Y) (e' : Z ≃ Y) (e'' : W ≃ X)
2024-07-26 15:41:26 -04:00
(h : cX = 𝓣.τ ∘ cY ∘ e) (h' : cZ = cY ∘ e') (h'' : cW = cX ∘ e'')
(x : 𝓣.Tensor cW) (y : 𝓣.Tensor cZ) :
(𝓣.tensoratorEquiv cX cY) ((𝓣.mapIso e'' h'' x) ⊗ₜ[R] (𝓣.mapIso e' h' y)) =
(𝓣.mapIso (Equiv.sumCongr e'' e') (𝓣.tensoratorEquiv_mapIso_cond h h' h''))
((𝓣.tensoratorEquiv cW cZ) (x ⊗ₜ y)) := by
2024-07-26 14:43:20 -04:00
rw [← tensoratorEquiv_mapIso_apply]
rfl
exact e
exact h
/-!
## Splitting tensors into tensor products
-/
2024-07-26 15:41:26 -04:00
/-- The decomposition of a set into a direct sum based on the image of an injection. -/
2024-07-26 14:43:20 -04:00
def decompEmbedSet (f : Y ↪ X) :
X ≃ {x // x ∈ (Finset.image f Finset.univ)ᶜ} ⊕ Y :=
(Equiv.Set.sumCompl (Set.range ⇑f)).symm.trans <|
(Equiv.sumComm _ _).trans <|
Equiv.sumCongr ((Equiv.subtypeEquivRight (by simp))) <|
(Function.Embedding.toEquivRange f).symm
2024-07-26 15:41:26 -04:00
/-- The restriction of a map from an indexing set to the space to the complement of the image
of an embedding. -/
def decompEmbedColorLeft (c : X → 𝓣.Color) (f : Y ↪ X) :
{x // x ∈ (Finset.image f Finset.univ)ᶜ} → 𝓣.Color :=
2024-07-26 14:43:20 -04:00
(c ∘ (decompEmbedSet f).symm) ∘ Sum.inl
2024-07-26 15:41:26 -04:00
/-- The restriction of a map from an indexing set to the space to the image
of an embedding. -/
2024-07-26 14:54:09 -04:00
def decompEmbedColorRight (c : X → 𝓣.Color) (f : Y ↪ X) : Y → 𝓣.Color :=
2024-07-26 14:43:20 -04:00
(c ∘ (decompEmbedSet f).symm) ∘ Sum.inr
2024-07-26 14:54:09 -04:00
lemma decompEmbed_cond (c : X → 𝓣.Color) (f : Y ↪ X) : c =
(Sum.elim (𝓣.decompEmbedColorLeft c f) (𝓣.decompEmbedColorRight c f)) ∘ decompEmbedSet f := by
2024-07-26 14:43:20 -04:00
simpa [decompEmbedColorLeft, decompEmbedColorRight] using (Equiv.comp_symm_eq _ _ _).mp rfl
2024-07-26 15:41:26 -04:00
/-- Decomposes a tensor into a tensor product of two tensors
one which has indices in the image of the given embedding, and the other has indices in
the complement of that image. -/
2024-07-26 14:43:20 -04:00
def decompEmbed (f : Y ↪ X) :
2024-07-26 15:41:26 -04:00
𝓣.Tensor cX ≃ₗ[R] 𝓣.Tensor (𝓣.decompEmbedColorLeft cX f) ⊗[R] 𝓣.Tensor (cX ∘ f) :=
(𝓣.mapIso (decompEmbedSet f) (𝓣.decompEmbed_cond cX f)) ≪≫ₗ
(𝓣.tensoratorEquiv (𝓣.decompEmbedColorLeft cX f) (𝓣.decompEmbedColorRight cX f)).symm
2024-07-26 14:43:20 -04:00
2024-07-25 16:57:57 -04:00
end TensorStructure
end