51 lines
1.6 KiB
Text
51 lines
1.6 KiB
Text
/-
|
||
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.Lorentz.ComplexTensor.Basis
|
||
/-!
|
||
|
||
## Lemmas related to complex Lorentz tensors.
|
||
|
||
-/
|
||
open IndexNotation
|
||
open CategoryTheory
|
||
open MonoidalCategory
|
||
open Matrix
|
||
open MatrixGroups
|
||
open Complex
|
||
open TensorProduct
|
||
open IndexNotation
|
||
open CategoryTheory
|
||
open TensorTree
|
||
open OverColor.Discrete
|
||
noncomputable section
|
||
|
||
namespace complexLorentzTensor
|
||
|
||
set_option maxRecDepth 5000 in
|
||
lemma antiSymm_contr_symm {A : complexLorentzTensor.F.obj (OverColor.mk ![Color.up, Color.up])}
|
||
{S : complexLorentzTensor.F.obj (OverColor.mk ![Color.down, Color.down])}
|
||
(hA : {A | μ ν = - (A | ν μ)}ᵀ) (hs : {S | μ ν = S | ν μ}ᵀ) :
|
||
{A | μ ν ⊗ S | μ ν = - A | μ ν ⊗ S | μ ν}ᵀ := by
|
||
conv =>
|
||
lhs
|
||
rw [contr_tensor_eq <| contr_tensor_eq <| prod_tensor_eq_fst <| hA]
|
||
rw [contr_tensor_eq <| contr_tensor_eq <| prod_tensor_eq_snd <| hs]
|
||
rw [contr_tensor_eq <| contr_tensor_eq <| prod_perm_left _ _ _ _]
|
||
rw [contr_tensor_eq <| contr_tensor_eq <| perm_tensor_eq <| prod_perm_right _ _ _ _]
|
||
rw [contr_tensor_eq <| contr_tensor_eq <| perm_perm _ _ _]
|
||
rw [contr_tensor_eq <| perm_contr_congr 1 2]
|
||
rw [perm_contr_congr 0 0]
|
||
rw [perm_tensor_eq <| contr_contr _ _ _]
|
||
rw [perm_perm]
|
||
rw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| neg_fst_prod _ _]
|
||
rw [perm_tensor_eq <| contr_tensor_eq <| neg_contr _]
|
||
rw [perm_tensor_eq <| neg_contr _]
|
||
apply perm_congr _ rfl
|
||
decide
|
||
|
||
end complexLorentzTensor
|
||
|
||
end
|