/- 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.Color import HepLean.SpaceTime.LorentzTensor.IndexNotation.Relations 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 toColorIndexList.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} instance : Coe 𝓣.TensorIndex (ColorIndexList 𝓣.toTensorColor) where coe T := T.toColorIndexList 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.colorMap' (cn ∘ Fin.cast hn.symm)) : 𝓣.TensorIndex where toColorIndexList := l tensor := 𝓣.mapIso (Equiv.refl _) (ColorMap.DualMap.split_dual' (by simpa using hd)) <| 𝓣.dualize (ColorMap.DualMap.split l.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 toColorIndexList := T.toColorIndexList.contr tensor := 𝓣.mapIso (Equiv.refl _) T.contrEquiv_colorMapIso <| 𝓣.contr T.toColorIndexList.contrEquiv T.contrEquiv_contrCond T.tensor /-- Applying contr to a tensor whose indices has no contracts does not do anything. -/ @[simp] lemma contr_of_withDual_empty (T : 𝓣.TensorIndex) (h : T.withDual = βˆ…) : T.contr = T := by refine ext ?_ ?_ Β· simp [contr, ColorIndexList.contr] have hx := T.contrIndexList_of_withDual_empty h apply ColorIndexList.ext simp [hx] Β· simp only [contr] 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 have hEm : IsEmpty { x // x ∈ i.withUniqueDualLT } := by rw [Finset.isEmpty_coe_sort] rw [Finset.eq_empty_iff_forall_not_mem] intro x hx have hx' : x ∈ i.withUniqueDual := by exact Finset.mem_of_mem_filter x hx rw [i.unique_duals] at h rw [h] at hx' simp_all erw [TensorStructure.contr_tprod_isEmpty] erw [mapIso_tprod] simp only [Equiv.refl_symm, Equiv.refl_apply, colorMap', mapIso_tprod, id_eq, OrderIso.toEquiv_symm, RelIso.coe_fn_toEquiv] 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 let hl := i.contrEquiv_on_withDual_empty l h exact let_value_heq f hl @[simp] lemma contr_contr (T : 𝓣.TensorIndex) : T.contr.contr = T.contr := T.contr.contr_of_withDual_empty (by simp [contr, ColorIndexList.contr]) @[simp] lemma contr_toColorIndexList (T : 𝓣.TensorIndex) : T.contr.toColorIndexList = T.toColorIndexList.contr := rfl @[simp] lemma contr_toIndexList (T : 𝓣.TensorIndex) : T.contr.toIndexList = T.toIndexList.contrIndexList := rfl /-! ## Scalar multiplication of -/ /-- The scalar multiplication of a `TensorIndex` by an element of `R`. -/ instance : SMul R 𝓣.TensorIndex where smul := fun r T => { toColorIndexList := T.toColorIndexList tensor := r β€’ T.tensor} @[simp] lemma smul_index (r : R) (T : 𝓣.TensorIndex) : (r β€’ T).toColorIndexList = T.toColorIndexList := 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₁.ContrPerm Tβ‚‚ ∧ βˆ€ (h : T₁.ContrPerm Tβ‚‚), T₁.contr.tensor = 𝓣.mapIso (contrPermEquiv h).symm (contrPermEquiv_colorMap_iso h) Tβ‚‚.contr.tensor namespace Rel /-- Rel is reflexive. -/ lemma refl (T : 𝓣.TensorIndex) : Rel T T := by apply And.intro simp simp /-- 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 (contrPermEquiv (h1.1.trans h2.1)).symm _) T₃.contr.tensor trans (𝓣.mapIso ((contrPermEquiv h1.1).trans (contrPermEquiv h2.1)).symm (by simp have h1 := contrPermEquiv_colorMap_iso (ContrPerm.symm (ContrPerm.trans h1.left h2.left)) rw [← ColorMap.MapIso.symm'] at h1 exact h1)) T₃.contr.tensor swap congr 1 simp only [contrPermEquiv_trans, contrPermEquiv_symm] 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 (contrPermEquiv h.1).symm (contrPermEquiv_colorMap_iso h.1) 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 intro h rw [tensor_eq_of_eq T.contr_contr] simp only [contr_toColorIndexList, colorMap', contrPermEquiv_self_contr, OrderIso.toEquiv_symm, Fin.symm_castOrderIso, mapIso_mapIso] trans 𝓣.mapIso (Equiv.refl _) (by rfl) T.contr.tensor simp only [contr_toColorIndexList, mapIso_refl, LinearEquiv.refl_apply] 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_toColorIndexList, smul_index, Fin.castOrderIso_refl, OrderIso.refl_toEquiv, mapIso_refl, smul_tensor, map_smul, LinearEquiv.refl_apply, contrPermEquiv_symm] 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₁.ContrPerm Tβ‚‚ namespace AddCond lemma to_PermContr {T₁ Tβ‚‚ : 𝓣.TensorIndex} (h : AddCond T₁ Tβ‚‚) : T₁.toColorIndexList.ContrPerm Tβ‚‚.toColorIndexList := 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 := ContrPerm.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 end AddCond /-- The equivalence between indices after contraction given a `AddCond`. -/ @[simp] def addCondEquiv {T₁ Tβ‚‚ : 𝓣.TensorIndex} (h : AddCond T₁ Tβ‚‚) : Fin T₁.contr.length ≃ Fin Tβ‚‚.contr.length := contrPermEquiv h lemma addCondEquiv_colorMap {T₁ Tβ‚‚ : 𝓣.TensorIndex} (h : AddCond T₁ Tβ‚‚) : ColorMap.MapIso (addCondEquiv h) T₁.contr.colorMap' Tβ‚‚.contr.colorMap' := contrPermEquiv_colorMap_iso' h /-- 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 toColorIndexList := Tβ‚‚.toColorIndexList.contr tensor := (𝓣.mapIso (addCondEquiv h) (addCondEquiv_colorMap h) 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] 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_toColorIndexList (T₁ Tβ‚‚ : 𝓣.TensorIndex) (h : AddCond T₁ Tβ‚‚) : (add T₁ Tβ‚‚ h).toColorIndexList = Tβ‚‚.toColorIndexList.contr := rfl @[simp] lemma add_tensor (T₁ Tβ‚‚ : 𝓣.TensorIndex) (h : AddCond T₁ Tβ‚‚) : (add T₁ Tβ‚‚ h).tensor = (𝓣.mapIso (addCondEquiv h) (addCondEquiv_colorMap h) 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_toColorIndexList, Fin.castOrderIso_refl, OrderIso.refl_toEquiv, mapIso_refl, smul_tensor, map_smul, LinearEquiv.refl_apply] lemma add_withDual_empty (T₁ Tβ‚‚ : 𝓣.TensorIndex) (h : AddCond T₁ Tβ‚‚) : (T₁ +[h] Tβ‚‚).withDual = βˆ… := by simp [contr] change Tβ‚‚.toColorIndexList.contr.withDual = βˆ… simp [ColorIndexList.contr] @[simp] lemma contr_add (T₁ Tβ‚‚ : 𝓣.TensorIndex) (h : AddCond T₁ Tβ‚‚) : (T₁ +[h] Tβ‚‚).contr = T₁ +[h] Tβ‚‚ := contr_of_withDual_empty (T₁ +[h] Tβ‚‚) (add_withDual_empty 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 (colormap_mapIso (by simp)) (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_toColorIndexList, add_toColorIndexList, contr_add_tensor, add_tensor, addCondEquiv, map_add, mapIso_mapIso, colorMap', contrPermEquiv_symm] rw [_root_.add_comm] congr 1 Β· apply congrFun apply congrArg congr 1 rw [← contrPermEquiv_self_contr, ← contrPermEquiv_self_contr, contrPermEquiv_trans, contrPermEquiv_trans] Β· apply congrFun apply congrArg congr 1 rw [← contrPermEquiv_self_contr, ← contrPermEquiv_self_contr, contrPermEquiv_trans, contrPermEquiv_trans] 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 ContrPerm.refl intro h simp only [contr_add_tensor, add_tensor, map_add] congr 1 rw [h'.to_eq] simp only [contr_toColorIndexList, add_toColorIndexList, colorMap', addCondEquiv, contrPermEquiv_symm, mapIso_mapIso, contrPermEquiv_trans, contrPermEquiv_refl, Equiv.refl_symm, mapIso_refl, LinearEquiv.refl_apply] simp only [contr_toColorIndexList, add_toColorIndexList, colorMap', contrPermEquiv_refl, Equiv.refl_symm, mapIso_refl, LinearEquiv.refl_apply] 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_toColorIndexList, ColorIndexList.contr_contr] simp only [add_toColorIndexList, add_tensor, contr_toColorIndexList, addCondEquiv, contr_add_tensor, map_add, mapIso_mapIso] rw [_root_.add_assoc] congr rw [← contrPermEquiv_self_contr, ← contrPermEquiv_self_contr] rw [contrPermEquiv_trans, contrPermEquiv_trans, contrPermEquiv_trans] erw [← contrPermEquiv_self_contr, contrPermEquiv_trans] 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 -/ def ProdCond (T₁ Tβ‚‚ : 𝓣.TensorIndex) : Prop := AppendCond T₁.toColorIndexList Tβ‚‚.toColorIndexList namespace ProdCond lemma to_AppendCond {T₁ Tβ‚‚ : 𝓣.TensorIndex} (h : ProdCond T₁ Tβ‚‚) : T₁.AppendCond Tβ‚‚ := h end ProdCond /-- The tensor product of two `TensorIndex`. -/ def prod (T₁ Tβ‚‚ : 𝓣.TensorIndex) (h : ProdCond T₁ Tβ‚‚) : 𝓣.TensorIndex where toColorIndexList := T₁ ++[h] Tβ‚‚ tensor := 𝓣.mapIso IndexList.appendEquiv (T₁.colorMap_sumELim Tβ‚‚) <| 𝓣.tensoratorEquiv _ _ (T₁.tensor βŠ—β‚œ[R] Tβ‚‚.tensor) @[simp] lemma prod_toColorIndexList (T₁ Tβ‚‚ : 𝓣.TensorIndex) (h : ProdCond T₁ Tβ‚‚) : (prod T₁ Tβ‚‚ h).toColorIndexList = T₁.toColorIndexList ++[h] Tβ‚‚.toColorIndexList := rfl @[simp] lemma prod_toIndexList (T₁ Tβ‚‚ : 𝓣.TensorIndex) (h : ProdCond T₁ Tβ‚‚) : (prod T₁ Tβ‚‚ h).toIndexList = T₁.toIndexList ++ Tβ‚‚.toIndexList := rfl end TensorIndex end end TensorStructure