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

@ -18,9 +18,6 @@ import HepLean.SpaceTime.PauliMatrices.AsTensor
## Complex Lorentz tensors
-/
namespace Fermion
open Matrix
open MatrixGroups
open Complex
@ -29,6 +26,8 @@ open IndexNotation
open CategoryTheory
open MonoidalCategory
namespace complexLorentzTensor
/-- The colors associated with complex representations of SL(2, ) of intrest to physics. -/
inductive Color
| upL : Color
@ -78,11 +77,13 @@ instance : DecidableEq Color := fun x y =>
| Color.down, Color.downR => isFalse fun h => Color.noConfusion h
| Color.down, Color.up => isFalse fun h => Color.noConfusion h
noncomputable section
end complexLorentzTensor
noncomputable section
open complexLorentzTensor in
/-- The tensor structure for complex Lorentz tensors. -/
def complexLorentzTensor : TensorSpecies where
C := Fermion.Color
C := complexLorentzTensor.Color
G := SL(2, )
G_group := inferInstance
k :=
@ -191,8 +192,8 @@ def complexLorentzTensor : TensorSpecies where
| Color.downR => by simpa using Fermion.altRightContraction_apply_metric
| Color.up => by simpa using Lorentz.contrCoContraction_apply_metric
| Color.down => by simpa using Lorentz.coContrContraction_apply_metric
instance : DecidableEq complexLorentzTensor.C := Fermion.instDecidableEqColor
namespace complexLorentzTensor
instance : DecidableEq complexLorentzTensor.C := complexLorentzTensor.instDecidableEqColor
lemma basis_contr (c : complexLorentzTensor.C) (i : Fin (complexLorentzTensor.repDim c))
(j : Fin (complexLorentzTensor.repDim (complexLorentzTensor.τ c))) :
@ -208,5 +209,6 @@ lemma basis_contr (c : complexLorentzTensor.C) (i : Fin (complexLorentzTensor.re
| Color.up => Lorentz.contrCoContraction_basis _ _
| Color.down => Lorentz.coContrContraction_basis _ _
end complexLorentzTensor
end
end Fermion

View file

@ -32,8 +32,8 @@ open IndexNotation
open CategoryTheory
open TensorTree
open OverColor.Discrete
open Fermion
noncomputable section
namespace Fermion
namespace complexLorentzTensor
/-- Basis vectors for complex Lorentz tensors. -/
@ -95,7 +95,7 @@ def contrBasisVectorMul {n : } {c : Fin n.succ.succ → complexLorentzTensor.
lemma contrBasisVectorMul_neg {n : } {c : Fin n.succ.succ → complexLorentzTensor.C}
{i : Fin n.succ.succ} {j : Fin n.succ} {b : Π k, Fin (complexLorentzTensor.repDim (c k))}
(h : ¬ (b i).val = (b (i.succAbove j)).val := by decide) :
(h : ¬ (b i).val = (b (i.succAbove j)).val) :
contrBasisVectorMul i j b = 0 := by
rw [contrBasisVectorMul]
simp only [ite_eq_else, one_ne_zero, imp_false]
@ -103,7 +103,7 @@ lemma contrBasisVectorMul_neg {n : } {c : Fin n.succ.succ → complexLorentzT
lemma contrBasisVectorMul_pos {n : } {c : Fin n.succ.succ → complexLorentzTensor.C}
{i : Fin n.succ.succ} {j : Fin n.succ} {b : Π k, Fin (complexLorentzTensor.repDim (c k))}
(h : (b i).val = (b (i.succAbove j)).val := by decide) :
(h : (b i).val = (b (i.succAbove j)).val) :
contrBasisVectorMul i j b = 1 := by
rw [contrBasisVectorMul]
simp only [ite_eq_then, zero_ne_one, imp_false, Decidable.not_not]
@ -156,7 +156,7 @@ lemma contr_basisVector_tree_neg {n : } {c : Fin n.succ.succ → complexLoren
(tensorNode 0).tensor := by
rw [contr_basisVector_tree, contrBasisVectorMul]
rw [if_neg hn]
simp only [Nat.succ_eq_add_one, smul_tensor, tensorNode_tensor, zero_smul]
simp only [Nat.succ_eq_add_one, smul_tensor, tensorNode_tensor, _root_.zero_smul]
/-- Equivalence of Fin types appearing in the product of two basis vectors. -/
def prodBasisVecEquiv {n m : } {c : Fin n → complexLorentzTensor.C}
@ -214,7 +214,7 @@ lemma eval_basisVector {n : } {c : Fin n.succ → complexLorentzTensor.C}
basisVector (c ∘ Fin.succAbove i) (fun k => b (i.succAbove k)) := by
rw [eval_tensor, basisVector, basisVector]
simp only [Nat.succ_eq_add_one, Functor.id_obj, OverColor.mk_hom, tensorNode_tensor,
Function.comp_apply, one_smul, zero_smul]
Function.comp_apply, one_smul, _root_.zero_smul]
erw [TensorSpecies.evalMap_tprod]
congr 1
have h1 : Fin.ofNat' ↑j (@Fin.size_pos' (complexLorentzTensor.repDim (c i)) _) = j := by
@ -448,5 +448,4 @@ lemma pauliMatrix_basis_expand_tree : {PauliMatrix.asConsTensor | μ α β}ᵀ.t
rfl
end complexLorentzTensor
end Fermion
end

View file

@ -36,8 +36,8 @@ open OverColor.Discrete
noncomputable section
namespace Fermion
open complexLorentzTensor
namespace complexLorentzTensor
open Fermion
/-!
@ -177,6 +177,25 @@ lemma basis_contr_pauliMatrix_basis_tree_expand' {n : } {c : Fin n → comple
<| contr_tensor_eq <| prod_basisVector_tree _ _]
rfl
def pauliMatrixBasisProdMap
{n : } {c : Fin n → complexLorentzTensor.C}
(b : Π k, Fin (complexLorentzTensor.repDim (c k))) (i1 i2 i3 : Fin 4) :
(i : Fin (n + (Nat.succ 0).succ.succ)) →
Fin (complexLorentzTensor.repDim (Sum.elim c ![Color.up, Color.upL, Color.upR]
(finSumFinEquiv.symm i))) := fun i => prodBasisVecEquiv (finSumFinEquiv.symm i)
((HepLean.PiTensorProduct.elimPureTensor b (fun | (0 : Fin 3) => i1 | 1 => i2 | 2 => i3))
(finSumFinEquiv.symm i))
def basisVectorContrPauli {n : } {c : Fin n → complexLorentzTensor.C}
(i : Fin (n + 3)) (j : Fin (n +2))
(b : Π k, Fin (complexLorentzTensor.repDim (c k)))
(i1 i2 i3 : Fin 4) :=
let c' := (Sum.elim c ![Color.up, Color.upL, Color.upR] ∘ finSumFinEquiv.symm)
∘ Fin.succAbove i ∘ Fin.succAbove j
let b' (i1 i2 i3 : Fin 4) := fun k => (pauliMatrixBasisProdMap b i1 i2 i3) (i.succAbove (j.succAbove k))
basisVector c' (b' i1 i2 i3)
lemma basis_contr_pauliMatrix_basis_tree_expand {n : } {c : Fin n → complexLorentzTensor.C}
(i : Fin (n + 3)) (j : Fin (n +2))
(h : (pauliMatrixContrMap c) (i.succAbove j) = complexLorentzTensor.τ
@ -184,30 +203,25 @@ lemma basis_contr_pauliMatrix_basis_tree_expand {n : } {c : Fin n → complex
(b : Π k, Fin (complexLorentzTensor.repDim (c k))) :
let c' := (Sum.elim c ![Color.up, Color.upL, Color.upR] ∘ finSumFinEquiv.symm)
∘ Fin.succAbove i ∘ Fin.succAbove j
let b'' (i1 i2 i3 : Fin 4) : (i : Fin (n + (Nat.succ 0).succ.succ)) →
Fin (complexLorentzTensor.repDim (Sum.elim c ![Color.up, Color.upL, Color.upR]
(finSumFinEquiv.symm i))) := fun i => prodBasisVecEquiv (finSumFinEquiv.symm i)
((HepLean.PiTensorProduct.elimPureTensor b (fun | (0 : Fin 3) => i1 | 1 => i2 | 2 => i3))
(finSumFinEquiv.symm i))
let b' (i1 i2 i3 : Fin 4) := fun k => (b'' i1 i2 i3) (i.succAbove (j.succAbove k))
let b' (i1 i2 i3 : Fin 4) := fun k => (pauliMatrixBasisProdMap b i1 i2 i3) (i.succAbove (j.succAbove k))
(contr i j h (TensorTree.prod (tensorNode (basisVector c b))
(constThreeNodeE complexLorentzTensor Color.up Color.upL Color.upR
PauliMatrix.asConsTensor))).tensor = (((
TensorTree.smul (contrBasisVectorMul i j (b'' 0 0 0))
TensorTree.smul (contrBasisVectorMul i j (pauliMatrixBasisProdMap b 0 0 0))
(tensorNode (basisVector c' (b' 0 0 0))))).add
(((TensorTree.smul (contrBasisVectorMul i j (b'' 0 1 1))
(((TensorTree.smul (contrBasisVectorMul i j (pauliMatrixBasisProdMap b 0 1 1))
(tensorNode (basisVector c' (b' 0 1 1))))).add
(((TensorTree.smul (contrBasisVectorMul i j (b'' 1 0 1))
(((TensorTree.smul (contrBasisVectorMul i j (pauliMatrixBasisProdMap b 1 0 1))
(tensorNode (basisVector c' (b' 1 0 1))))).add
(((TensorTree.smul (contrBasisVectorMul i j (b'' 1 1 0))
(((TensorTree.smul (contrBasisVectorMul i j (pauliMatrixBasisProdMap b 1 1 0))
(tensorNode (basisVector c' (b' 1 1 0))))).add
((TensorTree.smul (-I) ((TensorTree.smul (contrBasisVectorMul i j (b'' 2 0 1))
((TensorTree.smul (-I) ((TensorTree.smul (contrBasisVectorMul i j (pauliMatrixBasisProdMap b 2 0 1))
(tensorNode (basisVector c' (b' 2 0 1)))))).add
((TensorTree.smul I ((TensorTree.smul (contrBasisVectorMul i j (b'' 2 1 0))
((TensorTree.smul I ((TensorTree.smul (contrBasisVectorMul i j (pauliMatrixBasisProdMap b 2 1 0))
(tensorNode (basisVector c' (b' 2 1 0)))))).add
(((TensorTree.smul (contrBasisVectorMul i j (b'' 3 0 0))
(((TensorTree.smul (contrBasisVectorMul i j (pauliMatrixBasisProdMap b 3 0 0))
(tensorNode (basisVector c' (b' 3 0 0))))).add
(TensorTree.smul (-1) ((TensorTree.smul (contrBasisVectorMul i j (b'' 3 1 1)) (tensorNode
(TensorTree.smul (-1) ((TensorTree.smul (contrBasisVectorMul i j (pauliMatrixBasisProdMap b 3 1 1)) (tensorNode
(basisVector c' (b' 3 1 1))))))))))))).tensor := by
rw [basis_contr_pauliMatrix_basis_tree_expand']
/- Contracting basis vectors. -/
@ -227,7 +241,46 @@ lemma basis_contr_pauliMatrix_basis_tree_expand {n : } {c : Fin n → complex
rw [add_tensor_eq_snd <| add_tensor_eq_snd <| add_tensor_eq_snd <|
add_tensor_eq_snd <| add_tensor_eq_snd <| add_tensor_eq_snd <| add_tensor_eq_snd <|
smul_tensor_eq <| contr_basisVector_tree _]
rfl
end Fermion
def pauilMatrixBasisContrMap {n : } {c : Fin n → complexLorentzTensor.C}
(i : Fin (n + 3)) (j : Fin (n +2))
(b : Π k, Fin (complexLorentzTensor.repDim (c k))) (i1 i2 i3 : Fin 4) :
(k : Fin (n + 1)) →
Fin
(complexLorentzTensor.repDim
(Sum.elim c ![Color.up, Color.upL, Color.upR] (finSumFinEquiv.symm (i.succAbove (j.succAbove k))))) :=
fun k => (pauliMatrixBasisProdMap b i1 i2 i3) (i.succAbove (j.succAbove k))
lemma basis_contr_pauliMatrix_basis_tree_expand_tensor {n : } {c : Fin n → complexLorentzTensor.C}
(i : Fin (n + 3)) (j : Fin (n +2))
(h : (pauliMatrixContrMap c) (i.succAbove j) = complexLorentzTensor.τ
((pauliMatrixContrMap c) i))
(b : Π k, Fin (complexLorentzTensor.repDim (c k))) :
let c' := (Sum.elim c ![Color.up, Color.upL, Color.upR] ∘ finSumFinEquiv.symm)
∘ Fin.succAbove i ∘ Fin.succAbove j
let b' (i1 i2 i3 : Fin 4) := fun k => (pauliMatrixBasisProdMap b i1 i2 i3) (i.succAbove (j.succAbove k))
(contr i j h (TensorTree.prod (tensorNode (basisVector c b))
(constThreeNodeE complexLorentzTensor Color.up Color.upL Color.upR
PauliMatrix.asConsTensor))).tensor =
(contrBasisVectorMul i j (pauliMatrixBasisProdMap b 0 0 0)) • (basisVectorContrPauli i j b 0 0 0)
+ (contrBasisVectorMul i j (pauliMatrixBasisProdMap b 0 1 1)) • (basisVectorContrPauli i j b 0 1 1)
+ (contrBasisVectorMul i j (pauliMatrixBasisProdMap b 1 0 1)) • (basisVectorContrPauli i j b 1 0 1)
+ (contrBasisVectorMul i j (pauliMatrixBasisProdMap b 1 1 0)) • (basisVectorContrPauli i j b 1 1 0)
+ (-I) • (contrBasisVectorMul i j (pauliMatrixBasisProdMap b 2 0 1)) • (basisVectorContrPauli i j b 2 0 1)
+ I • (contrBasisVectorMul i j (pauliMatrixBasisProdMap b 2 1 0)) • (basisVectorContrPauli i j b 2 1 0)
+ (contrBasisVectorMul i j (pauliMatrixBasisProdMap b 3 0 0)) • (basisVectorContrPauli i j b 3 0 0)
+ (-1 : ) • (contrBasisVectorMul i j (pauliMatrixBasisProdMap b 3 1 1)) • (basisVectorContrPauli i j b 3 1 1) := by
rw [basis_contr_pauliMatrix_basis_tree_expand]
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, cons_val_one, head_cons, Fin.val_zero,
Nat.cast_zero, cons_val_two, Fin.val_one, Nat.cast_one, add_tensor, smul_tensor,
tensorNode_tensor, neg_smul, one_smul, Int.reduceNeg]
simp_all only [Function.comp_apply, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue]
rfl
end complexLorentzTensor
end

View file

@ -3,15 +3,7 @@ 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.Elab
import HepLean.Tensors.ComplexLorentz.Basic
import Mathlib.LinearAlgebra.TensorProduct.Basis
import HepLean.Tensors.Tree.NodeIdentities.Basic
import HepLean.Tensors.Tree.NodeIdentities.PermProd
import HepLean.Tensors.Tree.NodeIdentities.PermContr
import HepLean.Tensors.Tree.NodeIdentities.ProdComm
import HepLean.Tensors.Tree.NodeIdentities.ContrSwap
import HepLean.Tensors.Tree.NodeIdentities.ContrContr
import HepLean.Tensors.ComplexLorentz.PauliLower
/-!
## Bispinors
@ -30,18 +22,17 @@ open TensorTree
open OverColor.Discrete
open Fermion
noncomputable section
namespace Lorentz
namespace complexContr
namespace complexLorentzTensor
open Lorentz
/-- A bispinor `pᵃᵃ` created from a lorentz vector `p^μ`. -/
def bispinorUp (p : complexContr) :=
{p | μ ⊗ (Lorentz.coMetric | μ ν ⊗ PauliMatrix.asConsTensor | ν α β)}ᵀ.tensor
def contrBispinorUp (p : complexContr) :=
{p | μ ⊗ pauliCo | μ α β}ᵀ.tensor
/-- A bispinor `pₐₐ` created from a lorentz vector `p^μ`. -/
def bispinorDown (p : complexContr) :=
{Fermion.altRightMetric | β β' ⊗ Fermion.altLeftMetric | α α' ⊗
(complexContr.bispinorUp p) | α β}ᵀ.tensor
end complexContr
end Lorentz
end complexLorentzTensor
end

View file

@ -22,8 +22,8 @@ open TensorTree
open OverColor.Discrete
noncomputable section
namespace Fermion
open complexLorentzTensor
namespace complexLorentzTensor
open Fermion
set_option maxRecDepth 20000 in
lemma contr_rank_2_symm {T1 : (Lorentz.complexContr ⊗ Lorentz.complexContr).V}
{T2 : (Lorentz.complexCo ⊗ Lorentz.complexCo).V} :
@ -159,6 +159,6 @@ lemma leftMetric_mul_rightMetric_tree :
smul_tensor, neg_smul, one_smul]
rfl
end Fermion
end complexLorentzTensor
end

View file

@ -29,28 +29,52 @@ open TensorTree
open OverColor.Discrete
noncomputable section
namespace Fermion
open complexLorentzTensor
namespace complexLorentzTensor
open Fermion
/-- The map to colors one gets when contracting the 4-vector indices pauli matrices. -/
def pauliMatrixContrPauliMatrixMap := ((Sum.elim
((Sum.elim ![Color.down, Color.down] ![Color.up, Color.upL, Color.upR] ∘ ⇑finSumFinEquiv.symm) ∘
Fin.succAbove 0 ∘ Fin.succAbove 1) ![Color.up, Color.upL, Color.upR] ∘ ⇑finSumFinEquiv.symm) ∘
Fin.succAbove 1 ∘ Fin.succAbove 1) ![Color.up, Color.upL, Color.upR] ∘ ⇑finSumFinEquiv.symm) ∘
Fin.succAbove 0 ∘ Fin.succAbove 2)
lemma pauliMatrix_contr_lower_0_0_0 :
{(basisVector pauliMatrixLowerMap (fun | 0 => 0 | 1 => 0 | 2 => 0)) | μ α β ⊗
{(basisVector pauliCoMap (fun | 0 => 0 | 1 => 0 | 2 => 0)) | μ α β ⊗
PauliMatrix.asConsTensor | μ α' β'}ᵀ.tensor =
basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 0 | 1 => 0 | 2 => 0 | 3 => 0)
+ basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 0 | 1 => 0 | 2 => 1 | 3 => 1) := by
rw [basis_contr_pauliMatrix_basis_tree_expand]
rw [contrBasisVectorMul_pos, contrBasisVectorMul_pos,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg]
conv =>
lhs
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; rhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs
simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add]
rw [basisVectorContrPauli, basisVectorContrPauli]
/- Simplifying. -/
simp only [smul_tensor, add_tensor, tensorNode_tensor]
simp only [zero_smul, one_smul, smul_zero, add_zero, zero_add]
congr 1
· congr 1
funext k
@ -60,18 +84,42 @@ lemma pauliMatrix_contr_lower_0_0_0 :
fin_cases k <;> rfl
lemma pauliMatrix_contr_lower_0_1_1 :
{(basisVector pauliMatrixLowerMap (fun | 0 => 0 | 1 => 1 | 2 => 1)) | μ α β ⊗
{(basisVector pauliCoMap (fun | 0 => 0 | 1 => 1 | 2 => 1)) | μ α β ⊗
PauliMatrix.asConsTensor | μ α' β'}ᵀ.tensor =
basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 1 | 1 => 1 | 2 => 0 | 3 => 0)
+ basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 1 | 1 => 1 | 2 => 1 | 3 => 1) := by
rw [basis_contr_pauliMatrix_basis_tree_expand]
rw [contrBasisVectorMul_pos, contrBasisVectorMul_pos,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg]
conv =>
lhs
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; rhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs
simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add]
rw [basisVectorContrPauli, basisVectorContrPauli]
/- Simplifying. -/
simp only [smul_tensor, add_tensor, tensorNode_tensor]
simp only [zero_smul, one_smul, smul_zero, add_zero, zero_add]
congr 1
· congr 1
funext k
@ -81,18 +129,42 @@ lemma pauliMatrix_contr_lower_0_1_1 :
fin_cases k <;> rfl
lemma pauliMatrix_contr_lower_1_0_1 :
{(basisVector pauliMatrixLowerMap (fun | 0 => 1 | 1 => 0 | 2 => 1)) | μ α β ⊗
{(basisVector pauliCoMap (fun | 0 => 1 | 1 => 0 | 2 => 1)) | μ α β ⊗
PauliMatrix.asConsTensor | μ α' β'}ᵀ.tensor =
basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 0 | 1 => 1 | 2 => 0 | 3 => 1)
+ basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 0 | 1 => 1 | 2 => 1 | 3 => 0) := by
rw [basis_contr_pauliMatrix_basis_tree_expand]
rw [contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_pos, contrBasisVectorMul_pos,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg]
conv =>
lhs
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; rhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs
simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add]
rw [basisVectorContrPauli, basisVectorContrPauli]
/- Simplifying. -/
simp only [smul_tensor, add_tensor, tensorNode_tensor]
simp only [zero_smul, one_smul, smul_zero, add_zero, zero_add]
congr 1
· congr 1
funext k
@ -102,18 +174,42 @@ lemma pauliMatrix_contr_lower_1_0_1 :
fin_cases k <;> rfl
lemma pauliMatrix_contr_lower_1_1_0 :
{(basisVector pauliMatrixLowerMap (fun | 0 => 1 | 1 => 1 | 2 => 0)) | μ α β ⊗
{(basisVector pauliCoMap (fun | 0 => 1 | 1 => 1 | 2 => 0)) | μ α β ⊗
PauliMatrix.asConsTensor | μ α' β'}ᵀ.tensor =
basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 1 | 1 => 0 | 2 => 0 | 3 => 1)
+ basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 1 | 1 => 0 | 2 => 1 | 3 => 0) := by
rw [basis_contr_pauliMatrix_basis_tree_expand]
rw [contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_pos, contrBasisVectorMul_pos,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg]
conv =>
lhs
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; rhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs
simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add]
rw [basisVectorContrPauli, basisVectorContrPauli]
/- Simplifying. -/
simp only [smul_tensor, add_tensor, tensorNode_tensor]
simp only [zero_smul, one_smul, smul_zero, add_zero, zero_add]
congr 1
· congr 1
funext k
@ -123,19 +219,43 @@ lemma pauliMatrix_contr_lower_1_1_0 :
fin_cases k <;> rfl
lemma pauliMatrix_contr_lower_2_0_1 :
{(basisVector pauliMatrixLowerMap (fun | 0 => 2 | 1 => 0 | 2 => 1)) | μ α β ⊗
{(basisVector pauliCoMap (fun | 0 => 2 | 1 => 0 | 2 => 1)) | μ α β ⊗
PauliMatrix.asConsTensor | μ α' β'}ᵀ.tensor =
(-I) • basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 0 | 1 => 1 | 2 => 0 | 3 => 1)
+ (I) •
basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 0 | 1 => 1 | 2 => 1 | 3 => 0) := by
rw [basis_contr_pauliMatrix_basis_tree_expand]
rw [contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_pos, contrBasisVectorMul_pos,
contrBasisVectorMul_neg, contrBasisVectorMul_neg]
conv =>
lhs
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; rhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs
simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add]
rw [basisVectorContrPauli, basisVectorContrPauli]
/- Simplifying. -/
simp only [smul_tensor, add_tensor, tensorNode_tensor]
simp only [zero_smul, one_smul, smul_zero, add_zero, zero_add]
congr 1
· congr 2
funext k
@ -145,19 +265,43 @@ lemma pauliMatrix_contr_lower_2_0_1 :
fin_cases k <;> rfl
lemma pauliMatrix_contr_lower_2_1_0 :
{(basisVector pauliMatrixLowerMap (fun | 0 => 2 | 1 => 1 | 2 => 0)) | μ α β ⊗
{(basisVector pauliCoMap (fun | 0 => 2 | 1 => 1 | 2 => 0)) | μ α β ⊗
PauliMatrix.asConsTensor | μ α' β'}ᵀ.tensor =
(-I) • basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 1 | 1 => 0 | 2 => 0 | 3 => 1)
+ (I) •
basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 1 | 1 => 0 | 2 => 1 | 3 => 0) := by
rw [basis_contr_pauliMatrix_basis_tree_expand]
rw [contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_pos, contrBasisVectorMul_pos,
contrBasisVectorMul_neg, contrBasisVectorMul_neg]
conv =>
lhs
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; rhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs
simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add]
rw [basisVectorContrPauli, basisVectorContrPauli]
/- Simplifying. -/
simp only [smul_tensor, add_tensor, tensorNode_tensor]
simp only [zero_smul, one_smul, smul_zero, add_zero, zero_add]
congr 1
· congr 2
funext k
@ -167,19 +311,43 @@ lemma pauliMatrix_contr_lower_2_1_0 :
fin_cases k <;> rfl
lemma pauliMatrix_contr_lower_3_0_0 :
{(basisVector pauliMatrixLowerMap (fun | 0 => 3 | 1 => 0 | 2 => 0)) | μ α β ⊗
{(basisVector pauliCoMap (fun | 0 => 3 | 1 => 0 | 2 => 0)) | μ α β ⊗
PauliMatrix.asConsTensor | μ α' β'}ᵀ.tensor =
basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 0 | 1 => 0 | 2 => 0 | 3 => 0)
+ (-1 : ) • basisVector pauliMatrixContrPauliMatrixMap
(fun | 0 => 0 | 1 => 0 | 2 => 1 | 3 => 1) := by
rw [basis_contr_pauliMatrix_basis_tree_expand]
rw [contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_pos, contrBasisVectorMul_pos]
conv =>
lhs
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; rhs; lhs;
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; rhs; rhs; lhs;
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs
simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add]
rw [basisVectorContrPauli, basisVectorContrPauli]
/- Simplifying. -/
simp only [smul_tensor, add_tensor, tensorNode_tensor]
simp only [zero_smul, one_smul, smul_zero, add_zero, zero_add]
congr 1
· congr 2
funext k
@ -189,19 +357,43 @@ lemma pauliMatrix_contr_lower_3_0_0 :
fin_cases k <;> rfl
lemma pauliMatrix_contr_lower_3_1_1 :
{(basisVector pauliMatrixLowerMap (fun | 0 => 3 | 1 => 1 | 2 => 1)) | μ α β ⊗
{(basisVector pauliCoMap (fun | 0 => 3 | 1 => 1 | 2 => 1)) | μ α β ⊗
PauliMatrix.asConsTensor | μ α' β'}ᵀ.tensor =
basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 1 | 1 => 1 | 2 => 0 | 3 => 0)
+ (-1 : ) •
basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 1 | 1 => 1 | 2 => 1 | 3 => 1) := by
rw [basis_contr_pauliMatrix_basis_tree_expand]
rw [contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_pos, contrBasisVectorMul_pos]
conv =>
lhs
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; rhs; lhs;
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; rhs; rhs; lhs;
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs
simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add]
rw [basisVectorContrPauli, basisVectorContrPauli]
/- Simplifying. -/
simp only [smul_tensor, add_tensor, tensorNode_tensor]
simp only [zero_smul, one_smul, smul_zero, add_zero, zero_add]
congr 1
· congr 2
funext k
@ -210,35 +402,34 @@ lemma pauliMatrix_contr_lower_3_1_1 :
funext k
fin_cases k <;> rfl
/-! TODO: Work out why `pauliMatrix_lower_basis_expand_prod'` is needed. -/
/-- This lemma is exactly the same as `pauliMatrix_lower_basis_expand_prod'`.
/-! TODO: Work out why `pauliCo_prod_basis_expand'` is needed. -/
/-- This lemma is exactly the same as `pauliCo_prod_basis_expand`.
It is needed here for `pauliMatrix_contract_pauliMatrix_aux`. It is unclear why
`pauliMatrix_lower_basis_expand_prod` does not work. -/
private lemma pauliMatrix_lower_basis_expand_prod' {n : } {c : Fin n → complexLorentzTensor.C}
private lemma pauliCo_prod_basis_expand' {n : } {c : Fin n → complexLorentzTensor.C}
(t : TensorTree complexLorentzTensor c) :
(prod {Lorentz.coMetric | μ ν ⊗ PauliMatrix.asConsTensor | μ α β}ᵀ t).tensor =
(TensorTree.prod (tensorNode pauliCo) t).tensor =
((((tensorNode
(basisVector pauliMatrixLowerMap fun | 0 => 0 | 1 => 0 | 2 => 0)).prod t).add
(basisVector pauliCoMap fun | 0 => 0 | 1 => 0 | 2 => 0)).prod t).add
(((tensorNode
(basisVector pauliMatrixLowerMap fun | 0 => 0 | 1 => 1 | 2 => 1)).prod t).add
(basisVector pauliCoMap fun | 0 => 0 | 1 => 1 | 2 => 1)).prod t).add
((TensorTree.smul (-1) ((tensorNode
(basisVector pauliMatrixLowerMap fun | 0 => 1 | 1 => 0 | 2 => 1)).prod t)).add
(basisVector pauliCoMap fun | 0 => 1 | 1 => 0 | 2 => 1)).prod t)).add
((TensorTree.smul (-1) ((tensorNode
(basisVector pauliMatrixLowerMap fun | 0 => 1 | 1 => 1 | 2 => 0)).prod t)).add
(basisVector pauliCoMap fun | 0 => 1 | 1 => 1 | 2 => 0)).prod t)).add
((TensorTree.smul I ((tensorNode
(basisVector pauliMatrixLowerMap fun | 0 => 2 | 1 => 0 | 2 => 1)).prod t)).add
(basisVector pauliCoMap fun | 0 => 2 | 1 => 0 | 2 => 1)).prod t)).add
((TensorTree.smul (-I) ((tensorNode
(basisVector pauliMatrixLowerMap fun | 0 => 2 | 1 => 1 | 2 => 0)).prod t)).add
(basisVector pauliCoMap fun | 0 => 2 | 1 => 1 | 2 => 0)).prod t)).add
((TensorTree.smul (-1) ((tensorNode
(basisVector pauliMatrixLowerMap fun | 0 => 3 | 1 => 0 | 2 => 0)).prod t)).add
(basisVector pauliCoMap fun | 0 => 3 | 1 => 0 | 2 => 0)).prod t)).add
((tensorNode
(basisVector pauliMatrixLowerMap fun | 0 => 3 | 1 => 1 | 2 => 1)).prod
(basisVector pauliCoMap fun | 0 => 3 | 1 => 1 | 2 => 1)).prod
t))))))))).tensor := by
exact pauliMatrix_lower_basis_expand_prod _
exact pauliCo_prod_basis_expand _
lemma pauliMatrix_contract_pauliMatrix_aux :
{Lorentz.coMetric | μ ν ⊗ PauliMatrix.asConsTensor | μ α β ⊗
PauliMatrix.asConsTensor | ν α' β'}ᵀ.tensor
lemma pauliCo_contr_pauliContr_expand_aux :
{pauliCo | μ α β ⊗ PauliMatrix.asConsTensor | μ α' β'}ᵀ.tensor
= ((tensorNode
((basisVector pauliMatrixContrPauliMatrixMap fun | 0 => 0 | 1 => 0 | 2 => 0 | 3 => 0) +
basisVector pauliMatrixContrPauliMatrixMap fun | 0 => 0 | 1 => 0 | 2 => 1 | 3 => 1)).add
@ -266,7 +457,7 @@ lemma pauliMatrix_contract_pauliMatrix_aux :
((basisVector pauliMatrixContrPauliMatrixMap fun | 0 => 1 | 1 => 1 | 2 => 0 | 3 => 0) +
(-1 : ) • basisVector pauliMatrixContrPauliMatrixMap
fun | 0 => 1 | 1 => 1 | 2 => 1 | 3 => 1))))))))).tensor := by
rw [contr_tensor_eq <| pauliMatrix_lower_basis_expand_prod' _]
rw [contr_tensor_eq <| pauliCo_prod_basis_expand' _]
/- Moving contraction through addition. -/
rw [contr_add]
rw [add_tensor_eq_snd <| contr_add _ _]
@ -309,9 +500,8 @@ lemma pauliMatrix_contract_pauliMatrix_aux :
add_tensor_eq_snd <| add_tensor_eq_snd <| add_tensor_eq_snd <| eq_tensorNode_of_eq_tensor <|
pauliMatrix_contr_lower_3_1_1]
lemma pauliMatrix_contract_pauliMatrix_expand :
{Lorentz.coMetric | μ ν ⊗ PauliMatrix.asConsTensor | μ α β ⊗
PauliMatrix.asConsTensor | ν α' β'}ᵀ.tensor =
lemma pauliCo_contr_pauliContr_expand :
{pauliCo | ν α β ⊗ PauliMatrix.asConsTensor | ν α' β'}ᵀ.tensor =
2 • basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 0 | 1 => 0 | 2 => 1 | 3 => 1)
+ 2 • basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 1 | 1 => 1 | 2 => 0 | 3 => 0)
- 2 • basisVector pauliMatrixContrPauliMatrixMap (fun | 0 => 0 | 1 => 1 | 2 => 1 | 3 => 0)
@ -326,9 +516,8 @@ lemma pauliMatrix_contract_pauliMatrix_expand :
abel
/-- The statement that `η_{μν} σ^{μ α dot β} σ^{ν α' dot β'} = 2 ε^{αα'} ε^{dot β dot β'}`. -/
theorem pauliMatrix_contract_pauliMatrix :
{Lorentz.coMetric | μ ν ⊗ PauliMatrix.asConsTensor | μ α β ⊗
PauliMatrix.asConsTensor | ν α' β' =
theorem pauliCo_contr_pauliContr :
{pauliCo | ν α β ⊗ PauliMatrix.asConsTensor | ν α' β' =
2 •ₜ Fermion.leftMetric | α α' ⊗ Fermion.rightMetric | β β'}ᵀ := by
rw [pauliMatrix_contract_pauliMatrix_expand]
rw [perm_tensor_eq <| smul_tensor_eq <| leftMetric_mul_rightMetric_tree]
@ -370,4 +559,4 @@ theorem pauliMatrix_contract_pauliMatrix :
· funext i
fin_cases i <;> rfl
end Fermion
end complexLorentzTensor

View file

@ -26,209 +26,292 @@ namespace Fermion
open complexLorentzTensor
/-- The map to color one gets when lowering the indices of pauli matrices. -/
def pauliMatrixLowerMap := ((Sum.elim ![Color.down, Color.down] ![Color.up, Color.upL, Color.upR] ∘
⇑finSumFinEquiv.symm) ∘ Fin.succAbove 0 ∘ Fin.succAbove 1)
def pauliCoMap := ((Sum.elim ![Color.down, Color.down] ![Color.up, Color.upL, Color.upR] ∘
⇑finSumFinEquiv.symm) ∘ Fin.succAbove 1 ∘ Fin.succAbove 1)
lemma pauliMatrix_contr_down_0 :
(contr 0 1 rfl (((tensorNode (basisVector ![Color.down, Color.down] fun x => 0)).prod
(contr 1 1 rfl (((tensorNode (basisVector ![Color.down, Color.down] fun x => 0)).prod
(constThreeNodeE complexLorentzTensor Color.up Color.upL Color.upR
PauliMatrix.asConsTensor)))).tensor
= basisVector pauliMatrixLowerMap (fun | 0 => 0 | 1 => 0 | 2 => 0)
+ basisVector pauliMatrixLowerMap (fun | 0 => 0 | 1 => 1 | 2 => 1) := by
rw [basis_contr_pauliMatrix_basis_tree_expand]
rw [contrBasisVectorMul_pos, contrBasisVectorMul_pos,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg]
simp only [smul_tensor, add_tensor, tensorNode_tensor]
simp only [one_smul, zero_smul, smul_zero, add_zero]
= basisVector pauliCoMap (fun | 0 => 0 | 1 => 0 | 2 => 0)
+ basisVector pauliCoMap (fun | 0 => 0 | 1 => 1 | 2 => 1) := by
conv =>
lhs
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; rhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs
simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add]
congr 1
· congr 1
· rw [basisVectorContrPauli]
congr 1
funext k
fin_cases k <;> rfl
· congr 1
· rw [basisVectorContrPauli]
congr 1
funext k
fin_cases k <;> rfl
lemma pauliMatrix_contr_down_0_tree :
(contr 0 1 rfl (((tensorNode (basisVector ![Color.down, Color.down] fun x => 0)).prod
(constThreeNodeE complexLorentzTensor Color.up Color.upL Color.upR
PauliMatrix.asConsTensor)))).tensor
= (TensorTree.add (tensorNode
(basisVector pauliMatrixLowerMap (fun | 0 => 0 | 1 => 0 | 2 => 0)))
(tensorNode (basisVector pauliMatrixLowerMap (fun | 0 => 0 | 1 => 1 | 2 => 1)))).tensor := by
exact pauliMatrix_contr_down_0
lemma pauliMatrix_contr_down_1 :
{(basisVector ![Color.down, Color.down] fun x => 1) | μ ν
{(basisVector ![Color.down, Color.down] fun x => 1) | ν μ ⊗
PauliMatrix.asConsTensor | μ α β}ᵀ.tensor
= basisVector pauliMatrixLowerMap (fun | 0 => 1 | 1 => 0 | 2 => 1)
+ basisVector pauliMatrixLowerMap (fun | 0 => 1 | 1 => 1 | 2 => 0) := by
rw [basis_contr_pauliMatrix_basis_tree_expand]
rw [contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_pos, contrBasisVectorMul_pos,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg]
simp only [smul_tensor, add_tensor, tensorNode_tensor]
simp only [zero_smul, one_smul, smul_zero, add_zero, zero_add]
= basisVector pauliCoMap (fun | 0 => 1 | 1 => 0 | 2 => 1)
+ basisVector pauliCoMap (fun | 0 => 1 | 1 => 1 | 2 => 0) := by
conv =>
lhs
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; rhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs
simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add]
congr 1
· congr 1
· rw [basisVectorContrPauli]
congr 1
funext k
fin_cases k <;> rfl
· congr 1
· rw [basisVectorContrPauli]
congr 1
funext k
fin_cases k <;> rfl
lemma pauliMatrix_contr_down_1_tree :
{(basisVector ![Color.down, Color.down] fun x => 1) | μ ν
PauliMatrix.asConsTensor | μ α β}ᵀ.tensor
= (TensorTree.add (tensorNode
(basisVector pauliMatrixLowerMap (fun | 0 => 1 | 1 => 0 | 2 => 1)))
(tensorNode (basisVector pauliMatrixLowerMap (fun | 0 => 1 | 1 => 1 | 2 => 0)))).tensor := by
exact pauliMatrix_contr_down_1
lemma pauliMatrix_contr_down_2 :
{(basisVector ![Color.down, Color.down] fun x => 2) | μ ν
PauliMatrix.asConsTensor | μ α β}ᵀ.tensor
= (- I) • basisVector pauliMatrixLowerMap (fun | 0 => 2 | 1 => 0 | 2 => 1)
+ (I) • basisVector pauliMatrixLowerMap (fun | 0 => 2 | 1 => 1 | 2 => 0) := by
rw [basis_contr_pauliMatrix_basis_tree_expand]
rw [contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_pos, contrBasisVectorMul_pos,
contrBasisVectorMul_neg, contrBasisVectorMul_neg]
/- Simplifying. -/
simp only [smul_tensor, add_tensor, tensorNode_tensor]
simp only [zero_smul, one_smul, smul_zero, add_zero, zero_add]
PauliMatrix.asConsTensor | ν α β}ᵀ.tensor
= (- I) • basisVector pauliCoMap (fun | 0 => 2 | 1 => 0 | 2 => 1)
+ (I) • basisVector pauliCoMap (fun | 0 => 2 | 1 => 1 | 2 => 0) := by
conv =>
lhs
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; lhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; rhs; rhs; lhs;
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs
simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add]
congr 1
· congr 2
· rw [basisVectorContrPauli]
congr 2
funext k
fin_cases k <;> rfl
· congr 2
· rw [basisVectorContrPauli]
congr 2
funext k
fin_cases k <;> rfl
lemma pauliMatrix_contr_down_2_tree :
{(basisVector ![Color.down, Color.down] fun x => 2) | μ ν
PauliMatrix.asConsTensor | μ α β}ᵀ.tensor =
(TensorTree.add
(smul (- I) (tensorNode (basisVector pauliMatrixLowerMap (fun | 0 => 2 | 1 => 0 | 2 => 1))))
(smul I (tensorNode (basisVector
pauliMatrixLowerMap (fun | 0 => 2 | 1 => 1 | 2 => 0))))).tensor := by
exact pauliMatrix_contr_down_2
lemma pauliMatrix_contr_down_3 :
{(basisVector ![Color.down, Color.down] fun x => 3) | μ ν
PauliMatrix.asConsTensor | μ α β}ᵀ.tensor
= basisVector pauliMatrixLowerMap (fun | 0 => 3 | 1 => 0 | 2 => 0)
+ (- 1 : ) • basisVector pauliMatrixLowerMap (fun | 0 => 3 | 1 => 1 | 2 => 1) := by
rw [basis_contr_pauliMatrix_basis_tree_expand]
rw [contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_neg, contrBasisVectorMul_neg,
contrBasisVectorMul_pos, contrBasisVectorMul_pos]
/- Simplifying. -/
simp only [smul_tensor, add_tensor, tensorNode_tensor]
simp only [zero_smul, one_smul, smul_zero, add_zero, zero_add]
PauliMatrix.asConsTensor | ν α β}ᵀ.tensor
= basisVector pauliCoMap (fun | 0 => 3 | 1 => 0 | 2 => 0)
+ (- 1 : ) • basisVector pauliCoMap (fun | 0 => 3 | 1 => 1 | 2 => 1) := by
conv =>
lhs
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; lhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; lhs; rhs; rhs; lhs
rw [contrBasisVectorMul_neg (by decide)]
conv =>
lhs; lhs; rhs; lhs;
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs; rhs; rhs; lhs;
rw [contrBasisVectorMul_pos (by decide)]
conv =>
lhs
simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add]
congr 1
· congr 2
· rw [basisVectorContrPauli]
congr 1
funext k
fin_cases k <;> rfl
· congr 2
· rw [basisVectorContrPauli]
congr 1
congr 1
funext k
fin_cases k <;> rfl
lemma pauliMatrix_contr_down_3_tree : {(basisVector ![Color.down, Color.down] fun x => 3) | μ ν
PauliMatrix.asConsTensor | μ α β}ᵀ.tensor =
(TensorTree.add
((tensorNode (basisVector pauliMatrixLowerMap (fun | 0 => 3 | 1 => 0 | 2 => 0))))
(smul (-1) (tensorNode (basisVector pauliMatrixLowerMap
(fun | 0 => 3 | 1 => 1 | 2 => 1))))).tensor := by
exact pauliMatrix_contr_down_3
def pauliCo := {Lorentz.coMetric | μ ν ⊗ PauliMatrix.asConsTensor | ν α β}ᵀ.tensor
lemma pauliMatrix_lower : {Lorentz.coMetric | μ ν ⊗ PauliMatrix.asConsTensor | μ α β}ᵀ.tensor
= basisVector pauliMatrixLowerMap (fun | 0 => 0 | 1 => 0 | 2 => 0)
+ basisVector pauliMatrixLowerMap (fun | 0 => 0 | 1 => 1 | 2 => 1)
- basisVector pauliMatrixLowerMap (fun | 0 => 1 | 1 => 0 | 2 => 1)
- basisVector pauliMatrixLowerMap (fun | 0 => 1 | 1 => 1 | 2 => 0)
+ I • basisVector pauliMatrixLowerMap (fun | 0 => 2 | 1 => 0 | 2 => 1)
- I • basisVector pauliMatrixLowerMap (fun | 0 => 2 | 1 => 1 | 2 => 0)
- basisVector pauliMatrixLowerMap (fun | 0 => 3 | 1 => 0 | 2 => 0)
+ basisVector pauliMatrixLowerMap (fun | 0 => 3 | 1 => 1 | 2 => 1) := by
rw [contr_tensor_eq <| prod_tensor_eq_fst <| coMetric_basis_expand_tree]
/- Moving the prod through additions. -/
rw [contr_tensor_eq <| add_prod _ _ _]
rw [contr_tensor_eq <| add_tensor_eq_snd <| add_prod _ _ _]
rw [contr_tensor_eq <| add_tensor_eq_snd <| add_tensor_eq_snd <| add_prod _ _ _]
/- Moving the prod through smuls. -/
rw [contr_tensor_eq <| add_tensor_eq_snd <| add_tensor_eq_fst <| smul_prod _ _ _]
rw [contr_tensor_eq <| add_tensor_eq_snd <| add_tensor_eq_snd <| add_tensor_eq_fst
<| smul_prod _ _ _]
rw [contr_tensor_eq <| add_tensor_eq_snd <| add_tensor_eq_snd <| add_tensor_eq_snd
<| smul_prod _ _ _]
/- Moving contraction through addition. -/
rw [contr_add]
rw [add_tensor_eq_snd <| contr_add _ _]
rw [add_tensor_eq_snd <| add_tensor_eq_snd <| contr_add _ _]
/- Moving contraction through smul. -/
rw [add_tensor_eq_snd <| add_tensor_eq_fst <| contr_smul _ _]
rw [add_tensor_eq_snd <| add_tensor_eq_snd <| add_tensor_eq_fst <| contr_smul _ _]
rw [add_tensor_eq_snd <| add_tensor_eq_snd <| add_tensor_eq_snd <| contr_smul _ _]
/- Replacing the contractions. -/
rw [add_tensor_eq_fst <| pauliMatrix_contr_down_0_tree]
rw [add_tensor_eq_snd <| add_tensor_eq_fst <| smul_tensor_eq <| pauliMatrix_contr_down_1_tree]
rw [add_tensor_eq_snd <| add_tensor_eq_snd <| add_tensor_eq_fst <| smul_tensor_eq <|
pauliMatrix_contr_down_2_tree]
rw [add_tensor_eq_snd <| add_tensor_eq_snd <| add_tensor_eq_snd <| smul_tensor_eq <|
pauliMatrix_contr_down_3_tree]
/- Simplifying -/
simp only [add_tensor, smul_tensor, tensorNode_tensor, smul_add,_root_.smul_smul]
simp only [Nat.reduceAdd, Fin.isValue, neg_smul, one_smul, mul_neg, neg_mul, one_mul,
_root_.neg_neg, mul_one]
lemma tensoreNode_pauliCo : (tensorNode pauliCo).tensor =
{Lorentz.coMetric | μ ν ⊗ PauliMatrix.asConsTensor | ν α β}ᵀ.tensor := by
rw [pauliCo]
rfl
lemma pauliMatrix_lower_tree : {Lorentz.coMetric | μ ν ⊗ PauliMatrix.asConsTensor | μ α β}ᵀ.tensor
set_option profiler true
set_option profiler.threshold 10
lemma pauliCo_basis_expand : pauliCo
= basisVector pauliCoMap (fun | 0 => 0 | 1 => 0 | 2 => 0)
+ basisVector pauliCoMap (fun | 0 => 0 | 1 => 1 | 2 => 1)
- basisVector pauliCoMap (fun | 0 => 1 | 1 => 0 | 2 => 1)
- basisVector pauliCoMap (fun | 0 => 1 | 1 => 1 | 2 => 0)
+ I • basisVector pauliCoMap (fun | 0 => 2 | 1 => 0 | 2 => 1)
- I • basisVector pauliCoMap (fun | 0 => 2 | 1 => 1 | 2 => 0)
- basisVector pauliCoMap (fun | 0 => 3 | 1 => 0 | 2 => 0)
+ basisVector pauliCoMap (fun | 0 => 3 | 1 => 1 | 2 => 1) := by
conv =>
lhs
rw [pauliCo]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| coMetric_basis_expand_tree]
/- Moving the prod through additions. -/
rw [contr_tensor_eq <| add_prod _ _ _]
rw [contr_tensor_eq <| add_tensor_eq_snd <| add_prod _ _ _]
rw [contr_tensor_eq <| add_tensor_eq_snd <| add_tensor_eq_snd <| add_prod _ _ _]
/- Moving the prod through smuls. -/
rw [contr_tensor_eq <| add_tensor_eq_snd <| add_tensor_eq_fst <| smul_prod _ _ _]
rw [contr_tensor_eq <| add_tensor_eq_snd <| add_tensor_eq_snd <| add_tensor_eq_fst
<| smul_prod _ _ _]
rw [contr_tensor_eq <| add_tensor_eq_snd <| add_tensor_eq_snd <| add_tensor_eq_snd
<| smul_prod _ _ _]
/- Moving contraction through addition. -/
rw [contr_add]
rw [add_tensor_eq_snd <| contr_add _ _]
rw [add_tensor_eq_snd <| add_tensor_eq_snd <| contr_add _ _]
/- Moving contraction through smul. -/
rw [add_tensor_eq_snd <| add_tensor_eq_fst <| contr_smul _ _]
rw [add_tensor_eq_snd <| add_tensor_eq_snd <| add_tensor_eq_fst <| contr_smul _ _]
rw [add_tensor_eq_snd <| add_tensor_eq_snd <| add_tensor_eq_snd <| contr_smul _ _]
simp only [tensorNode_tensor, add_tensor, smul_tensor]
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, neg_smul, one_smul]
conv =>
lhs; lhs;
rw [pauliMatrix_contr_down_0]
conv =>
lhs; rhs; lhs; rhs;
rw [pauliMatrix_contr_down_1]
conv =>
lhs; rhs; rhs; lhs; rhs;
rw [pauliMatrix_contr_down_2]
conv =>
lhs; rhs; rhs; rhs; rhs;
rw [pauliMatrix_contr_down_3]
simp only [neg_smul, one_smul]
abel
lemma pauliCo_basis_expand_tree : pauliCo
= (TensorTree.add (tensorNode
(basisVector pauliMatrixLowerMap (fun | 0 => 0 | 1 => 0 | 2 => 0))) <|
(basisVector pauliCoMap (fun | 0 => 0 | 1 => 0 | 2 => 0))) <|
TensorTree.add (tensorNode
(basisVector pauliMatrixLowerMap (fun | 0 => 0 | 1 => 1 | 2 => 1))) <|
(basisVector pauliCoMap (fun | 0 => 0 | 1 => 1 | 2 => 1))) <|
TensorTree.add (TensorTree.smul (-1) (tensorNode
(basisVector pauliMatrixLowerMap (fun | 0 => 1 | 1 => 0 | 2 => 1)))) <|
(basisVector pauliCoMap (fun | 0 => 1 | 1 => 0 | 2 => 1)))) <|
TensorTree.add (TensorTree.smul (-1) (tensorNode
(basisVector pauliMatrixLowerMap (fun | 0 => 1 | 1 => 1 | 2 => 0)))) <|
(basisVector pauliCoMap (fun | 0 => 1 | 1 => 1 | 2 => 0)))) <|
TensorTree.add (TensorTree.smul I (tensorNode
(basisVector pauliMatrixLowerMap (fun | 0 => 2 | 1 => 0 | 2 => 1)))) <|
(basisVector pauliCoMap (fun | 0 => 2 | 1 => 0 | 2 => 1)))) <|
TensorTree.add (TensorTree.smul (-I) (tensorNode
(basisVector pauliMatrixLowerMap (fun | 0 => 2 | 1 => 1 | 2 => 0)))) <|
(basisVector pauliCoMap (fun | 0 => 2 | 1 => 1 | 2 => 0)))) <|
TensorTree.add (TensorTree.smul (-1) (tensorNode
(basisVector pauliMatrixLowerMap (fun | 0 => 3 | 1 => 0 | 2 => 0)))) <|
(tensorNode (basisVector pauliMatrixLowerMap (fun | 0 => 3 | 1 => 1 | 2 => 1)))).tensor := by
rw [pauliMatrix_lower]
simp only [Nat.reduceAdd, Fin.isValue, add_tensor,
tensorNode_tensor, smul_tensor, neg_smul, one_smul]
(basisVector pauliCoMap (fun | 0 => 3 | 1 => 0 | 2 => 0)))) <|
(tensorNode (basisVector pauliCoMap (fun | 0 => 3 | 1 => 1 | 2 => 1)))).tensor := by
rw [pauliCo_basis_expand]
simp only [Nat.reduceAdd, Fin.isValue, add_tensor, tensorNode_tensor, smul_tensor, neg_smul,
one_smul]
rfl
lemma pauliMatrix_lower_basis_expand_prod {n : } {c : Fin n → complexLorentzTensor.C}
lemma pauliCo_prod_basis_expand {n : } {c : Fin n → complexLorentzTensor.C}
(t : TensorTree complexLorentzTensor c) :
(prod {Lorentz.coMetric | μ ν ⊗ PauliMatrix.asConsTensor | μ α β}ᵀ t).tensor =
(prod (tensorNode pauliCo) t).tensor =
(((tensorNode
(basisVector pauliMatrixLowerMap fun | 0 => 0 | 1 => 0 | 2 => 0)).prod t).add
(basisVector pauliCoMap fun | 0 => 0 | 1 => 0 | 2 => 0)).prod t).add
(((tensorNode
(basisVector pauliMatrixLowerMap fun | 0 => 0 | 1 => 1 | 2 => 1)).prod t).add
(basisVector pauliCoMap fun | 0 => 0 | 1 => 1 | 2 => 1)).prod t).add
((TensorTree.smul (-1) ((tensorNode
(basisVector pauliMatrixLowerMap fun | 0 => 1 | 1 => 0 | 2 => 1)).prod t)).add
(basisVector pauliCoMap fun | 0 => 1 | 1 => 0 | 2 => 1)).prod t)).add
((TensorTree.smul (-1) ((tensorNode
(basisVector pauliMatrixLowerMap fun | 0 => 1 | 1 => 1 | 2 => 0)).prod t)).add
(basisVector pauliCoMap fun | 0 => 1 | 1 => 1 | 2 => 0)).prod t)).add
((TensorTree.smul I ((tensorNode
(basisVector pauliMatrixLowerMap fun | 0 => 2 | 1 => 0 | 2 => 1)).prod t)).add
(basisVector pauliCoMap fun | 0 => 2 | 1 => 0 | 2 => 1)).prod t)).add
((TensorTree.smul (-I) ((tensorNode
(basisVector pauliMatrixLowerMap fun | 0 => 2 | 1 => 1 | 2 => 0)).prod t)).add
(basisVector pauliCoMap fun | 0 => 2 | 1 => 1 | 2 => 0)).prod t)).add
((TensorTree.smul (-1) ((tensorNode
(basisVector pauliMatrixLowerMap fun | 0 => 3 | 1 => 0 | 2 => 0)).prod t)).add
(basisVector pauliCoMap fun | 0 => 3 | 1 => 0 | 2 => 0)).prod t)).add
((tensorNode
(basisVector pauliMatrixLowerMap fun | 0 => 3 | 1 => 1 | 2 => 1)).prod
(basisVector pauliCoMap fun | 0 => 3 | 1 => 1 | 2 => 1)).prod
t)))))))).tensor := by
rw [prod_tensor_eq_fst <| pauliMatrix_lower_tree]
rw [prod_tensor_eq_fst <| tensoreNode_pauliCo]
rw [prod_tensor_eq_fst <| pauliCo_basis_expand_tree]
/- Moving the prod through additions. -/
rw [add_prod _ _ _]
rw [add_tensor_eq_snd <| add_prod _ _ _]

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