2024-10-17 11:43:33 +00:00
|
|
|
|
/-
|
|
|
|
|
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
|
|
|
|
|
-/
|
2024-11-09 17:46:52 +00:00
|
|
|
|
import HepLean.Lorentz.ComplexTensor.Basis
|
2024-10-31 14:03:31 +00:00
|
|
|
|
import HepLean.Tensors.Tree.NodeIdentities.PermProd
|
2024-10-17 11:43:33 +00:00
|
|
|
|
/-!
|
|
|
|
|
|
|
|
|
|
## 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
|
|
|
|
|
|
2024-10-25 13:50:19 +00:00
|
|
|
|
namespace complexLorentzTensor
|
2024-10-21 05:38:22 +00:00
|
|
|
|
|
2024-10-31 14:03:31 +00:00
|
|
|
|
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])}
|
2024-10-21 14:05:54 +00:00
|
|
|
|
(hA : {A | μ ν = - (A | ν μ)}ᵀ) (hs : {S | μ ν = S | ν μ}ᵀ) :
|
2024-10-31 14:03:31 +00:00
|
|
|
|
{A | μ ν ⊗ S | μ ν = - A | μ ν ⊗ S | μ ν}ᵀ := by
|
|
|
|
|
conv =>
|
|
|
|
|
lhs
|
2024-10-31 15:48:18 +00:00
|
|
|
|
rw [contr_tensor_eq <| contr_tensor_eq <| prod_tensor_eq_fst <| hA]
|
|
|
|
|
rw [contr_tensor_eq <| contr_tensor_eq <| prod_tensor_eq_snd <| hs]
|
2024-10-31 14:03:31 +00:00
|
|
|
|
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]
|
2024-10-31 15:48:18 +00:00
|
|
|
|
rw [perm_contr_congr 0 0]
|
2024-10-31 14:03:31 +00:00
|
|
|
|
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
|
2024-10-22 11:49:58 +00:00
|
|
|
|
|
2024-10-25 13:50:19 +00:00
|
|
|
|
end complexLorentzTensor
|
2024-10-17 11:43:33 +00:00
|
|
|
|
|
|
|
|
|
end
|