/- 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 IndexListColor 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 -/ /-- 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 /-! TODO: Show that contracting twice is the same as contracting once. -/ /-! ## Product of `TensorIndex` allowed -/ /-- 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) @[simp] lemma prod_index (Tโ‚ Tโ‚‚ : ๐“ฃ.TensorIndex) (h : IndexListColorProp ๐“ฃ.toTensorColor (Tโ‚.index.1 ++ Tโ‚‚.index.1)) : (prod Tโ‚ Tโ‚‚ h).index = Tโ‚.index.prod Tโ‚‚.index h := rfl /-! ## Scalar multiplication of -/ /-- 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 /-! ## Addition of allowed `TensorIndex` -/ /-- 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 /-! ## 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 equivalence : 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 end TensorIndex end end TensorStructure