/- 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.IndexListColor 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 -/ 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 where /-- The list of indices. -/ index : IndexListColor ๐“ฃ.toTensorColor /-- The underlying tensor. -/ tensor : ๐“ฃ.Tensor index.1.colorMap namespace TensorIndex open TensorColor variable {๐“ฃ : TensorStructure R} [IndexNotation ๐“ฃ.Color] [Fintype ๐“ฃ.Color] [DecidableEq ๐“ฃ.Color] variable {n m : โ„•} {cn : Fin n โ†’ ๐“ฃ.Color} {cm : Fin m โ†’ ๐“ฃ.Color} lemma index_eq_colorMap_eq {Tโ‚ Tโ‚‚ : ๐“ฃ.TensorIndex} (hi : Tโ‚.index = Tโ‚‚.index) : (Tโ‚‚.index).1.colorMap = (Tโ‚.index).1.colorMap โˆ˜ (Fin.castOrderIso (by rw [hi])).toEquiv := by funext i congr 1 rw [hi] simp only [RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply] exact (Fin.heq_ext_iff (congrArg IndexList.numIndices (congrArg Subtype.val (id (Eq.symm hi))))).mpr rfl lemma ext (Tโ‚ Tโ‚‚ : ๐“ฃ.TensorIndex) (hi : Tโ‚.index = Tโ‚‚.index) (h : Tโ‚.tensor = ๐“ฃ.mapIso (Fin.castOrderIso (by rw [hi])).toEquiv (index_eq_colorMap_eq hi) Tโ‚‚.tensor) : Tโ‚ = Tโ‚‚ := by cases Tโ‚; cases Tโ‚‚ simp at hi subst hi simp_all /-- The construction of a `TensorIndex` from a tensor, a IndexListColor, and a condition on the dual maps. -/ def mkDualMap (T : ๐“ฃ.Tensor cn) (l : IndexListColor ๐“ฃ.toTensorColor) (hn : n = l.1.length) (hd : DualMap l.1.colorMap (cn โˆ˜ Fin.cast hn.symm)) : ๐“ฃ.TensorIndex where index := l tensor := ๐“ฃ.mapIso (Equiv.refl _) (DualMap.split_dual' (by simp [hd])) <| ๐“ฃ.dualize (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 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 /-- The tensor product of two `TensorIndex`. -/ def prod (Tโ‚ Tโ‚‚ : ๐“ฃ.TensorIndex) (h : IndexListColorProp ๐“ฃ.toTensorColor (Tโ‚.index.1 ++ Tโ‚‚.index.1)) : ๐“ฃ.TensorIndex where index := Tโ‚.index.prod Tโ‚‚.index h tensor := ๐“ฃ.mapIso ((Fin.castOrderIso (IndexListColor.prod_numIndices)).toEquiv.trans (finSumFinEquiv.symm)).symm (IndexListColor.prod_colorMap h) <| ๐“ฃ.tensoratorEquiv _ _ (Tโ‚.tensor โŠ—โ‚œ[R] Tโ‚‚.tensor) /-- The scalar multiplication of a `TensorIndex` by an element of `R`. -/ def smul (r : R) (T : ๐“ฃ.TensorIndex) : ๐“ฃ.TensorIndex where index := T.index tensor := r โ€ข T.tensor /-- The addition of two `TensorIndex` given the condition that, after contraction, their index lists are the same. -/ def add (Tโ‚ Tโ‚‚ : ๐“ฃ.TensorIndex) (h : IndexListColor.PermContr Tโ‚.index Tโ‚‚.index) : ๐“ฃ.TensorIndex where index := Tโ‚.index.contr tensor := let T1 := Tโ‚.contr.tensor let T2 :๐“ฃ.Tensor (Tโ‚.contr.index).1.colorMap := ๐“ฃ.mapIso h.toEquiv.symm h.toEquiv_colorMap Tโ‚‚.contr.tensor T1 + T2 /-- An (equivalence) relation on two `TensorIndex` given that after contraction, the two underlying tensors are the equal. -/ 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 end TensorIndex end end TensorStructure