feat: Important fix to ContrContr
This commit is contained in:
parent
83ff8f5358
commit
4756466001
7 changed files with 184 additions and 2 deletions
|
@ -4,7 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE.
|
|||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.Tensors.Tree.NodeIdentities.ProdAssoc
|
||||
import HepLean.Tensors.Tree.NodeIdentities.ProdComm
|
||||
import HepLean.Tensors.Tree.NodeIdentities.ProdContr
|
||||
import HepLean.Tensors.Tree.NodeIdentities.ContrContr
|
||||
import HepLean.Tensors.Tree.NodeIdentities.ContrSwap
|
||||
import HepLean.Tensors.Tree.NodeIdentities.PermContr
|
||||
import HepLean.Tensors.Tree.NodeIdentities.Congr
|
||||
/-!
|
||||
|
||||
|
@ -75,4 +79,61 @@ lemma tensorNode_pauliContrDown : {pauliContrDown | μ α β}ᵀ.tensor =
|
|||
Fermion.altRightMetric | β β'}ᵀ.tensor := by
|
||||
rfl
|
||||
|
||||
/-!
|
||||
|
||||
## Basic equalities
|
||||
|
||||
-/
|
||||
|
||||
set_option maxRecDepth 5000 in
|
||||
/-- A rearanging of `pauliCoDown` to place the pauli matrices on the right. -/
|
||||
lemma pauliCoDown_eq_metric_mul_pauliCo :
|
||||
{pauliCoDown | μ α' β' = Fermion.altLeftMetric | α α' ⊗
|
||||
Fermion.altRightMetric | β β' ⊗ pauliCo | μ α β}ᵀ := by
|
||||
conv =>
|
||||
lhs
|
||||
rw [tensorNode_pauliCoDown]
|
||||
rw [contr_tensor_eq <| contr_prod _ _ _]
|
||||
rw [perm_contr]
|
||||
erw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| perm_eq_id _ rfl _]
|
||||
rw [perm_tensor_eq <| contr_congr 1 2]
|
||||
rw [perm_perm]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| contr_congr 1 2]
|
||||
rw [perm_tensor_eq <| perm_contr _ _]
|
||||
rw [perm_perm]
|
||||
rw [perm_tensor_eq <| contr_congr 1 2]
|
||||
rw [perm_perm]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| prod_assoc' _ _ _ _ _ _]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| perm_contr _ _]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| perm_tensor_eq <| contr_congr 1 2]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| perm_perm _ _ _]
|
||||
rw [perm_tensor_eq <| perm_contr _ _]
|
||||
rw [perm_perm]
|
||||
rw [perm_tensor_eq <| contr_congr 1 2]
|
||||
rw [perm_perm]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| prod_comm _ _ _ _]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| perm_contr _ _]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| perm_tensor_eq <| contr_congr 5 0]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| perm_perm _ _ _]
|
||||
rw [perm_tensor_eq <| perm_contr _ _]
|
||||
rw [perm_perm]
|
||||
rw [perm_tensor_eq <| contr_congr 4 1]
|
||||
rw [perm_perm]
|
||||
conv =>
|
||||
rhs
|
||||
rw [perm_tensor_eq <| contr_swap _ _]
|
||||
rw [perm_perm]
|
||||
erw [perm_tensor_eq <| contr_congr 4 1]
|
||||
rw [perm_perm]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| contr_swap _ _]
|
||||
erw [perm_tensor_eq <| contr_tensor_eq <| perm_tensor_eq <| contr_congr 5 0]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| perm_perm _ _ _]
|
||||
rw [perm_tensor_eq <| perm_contr _ _]
|
||||
rw [perm_perm]
|
||||
rw [perm_tensor_eq <| contr_congr 4 1]
|
||||
rw [perm_perm]
|
||||
apply perm_congr _ rfl
|
||||
decide
|
||||
|
||||
|
||||
end complexLorentzTensor
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue