510 lines
19 KiB
Text
510 lines
19 KiB
Text
/-
|
||
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 HepLean.SpaceTime.LorentzTensor.IndexNotation.Indices.Color
|
||
import HepLean.SpaceTime.LorentzTensor.Basic
|
||
import HepLean.SpaceTime.LorentzTensor.RisingLowering
|
||
import HepLean.SpaceTime.LorentzTensor.Contraction
|
||
/-!
|
||
|
||
# The structure of a tensor with a string of indices
|
||
|
||
-/
|
||
|
||
/-! TODO: Introduce a way to change an index from e.g. `įµĀ¹` to `įµĀ²`.
|
||
Would be nice to have a tactic that did this automatically. -/
|
||
|
||
namespace TensorStructure
|
||
noncomputable section
|
||
|
||
open TensorColor
|
||
open IndexNotation
|
||
|
||
variable {R : Type} [CommSemiring R] (š£ : TensorStructure R)
|
||
|
||
variable {d : ā} {X Y Y' Z W : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||
[Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z] [Fintype W] [DecidableEq W]
|
||
{cX cX2 : X ā š£.Color} {cY : Y ā š£.Color} {cZ : Z ā š£.Color}
|
||
{cW : W ā š£.Color} {cY' : Y' ā š£.Color} {μ ν Ī· : š£.Color}
|
||
|
||
variable [IndexNotation š£.Color] [Fintype š£.Color] [DecidableEq š£.Color]
|
||
|
||
/-- The structure an tensor with a index specification e.g. `įµĀ¹įµ¤ā`. -/
|
||
structure TensorIndex extends ColorIndexList š£.toTensorColor where
|
||
/-- The underlying tensor. -/
|
||
tensor : š£.Tensor toIndexList.colorMap
|
||
|
||
namespace TensorIndex
|
||
|
||
open TensorColor ColorIndexList
|
||
|
||
variable {š£ : TensorStructure R} [IndexNotation š£.Color] [Fintype š£.Color] [DecidableEq š£.Color]
|
||
variable {n m : ā} {cn : Fin n ā š£.Color} {cm : Fin m ā š£.Color}
|
||
|
||
lemma colormap_mapIso {Tā Tā : š£.TensorIndex} (hi : Tā.toColorIndexList = Tā.toColorIndexList) :
|
||
ColorMap.MapIso (Fin.castOrderIso (by simp [IndexList.length, hi])).toEquiv
|
||
Tā.colorMap Tā.colorMap := by
|
||
cases Tā; cases Tā
|
||
simp [ColorMap.MapIso]
|
||
simp at hi
|
||
rename_i a b c d
|
||
cases a
|
||
cases c
|
||
rename_i a1 a2 a3 a4 a5 a6
|
||
cases a1
|
||
cases a4
|
||
simp_all
|
||
simp at hi
|
||
subst hi
|
||
rfl
|
||
|
||
lemma ext {Tā Tā : š£.TensorIndex} (hi : Tā.toColorIndexList = Tā.toColorIndexList)
|
||
(h : Tā.tensor = š£.mapIso (Fin.castOrderIso (by simp [IndexList.length, hi])).toEquiv
|
||
(colormap_mapIso hi.symm) Tā.tensor) : Tā = Tā := by
|
||
cases Tā; cases Tā
|
||
simp at h
|
||
simp_all
|
||
simp at hi
|
||
subst hi
|
||
simp_all
|
||
|
||
|
||
lemma index_eq_of_eq {Tā Tā : š£.TensorIndex} (h : Tā = Tā) :
|
||
Tā.toColorIndexList = Tā.toColorIndexList := by
|
||
cases h
|
||
rfl
|
||
|
||
@[simp]
|
||
lemma tensor_eq_of_eq {Tā Tā : š£.TensorIndex} (h : Tā = Tā) : Tā.tensor =
|
||
š£.mapIso (Fin.castOrderIso (by rw [index_eq_of_eq h])).toEquiv
|
||
(colormap_mapIso (index_eq_of_eq h).symm) Tā.tensor := by
|
||
have hi := index_eq_of_eq h
|
||
cases Tā
|
||
cases Tā
|
||
simp at hi
|
||
subst hi
|
||
simpa using h
|
||
|
||
/-- The construction of a `TensorIndex` from a tensor, a IndexListColor, and a condition
|
||
on the dual maps. -/
|
||
def mkDualMap (T : š£.Tensor cn) (l : ColorIndexList š£.toTensorColor) (hn : n = l.1.length)
|
||
(hd : ColorMap.DualMap l.1.colorMap (cn ā Fin.cast hn.symm)) :
|
||
š£.TensorIndex where
|
||
toColorIndexList := l
|
||
tensor :=
|
||
š£.mapIso (Equiv.refl _) (ColorMap.DualMap.split_dual' (by simp [hd])) <|
|
||
š£.dualize (ColorMap.DualMap.split l.1.colorMap (cn ā Fin.cast hn.symm)) <|
|
||
(š£.mapIso (Fin.castOrderIso hn).toEquiv rfl T : š£.Tensor (cn ā Fin.cast hn.symm))
|
||
|
||
/-!
|
||
|
||
## The contraction of indices
|
||
|
||
-/
|
||
|
||
/-- The contraction of indices in a `TensorIndex`. -/
|
||
def contr (T : š£.TensorIndex) : š£.TensorIndex where
|
||
index := T.index.contr
|
||
tensor :=
|
||
š£.mapIso (Fin.castOrderIso T.index.contr_numIndices.symm).toEquiv
|
||
T.index.contr_colorMap <|
|
||
š£.contr (T.index.splitContr).symm T.index.splitContr_map T.tensor
|
||
|
||
/-- Applying contr to a tensor whose indices has no contracts does not do anything. -/
|
||
@[simp]
|
||
lemma contr_of_hasNoContr (T : š£.TensorIndex) (h : T.index.1.HasNoContr) :
|
||
T.contr = T := by
|
||
refine ext _ _ ?_ ?_
|
||
exact Subtype.eq (T.index.1.contr_of_hasNoContr h)
|
||
simp only [contr]
|
||
have h1 : IsEmpty T.index.1.contrPairSet := T.index.1.contrPairSet_isEmpty_of_hasNoContr h
|
||
cases T
|
||
rename_i i T
|
||
simp only
|
||
refine PiTensorProduct.induction_on' T ?_ (by
|
||
intro a b hx hy
|
||
simp [map_add, add_mul, hx, hy])
|
||
intro r f
|
||
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, LinearMapClass.map_smul, mapIso_tprod, id_eq,
|
||
eq_mpr_eq_cast, OrderIso.toEquiv_symm, RelIso.coe_fn_toEquiv]
|
||
apply congrArg
|
||
erw [TensorStructure.contr_tprod_isEmpty]
|
||
erw [mapIso_tprod]
|
||
apply congrArg
|
||
funext l
|
||
rw [ā LinearEquiv.symm_apply_eq]
|
||
simp only [colorModuleCast, Equiv.cast_symm, OrderIso.toEquiv_symm, RelIso.coe_fn_toEquiv,
|
||
Function.comp_apply, LinearEquiv.coe_mk, Equiv.cast_apply, LinearEquiv.coe_symm_mk, cast_cast]
|
||
apply cast_eq_iff_heq.mpr
|
||
rw [splitContr_symm_apply_of_hasNoContr _ h]
|
||
rfl
|
||
|
||
@[simp]
|
||
lemma contr_contr (T : š£.TensorIndex) : T.contr.contr = T.contr :=
|
||
T.contr.contr_of_hasNoContr T.index.1.contr_hasNoContr
|
||
|
||
@[simp]
|
||
lemma contr_index (T : š£.TensorIndex) : T.contr.index = T.index.contr := rfl
|
||
|
||
|
||
/-!
|
||
|
||
## Scalar multiplication of
|
||
|
||
-/
|
||
|
||
/-- The scalar multiplication of a `TensorIndex` by an element of `R`. -/
|
||
instance : SMul R š£.TensorIndex where
|
||
smul := fun r T => {
|
||
index := T.index
|
||
tensor := r ⢠T.tensor}
|
||
|
||
@[simp]
|
||
lemma smul_index (r : R) (T : š£.TensorIndex) : (r ⢠T).index = T.index := rfl
|
||
|
||
@[simp]
|
||
lemma smul_tensor (r : R) (T : š£.TensorIndex) : (r ⢠T).tensor = r ⢠T.tensor := rfl
|
||
|
||
@[simp]
|
||
lemma smul_contr (r : R) (T : š£.TensorIndex) : (r ⢠T).contr = r ⢠T.contr := by
|
||
refine ext _ _ rfl ?_
|
||
simp only [contr, smul_index, smul_tensor, LinearMapClass.map_smul, Fin.castOrderIso_refl,
|
||
OrderIso.refl_toEquiv, mapIso_refl, LinearEquiv.refl_apply]
|
||
|
||
/-!
|
||
|
||
## Equivalence relation on `TensorIndex`
|
||
|
||
-/
|
||
|
||
/-- An (equivalence) relation on two `TensorIndex`.
|
||
The point in this equivalence relation is that certain things (like the
|
||
permutation of indices, the contraction of indices, or rising or lowering indices) can be placed
|
||
in the indices or moved to the tensor itself. These two descriptions are equivalent. -/
|
||
def Rel (Tā Tā : š£.TensorIndex) : Prop :=
|
||
Tā.index.PermContr Tā.index ā§ ā (h : Tā.index.PermContr Tā.index),
|
||
Tā.contr.tensor = š£.mapIso h.toEquiv.symm h.toEquiv_colorMap Tā.contr.tensor
|
||
|
||
namespace Rel
|
||
|
||
/-- Rel is reflexive. -/
|
||
lemma refl (T : š£.TensorIndex) : Rel T T := by
|
||
apply And.intro
|
||
exact IndexListColor.PermContr.refl T.index
|
||
intro h
|
||
simp [PermContr.toEquiv_refl']
|
||
|
||
/-- Rel is symmetric. -/
|
||
lemma symm {Tā Tā : š£.TensorIndex} (h : Rel Tā Tā) : Rel Tā Tā := by
|
||
apply And.intro h.1.symm
|
||
intro h'
|
||
rw [ā mapIso_symm]
|
||
symm
|
||
erw [LinearEquiv.symm_apply_eq]
|
||
rw [h.2]
|
||
apply congrFun
|
||
congr
|
||
exact h'.symm
|
||
|
||
/-- Rel is transitive. -/
|
||
lemma trans {Tā Tā Tā : š£.TensorIndex} (h1 : Rel Tā Tā) (h2 : Rel Tā Tā) : Rel Tā Tā := by
|
||
apply And.intro (h1.1.trans h2.1)
|
||
intro h
|
||
change _ = (š£.mapIso (h1.1.trans h2.1).toEquiv.symm _) Tā.contr.tensor
|
||
trans (š£.mapIso ((h1.1).toEquiv.trans (h2.1).toEquiv).symm (by
|
||
rw [PermContr.toEquiv_trans]
|
||
exact proof_2 Tā Tā h)) Tā.contr.tensor
|
||
swap
|
||
congr
|
||
rw [PermContr.toEquiv_trans]
|
||
erw [ā mapIso_trans]
|
||
simp only [LinearEquiv.trans_apply]
|
||
apply (h1.2 h1.1).trans
|
||
apply congrArg
|
||
exact h2.2 h2.1
|
||
|
||
/-- Rel forms an equivalence relation. -/
|
||
lemma isEquivalence : Equivalence (@Rel _ _ š£ _) where
|
||
refl := Rel.refl
|
||
symm := Rel.symm
|
||
trans := Rel.trans
|
||
|
||
/-- The equality of tensors corresponding to related tensor indices. -/
|
||
lemma to_eq {Tā Tā : š£.TensorIndex} (h : Rel Tā Tā) :
|
||
Tā.contr.tensor = š£.mapIso h.1.toEquiv.symm h.1.toEquiv_colorMap Tā.contr.tensor := h.2 h.1
|
||
|
||
end Rel
|
||
|
||
/-- The structure of a Setoid on `š£.TensorIndex` induced by `Rel`. -/
|
||
instance asSetoid : Setoid š£.TensorIndex := āØRel, Rel.isEquivalenceā©
|
||
|
||
/-- A tensor index is equivalent to its contraction. -/
|
||
lemma rel_contr (T : š£.TensorIndex) : T ā T.contr := by
|
||
apply And.intro
|
||
simp only [PermContr, contr_index, IndexListColor.contr_contr, List.Perm.refl, true_and]
|
||
rw [IndexListColor.contr_contr]
|
||
exact T.index.contr.1.hasNoContr_color_eq_of_id_eq T.index.1.contr_hasNoContr
|
||
intro h
|
||
rw [tensor_eq_of_eq T.contr_contr]
|
||
simp only [contr_index, mapIso_mapIso]
|
||
trans š£.mapIso (Equiv.refl _) (by rfl) T.contr.tensor
|
||
simp only [contr_index, mapIso_refl, LinearEquiv.refl_apply]
|
||
congr
|
||
apply Equiv.ext
|
||
intro x
|
||
rw [PermContr.toEquiv_contr_eq T.index.contr_contr.symm]
|
||
rfl
|
||
|
||
lemma smul_equiv {Tā Tā : š£.TensorIndex} (h : Tā ā Tā) (r : R) : r ⢠Tā ā r ⢠Tā := by
|
||
apply And.intro h.1
|
||
intro h1
|
||
rw [tensor_eq_of_eq (smul_contr r Tā), tensor_eq_of_eq (smul_contr r Tā)]
|
||
simp only [contr_index, smul_index, Fin.castOrderIso_refl, OrderIso.refl_toEquiv, mapIso_refl,
|
||
smul_tensor, LinearMapClass.map_smul, LinearEquiv.refl_apply]
|
||
apply congrArg
|
||
exact h.2 h1
|
||
|
||
/-!
|
||
|
||
## Addition of allowed `TensorIndex`
|
||
|
||
-/
|
||
|
||
/-- The condition on tensors with indices for their addition to exists.
|
||
This condition states that the the indices of one tensor are exact permutations of indices
|
||
of another after contraction. This includes the id of the index and the color.
|
||
|
||
This condition is general enough to allow addition of e.g. `Ļᵤāᵤā + Ļᵤāᵤā`, but
|
||
will NOT allow e.g. `Ļᵤāᵤā + ĻįµĀ²įµ¤ā`. -/
|
||
def AddCond (Tā Tā : š£.TensorIndex) : Prop :=
|
||
Tā.index.PermContr Tā.index
|
||
|
||
namespace AddCond
|
||
|
||
lemma to_PermContr {Tā Tā : š£.TensorIndex} (h : AddCond Tā Tā) : Tā.index.PermContr Tā.index := h
|
||
|
||
@[symm]
|
||
lemma symm {Tā Tā : š£.TensorIndex} (h : AddCond Tā Tā) : AddCond Tā Tā := by
|
||
rw [AddCond] at h
|
||
exact h.symm
|
||
|
||
lemma refl (T : š£.TensorIndex) : AddCond T T := by
|
||
exact PermContr.refl _
|
||
|
||
lemma trans {Tā Tā Tā : š£.TensorIndex} (h1 : AddCond Tā Tā) (h2 : AddCond Tā Tā) :
|
||
AddCond Tā Tā := by
|
||
rw [AddCond] at h1 h2
|
||
exact h1.trans h2
|
||
|
||
lemma rel_left {Tā Tā' Tā : š£.TensorIndex} (h : AddCond Tā Tā) (h' : Tā ā Tā') :
|
||
AddCond Tā' Tā := h'.1.symm.trans h
|
||
|
||
lemma rel_right {Tā Tā Tā' : š£.TensorIndex} (h : AddCond Tā Tā) (h' : Tā ā Tā') :
|
||
AddCond Tā Tā' := h.trans h'.1
|
||
|
||
/-- The equivalence between indices after contraction given a `AddCond`. -/
|
||
@[simp]
|
||
def toEquiv {Tā Tā : š£.TensorIndex} (h : AddCond Tā Tā) :
|
||
Fin Tā.contr.index.1.length ā Fin Tā.contr.index.1.length := h.to_PermContr.toEquiv
|
||
|
||
lemma toEquiv_colorMap {Tā Tā : š£.TensorIndex} (h : AddCond Tā Tā) :
|
||
ColorMap.MapIso h.toEquiv (Tā.contr.index).1.colorMap (Tā.contr.index).1.colorMap :=
|
||
h.to_PermContr.toEquiv_colorMap'
|
||
|
||
end AddCond
|
||
|
||
/-- The addition of two `TensorIndex` given the condition that, after contraction,
|
||
their index lists are the same. -/
|
||
def add (Tā Tā : š£.TensorIndex) (h : AddCond Tā Tā) :
|
||
š£.TensorIndex where
|
||
index := Tā.index.contr
|
||
tensor := (š£.mapIso h.toEquiv h.toEquiv_colorMap Tā.contr.tensor) + Tā.contr.tensor
|
||
|
||
/-- Notation for addition of tensor indices. -/
|
||
notation:71 Tā "+["h"]" Tā:72 => add Tā Tā h
|
||
|
||
namespace AddCond
|
||
|
||
lemma add_right {Tā Tā Tā : š£.TensorIndex} (h : AddCond Tā Tā) (h' : AddCond Tā Tā) :
|
||
AddCond Tā (Tā +[h'] Tā) := by
|
||
simpa only [AddCond, add, contr_index] using h.rel_right Tā.rel_contr
|
||
|
||
lemma add_left {Tā Tā Tā : š£.TensorIndex} (h : AddCond Tā Tā) (h' : AddCond Tā Tā) :
|
||
AddCond (Tā +[h] Tā) Tā :=
|
||
(add_right h'.symm h).symm
|
||
|
||
lemma of_add_right' {Tā Tā Tā : š£.TensorIndex} {h' : AddCond Tā Tā} (h : AddCond Tā (Tā +[h'] Tā)) :
|
||
AddCond Tā Tā := by
|
||
change Tā.AddCond Tā.contr at h
|
||
exact h.rel_right Tā.rel_contr.symm
|
||
|
||
lemma of_add_right {Tā Tā Tā : š£.TensorIndex} {h' : AddCond Tā Tā} (h : AddCond Tā (Tā +[h'] Tā)) :
|
||
AddCond Tā Tā := h.of_add_right'.trans h'.symm
|
||
|
||
lemma of_add_left {Tā Tā Tā : š£.TensorIndex} {h' : AddCond Tā Tā}
|
||
(h : AddCond (Tā +[h'] Tā) Tā) : AddCond Tā Tā :=
|
||
(of_add_right' h.symm).symm
|
||
|
||
lemma of_add_left' {Tā Tā Tā : š£.TensorIndex} {h' : AddCond Tā Tā}
|
||
(h : AddCond (Tā +[h'] Tā) Tā) : AddCond Tā Tā :=
|
||
(of_add_right h.symm).symm
|
||
|
||
lemma add_left_of_add_right {Tā Tā Tā : š£.TensorIndex} {h' : AddCond Tā Tā}
|
||
(h : AddCond Tā (Tā +[h'] Tā)) : AddCond (Tā +[of_add_right h] Tā) Tā := by
|
||
have h0 := ((of_add_right' h).trans h'.symm)
|
||
exact (h'.symm.add_right h0).symm
|
||
|
||
lemma add_right_of_add_left {Tā Tā Tā : š£.TensorIndex} {h' : AddCond Tā Tā}
|
||
(h : AddCond (Tā +[h'] Tā) Tā) : AddCond Tā (Tā +[of_add_left h] Tā) :=
|
||
(add_left (of_add_left h) (of_add_left' h).symm).symm
|
||
|
||
lemma add_comm {Tā Tā : š£.TensorIndex} (h : AddCond Tā Tā) :
|
||
AddCond (Tā +[h] Tā) (Tā +[h.symm] Tā) := by
|
||
apply add_right
|
||
apply add_left
|
||
exact h.symm
|
||
|
||
end AddCond
|
||
|
||
@[simp]
|
||
lemma add_index (Tā Tā : š£.TensorIndex) (h : AddCond Tā Tā) :
|
||
(add Tā Tā h).index = Tā.index.contr := rfl
|
||
|
||
@[simp]
|
||
lemma add_tensor (Tā Tā : š£.TensorIndex) (h : AddCond Tā Tā) :
|
||
(add Tā Tā h).tensor =
|
||
(š£.mapIso h.toEquiv h.toEquiv_colorMap Tā.contr.tensor) + Tā.contr.tensor := by rfl
|
||
|
||
/-- Scalar multiplication commutes with addition. -/
|
||
lemma smul_add (r : R) (Tā Tā : š£.TensorIndex) (h : AddCond Tā Tā) :
|
||
r ⢠(Tā +[h] Tā) = r ⢠Tā +[h] r ⢠Tā := by
|
||
refine ext _ _ rfl ?_
|
||
simp [add]
|
||
rw [tensor_eq_of_eq (smul_contr r Tā), tensor_eq_of_eq (smul_contr r Tā)]
|
||
simp only [smul_index, contr_index, Fin.castOrderIso_refl, OrderIso.refl_toEquiv, mapIso_refl,
|
||
smul_tensor, AddCond.toEquiv, LinearMapClass.map_smul, LinearEquiv.refl_apply]
|
||
|
||
lemma add_hasNoContr (Tā Tā : š£.TensorIndex) (h : AddCond Tā Tā) :
|
||
(Tā +[h] Tā).index.1.HasNoContr := by
|
||
simpa using Tā.index.1.contr_hasNoContr
|
||
|
||
@[simp]
|
||
lemma contr_add (Tā Tā : š£.TensorIndex) (h : AddCond Tā Tā) :
|
||
(Tā +[h] Tā).contr = Tā +[h] Tā :=
|
||
contr_of_hasNoContr (Tā +[h] Tā) (add_hasNoContr Tā Tā h)
|
||
|
||
@[simp]
|
||
lemma contr_add_tensor (Tā Tā : š£.TensorIndex) (h : AddCond Tā Tā) :
|
||
(Tā +[h] Tā).contr.tensor =
|
||
š£.mapIso (Fin.castOrderIso (by rw [index_eq_of_eq (contr_add Tā Tā h)])).toEquiv
|
||
(index_eq_colorMap_eq (index_eq_of_eq (contr_add Tā Tā h))) (Tā +[h] Tā).tensor :=
|
||
tensor_eq_of_eq (contr_add Tā Tā h)
|
||
|
||
lemma add_comm {Tā Tā : š£.TensorIndex} (h : AddCond Tā Tā) : Tā +[h] Tā ā Tā +[h.symm] Tā := by
|
||
apply And.intro h.add_comm
|
||
intro h
|
||
simp only [contr_index, add_index, contr_add_tensor, add_tensor, AddCond.toEquiv, map_add,
|
||
mapIso_mapIso]
|
||
rw [_root_.add_comm]
|
||
congr 1
|
||
all_goals
|
||
apply congrFun
|
||
apply congrArg
|
||
congr 1
|
||
rw [ā PermContr.toEquiv_contr_eq, ā PermContr.toEquiv_contr_eq,
|
||
PermContr.toEquiv_trans, PermContr.toEquiv_symm, PermContr.toEquiv_trans]
|
||
simp only [IndexListColor.contr_contr]
|
||
simp only [IndexListColor.contr_contr]
|
||
|
||
open AddCond in
|
||
lemma add_rel_left {Tā Tā' Tā : š£.TensorIndex} (h : AddCond Tā Tā) (h' : Tā ā Tā') :
|
||
Tā +[h] Tā ā Tā' +[h.rel_left h'] Tā := by
|
||
apply And.intro (PermContr.refl _)
|
||
intro h
|
||
simp only [contr_index, add_index, contr_add_tensor, add_tensor, toEquiv, map_add, mapIso_mapIso,
|
||
PermContr.toEquiv_refl, Equiv.refl_symm, mapIso_refl, LinearEquiv.refl_apply]
|
||
congr 1
|
||
rw [h'.to_eq]
|
||
simp only [mapIso_mapIso]
|
||
congr 1
|
||
congr 1
|
||
rw [PermContr.toEquiv_symm, ā PermContr.toEquiv_contr_eq, PermContr.toEquiv_trans,
|
||
PermContr.toEquiv_trans, PermContr.toEquiv_trans]
|
||
simp only [IndexListColor.contr_contr]
|
||
|
||
open AddCond in
|
||
lemma add_rel_right {Tā Tā Tā' : š£.TensorIndex} (h : AddCond Tā Tā) (h' : Tā ā Tā') :
|
||
Tā +[h] Tā ā Tā +[h.rel_right h'] Tā' :=
|
||
(add_comm _).trans ((add_rel_left _ h').trans (add_comm _))
|
||
|
||
open AddCond in
|
||
lemma add_assoc' {Tā Tā Tā : š£.TensorIndex} {h' : AddCond Tā Tā} (h : AddCond Tā (Tā +[h'] Tā)) :
|
||
Tā +[h] (Tā +[h'] Tā) = Tā +[h'.of_add_right h] Tā +[h'.add_left_of_add_right h] Tā := by
|
||
refine ext _ _ ?_ ?_
|
||
simp only [add_index, IndexListColor.contr_contr]
|
||
simp only [add_index, add_tensor, contr_index, toEquiv, contr_add_tensor, map_add, mapIso_mapIso]
|
||
rw [_root_.add_assoc]
|
||
congr
|
||
rw [ā PermContr.toEquiv_contr_eq, ā PermContr.toEquiv_contr_eq]
|
||
rw [PermContr.toEquiv_trans, PermContr.toEquiv_trans, PermContr.toEquiv_trans]
|
||
simp only [IndexListColor.contr_contr]
|
||
simp only [IndexListColor.contr_contr]
|
||
rw [ā PermContr.toEquiv_contr_eq, PermContr.toEquiv_trans]
|
||
simp only [IndexListColor.contr_contr]
|
||
|
||
open AddCond in
|
||
lemma add_assoc {Tā Tā Tā : š£.TensorIndex} {h' : AddCond Tā Tā} (h : AddCond (Tā +[h'] Tā) Tā) :
|
||
Tā +[h'] Tā +[h] Tā = Tā +[h'.add_right_of_add_left h] (Tā +[h'.of_add_left h] Tā) := by
|
||
rw [add_assoc']
|
||
|
||
/-! TODO: Show that the product is well defined with respect to Rel. -/
|
||
|
||
/-!
|
||
|
||
## Product of `TensorIndex` allowed
|
||
|
||
-/
|
||
|
||
/-- The condition on two tensors with indices determining if it possible to
|
||
take their product.
|
||
|
||
This condition says that the indices of the two tensors can contract nicely,
|
||
after the contraction of indivdual indices has taken place. Note that
|
||
it is required to take the contraction of indivdual tensors before taking the product
|
||
to ensure that the product is well-defined under the `Rel` equivalence relation.
|
||
|
||
For example, indices with the same id have dual colors, and no more then two indices
|
||
have the same id (after contraction). For example, the product of `ĻįµĀ¹įµ¤āįµĀ²` could be taken with
|
||
`ĻᵤāᵤāįµĀ³` or `ĻᵤāᵤāįµĀ³` or `ĻᵤāᵤāįµĀ²` or `ĻᵤāᵤāįµĀ¹`
|
||
(since contraction is done before taking the product)
|
||
but not with `ĻᵤāᵤāįµĀ³` or `ĻᵤāᵤāįµĀ²` or `ĻᵤāᵤāįµĀ²`. -/
|
||
def ProdCond (Tā Tā : š£.TensorIndex) : Prop :=
|
||
IndexListColorProp š£.toTensorColor (Tā.contr.index.1 ++ Tā.contr.index.1)
|
||
|
||
namespace ProdCond
|
||
|
||
lemma to_indexListColorProp {Tā Tā : š£.TensorIndex} (h : ProdCond Tā Tā) :
|
||
IndexListColorProp š£.toTensorColor (Tā.contr.index.1 ++ Tā.contr.index.1) := h
|
||
|
||
end ProdCond
|
||
|
||
/-- The tensor product of two `TensorIndex`. -/
|
||
def prod (Tā Tā : š£.TensorIndex)
|
||
(h : ProdCond Tā Tā) : š£.TensorIndex where
|
||
index := Tā.contr.index.prod Tā.contr.index h.to_indexListColorProp
|
||
tensor :=
|
||
š£.mapIso ((Fin.castOrderIso (IndexListColor.prod_numIndices)).toEquiv.trans
|
||
(finSumFinEquiv.symm)).symm
|
||
(IndexListColor.prod_colorMap h) <|
|
||
š£.tensoratorEquiv _ _ (Tā.contr.tensor āā[R] Tā.contr.tensor)
|
||
|
||
@[simp]
|
||
lemma prod_index (Tā Tā : š£.TensorIndex) (h : ProdCond Tā Tā) :
|
||
(prod Tā Tā h).index = Tā.contr.index.prod Tā.contr.index h.to_indexListColorProp := rfl
|
||
|
||
|
||
end TensorIndex
|
||
end
|
||
end TensorStructure
|