refactor: Start of major refactor of index notation

This commit is contained in:
jstoobysmith 2024-10-07 12:20:53 +00:00
parent 96e911ec27
commit 25a1d84c91
5 changed files with 567 additions and 2 deletions

View file

@ -0,0 +1,198 @@
/-
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.Tensors.ColorCat.Basic
/-!
## Monodial functor from color cat.
-/
namespace Fermion
noncomputable section
open Matrix
open MatrixGroups
open Complex
open TensorProduct
open IndexNotation
open CategoryTheory
/-- The colors associated with complex representations of SL(2, ) of intrest to physics. -/
inductive Color
| upL : Color
| downL : Color
| upR : Color
| downR : Color
| up : Color
| down : Color
/-- The corresponding representations associated with a color. -/
def colorToRep (c : Color) : Rep SL(2, ) :=
match c with
| Color.upL => altLeftHanded
| Color.downL => leftHanded
| Color.upR => altRightHanded
| Color.downR => rightHanded
| Color.up => Lorentz.complexContr
| Color.down => Lorentz.complexCo
/-- The linear equivalence between `colorToRep c1` and `colorToRep c2` when `c1 = c2`. -/
def colorToRepCongr {c1 c2 : Color} (h : c1 = c2) : colorToRep c1 ≃ₗ[] colorToRep c2 where
toFun := Equiv.cast (congrArg (CoeSort.coe ∘ colorToRep) h)
invFun := (Equiv.cast (congrArg (CoeSort.coe ∘ colorToRep) h)).symm
map_add' x y := by
subst h
rfl
map_smul' x y := by
subst h
rfl
left_inv x := Equiv.symm_apply_apply (Equiv.cast (congrArg (CoeSort.coe ∘ colorToRep) h)) x
right_inv x := Equiv.apply_symm_apply (Equiv.cast (congrArg (CoeSort.coe ∘ colorToRep) h)) x
lemma colorToRepCongr_comm_ρ {c1 c2 : Color} (h : c1 = c2) (M : SL(2, )) (x : (colorToRep c1)) :
(colorToRepCongr h) ((colorToRep c1).ρ M x) = (colorToRep c2).ρ M ((colorToRepCongr h) x) := by
subst h
rfl
namespace colorFun
/-- Given a object in `OverColor Color` the correpsonding tensor product of representations. -/
def obj' (f : OverColor Color) : Rep SL(2, ) := Rep.of {
toFun := fun M => PiTensorProduct.map (fun x => (colorToRep (f.hom x)).ρ M),
map_one' := by
simp
map_mul' := fun M N => by
simp only [CategoryTheory.Functor.id_obj, _root_.map_mul]
ext x : 2
simp only [LinearMap.compMultilinearMap_apply, PiTensorProduct.map_tprod, LinearMap.mul_apply]}
lemma obj'_ρ (f : OverColor Color) (M : SL(2, )) : (obj' f).ρ M =
PiTensorProduct.map (fun x => (colorToRep (f.hom x)).ρ M) := rfl
lemma obj'_ρ_tprod (f : OverColor Color) (M : SL(2, ))
(x : (i : f.left) → CoeSort.coe (colorToRep (f.hom i))) :
(obj' f).ρ M ((PiTensorProduct.tprod ) x) =
PiTensorProduct.tprod (fun i => (colorToRep (f.hom i)).ρ M (x i)) := by
rw [obj'_ρ]
change (PiTensorProduct.map fun x => (colorToRep (f.hom x)).ρ M) ((PiTensorProduct.tprod ) x) =
(PiTensorProduct.tprod ) fun i => ((colorToRep (f.hom i)).ρ M) (x i)
rw [PiTensorProduct.map_tprod]
/-- Given a morphism in `OverColor Color` the corresopnding linear equivalence between `obj' _`
induced by reindexing. -/
def mapToLinearEquiv' {f g : OverColor Color} (m : f ⟶ g) : (obj' f).V ≃ₗ[] (obj' g).V :=
(PiTensorProduct.reindex (fun x => colorToRep (f.hom x)) (OverColor.Hom.toEquiv m)).trans
(PiTensorProduct.congr (fun i => colorToRepCongr (OverColor.Hom.toEquiv_symm_apply m i)))
lemma mapToLinearEquiv'_tprod {f g : OverColor Color} (m : f ⟶ g)
(x : (i : f.left) → CoeSort.coe (colorToRep (f.hom i))) :
mapToLinearEquiv' m (PiTensorProduct.tprod x) =
PiTensorProduct.tprod (fun i => (colorToRepCongr (OverColor.Hom.toEquiv_symm_apply m i))
(x ((OverColor.Hom.toEquiv m).symm i))) := by
rw [mapToLinearEquiv']
simp only [CategoryTheory.Functor.id_obj, LinearEquiv.trans_apply]
change (PiTensorProduct.congr fun i => colorToRepCongr _)
((PiTensorProduct.reindex (fun x => CoeSort.coe (colorToRep (f.hom x)))
(OverColor.Hom.toEquiv m)) ((PiTensorProduct.tprod ) x)) = _
rw [PiTensorProduct.reindex_tprod, PiTensorProduct.congr_tprod]
rfl
/-- Given a morphism in `OverColor Color` the corresopnding map of representations induced by
reindexing. -/
def map' {f g : OverColor Color} (m : f ⟶ g) : obj' f ⟶ obj' g where
hom := (mapToLinearEquiv' m).toLinearMap
comm M := by
ext x : 2
refine PiTensorProduct.induction_on' x ?_ (by
intro x y hx hy
simp only [CategoryTheory.Functor.id_obj, map_add, hx, ModuleCat.coe_comp,
Function.comp_apply, hy])
intro r x
simp only [CategoryTheory.Functor.id_obj, PiTensorProduct.tprodCoeff_eq_smul_tprod,
_root_.map_smul, ModuleCat.coe_comp, Function.comp_apply]
apply congrArg
change (mapToLinearEquiv' m) (((obj' f).ρ M) ((PiTensorProduct.tprod ) x)) =
((obj' g).ρ M) ((mapToLinearEquiv' m) ((PiTensorProduct.tprod ) x))
rw [mapToLinearEquiv'_tprod, obj'_ρ_tprod]
erw [mapToLinearEquiv'_tprod, obj'_ρ_tprod]
apply congrArg
funext i
rw [colorToRepCongr_comm_ρ]
end colorFun
/-- The functor between `OverColor Color` and `Rep SL(2, )` taking a map of colors
to the corresponding tensor product representation. -/
@[simps!]
def colorFun : OverColor Color ⥤ Rep SL(2, ) where
obj := colorFun.obj'
map := colorFun.map'
map_id f := by
ext x
refine PiTensorProduct.induction_on' x (fun r x => ?_) (fun x y hx hy => by
simp only [CategoryTheory.Functor.id_obj, map_add, hx, ModuleCat.coe_comp,
Function.comp_apply, hy])
simp only [CategoryTheory.Functor.id_obj, PiTensorProduct.tprodCoeff_eq_smul_tprod,
_root_.map_smul, Action.id_hom, ModuleCat.id_apply]
apply congrArg
erw [colorFun.mapToLinearEquiv'_tprod]
exact congrArg _ (funext (fun i => rfl))
map_comp {X Y Z} f g := by
ext x
refine PiTensorProduct.induction_on' x (fun r x => ?_) (fun x y hx hy => by
simp only [CategoryTheory.Functor.id_obj, map_add, hx, ModuleCat.coe_comp,
Function.comp_apply, hy])
simp only [Functor.id_obj, PiTensorProduct.tprodCoeff_eq_smul_tprod, _root_.map_smul,
Action.comp_hom, ModuleCat.coe_comp, Function.comp_apply]
apply congrArg
rw [colorFun.map', colorFun.map', colorFun.map']
change (colorFun.mapToLinearEquiv' (CategoryTheory.CategoryStruct.comp f g))
((PiTensorProduct.tprod ) x) =
(colorFun.mapToLinearEquiv' g) ((colorFun.mapToLinearEquiv' f) ((PiTensorProduct.tprod ) x))
rw [colorFun.mapToLinearEquiv'_tprod, colorFun.mapToLinearEquiv'_tprod]
erw [colorFun.mapToLinearEquiv'_tprod]
refine congrArg _ (funext (fun i => ?_))
simp only [colorToRepCongr, Function.comp_apply, Equiv.cast_symm, LinearEquiv.coe_mk,
Equiv.cast_apply, cast_cast, cast_inj]
rfl
namespace colorFun
open CategoryTheory
open MonoidalCategory
@[simp]
lemma obj_ρ_empty (g : SL(2, )) : (colorFun.obj (𝟙_ (OverColor Color))).ρ g = LinearMap.id := by
erw [colorFun.obj'_ρ]
ext x
refine PiTensorProduct.induction_on' x (fun r x => ?_) <| fun x y hx hy => by
simp only [CategoryTheory.Functor.id_obj, map_add, hx, ModuleCat.coe_comp,
Function.comp_apply, hy]
erw [hx, hy]
rfl
simp only [OverColor.instMonoidalCategoryStruct_tensorUnit_left, Functor.id_obj,
OverColor.instMonoidalCategoryStruct_tensorUnit_hom, PiTensorProduct.tprodCoeff_eq_smul_tprod,
_root_.map_smul, PiTensorProduct.map_tprod, LinearMap.id_coe, id_eq]
apply congrArg
apply congrArg
funext i
exact Empty.elim i
/-- The unit natural transformation. -/
def ε : 𝟙_ (Rep SL(2, )) ⟶ colorFun.obj (𝟙_ (OverColor Color)) where
hom := (PiTensorProduct.isEmptyEquiv Empty).symm.toLinearMap
comm M := by
refine LinearMap.ext (fun x => ?_)
simp only [colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorUnit_left,
OverColor.instMonoidalCategoryStruct_tensorUnit_hom, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', Functor.id_obj, Category.id_comp, LinearEquiv.coe_coe]
erw [obj_ρ_empty M]
rfl
end colorFun
end
end Fermion

View file

@ -541,8 +541,10 @@ def domCoprod : MultilinearMap R (fun x => 𝓣.ColorModule (Sum.elim cX cY x))
(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]
| Sum.inl x => simp only [Sum.elim_inl, inlPureTensor_update_left, MultilinearMap.map_add,
inrPureTensor_update_left, ← add_tmul]
| Sum.inr y => simp only [Sum.elim_inr, inlPureTensor_update_right, inrPureTensor_update_right,
MultilinearMap.map_add, ← tmul_add]
map_smul' f xy r p := by
match xy with
| Sum.inl x => simp [TensorProduct.tmul_smul, TensorProduct.smul_tmul]

View file

@ -0,0 +1,249 @@
/-
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.CategoryTheory.Category.Basic
import Mathlib.CategoryTheory.Types
import Mathlib.CategoryTheory.Monoidal.Category
import Mathlib.CategoryTheory.Comma.Over
import Mathlib.CategoryTheory.Core
import Mathlib.CategoryTheory.Monoidal.Braided.Basic
import HepLean.SpaceTime.WeylFermion.Basic
import HepLean.SpaceTime.LorentzVector.Complex
/-!
## Over category.
-/
namespace IndexNotation
open CategoryTheory
/-- The core of the category of Types over C. -/
def OverColor (C : Type) := CategoryTheory.Core (CategoryTheory.Over C)
/-- The instance of `OverColor C` as a groupoid. -/
instance (C : Type) : Groupoid (OverColor C) := coreCategory
namespace OverColor
namespace Hom
variable {C : Type} {f g h : OverColor C}
/-- Given a hom in `OverColor C` the underlying equivalence between types. -/
def toEquiv (m : f ⟶ g) : f.left ≃ g.left where
toFun := m.hom.left
invFun := m.inv.left
left_inv := by
simpa only [Over.comp_left] using congrFun (congrArg (fun x => x.left) m.hom_inv_id)
right_inv := by
simpa only [Over.comp_left] using congrFun (congrArg (fun x => x.left) m.inv_hom_id)
@[simp]
lemma toEquiv_id (f : OverColor C) : toEquiv (𝟙 f) = Equiv.refl f.left := by
ext x
simp [toEquiv]
rfl
@[simp]
lemma toEquiv_comp (m : f ⟶ g) (n : g ⟶ h) : toEquiv (m ≫ n) = (toEquiv m).trans (toEquiv n) := by
ext x
simp [toEquiv]
rfl
lemma toEquiv_symm_apply (m : f ⟶ g) (i : g.left) :
f.hom ((toEquiv m).symm i) = g.hom i := by
simpa [toEquiv, types_comp] using congrFun m.inv.w i
lemma toEquiv_comp_hom (m : f ⟶ g) : g.hom ∘ (toEquiv m) = f.hom := by
ext x
simpa [types_comp, toEquiv] using congrFun m.hom.w x
end Hom
section monoidal
/-!
## The monoidal structure on `OverColor C`.
The category `OverColor C` can, through the disjoint union, be given the structure of a
symmetric monoidal category.
-/
@[simps!]
instance (C : Type) : MonoidalCategoryStruct (OverColor C) where
tensorObj f g := Over.mk (Sum.elim f.hom g.hom)
tensorUnit := Over.mk Empty.elim
whiskerLeft X Y1 Y2 m := Over.isoMk (Equiv.sumCongr (Equiv.refl X.left) (Hom.toEquiv m)).toIso
(by
ext x
simp only [Functor.id_obj, Functor.const_obj_obj, Over.mk_left, Equiv.toIso_hom, Over.mk_hom,
types_comp_apply, Equiv.sumCongr_apply, Equiv.coe_refl]
rw [Sum.elim_map, Hom.toEquiv_comp_hom]
rfl)
whiskerRight m X := Over.isoMk (Equiv.sumCongr (Hom.toEquiv m) (Equiv.refl X.left)).toIso
(by
ext x
simp only [Functor.id_obj, Functor.const_obj_obj, Over.mk_left, Equiv.toIso_hom, Over.mk_hom,
types_comp_apply, Equiv.sumCongr_apply, Equiv.coe_refl]
rw [Sum.elim_map, Hom.toEquiv_comp_hom]
rfl)
associator X Y Z := {
hom := Over.isoMk (Equiv.sumAssoc X.left Y.left Z.left).toIso (by
ext x
match x with
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => rfl
| Sum.inr x => rfl),
inv := (Over.isoMk (Equiv.sumAssoc X.left Y.left Z.left).toIso (by
ext x
match x with
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => rfl
| Sum.inr x => rfl)).symm,
hom_inv_id := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => rfl
| Sum.inr x => rfl,
inv_hom_id := by
apply CategoryTheory.Iso.ext
erw [CategoryTheory.Iso.trans_hom]
simp only [Functor.id_obj, Over.mk_left, Over.mk_hom, Iso.symm_hom, Iso.inv_hom_id]
rfl}
leftUnitor X := {
hom := Over.isoMk (Equiv.emptySum Empty X.left).toIso
inv := (Over.isoMk (Equiv.emptySum Empty X.left).toIso).symm
hom_inv_id := by
apply CategoryTheory.Iso.ext
erw [CategoryTheory.Iso.trans_hom]
simp only [Functor.id_obj, Over.mk_left, Over.mk_hom, Iso.symm_hom, Iso.hom_inv_id]
rfl,
inv_hom_id := by
apply CategoryTheory.Iso.ext
erw [CategoryTheory.Iso.trans_hom]}
rightUnitor X := {
hom := Over.isoMk (Equiv.sumEmpty X.left Empty).toIso
inv := (Over.isoMk (Equiv.sumEmpty X.left Empty).toIso).symm
hom_inv_id := by
apply CategoryTheory.Iso.ext
erw [CategoryTheory.Iso.trans_hom]
simp only [Functor.id_obj, Over.mk_left, Over.mk_hom, Iso.symm_hom, Iso.hom_inv_id]
rfl,
inv_hom_id := by
apply CategoryTheory.Iso.ext
erw [CategoryTheory.Iso.trans_hom]}
instance (C : Type) : MonoidalCategory (OverColor C) where
tensorHom_def f g := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => rfl
tensor_id X Y := CategoryTheory.Iso.ext <| (Iso.eq_inv_comp _).mp rfl
tensor_comp f1 f2 g1 g2 := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl
whiskerLeft_id X Y := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl
id_whiskerRight X Y := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl
associator_naturality {X1 X2 X3 Y1 Y2 Y3} f1 f2 f3 :=
CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => rfl
| Sum.inr x => rfl
leftUnitor_naturality f :=
CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => exact Empty.elim x
| Sum.inr x => rfl
rightUnitor_naturality f :=
CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => rfl
| Sum.inr x => exact Empty.elim x
pentagon f g h i := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl (Sum.inl (Sum.inl x)) => rfl
| Sum.inl (Sum.inl (Sum.inr x)) => rfl
| Sum.inl (Sum.inr x) => rfl
| Sum.inr x => rfl
triangle f g := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => exact Empty.elim x
| Sum.inr x => rfl
instance (C : Type) : BraidedCategory (OverColor C) where
braiding f g := {
hom := Over.isoMk (Equiv.sumComm f.left g.left).toIso
inv := (Over.isoMk (Equiv.sumComm f.left g.left).toIso).symm
hom_inv_id := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl,
inv_hom_id := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl}
braiding_naturality_right X Y1 Y2 f := CategoryTheory.Iso.ext <| Over.OverMorphism.ext
<| funext fun x => by
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl
braiding_naturality_left X f := CategoryTheory.Iso.ext <| Over.OverMorphism.ext
<| funext fun x => by
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl
hexagon_forward X1 X2 X3 := CategoryTheory.Iso.ext <| Over.OverMorphism.ext
<| funext fun x => by
match x with
| Sum.inl (Sum.inl x) => rfl
| Sum.inl (Sum.inr x) => rfl
| Sum.inr x => rfl
hexagon_reverse X1 X2 X3 := CategoryTheory.Iso.ext <| Over.OverMorphism.ext
<| funext fun x => by
match x with
| Sum.inr (Sum.inl x) => rfl
| Sum.inr (Sum.inr x) => rfl
| Sum.inl x => rfl
instance (C : Type) : SymmetricCategory (OverColor C) where
toBraidedCategory := instBraidedCategory C
symmetry X Y := CategoryTheory.Iso.ext <| Over.OverMorphism.ext <| funext fun x => by
match x with
| Sum.inl x => rfl
| Sum.inr x => rfl
end monoidal
def mk (f : X → C) : OverColor C := Over.mk f
open MonoidalCategory
def equivToIso {c : X → C} (e : X ≃ Y) : mk c ≅ mk (c ∘ e.symm) := {
hom := Over.isoMk e.toIso ((Iso.eq_inv_comp e.toIso).mp rfl),
inv := (Over.isoMk e.toIso ((Iso.eq_inv_comp e.toIso).mp rfl)).symm,
hom_inv_id := by
apply CategoryTheory.Iso.ext
erw [CategoryTheory.Iso.trans_hom]
simp only [Functor.id_obj, Over.mk_left, Over.mk_hom, Iso.symm_hom, Iso.hom_inv_id]
rfl,
inv_hom_id := by
apply CategoryTheory.Iso.ext
erw [CategoryTheory.Iso.trans_hom]
simp only [Iso.symm_hom, Iso.inv_hom_id]
rfl}
end OverColor
end IndexNotation

View file

@ -0,0 +1,89 @@
/-
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.Tensors.ColorCat.Basic
/-!
## Tensor trees
-/
open IndexNotation
open CategoryTheory
structure TensorStruct where
C : Type
G : Type
G_group : Group G
k : Type
k_commRing : CommRing k
F : MonoidalFunctor (OverColor C) (Rep k G)
τ : C → C
evalNo : C → N
namespace TensorStruct
variable (S : TensorStruct)
instance : CommRing S.k := S.k_commRing
instance : Group S.G := S.G_group
end TensorStruct
inductive TensorTree (S : TensorStruct) : ∀ {n : }, (Fin n → S.C) → Type where
| tensorNode {n : } {c : Fin n → S.C} : S.F.obj (OverColor.mk c) → TensorTree S c
| add {n : } {c : Fin n → S.C} : TensorTree S c → TensorTree S c → TensorTree S c
| perm {n m : } {c : Fin n → S.C} {c1 : Fin m → S.C}
(σ : (OverColor.mk c) ⟶ (OverColor.mk c1)) (t : TensorTree S c) : TensorTree S c1
| prod {n m : } {c : Fin n → S.C} {c1 : Fin m → S.C}
(t : TensorTree S c) (t1 : TensorTree S c1) : TensorTree S (Sum.elim c c1 ∘ finSumFinEquiv.symm)
| scale {n : } {c : Fin n → S.C} : S.k → TensorTree S c → TensorTree S c
| mult {n m : } {c : Fin n.succ → S.C} {c1 : Fin m.succ → S.C} :
(i : Fin n.succ) → (j : Fin m.succ) → TensorTree S c → TensorTree S c1 →
TensorTree S (Sum.elim (c ∘ Fin.succAbove i) (c1 ∘ Fin.succAbove j) ∘ finSumFinEquiv.symm)
| contr {n : } {c : Fin n.succ.succ → S.C} : (i : Fin n.succ.succ) →
(j : Fin n.succ) → TensorTree S c → TensorTree S (c ∘ Fin.succAbove i ∘ Fin.succAbove j)
| jiggle {n : } {c : Fin n → S.C} : (i : Fin n) → TensorTree S c →
TensorTree S (Function.update c i (S.τ (c i)))
| eval {n : } {c : Fin n.succ → S.C} : (i : Fin n.succ) → (x : Fin (S.evalNo (c i))) →
TensorTree S c → TensorTree S (c ∘ Fin.succAbove i)
namespace TensorTree
variable {S : TensorStruct} {n : } {c : Fin n → S.C} (T : TensorTree S c)
open MonoidalCategory
open TensorProduct
def size : ∀ {n : } {c : Fin n → S.C}, TensorTree S c → := fun
| tensorNode _ => 1
| add t1 t2 => t1.size + t2.size + 1
| perm _ t => t.size + 1
| scale _ t => t.size + 1
| prod t1 t2 => t1.size + t2.size + 1
| mult _ _ t1 t2 => t1.size + t2.size + 1
| contr _ _ t => t.size + 1
| jiggle _ t => t.size + 1
| eval _ _ t => t.size + 1
noncomputable section
def tensor : ∀ {n : } {c : Fin n → S.C}, TensorTree S c → S.F.obj (OverColor.mk c) := fun
| tensorNode t => t
| add t1 t2 => t1.tensor + t2.tensor
| perm σ t => (S.F.map σ).hom t.tensor
| scale a t => a • t.tensor
| prod t1 t2 => (S.F.map (OverColor.equivToIso finSumFinEquiv).hom).hom
((S.F.μ _ _).hom (t1.tensor ⊗ₜ t2.tensor))
| _ => 0
lemma tensor_tensorNode {c : Fin n → S.C} (T : S.F.obj (OverColor.mk c)) :
(tensorNode T).tensor = T := rfl
end
end TensorTree

View file

@ -0,0 +1,27 @@
/-
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.Tensors.Tree.Basic
import Lean.Elab.Term
/-!
## Elaboration of tensor trees
This file turns
-/
open Lean
open Lean.Elab.Term
open Lean
open Lean.Meta
open Lean.Elab
open Lean.Elab.Term
declare_syntax_cat tensorExpr
syntax ident (ppSpace term)* : tensorExpr
syntax tensorExpr "⊗" tensorExpr : tensorExpr