feat: Fix pauli matrices as tensors. Speedup

This commit is contained in:
jstoobysmith 2024-10-25 13:50:19 +00:00
parent c565e7ea1c
commit e6ef68d7e6
9 changed files with 692 additions and 322 deletions

View file

@ -770,10 +770,63 @@ lemma smul_tensor_eq {T1 T2 : TensorTree S c} {a : S.k} (h : T1.tensor = T2.tens
simp only [smul_tensor]
rw [h]
lemma smul_mul_eq {T1 : TensorTree S c} {a b : S.k} (h : a = b) :
(smul a T1).tensor = (smul b T1).tensor := by
rw [h]
lemma eq_tensorNode_of_eq_tensor {T1 : TensorTree S c} {t : S.F.obj (OverColor.mk c)}
(h : T1.tensor = t) : T1.tensor = (tensorNode t).tensor := by
simpa using h
/-!
## The zero tensor tree
-/
/-- The zero tensor tree. -/
def zeroTree {n : } {c : Fin n → S.C} : TensorTree S c := tensorNode 0
@[simp]
lemma zeroTree_tensor {n : } {c : Fin n → S.C} : (zeroTree (c := c)).tensor = 0 := by
rfl
lemma zero_smul {T1 : TensorTree S c} :
(smul 0 T1).tensor = zeroTree.tensor := by
simp only [smul_tensor, _root_.zero_smul, zeroTree_tensor]
lemma smul_zero {a : S.k} : (smul a (zeroTree (c :=c ))).tensor = zeroTree.tensor := by
simp only [smul_tensor, zeroTree_tensor, _root_.smul_zero]
lemma zero_add {T1 : TensorTree S c} : (add zeroTree T1).tensor = T1.tensor := by
simp only [add_tensor, zeroTree_tensor, _root_.zero_add]
lemma add_zero {T1 : TensorTree S c} : (add T1 zeroTree).tensor = T1.tensor := by
simp only [add_tensor, zeroTree_tensor, _root_.add_zero]
lemma perm_zero {n m : } {c : Fin n → S.C} {c1 : Fin m → S.C} (σ : (OverColor.mk c) ⟶
(OverColor.mk c1)) : (perm σ zeroTree).tensor = zeroTree.tensor := by
simp only [perm_tensor, zeroTree_tensor, map_zero]
lemma neg_zero : (neg (zeroTree (c := c))).tensor = zeroTree.tensor := by
simp only [neg_tensor, zeroTree_tensor, _root_.neg_zero]
lemma contr_zero {n : } {c : Fin n.succ.succ → S.C} {i : Fin n.succ.succ} {j : Fin n.succ}
{h : c (i.succAbove j) = S.τ (c i)} : (contr i j h zeroTree).tensor = zeroTree.tensor := by
simp only [contr_tensor, zeroTree_tensor, map_zero]
lemma zero_prod {n m : } {c : Fin n → S.C} {c1 : Fin m → S.C} (t : TensorTree S c1) :
(prod (zeroTree (c := c)) t).tensor = zeroTree.tensor := by
simp only [prod_tensor, Functor.id_obj, OverColor.mk_hom, Action.instMonoidalCategory_tensorObj_V,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, zeroTree_tensor, zero_tmul, map_zero]
lemma prod_zero {n m : } {c : Fin n → S.C} {c1 : Fin m → S.C} (t : TensorTree S c) :
(prod t (zeroTree (c := c1))).tensor = zeroTree.tensor := by
simp only [prod_tensor, Functor.id_obj, OverColor.mk_hom, Action.instMonoidalCategory_tensorObj_V,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, zeroTree_tensor, tmul_zero, map_zero]
/-- A structure containing a pair of indices (i, j) to be contracted in a tensor.
This is used in some proofs of node identities for tensor trees. -/
structure ContrPair {n : } (c : Fin n.succ.succ → S.C) where

View file

@ -23,7 +23,7 @@ open Lean.Elab
open Lean.Elab.Term
open Lean Meta Elab Tactic
open IndexNotation
open complexLorentzTensor
namespace TensorTree
/-!
@ -182,60 +182,60 @@ def stringToTerm (str : String) : TermElabM Term := do
/-- Specific types of tensors which appear which we want to elaborate in specific ways. -/
def specialTypes : List (String × (Term → Term)) := [
("CoeSort.coe Lorentz.complexCo", fun T =>
Syntax.mkApp (mkIdent ``TensorTree.vecNodeE) #[mkIdent ``Fermion.complexLorentzTensor,
mkIdent ``Fermion.Color.down, T]),
Syntax.mkApp (mkIdent ``TensorTree.vecNodeE) #[mkIdent ``complexLorentzTensor,
mkIdent ``complexLorentzTensor.Color.down, T]),
("CoeSort.coe Lorentz.complexContr", fun T =>
Syntax.mkApp (mkIdent ``TensorTree.vecNodeE) #[mkIdent ``Fermion.complexLorentzTensor,
mkIdent ``Fermion.Color.up, T]),
Syntax.mkApp (mkIdent ``TensorTree.vecNodeE) #[mkIdent ``complexLorentzTensor,
mkIdent ``complexLorentzTensor.Color.up, T]),
("ModuleCat.carrier (Lorentz.complexContr ⊗ Lorentz.complexCo).V", fun T =>
Syntax.mkApp (mkIdent ``TensorTree.twoNodeE) #[mkIdent ``Fermion.complexLorentzTensor,
mkIdent ``Fermion.Color.up, mkIdent ``Fermion.Color.down, T]),
Syntax.mkApp (mkIdent ``TensorTree.twoNodeE) #[mkIdent ``complexLorentzTensor,
mkIdent ``complexLorentzTensor.Color.up, mkIdent ``complexLorentzTensor.Color.down, T]),
("ModuleCat.carrier (Lorentz.complexContr ⊗ Lorentz.complexContr).V", fun T =>
Syntax.mkApp (mkIdent ``TensorTree.twoNodeE) #[mkIdent ``Fermion.complexLorentzTensor,
mkIdent ``Fermion.Color.up, mkIdent ``Fermion.Color.up, T]),
Syntax.mkApp (mkIdent ``TensorTree.twoNodeE) #[mkIdent ``complexLorentzTensor,
mkIdent ``complexLorentzTensor.Color.up, mkIdent ``complexLorentzTensor.Color.up, T]),
("ModuleCat.carrier (Lorentz.complexCo ⊗ Lorentz.complexCo).V", fun T =>
Syntax.mkApp (mkIdent ``TensorTree.twoNodeE) #[mkIdent ``Fermion.complexLorentzTensor,
mkIdent ``Fermion.Color.down, mkIdent ``Fermion.Color.down, T]),
Syntax.mkApp (mkIdent ``TensorTree.twoNodeE) #[mkIdent ``complexLorentzTensor,
mkIdent ``complexLorentzTensor.Color.down, mkIdent ``complexLorentzTensor.Color.down, T]),
("ModuleCat.carrier (Lorentz.complexCo ⊗ Lorentz.complexContr).V", fun T =>
Syntax.mkApp (mkIdent ``TensorTree.twoNodeE) #[
mkIdent ``Fermion.complexLorentzTensor,
mkIdent ``Fermion.Color.down,
mkIdent ``Fermion.Color.up, T]),
mkIdent ``complexLorentzTensor,
mkIdent ``complexLorentzTensor.Color.down,
mkIdent ``complexLorentzTensor.Color.up, T]),
("𝟙_ (Rep SL(2, )) ⟶ Lorentz.complexCo ⊗ Lorentz.complexCo", fun T =>
Syntax.mkApp (mkIdent ``TensorTree.constTwoNodeE) #[
mkIdent ``Fermion.complexLorentzTensor,
mkIdent ``Fermion.Color.down,
mkIdent ``Fermion.Color.down, T]),
mkIdent ``complexLorentzTensor,
mkIdent ``complexLorentzTensor.Color.down,
mkIdent ``complexLorentzTensor.Color.down, T]),
("𝟙_ (Rep SL(2, )) ⟶ Lorentz.complexContr ⊗ Lorentz.complexContr", fun T =>
Syntax.mkApp (mkIdent ``TensorTree.constTwoNodeE) #[
mkIdent ``Fermion.complexLorentzTensor,
mkIdent ``Fermion.Color.up,
mkIdent ``Fermion.Color.up, T]),
mkIdent ``complexLorentzTensor,
mkIdent ``complexLorentzTensor.Color.up,
mkIdent ``complexLorentzTensor.Color.up, T]),
("𝟙_ (Rep SL(2, )) ⟶ Lorentz.complexContr ⊗ Fermion.leftHanded ⊗ Fermion.rightHanded", fun T =>
Syntax.mkApp (mkIdent ``TensorTree.constThreeNodeE) #[
mkIdent ``Fermion.complexLorentzTensor, mkIdent ``Fermion.Color.up,
mkIdent ``Fermion.Color.upL,
mkIdent ``Fermion.Color.upR, T]),
mkIdent ``complexLorentzTensor, mkIdent ``complexLorentzTensor.Color.up,
mkIdent ``complexLorentzTensor.Color.upL,
mkIdent ``complexLorentzTensor.Color.upR, T]),
("𝟙_ (Rep SL(2, )) ⟶ Fermion.leftHanded ⊗ Fermion.leftHanded", fun T =>
Syntax.mkApp (mkIdent ``TensorTree.constTwoNodeE) #[
mkIdent ``Fermion.complexLorentzTensor,
mkIdent ``Fermion.Color.upL,
mkIdent ``Fermion.Color.upL, T]),
mkIdent ``complexLorentzTensor,
mkIdent ``complexLorentzTensor.Color.upL,
mkIdent ``complexLorentzTensor.Color.upL, T]),
("𝟙_ (Rep SL(2, )) ⟶ Fermion.altLeftHanded ⊗ Fermion.altLeftHanded", fun T =>
Syntax.mkApp (mkIdent ``TensorTree.constTwoNodeE) #[
mkIdent ``Fermion.complexLorentzTensor,
mkIdent ``Fermion.Color.downL,
mkIdent ``Fermion.Color.downL, T]),
mkIdent ``complexLorentzTensor,
mkIdent ``complexLorentzTensor.Color.downL,
mkIdent ``complexLorentzTensor.Color.downL, T]),
("𝟙_ (Rep SL(2, )) ⟶ Fermion.altRightHanded ⊗ Fermion.altRightHanded", fun T =>
Syntax.mkApp (mkIdent ``TensorTree.constTwoNodeE) #[
mkIdent ``Fermion.complexLorentzTensor,
mkIdent ``Fermion.Color.downR,
mkIdent ``Fermion.Color.downR, T]),
mkIdent ``complexLorentzTensor,
mkIdent ``complexLorentzTensor.Color.downR,
mkIdent ``complexLorentzTensor.Color.downR, T]),
("𝟙_ (Rep SL(2, )) ⟶ Fermion.rightHanded ⊗ Fermion.rightHanded", fun T =>
Syntax.mkApp (mkIdent ``TensorTree.constTwoNodeE) #[
mkIdent ``Fermion.complexLorentzTensor,
mkIdent ``Fermion.Color.upR,
mkIdent ``Fermion.Color.upR, T])]
mkIdent ``complexLorentzTensor,
mkIdent ``complexLorentzTensor.Color.upR,
mkIdent ``complexLorentzTensor.Color.upR, T])]
/-- The syntax associated with a terminal node of a tensor tree. -/
def termNodeSyntax (T : Term) : TermElabM Term := do