feat: Lemmas regarding bispinors

This commit is contained in:
jstoobysmith 2024-10-28 14:05:38 +00:00
parent 7e29b5470d
commit 9cae20c114
5 changed files with 204 additions and 1 deletions

View file

@ -4,6 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith Authors: Joseph Tooby-Smith
-/ -/
import HepLean.Tensors.ComplexLorentz.PauliLower import HepLean.Tensors.ComplexLorentz.PauliLower
import HepLean.Tensors.Tree.NodeIdentities.ProdContr
import HepLean.Tensors.Tree.NodeIdentities.Congr
import HepLean.Tensors.Tree.NodeIdentities.ProdAssoc
/-! /-!
## Bispinors ## Bispinors
@ -33,6 +36,54 @@ lemma tensorNode_contrBispinorUp (p : complexContr) :
(tensorNode (contrBispinorUp p)).tensor = {p | μ ⊗ pauliCo | μ α β}ᵀ.tensor := by (tensorNode (contrBispinorUp p)).tensor = {p | μ ⊗ pauliCo | μ α β}ᵀ.tensor := by
rw [contrBispinorUp, tensorNode_tensor] rw [contrBispinorUp, tensorNode_tensor]
/-- An up-bispinor is equal to `pauliCo | μ α β ⊗ p | μ`-/
lemma contrBispinorUp_eq_pauliCo_self (p : complexContr) :
{contrBispinorUp p | α β = pauliCo | μ α β ⊗ p | μ}ᵀ := by
rw [tensorNode_contrBispinorUp]
conv_lhs =>
rw [contr_tensor_eq <| prod_comm _ _ _ _]
rw [perm_contr]
rw [perm_tensor_eq <| contr_swap _ _]
rw [perm_perm]
apply perm_congr
· apply OverColor.Hom.ext
ext x
match x with
| (0 : Fin 2) => rfl
| (1 : Fin 2) => rfl
· rfl
set_option maxRecDepth 2000 in
lemma altRightMetric_contr_contrBispinorUp_assoc (p : complexContr) :
{Fermion.altRightMetric | β β' ⊗ contrBispinorUp p | α β =
Fermion.altRightMetric | β β' ⊗ pauliCo | μ α β ⊗ p | μ}ᵀ := by
conv_lhs =>
rw [contr_tensor_eq <| prod_tensor_eq_snd <| contrBispinorUp_eq_pauliCo_self _]
rw [contr_tensor_eq <| prod_perm_right _ _ _ _]
rw [perm_contr]
rw [perm_tensor_eq <| contr_tensor_eq <| prod_contr _ _ _]
rw [perm_tensor_eq <| perm_contr _ _]
rw [perm_perm]
erw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| prod_assoc _ _ _ _ _ _]
rw [perm_tensor_eq <| contr_tensor_eq <| perm_contr _ _]
rw [perm_tensor_eq <| perm_contr _ _]
rw [perm_perm]
conv_rhs =>
rw [perm_tensor_eq <| contr_tensor_eq <| contr_prod _ _ _]
rw [perm_tensor_eq <| perm_contr _ _]
rw [perm_perm]
erw [perm_tensor_eq <| contr_tensor_eq <| perm_contr _ _]
rw [perm_tensor_eq <| perm_contr _ _]
rw [perm_perm]
rw [perm_tensor_eq <| contr_contr _ _ _]
rw [perm_perm]
apply perm_congr (_) rfl
· apply OverColor.Hom.fin_ext
intro i
fin_cases i
exact rfl
exact rfl
/-- A bispinor `pₐₐ` created from a lorentz vector `p^μ`. -/ /-- A bispinor `pₐₐ` created from a lorentz vector `p^μ`. -/
def contrBispinorDown (p : complexContr) := def contrBispinorDown (p : complexContr) :=
{Fermion.altLeftMetric | α α' ⊗ Fermion.altRightMetric | β β' ⊗ {Fermion.altLeftMetric | α α' ⊗ Fermion.altRightMetric | β β' ⊗
@ -45,9 +96,70 @@ lemma tensorNode_contrBispinorDown (p : complexContr) :
Fermion.altRightMetric | β β' ⊗ (contrBispinorUp p) | α β}ᵀ.tensor := by Fermion.altRightMetric | β β' ⊗ (contrBispinorUp p) | α β}ᵀ.tensor := by
rw [contrBispinorDown, tensorNode_tensor] rw [contrBispinorDown, tensorNode_tensor]
set_option maxRecDepth 10000 in
lemma contrBispinorDown_eq_metric_contr_contrBispinorUp (p : complexContr) :
{contrBispinorDown p | α' β' = Fermion.altLeftMetric | α α' ⊗
(Fermion.altRightMetric | β β' ⊗ contrBispinorUp p | α β)}ᵀ := by
rw [tensorNode_contrBispinorDown]
conv_lhs =>
rw [contr_tensor_eq <| contr_tensor_eq <| prod_assoc' _ _ _ _ _ _]
rw [contr_tensor_eq <| perm_contr _ _]
rw [perm_contr]
rw [perm_tensor_eq <| contr_contr _ _ _]
rw [perm_perm]
conv_rhs =>
rw [perm_tensor_eq <| contr_tensor_eq <| prod_contr _ _ _ ]
rw [perm_tensor_eq <| perm_contr _ _]
rw [perm_perm]
apply perm_congr
· apply OverColor.Hom.ext
ext x
match x with
| (0 : Fin 2) => rfl
| (1 : Fin 2) => rfl
· rfl
set_option maxHeartbeats 400000 in
set_option maxRecDepth 2000 in
lemma contrBispinorDown_eq_contr_with_self (p : complexContr) :
{contrBispinorDown p | α' β' = (Fermion.altLeftMetric | α α' ⊗
(Fermion.altRightMetric | β β' ⊗ pauliCo | μ α β)) ⊗ p | μ}ᵀ := by
rw [contrBispinorDown_eq_metric_contr_contrBispinorUp]
conv_lhs =>
rw [perm_tensor_eq <| contr_tensor_eq <| prod_tensor_eq_snd <| altRightMetric_contr_contrBispinorUp_assoc _]
rw [perm_tensor_eq <| contr_tensor_eq <| prod_perm_right _ _ _ _]
rw [perm_tensor_eq <| perm_contr _ _ ]
rw [perm_perm]
rw [perm_tensor_eq <| contr_tensor_eq <| prod_contr _ _ _]
rw [perm_tensor_eq <| perm_contr _ _]
rw [perm_perm]
erw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <|
prod_assoc _ _ _ _ _ _]
rw [perm_tensor_eq <| contr_tensor_eq <| perm_contr _ _]
rw [perm_tensor_eq <| perm_contr _ _]
rw [perm_perm]
conv =>
rhs
rw [perm_tensor_eq <| contr_tensor_eq <| contr_prod _ _ _]
rw [perm_tensor_eq <| perm_contr _ _]
rw [perm_perm]
erw [perm_tensor_eq <| contr_tensor_eq <| perm_contr _ _]
rw [perm_tensor_eq <| perm_contr _ _]
rw [perm_perm]
rw [perm_tensor_eq <| contr_contr _ _ _]
rw [perm_perm]
apply congrArg
apply congrFun
apply congrArg
apply OverColor.Hom.fin_ext
intro i
fin_cases i
exact rfl
exact rfl
/-- Expansion of a `contrBispinorDown` into the original contravariant tensor nested /-- Expansion of a `contrBispinorDown` into the original contravariant tensor nested
between pauli matrices and metrics. -/ between pauli matrices and metrics. -/
lemma contrBispinorDown_full_nested (p : complexContr) : lemma contrBispinorDown_eq_metric_mul_self_mul_pauli (p : complexContr) :
{contrBispinorDown p | α β}ᵀ.tensor = {Fermion.altLeftMetric | α α' ⊗ {contrBispinorDown p | α β}ᵀ.tensor = {Fermion.altLeftMetric | α α' ⊗
Fermion.altRightMetric | β β' ⊗ (p | μ ⊗ pauliCo | μ α β)}ᵀ.tensor := by Fermion.altRightMetric | β β' ⊗ (p | μ ⊗ pauliCo | μ α β)}ᵀ.tensor := by
conv => conv =>

View file

@ -45,6 +45,7 @@ variable {C : Type} {f g h : OverColor C}
lemma ext (m n : f ⟶ g) (h : m.hom = n.hom) : m = n := by lemma ext (m n : f ⟶ g) (h : m.hom = n.hom) : m = n := by
apply CategoryTheory.Iso.ext h apply CategoryTheory.Iso.ext h
/-- Given a hom in `OverColor C` the underlying equivalence between types. -/ /-- Given a hom in `OverColor C` the underlying equivalence between types. -/
def toEquiv (m : f ⟶ g) : f.left ≃ g.left where def toEquiv (m : f ⟶ g) : f.left ≃ g.left where
toFun := m.hom.left toFun := m.hom.left
@ -258,6 +259,12 @@ def mk (f : X → C) : OverColor C := Over.mk f
lemma mk_hom (f : X → C) : (mk f).hom = f := rfl lemma mk_hom (f : X → C) : (mk f).hom = f := rfl
open MonoidalCategory open MonoidalCategory
lemma Hom.fin_ext {n : } {f g : Fin n → C} (σ σ' : OverColor.mk f ⟶ OverColor.mk g)
(h : ∀ (i : Fin n), σ.hom.left i = σ'.hom.left i) : σ = σ' := by
apply Hom.ext
ext i
apply h
end OverColor end OverColor
end IndexNotation end IndexNotation

View file

@ -61,6 +61,10 @@ def mkIso {c1 c2 : X → C} (h : c1 = c2) : mk c1 ≅ mk c2 :=
subst h subst h
rfl)) rfl))
lemma mkIso_refl_hom {c : X → C} : (mkIso (by rfl : c =c)).hom = 𝟙 _ := by
simp [mkIso]
rfl
@[simp] @[simp]
lemma equivToIso_mkIso_hom {c1 c2 : X → C} (h : c1 = c2) : lemma equivToIso_mkIso_hom {c1 c2 : X → C} (h : c1 = c2) :
Hom.toEquiv (mkIso h).hom = Equiv.refl _ := by Hom.toEquiv (mkIso h).hom = Equiv.refl _ := by

View file

@ -115,6 +115,33 @@ lemma perm_eq_of_eq_perm {n m : } {c : Fin n → S.C} {c1 : Fin m → S.C}
change _ = (S.F.map σ.hom ≫ S.F.map σ.inv).hom _ change _ = (S.F.map σ.hom ≫ S.F.map σ.inv).hom _
simp only [Iso.map_hom_inv_id, Action.id_hom, ModuleCat.id_apply] simp only [Iso.map_hom_inv_id, Action.id_hom, ModuleCat.id_apply]
lemma perm_eq_iff_eq_perm {n m : } {c : Fin n → S.C} {c1 : Fin m → S.C}
(σ : (OverColor.mk c) ⟶ (OverColor.mk c1))
{t : TensorTree S c} {t2 : TensorTree S c1} :
(perm σ t).tensor = t2.tensor ↔ t.tensor =
(perm (equivToHomEq (Hom.toEquiv σ).symm (fun x => Hom.toEquiv_comp_apply σ x)) t2).tensor := by
refine Iff.intro (fun h => ?_) (fun h => ?_)
· simp [perm_tensor, ← h]
change _ = (S.F.map _ ≫ S.F.map _).hom _
rw [← S.F.map_comp]
have h1 : (σ ≫ equivToHomEq (Hom.toEquiv σ).symm (fun x => Hom.toEquiv_comp_apply σ x)) = 𝟙 _ := by
apply Hom.ext
ext x
change (Hom.toEquiv σ).symm ((Hom.toEquiv σ) x) = x
simp
rw [h1]
simp
· rw [perm_tensor, h]
change (S.F.map _ ≫ S.F.map _).hom _ = _
rw [← S.F.map_comp]
have h1 : (equivToHomEq (Hom.toEquiv σ).symm (fun x => Hom.toEquiv_comp_apply σ x) ≫ σ) = 𝟙 _ := by
apply Hom.ext
ext x
change (Hom.toEquiv σ) ((Hom.toEquiv σ).symm x) = x
simp
rw [h1]
simp
/-! /-!
## Vector based identities ## Vector based identities

View file

@ -0,0 +1,53 @@
/-
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
/-!
## Congr results
-/
open IndexNotation
open CategoryTheory
open MonoidalCategory
open OverColor
open HepLean.Fin
open TensorProduct
namespace TensorTree
variable {S : TensorSpecies}
/-!
## Congr results
-/
variable {n m : }
lemma perm_congr {c1 : Fin n → S.C} {c2 : Fin m → S.C} {T T' : TensorTree S c1}
{σ σ' : OverColor.mk c1 ⟶ OverColor.mk c2}
(h : σ = σ') (hT : T.tensor = T'.tensor):
(perm σ T).tensor = (perm σ' T').tensor := by
rw [h]
simp only [perm_tensor, hT]
lemma perm_update {c1 : Fin n → S.C} {c2 : Fin m → S.C} {T : TensorTree S c1}
{σ : OverColor.mk c1 ⟶ OverColor.mk c2} (σ' : OverColor.mk c1 ⟶ OverColor.mk c2)
(h : σ = σ') :
(perm σ T).tensor = (perm σ' T).tensor := by rw [h]
lemma contr_congr {n : } {c : Fin n.succ.succ → S.C} {i : Fin n.succ.succ}
(i' : Fin n.succ.succ) {j : Fin n.succ} (j' : Fin n.succ)
{h : c (i.succAbove j) = S.τ (c i)} {t : TensorTree S c}
(hi : i = i' := by decide) (hj : j = j' := by decide)
:(contr i j h t).tensor =
(perm (mkIso (by rw [hi, hj])).hom (contr i' j' (by rw [← hi, ← hj, h]) t)).tensor := by
subst hi
subst hj
simp [perm_tensor, mkIso_refl_hom]
end TensorTree