Merge branch 'master' into IndexNotation

This commit is contained in:
jstoobysmith 2024-10-16 12:02:26 +00:00
commit d9f6760541
23 changed files with 1855 additions and 272 deletions

View file

@ -82,16 +82,27 @@ import HepLean.SpaceTime.LorentzTensor.Real.Basic
import HepLean.SpaceTime.LorentzTensor.Real.IndexNotation
import HepLean.SpaceTime.LorentzVector.AsSelfAdjointMatrix
import HepLean.SpaceTime.LorentzVector.Basic
import HepLean.SpaceTime.LorentzVector.Complex
import HepLean.SpaceTime.LorentzVector.Complex.Basic
import HepLean.SpaceTime.LorentzVector.Complex.Contraction
import HepLean.SpaceTime.LorentzVector.Complex.Metric
import HepLean.SpaceTime.LorentzVector.Complex.Two
import HepLean.SpaceTime.LorentzVector.Complex.Unit
import HepLean.SpaceTime.LorentzVector.Contraction
import HepLean.SpaceTime.LorentzVector.Covariant
import HepLean.SpaceTime.LorentzVector.LorentzAction
import HepLean.SpaceTime.LorentzVector.Modules
import HepLean.SpaceTime.LorentzVector.NormOne
import HepLean.SpaceTime.MinkowskiMetric
import HepLean.SpaceTime.PauliMatrices.AsTensor
import HepLean.SpaceTime.PauliMatrices.Basic
import HepLean.SpaceTime.PauliMatrices.SelfAdjoint
import HepLean.SpaceTime.SL2C.Basic
import HepLean.SpaceTime.WeylFermion.Basic
import HepLean.SpaceTime.WeylFermion.Contraction
import HepLean.SpaceTime.WeylFermion.Metric
import HepLean.SpaceTime.WeylFermion.Modules
import HepLean.SpaceTime.WeylFermion.Two
import HepLean.SpaceTime.WeylFermion.Unit
import HepLean.StandardModel.Basic
import HepLean.StandardModel.HiggsBoson.Basic
import HepLean.StandardModel.HiggsBoson.GaugeAction

View file

@ -36,7 +36,7 @@ open minkowskiMetric in
/-- The Lorentz group is the subset of matrices which preserve the minkowski metric. -/
def LorentzGroup (d : ) : Set (Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) :=
{Λ : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) |
∀ (x y : LorentzVector d), ⟪Λ *ᵥ x, Λ *ᵥ y⟫ₘ = ⟪x, y⟫ₘ}
∀ (x y : LorentzVector d), ⟪Λ *ᵥ x, Λ *ᵥ y⟫ₘ = ⟪x, y⟫ₘ}
namespace LorentzGroup
/-- Notation for the Lorentz group. -/
@ -286,5 +286,69 @@ lemma timeComp_mul (Λ Λ' : LorentzGroup d) : timeComp (Λ * Λ') =
erw [Pi.basisFun_apply, Matrix.mulVec_single_one]
simp
/-!
## To Complex matrices
-/
/-- The monoid homomorphisms taking the lorentz group to complex matrices. -/
def toComplex : LorentzGroup d →* Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) where
toFun Λ := Λ.1.map ofReal
map_one' := by
ext i j
simp only [lorentzGroupIsGroup_one_coe, map_apply, ofReal_eq_coe]
simp only [Matrix.one_apply, ofReal_one, ofReal_zero]
split_ifs
· rfl
· rfl
map_mul' Λ Λ' := by
ext i j
simp only [lorentzGroupIsGroup_mul_coe, map_apply, ofReal_eq_coe]
simp only [← Matrix.map_mul, RingHom.map_matrix_mul]
rfl
instance (M : LorentzGroup d) : Invertible (toComplex M) where
invOf := toComplex M⁻¹
invOf_mul_self := by
rw [← toComplex.map_mul, Group.inv_mul_cancel]
simp
mul_invOf_self := by
rw [← toComplex.map_mul]
rw [@mul_inv_cancel]
simp
lemma toComplex_inv (Λ : LorentzGroup d) : (toComplex Λ)⁻¹ = toComplex Λ⁻¹ := by
refine inv_eq_right_inv ?h
rw [← toComplex.map_mul, mul_inv_cancel]
simp
@[simp]
lemma toComplex_mul_minkowskiMatrix_mul_transpose (Λ : LorentzGroup d) :
toComplex Λ * minkowskiMatrix.map ofReal * (toComplex Λ)ᵀ = minkowskiMatrix.map ofReal := by
simp only [toComplex, MonoidHom.coe_mk, OneHom.coe_mk]
have h1 : ((Λ.1).map ⇑ofReal)ᵀ = (Λ.1ᵀ).map ofReal := rfl
rw [h1]
trans (Λ.1 * minkowskiMatrix * Λ.1ᵀ).map ofReal
· simp only [Matrix.map_mul]
simp only [mul_minkowskiMatrix_mul_transpose]
@[simp]
lemma toComplex_transpose_mul_minkowskiMatrix_mul_self (Λ : LorentzGroup d) :
(toComplex Λ)ᵀ * minkowskiMatrix.map ofReal * toComplex Λ = minkowskiMatrix.map ofReal := by
simp only [toComplex, MonoidHom.coe_mk, OneHom.coe_mk]
have h1 : ((Λ.1).map ⇑ofReal)ᵀ = (Λ.1ᵀ).map ofReal := rfl
rw [h1]
trans (Λ.1ᵀ * minkowskiMatrix * Λ.1).map ofReal
· simp only [Matrix.map_mul]
simp only [transpose_mul_minkowskiMatrix_mul_self]
lemma toComplex_mulVec_ofReal (v : Fin 1 ⊕ Fin d → ) (Λ : LorentzGroup d) :
toComplex Λ *ᵥ (ofReal ∘ v) = ofReal ∘ (Λ *ᵥ v) := by
simp only [toComplex, MonoidHom.coe_mk, OneHom.coe_mk]
funext i
rw [← RingHom.map_mulVec]
rfl
end
end LorentzGroup

View file

@ -4,7 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.MinkowskiMetric
import HepLean.SpaceTime.PauliMatrices.SelfAdjoint
import Mathlib.LinearAlgebra.Matrix.SpecialLinearGroup
import Mathlib.Tactic.Polyrith
/-!
# Lorentz vector as a self-adjoint matrix
@ -20,150 +22,86 @@ namespace SpaceTime
open Matrix
open MatrixGroups
open Complex
/-- A 2×2-complex matrix formed from a Lorentz vector point. -/
@[simp]
def toMatrix (x : LorentzVector 3) : Matrix (Fin 2) (Fin 2) :=
!![x.time + x.space 2, x.space 0 - x.space 1 * I; x.space 0 + x.space 1 * I, x.time - x.space 2]
/-- The matrix `x.toMatrix` for `x ∈ spaceTime` is self adjoint. -/
lemma toMatrix_isSelfAdjoint (x : LorentzVector 3) : IsSelfAdjoint (toMatrix x) := by
rw [isSelfAdjoint_iff, star_eq_conjTranspose, ← Matrix.ext_iff]
intro i j
fin_cases i <;> fin_cases j <;>
simp only [toMatrix, LorentzVector.time, Fin.isValue, LorentzVector.space, Function.comp_apply,
Fin.zero_eta, conjTranspose_apply, of_apply, cons_val', cons_val_zero, empty_val',
cons_val_fin_one, star_add, RCLike.star_def, conj_ofReal, Fin.mk_one, cons_val_one,
head_fin_const, star_mul', conj_I, mul_neg, head_cons, star_sub, sub_neg_eq_add]
rfl
/-- A self-adjoint matrix formed from a Lorentz vector point. -/
@[simps!]
def toSelfAdjointMatrix' (x : LorentzVector 3) : selfAdjoint (Matrix (Fin 2) (Fin 2) ) :=
⟨toMatrix x, toMatrix_isSelfAdjoint x⟩
/-- A self-adjoint matrix formed from a Lorentz vector point. -/
@[simp]
noncomputable def fromSelfAdjointMatrix' (x : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
LorentzVector 3 := fun i =>
match i with
| Sum.inl 0 => 1/2 * (x.1 0 0 + x.1 1 1).re
| Sum.inr 0 => (x.1 1 0).re
| Sum.inr 1 => (x.1 1 0).im
| Sum.inr 2 => 1/2 * (x.1 0 0 - x.1 1 1).re
noncomputable section
/-- The linear equivalence between the vector-space `spaceTime` and self-adjoint
2×2-complex matrices. -/
noncomputable def toSelfAdjointMatrix :
LorentzVector 3 ≃ₗ[] selfAdjoint (Matrix (Fin 2) (Fin 2) ) where
toFun := toSelfAdjointMatrix'
invFun := fromSelfAdjointMatrix'
left_inv x := by
funext i
match i with
| Sum.inl 0 =>
simp only [fromSelfAdjointMatrix', one_div, toSelfAdjointMatrix', toMatrix,
LorentzVector.time, Fin.isValue, LorentzVector.space, Function.comp_apply, of_apply,
cons_val', cons_val_zero, empty_val', cons_val_fin_one, cons_val_one, head_cons,
head_fin_const, add_add_sub_cancel, add_re, ofReal_re]
ring_nf
| Sum.inr 0 =>
simp [fromSelfAdjointMatrix', toSelfAdjointMatrix', toMatrix, toMatrix_isSelfAdjoint]
| Sum.inr 1 =>
simp only [fromSelfAdjointMatrix', toSelfAdjointMatrix', toMatrix, LorentzVector.time,
Fin.isValue, LorentzVector.space, Function.comp_apply, of_apply, cons_val', cons_val_zero,
empty_val', cons_val_fin_one, cons_val_one, head_fin_const, add_im, ofReal_im, mul_im,
ofReal_re, I_im, mul_one, I_re, mul_zero, add_zero, zero_add]
| Sum.inr 2 =>
simp only [fromSelfAdjointMatrix', one_div, toSelfAdjointMatrix', toMatrix,
LorentzVector.time, Fin.isValue, LorentzVector.space, Function.comp_apply, of_apply,
cons_val', cons_val_zero, empty_val', cons_val_fin_one, cons_val_one, head_cons,
head_fin_const, add_sub_sub_cancel, add_re, ofReal_re]
ring
right_inv x := by
simp only [toSelfAdjointMatrix', toMatrix, LorentzVector.time, fromSelfAdjointMatrix', one_div,
Fin.isValue, add_re, ofReal_mul, ofReal_inv, ofReal_ofNat, ofReal_add, LorentzVector.space,
Function.comp_apply, sub_re, ofReal_sub, re_add_im]
ext i j
fin_cases i <;> fin_cases j <;>
field_simp [fromSelfAdjointMatrix', toMatrix, conj_ofReal]
· exact conj_eq_iff_re.mp (congrArg (fun M => M 0 0) $ selfAdjoint.mem_iff.mp x.2)
· have h01 := congrArg (fun M => M 0 1) $ selfAdjoint.mem_iff.mp x.2
simp only [Fin.isValue, star_apply, RCLike.star_def] at h01
rw [← h01, RCLike.conj_eq_re_sub_im]
rfl
· exact conj_eq_iff_re.mp (congrArg (fun M => M 1 1) $ selfAdjoint.mem_iff.mp x.2)
map_add' x y := by
ext i j : 2
simp only [toSelfAdjointMatrix'_coe, add_apply, ofReal_add, of_apply, cons_val', empty_val',
cons_val_fin_one, AddSubmonoid.coe_add, AddSubgroup.coe_toAddSubmonoid, Matrix.add_apply]
fin_cases i <;> fin_cases j
· rw [show (x + y) (Sum.inl 0) = x (Sum.inl 0) + y (Sum.inl 0) from rfl]
rw [show (x + y) (Sum.inr 2) = x (Sum.inr 2) + y (Sum.inr 2) from rfl]
simp only [Fin.isValue, ofReal_add, Fin.zero_eta, cons_val_zero]
simp only [Fin.isValue, toSelfAdjointMatrix', toMatrix, LorentzVector.time,
LorentzVector.space, Function.comp_apply, AddMemClass.mk_add_mk, of_add_of, add_cons,
head_cons, tail_cons, empty_add_empty, of_apply, cons_val', cons_val_zero, empty_val',
cons_val_fin_one]
ring
· rw [show (x + y) (Sum.inr 0) = x (Sum.inr 0) + y (Sum.inr 0) from rfl]
rw [show (x + y) (Sum.inr 1) = x (Sum.inr 1) + y (Sum.inr 1) from rfl]
simp only [Fin.isValue, ofReal_add, Fin.mk_one, cons_val_one, head_cons, Fin.zero_eta,
cons_val_zero]
simp only [Fin.isValue, toSelfAdjointMatrix', toMatrix, LorentzVector.time,
LorentzVector.space, Function.comp_apply, AddMemClass.mk_add_mk, of_add_of, add_cons,
head_cons, tail_cons, empty_add_empty, of_apply, cons_val', cons_val_one, empty_val',
cons_val_fin_one, cons_val_zero]
ring
· rw [show (x + y) (Sum.inr 0) = x (Sum.inr 0) + y (Sum.inr 0) from rfl]
rw [show (x + y) (Sum.inr 1) = x (Sum.inr 1) + y (Sum.inr 1) from rfl]
simp only [Fin.isValue, ofReal_add, Fin.zero_eta, cons_val_zero, Fin.mk_one, cons_val_one,
head_fin_const]
simp only [Fin.isValue, toSelfAdjointMatrix', toMatrix, LorentzVector.time,
LorentzVector.space, Function.comp_apply, AddMemClass.mk_add_mk, of_add_of, add_cons,
head_cons, tail_cons, empty_add_empty, of_apply, cons_val', cons_val_zero, empty_val',
cons_val_fin_one, cons_val_one, head_fin_const]
ring
· rw [show (x + y) (Sum.inl 0) = x (Sum.inl 0) + y (Sum.inl 0) from rfl]
rw [show (x + y) (Sum.inr 2) = x (Sum.inr 2) + y (Sum.inr 2) from rfl]
simp only [Fin.isValue, ofReal_add, Fin.mk_one, cons_val_one, head_cons, head_fin_const]
simp only [Fin.isValue, toSelfAdjointMatrix', toMatrix, LorentzVector.time,
LorentzVector.space, Function.comp_apply, AddMemClass.mk_add_mk, of_add_of, add_cons,
head_cons, tail_cons, empty_add_empty, of_apply, cons_val', cons_val_one, empty_val',
cons_val_fin_one, head_fin_const]
ring
map_smul' r x := by
ext i j : 2
simp only [toSelfAdjointMatrix'_coe, Fin.isValue, of_apply, cons_val', empty_val',
cons_val_fin_one, RingHom.id_apply, selfAdjoint.val_smul, smul_apply, real_smul]
fin_cases i <;> fin_cases j
· rw [show (r • x) (Sum.inl 0) = r * x (Sum.inl 0) from rfl]
rw [show (r • x) (Sum.inr 2) = r * x (Sum.inr 2) from rfl]
simp only [Fin.isValue, ofReal_mul, Fin.zero_eta, cons_val_zero]
ring
· rw [show (r • x) (Sum.inr 0) = r * x (Sum.inr 0) from rfl]
rw [show (r • x) (Sum.inr 1) = r * x (Sum.inr 1) from rfl]
simp only [Fin.isValue, ofReal_mul, Fin.mk_one, cons_val_one, head_cons, Fin.zero_eta,
cons_val_zero]
ring
· rw [show (r • x) (Sum.inr 0) = r * x (Sum.inr 0) from rfl]
rw [show (r • x) (Sum.inr 1) = r * x (Sum.inr 1) from rfl]
simp only [Fin.isValue, ofReal_mul, Fin.zero_eta, cons_val_zero, Fin.mk_one, cons_val_one,
head_fin_const]
ring
· rw [show (r • x) (Sum.inl 0) = r * x (Sum.inl 0) from rfl]
rw [show (r • x) (Sum.inr 2) = r * x (Sum.inr 2) from rfl]
simp only [Fin.isValue, ofReal_mul, Fin.mk_one, cons_val_one, head_cons, head_fin_const]
ring
`2×2`-complex matrices. -/
def toSelfAdjointMatrix : LorentzVector 3 ≃ₗ[] selfAdjoint (Matrix (Fin 2) (Fin 2) ) :=
(Finsupp.linearEquivFunOnFinite (Fin 1 ⊕ Fin 3)).symm ≪≫ₗ PauliMatrix.σSAL.repr.symm
lemma toSelfAdjointMatrix_apply (x : LorentzVector 3) : toSelfAdjointMatrix x =
x (Sum.inl 0) • ⟨PauliMatrix.σ0, PauliMatrix.σ0_selfAdjoint⟩
- x (Sum.inr 0) • ⟨PauliMatrix.σ1, PauliMatrix.σ1_selfAdjoint⟩
- x (Sum.inr 1) • ⟨PauliMatrix.σ2, PauliMatrix.σ2_selfAdjoint⟩
- x (Sum.inr 2) • ⟨PauliMatrix.σ3, PauliMatrix.σ3_selfAdjoint⟩ := by
simp only [toSelfAdjointMatrix, PauliMatrix.σSAL, LinearEquiv.trans_apply, Basis.repr_symm_apply,
Basis.coe_mk, Fin.isValue]
rw [Finsupp.linearCombination_apply_of_mem_supported (s := Finset.univ)]
· change (∑ i : Fin 1 ⊕ Fin 3, x i • PauliMatrix.σSAL' i) = _
simp only [PauliMatrix.σSAL', Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero,
Fin.isValue, Finset.sum_singleton, Fin.sum_univ_three]
apply Subtype.ext
simp only [Fin.isValue, AddSubgroup.coe_add, selfAdjoint.val_smul, smul_neg,
AddSubgroupClass.coe_sub]
simp only [neg_add, add_assoc, sub_eq_add_neg]
· simp_all only [Finset.coe_univ, Finsupp.supported_univ, Submodule.mem_top]
lemma toSelfAdjointMatrix_apply_coe (x : LorentzVector 3) : (toSelfAdjointMatrix x).1 =
x (Sum.inl 0) • PauliMatrix.σ0
- x (Sum.inr 0) • PauliMatrix.σ1
- x (Sum.inr 1) • PauliMatrix.σ2
- x (Sum.inr 2) • PauliMatrix.σ3 := by
rw [toSelfAdjointMatrix_apply]
simp only [Fin.isValue, AddSubgroupClass.coe_sub, selfAdjoint.val_smul]
lemma toSelfAdjointMatrix_stdBasis (i : Fin 1 ⊕ Fin 3) :
toSelfAdjointMatrix (LorentzVector.stdBasis i) = PauliMatrix.σSAL i := by
rw [toSelfAdjointMatrix_apply]
match i with
| Sum.inl 0 =>
simp only [LorentzVector.stdBasis, Fin.isValue]
erw [Pi.basisFun_apply]
simp [PauliMatrix.σSAL, PauliMatrix.σSAL']
| Sum.inr 0 =>
simp only [LorentzVector.stdBasis, Fin.isValue]
erw [Pi.basisFun_apply]
simp only [Fin.isValue, ne_eq, reduceCtorEq, not_false_eq_true, Pi.single_eq_of_ne, zero_smul,
Pi.single_eq_same, one_smul, zero_sub, Sum.inr.injEq, one_ne_zero, sub_zero, Fin.reduceEq,
PauliMatrix.σSAL, Basis.coe_mk, PauliMatrix.σSAL']
refine Eq.symm (PauliMatrix.selfAdjoint_ext rfl rfl rfl rfl)
| Sum.inr 1 =>
simp only [LorentzVector.stdBasis, Fin.isValue]
erw [Pi.basisFun_apply]
simp only [Fin.isValue, ne_eq, reduceCtorEq, not_false_eq_true, Pi.single_eq_of_ne, zero_smul,
Sum.inr.injEq, zero_ne_one, sub_self, Pi.single_eq_same, one_smul, zero_sub, Fin.reduceEq,
sub_zero, PauliMatrix.σSAL, Basis.coe_mk, PauliMatrix.σSAL']
refine Eq.symm (PauliMatrix.selfAdjoint_ext rfl rfl rfl rfl)
| Sum.inr 2 =>
simp only [LorentzVector.stdBasis, Fin.isValue]
erw [Pi.basisFun_apply]
simp only [Fin.isValue, ne_eq, reduceCtorEq, not_false_eq_true, Pi.single_eq_of_ne, zero_smul,
Sum.inr.injEq, Fin.reduceEq, sub_self, Pi.single_eq_same, one_smul, zero_sub,
PauliMatrix.σSAL, Basis.coe_mk, PauliMatrix.σSAL']
refine Eq.symm (PauliMatrix.selfAdjoint_ext rfl rfl rfl rfl)
@[simp]
lemma toSelfAdjointMatrix_symm_basis (i : Fin 1 ⊕ Fin 3) :
toSelfAdjointMatrix.symm (PauliMatrix.σSAL i) = (LorentzVector.stdBasis i) := by
refine (LinearEquiv.symm_apply_eq toSelfAdjointMatrix).mpr ?_
rw [toSelfAdjointMatrix_stdBasis]
open minkowskiMetric in
lemma det_eq_ηLin (x : LorentzVector 3) : det (toSelfAdjointMatrix x).1 = ⟪x, x⟫ₘ := by
simp only [toSelfAdjointMatrix, LinearEquiv.coe_mk, toSelfAdjointMatrix'_coe, Fin.isValue,
det_fin_two_of, eq_time_minus_inner_prod, LorentzVector.time, LorentzVector.space,
rw [toSelfAdjointMatrix_apply_coe]
simp only [Fin.isValue, eq_time_minus_inner_prod, LorentzVector.time, LorentzVector.space,
PiLp.inner_apply, Function.comp_apply, RCLike.inner_apply, conj_trivial, Fin.sum_univ_three,
ofReal_sub, ofReal_mul, ofReal_add]
simp only [Fin.isValue, PauliMatrix.σ0, smul_of, smul_cons, real_smul, mul_one, smul_zero,
smul_empty, PauliMatrix.σ1, of_sub_of, sub_cons, head_cons, sub_zero, tail_cons, zero_sub,
sub_self, zero_empty, PauliMatrix.σ2, smul_neg, sub_neg_eq_add, PauliMatrix.σ3, det_fin_two_of]
ring_nf
simp only [Fin.isValue, I_sq, mul_neg, mul_one]
ring
end
end SpaceTime

View file

@ -1,39 +0,0 @@
/-
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 Mathlib.Data.Complex.Exponential
import Mathlib.Analysis.InnerProductSpace.PiL2
import HepLean.SpaceTime.SL2C.Basic
import HepLean.SpaceTime.LorentzVector.Modules
import HepLean.Meta.Informal
import Mathlib.RepresentationTheory.Rep
import HepLean.Tensors.Basic
/-!
# Complex Lorentz vectors
We define complex Lorentz vectors in 4d space-time as representations of SL(2, C).
-/
noncomputable section
open Matrix
open MatrixGroups
open Complex
open TensorProduct
namespace Lorentz
/-- The representation of `SL(2, )` on complex vectors corresponding to contravariant
Lorentz vectors. -/
def complexContr : Rep SL(2, ) := Rep.of ContrModule.SL2CRep
/-- The representation of `SL(2, )` on complex vectors corresponding to contravariant
Lorentz vectors. -/
def complexCo : Rep SL(2, ) := Rep.of CoModule.SL2CRep
end Lorentz
end

View file

@ -0,0 +1,132 @@
/-
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 Mathlib.Data.Complex.Exponential
import Mathlib.Analysis.InnerProductSpace.PiL2
import HepLean.SpaceTime.SL2C.Basic
import HepLean.SpaceTime.LorentzVector.Modules
import HepLean.Meta.Informal
import Mathlib.RepresentationTheory.Rep
import HepLean.Tensors.Basic
import HepLean.SpaceTime.PauliMatrices.SelfAdjoint
/-!
# Complex Lorentz vectors
We define complex Lorentz vectors in 4d space-time as representations of SL(2, C).
-/
noncomputable section
open Matrix
open MatrixGroups
open Complex
open TensorProduct
open SpaceTime
namespace Lorentz
/-- The representation of `SL(2, )` on complex vectors corresponding to contravariant
Lorentz vectors. In index notation these have an up index `ψⁱ`. -/
def complexContr : Rep SL(2, ) := Rep.of ContrModule.SL2CRep
/-- The representation of `SL(2, )` on complex vectors corresponding to contravariant
Lorentz vectors. In index notation these have a down index `ψⁱ`. -/
def complexCo : Rep SL(2, ) := Rep.of CoModule.SL2CRep
/-- The standard basis of complex contravariant Lorentz vectors. -/
def complexContrBasis : Basis (Fin 1 ⊕ Fin 3) complexContr := Basis.ofEquivFun
(Equiv.linearEquiv ContrModule.toFin13Fun)
@[simp]
lemma complexContrBasis_ρ_apply (M : SL(2,)) (i j : Fin 1 ⊕ Fin 3) :
(LinearMap.toMatrix complexContrBasis complexContrBasis) (complexContr.ρ M) i j =
(LorentzGroup.toComplex (SL2C.toLorentzGroup M)) i j := by
rw [LinearMap.toMatrix_apply]
simp only [complexContrBasis, Basis.coe_ofEquivFun, Basis.ofEquivFun_repr_apply, transpose_apply]
change (((LorentzGroup.toComplex (SL2C.toLorentzGroup M))) *ᵥ (Pi.single j 1)) i = _
simp only [mulVec_single, transpose_apply, mul_one]
lemma complexContrBasis_ρ_val (M : SL(2,)) (v : complexContr) :
((complexContr.ρ M) v).val =
LorentzGroup.toComplex (SL2C.toLorentzGroup M) *ᵥ v.val := by
rfl
/-- The standard basis of complex covariant Lorentz vectors. -/
def complexCoBasis : Basis (Fin 1 ⊕ Fin 3) complexCo := Basis.ofEquivFun
(Equiv.linearEquiv CoModule.toFin13Fun)
@[simp]
lemma complexCoBasis_ρ_apply (M : SL(2,)) (i j : Fin 1 ⊕ Fin 3) :
(LinearMap.toMatrix complexCoBasis complexCoBasis) (complexCo.ρ M) i j =
(LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ᵀ i j := by
rw [LinearMap.toMatrix_apply]
simp only [complexCoBasis, Basis.coe_ofEquivFun, Basis.ofEquivFun_repr_apply, transpose_apply]
change ((LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ᵀ *ᵥ (Pi.single j 1)) i = _
simp only [mulVec_single, transpose_apply, mul_one]
/-!
## Relation to real
-/
/-- The semilinear map including real Lorentz vectors into complex contravariant
lorentz vectors. -/
def inclCongrRealLorentz : LorentzVector 3 →ₛₗ[Complex.ofReal] complexContr where
toFun v := {val := ofReal ∘ v}
map_add' x y := by
apply Lorentz.ContrModule.ext
rw [Lorentz.ContrModule.val_add]
funext i
simp only [Function.comp_apply, ofReal_eq_coe, Pi.add_apply]
change ofReal (x i + y i) = _
simp only [ofReal_eq_coe, ofReal_add]
map_smul' c x := by
apply Lorentz.ContrModule.ext
rw [Lorentz.ContrModule.val_smul]
funext i
simp only [Function.comp_apply, ofReal_eq_coe, Pi.smul_apply]
change ofReal (c • x i) = _
simp only [smul_eq_mul, ofReal_eq_coe, ofReal_mul]
lemma inclCongrRealLorentz_val (v : LorentzVector 3) :
(inclCongrRealLorentz v).val = ofReal ∘ v := rfl
lemma complexContrBasis_of_real (i : Fin 1 ⊕ Fin 3) :
(complexContrBasis i) = inclCongrRealLorentz (LorentzVector.stdBasis i) := by
apply Lorentz.ContrModule.ext
simp only [complexContrBasis, Basis.coe_ofEquivFun, inclCongrRealLorentz, LorentzVector.stdBasis,
LinearMap.coe_mk, AddHom.coe_mk]
ext j
simp only [Function.comp_apply, ofReal_eq_coe]
erw [Pi.basisFun_apply]
change (Pi.single i 1) j = _
exact Eq.symm (Pi.apply_single (fun _ => ofReal') (congrFun rfl) i 1 j)
lemma inclCongrRealLorentz_ρ (M : SL(2, )) (v : LorentzVector 3) :
(complexContr.ρ M) (inclCongrRealLorentz v) =
inclCongrRealLorentz (SL2C.repLorentzVector M v) := by
apply Lorentz.ContrModule.ext
rw [complexContrBasis_ρ_val, inclCongrRealLorentz_val, inclCongrRealLorentz_val]
rw [LorentzGroup.toComplex_mulVec_ofReal]
apply congrArg
simp only [SL2C.toLorentzGroup_apply_coe]
rw [SL2C.repLorentzVector_apply_eq_mulVec]
rfl
lemma SL2CRep_ρ_basis (M : SL(2, )) (i : Fin 1 ⊕ Fin 3) :
(complexContr.ρ M) (complexContrBasis i) =
∑ j, (SL2C.toLorentzGroup M).1 j i •
complexContrBasis j := by
rw [complexContrBasis_of_real, inclCongrRealLorentz_ρ, SL2C.repLorentzVector_stdBasis, map_sum]
apply congrArg
funext j
simp only [LinearMap.map_smulₛₗ, ofReal_eq_coe, coe_smul]
rw [complexContrBasis_of_real]
end Lorentz
end

View file

@ -0,0 +1,100 @@
/-
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.SpaceTime.LorentzVector.Complex.Basic
/-!
# Contraction of Lorentz vectors
-/
noncomputable section
open Matrix
open MatrixGroups
open Complex
open TensorProduct
open SpaceTime
open CategoryTheory.MonoidalCategory
namespace Lorentz
/-- The bi-linear map corresponding to contraction of a contravariant Lorentz vector with a
covariant Lorentz vector. -/
def contrCoContrBi : complexContr →ₗ[] complexCo →ₗ[] where
toFun ψ := {
toFun := fun φ => ψ.toFin13 ⬝ᵥ φ.toFin13,
map_add' := by
intro φ φ'
simp only [map_add]
rw [dotProduct_add]
map_smul' := by
intro r φ
simp only [LinearEquiv.map_smul]
rw [dotProduct_smul]
rfl}
map_add' ψ ψ':= by
refine LinearMap.ext (fun φ => ?_)
simp only [map_add, LinearMap.coe_mk, AddHom.coe_mk, LinearMap.add_apply]
rw [add_dotProduct]
map_smul' r ψ := by
refine LinearMap.ext (fun φ => ?_)
simp only [LinearEquiv.map_smul, LinearMap.coe_mk, AddHom.coe_mk]
rw [smul_dotProduct]
rfl
/-- The bi-linear map corresponding to contraction of a covariant Lorentz vector with a
contravariant Lorentz vector. -/
def contrContrCoBi : complexCo →ₗ[] complexContr →ₗ[] where
toFun φ := {
toFun := fun ψ => φ.toFin13 ⬝ᵥ ψ.toFin13,
map_add' := by
intro ψ ψ'
simp only [map_add]
rw [dotProduct_add]
map_smul' := by
intro r ψ
simp only [LinearEquiv.map_smul]
rw [dotProduct_smul]
rfl}
map_add' φ φ' := by
refine LinearMap.ext (fun ψ => ?_)
simp only [map_add, LinearMap.coe_mk, AddHom.coe_mk, LinearMap.add_apply]
rw [add_dotProduct]
map_smul' r φ := by
refine LinearMap.ext (fun ψ => ?_)
simp only [LinearEquiv.map_smul, LinearMap.coe_mk, AddHom.coe_mk]
rw [smul_dotProduct]
rfl
/-- The linear map from complexContr ⊗ complexCo to given by
summing over components of contravariant Lorentz vector and
covariant Lorentz vector in the
standard basis (i.e. the dot product).
In terms of index notation this is the contraction is ψⁱ φᵢ. -/
def contrCoContraction : complexContr ⊗ complexCo ⟶ 𝟙_ (Rep SL(2,)) where
hom := TensorProduct.lift contrCoContrBi
comm M := TensorProduct.ext' fun ψ φ => by
change ((LorentzGroup.toComplex (SL2C.toLorentzGroup M)) *ᵥ ψ.toFin13) ⬝ᵥ
((LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ᵀ *ᵥ φ.toFin13) = ψ.toFin13 ⬝ᵥ φ.toFin13
rw [dotProduct_mulVec, vecMul_transpose, mulVec_mulVec]
rw [inv_mul_of_invertible (LorentzGroup.toComplex (SL2C.toLorentzGroup M))]
simp
/-- The linear map from complexCo ⊗ complexContr to given by
summing over components of covariant Lorentz vector and
contravariant Lorentz vector in the
standard basis (i.e. the dot product).
In terms of index notation this is the contraction is φᵢ ψⁱ. -/
def coContrContraction : complexCo ⊗ complexContr ⟶ 𝟙_ (Rep SL(2,)) where
hom := TensorProduct.lift contrContrCoBi
comm M := TensorProduct.ext' fun φ ψ => by
change ((LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ᵀ *ᵥ φ.toFin13) ⬝ᵥ
((LorentzGroup.toComplex (SL2C.toLorentzGroup M)) *ᵥ ψ.toFin13) = φ.toFin13 ⬝ᵥ ψ.toFin13
rw [dotProduct_mulVec, mulVec_transpose, vecMul_vecMul]
rw [inv_mul_of_invertible (LorentzGroup.toComplex (SL2C.toLorentzGroup M))]
simp
end Lorentz
end

View file

@ -0,0 +1,88 @@
/-
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.SpaceTime.LorentzVector.Complex.Two
import HepLean.SpaceTime.MinkowskiMetric
/-!
# Metric for complex Lorentz vectors
-/
noncomputable section
open Matrix
open MatrixGroups
open Complex
open TensorProduct
open SpaceTime
open CategoryTheory.MonoidalCategory
namespace Lorentz
/-- The metric `ηᵃᵃ` as an element of `(complexContr ⊗ complexContr).V`. -/
def contrMetricVal : (complexContr ⊗ complexContr).V :=
contrContrToMatrix.symm ((@minkowskiMatrix 3).map ofReal)
/-- The metric `ηᵃᵃ` as a morphism `𝟙_ (Rep SL(2,)) ⟶ complexContr ⊗ complexContr`,
making its invariance under the action of `SL(2,)`. -/
def contrMetric : 𝟙_ (Rep SL(2,)) ⟶ complexContr ⊗ complexContr where
hom := {
toFun := fun a =>
let a' : := a
a' • contrMetricVal,
map_add' := fun x y => by
simp only [add_smul],
map_smul' := fun m x => by
simp only [smul_smul]
rfl}
comm M := by
ext x : 2
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • contrMetricVal =
(TensorProduct.map (complexContr.ρ M) (complexContr.ρ M)) (x' • contrMetricVal)
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
simp only [Action.instMonoidalCategory_tensorObj_V, contrMetricVal]
erw [contrContrToMatrix_ρ_symm]
apply congrArg
simp only [LorentzGroup.toComplex_mul_minkowskiMatrix_mul_transpose]
/-- The metric `ηᵢᵢ` as an element of `(complexCo ⊗ complexCo).V`. -/
def coMetricVal : (complexCo ⊗ complexCo).V :=
coCoToMatrix.symm ((@minkowskiMatrix 3).map ofReal)
/-- The metric `ηᵢᵢ` as a morphism `𝟙_ (Rep SL(2,)) ⟶ complexCo ⊗ complexCo`,
making its invariance under the action of `SL(2,)`. -/
def coMetric : 𝟙_ (Rep SL(2,)) ⟶ complexCo ⊗ complexCo where
hom := {
toFun := fun a =>
let a' : := a
a' • coMetricVal,
map_add' := fun x y => by
simp only [add_smul],
map_smul' := fun m x => by
simp only [smul_smul]
rfl}
comm M := by
ext x : 2
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • coMetricVal =
(TensorProduct.map (complexCo.ρ M) (complexCo.ρ M)) (x' • coMetricVal)
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
simp only [Action.instMonoidalCategory_tensorObj_V, coMetricVal]
erw [coCoToMatrix_ρ_symm]
apply congrArg
rw [LorentzGroup.toComplex_inv]
simp only [lorentzGroupIsGroup_inv, SL2C.toLorentzGroup_apply_coe,
LorentzGroup.toComplex_transpose_mul_minkowskiMatrix_mul_self]
end Lorentz
end

View file

@ -0,0 +1,270 @@
/-
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.SpaceTime.LorentzVector.Complex.Basic
import Mathlib.LinearAlgebra.TensorProduct.Matrix
/-!
# Tensor products of two complex Lorentz vectors
-/
noncomputable section
open Matrix
open MatrixGroups
open Complex
open TensorProduct
open SpaceTime
open CategoryTheory.MonoidalCategory
namespace Lorentz
/-- Equivalence of `complexContr ⊗ complexContr` to `4 x 4` complex matrices. -/
def contrContrToMatrix : (complexContr ⊗ complexContr).V ≃ₗ[]
Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) :=
(Basis.tensorProduct complexContrBasis complexContrBasis).repr ≪≫ₗ
Finsupp.linearEquivFunOnFinite ((Fin 1 ⊕ Fin 3) × (Fin 1 ⊕ Fin 3)) ≪≫ₗ
LinearEquiv.curry (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3)
/-- Equivalence of `complexCo ⊗ complexCo` to `4 x 4` complex matrices. -/
def coCoToMatrix : (complexCo ⊗ complexCo).V ≃ₗ[]
Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) :=
(Basis.tensorProduct complexCoBasis complexCoBasis).repr ≪≫ₗ
Finsupp.linearEquivFunOnFinite ((Fin 1 ⊕ Fin 3) × (Fin 1 ⊕ Fin 3)) ≪≫ₗ
LinearEquiv.curry (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3)
/-- Equivalence of `complexContr ⊗ complexCo` to `4 x 4` complex matrices. -/
def contrCoToMatrix : (complexContr ⊗ complexCo).V ≃ₗ[]
Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) :=
(Basis.tensorProduct complexContrBasis complexCoBasis).repr ≪≫ₗ
Finsupp.linearEquivFunOnFinite ((Fin 1 ⊕ Fin 3) × (Fin 1 ⊕ Fin 3)) ≪≫ₗ
LinearEquiv.curry (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3)
/-- Equivalence of `complexCo ⊗ complexContr` to `4 x 4` complex matrices. -/
def coContrToMatrix : (complexCo ⊗ complexContr).V ≃ₗ[]
Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) :=
(Basis.tensorProduct complexCoBasis complexContrBasis).repr ≪≫ₗ
Finsupp.linearEquivFunOnFinite ((Fin 1 ⊕ Fin 3) × (Fin 1 ⊕ Fin 3)) ≪≫ₗ
LinearEquiv.curry (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3)
/-!
## Group actions
-/
lemma contrContrToMatrix_ρ (v : (complexContr ⊗ complexContr).V) (M : SL(2,)) :
contrContrToMatrix (TensorProduct.map (complexContr.ρ M) (complexContr.ρ M) v) =
(LorentzGroup.toComplex (SL2C.toLorentzGroup M)) * contrContrToMatrix v *
(LorentzGroup.toComplex (SL2C.toLorentzGroup M))ᵀ := by
nth_rewrite 1 [contrContrToMatrix]
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.trans_apply]
trans (LinearEquiv.curry (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3)) ((LinearMap.toMatrix
(complexContrBasis.tensorProduct complexContrBasis)
(complexContrBasis.tensorProduct complexContrBasis)
(TensorProduct.map (complexContr.ρ M) (complexContr.ρ M)))
*ᵥ ((Finsupp.linearEquivFunOnFinite ((Fin 1 ⊕ Fin 3) × (Fin 1 ⊕ Fin 3)))
((complexContrBasis.tensorProduct complexContrBasis).repr v)))
· apply congrArg
have h1 := (LinearMap.toMatrix_mulVec_repr (complexContrBasis.tensorProduct complexContrBasis)
(complexContrBasis.tensorProduct complexContrBasis)
(TensorProduct.map (complexContr.ρ M) (complexContr.ρ M)) v)
erw [h1]
rfl
rw [TensorProduct.toMatrix_map]
funext i j
change ∑ k, ((kroneckerMap (fun x1 x2 => x1 * x2)
((LinearMap.toMatrix complexContrBasis complexContrBasis) (complexContr.ρ M))
((LinearMap.toMatrix complexContrBasis complexContrBasis) (complexContr.ρ M)) (i, j) k)
* contrContrToMatrix v k.1 k.2) = _
erw [Finset.sum_product]
simp_rw [kroneckerMap_apply, Matrix.mul_apply, Matrix.transpose_apply]
have h1 : ∑ x, (∑ x1, LorentzGroup.toComplex (SL2C.toLorentzGroup M) i x1 *
contrContrToMatrix v x1 x) * LorentzGroup.toComplex (SL2C.toLorentzGroup M) j x
= ∑ x, ∑ x1, (LorentzGroup.toComplex (SL2C.toLorentzGroup M) i x1
* contrContrToMatrix v x1 x) * LorentzGroup.toComplex (SL2C.toLorentzGroup M) j x := by
congr
funext x
rw [Finset.sum_mul]
erw [h1]
rw [Finset.sum_comm]
congr
funext x
congr
funext x1
simp only [complexContrBasis_ρ_apply, transpose_apply, Action.instMonoidalCategory_tensorObj_V]
ring
lemma coCoToMatrix_ρ (v : (complexCo ⊗ complexCo).V) (M : SL(2,)) :
coCoToMatrix (TensorProduct.map (complexCo.ρ M) (complexCo.ρ M) v) =
(LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ᵀ * coCoToMatrix v *
(LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ := by
nth_rewrite 1 [coCoToMatrix]
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.trans_apply]
trans (LinearEquiv.curry (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3)) ((LinearMap.toMatrix
(complexCoBasis.tensorProduct complexCoBasis)
(complexCoBasis.tensorProduct complexCoBasis)
(TensorProduct.map (complexCo.ρ M) (complexCo.ρ M))
*ᵥ ((Finsupp.linearEquivFunOnFinite ((Fin 1 ⊕ Fin 3) × (Fin 1 ⊕ Fin 3)))
((complexCoBasis.tensorProduct complexCoBasis).repr v))))
· apply congrArg
have h1 := (LinearMap.toMatrix_mulVec_repr (complexCoBasis.tensorProduct complexCoBasis)
(complexCoBasis.tensorProduct complexCoBasis)
(TensorProduct.map (complexCo.ρ M) (complexCo.ρ M)) v)
erw [h1]
rfl
rw [TensorProduct.toMatrix_map]
funext i j
change ∑ k, ((kroneckerMap (fun x1 x2 => x1 * x2)
((LinearMap.toMatrix complexCoBasis complexCoBasis) (complexCo.ρ M))
((LinearMap.toMatrix complexCoBasis complexCoBasis) (complexCo.ρ M)) (i, j) k)
* coCoToMatrix v k.1 k.2) = _
erw [Finset.sum_product]
simp_rw [kroneckerMap_apply, Matrix.mul_apply, Matrix.transpose_apply]
have h1 : ∑ x, (∑ x1, (LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ x1 i *
coCoToMatrix v x1 x) * (LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ x j
= ∑ x, ∑ x1, ((LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ x1 i
* coCoToMatrix v x1 x) * (LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ x j := by
congr
funext x
rw [Finset.sum_mul]
erw [h1]
rw [Finset.sum_comm]
congr
funext x
congr
funext x1
simp only [complexCoBasis_ρ_apply, transpose_apply, Action.instMonoidalCategory_tensorObj_V]
ring
lemma contrCoToMatrix_ρ (v : (complexContr ⊗ complexCo).V) (M : SL(2,)) :
contrCoToMatrix (TensorProduct.map (complexContr.ρ M) (complexCo.ρ M) v) =
(LorentzGroup.toComplex (SL2C.toLorentzGroup M)) * contrCoToMatrix v *
(LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ := by
nth_rewrite 1 [contrCoToMatrix]
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.trans_apply]
trans (LinearEquiv.curry (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3)) ((LinearMap.toMatrix
(complexContrBasis.tensorProduct complexCoBasis)
(complexContrBasis.tensorProduct complexCoBasis)
(TensorProduct.map (complexContr.ρ M) (complexCo.ρ M))
*ᵥ ((Finsupp.linearEquivFunOnFinite ((Fin 1 ⊕ Fin 3) × (Fin 1 ⊕ Fin 3)))
((complexContrBasis.tensorProduct complexCoBasis).repr v))))
· apply congrArg
have h1 := (LinearMap.toMatrix_mulVec_repr (complexContrBasis.tensorProduct complexCoBasis)
(complexContrBasis.tensorProduct complexCoBasis)
(TensorProduct.map (complexContr.ρ M) (complexCo.ρ M)) v)
erw [h1]
rfl
rw [TensorProduct.toMatrix_map]
funext i j
change ∑ k, ((kroneckerMap (fun x1 x2 => x1 * x2)
((LinearMap.toMatrix complexContrBasis complexContrBasis) (complexContr.ρ M))
((LinearMap.toMatrix complexCoBasis complexCoBasis) (complexCo.ρ M)) (i, j) k)
* contrCoToMatrix v k.1 k.2) = _
erw [Finset.sum_product]
simp_rw [kroneckerMap_apply, Matrix.mul_apply]
have h1 : ∑ x, (∑ x1, LorentzGroup.toComplex (SL2C.toLorentzGroup M) i x1 *
contrCoToMatrix v x1 x) * (LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ x j
= ∑ x, ∑ x1, (LorentzGroup.toComplex (SL2C.toLorentzGroup M) i x1
* contrCoToMatrix v x1 x) * (LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ x j := by
congr
funext x
rw [Finset.sum_mul]
erw [h1]
rw [Finset.sum_comm]
congr
funext x
congr
funext x1
simp only [complexContrBasis_ρ_apply, complexCoBasis_ρ_apply, transpose_apply,
Action.instMonoidalCategory_tensorObj_V]
ring
lemma coContrToMatrix_ρ (v : (complexCo ⊗ complexContr).V) (M : SL(2,)) :
coContrToMatrix (TensorProduct.map (complexCo.ρ M) (complexContr.ρ M) v) =
(LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ᵀ * coContrToMatrix v *
(LorentzGroup.toComplex (SL2C.toLorentzGroup M))ᵀ := by
nth_rewrite 1 [coContrToMatrix]
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.trans_apply]
trans (LinearEquiv.curry (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3)) ((LinearMap.toMatrix
(complexCoBasis.tensorProduct complexContrBasis)
(complexCoBasis.tensorProduct complexContrBasis)
(TensorProduct.map (complexCo.ρ M) (complexContr.ρ M))
*ᵥ ((Finsupp.linearEquivFunOnFinite ((Fin 1 ⊕ Fin 3) × (Fin 1 ⊕ Fin 3)))
((complexCoBasis.tensorProduct complexContrBasis).repr v))))
· apply congrArg
have h1 := (LinearMap.toMatrix_mulVec_repr (complexCoBasis.tensorProduct complexContrBasis)
(complexCoBasis.tensorProduct complexContrBasis)
(TensorProduct.map (complexCo.ρ M) (complexContr.ρ M)) v)
erw [h1]
rfl
rw [TensorProduct.toMatrix_map]
funext i j
change ∑ k, ((kroneckerMap (fun x1 x2 => x1 * x2)
((LinearMap.toMatrix complexCoBasis complexCoBasis) (complexCo.ρ M))
((LinearMap.toMatrix complexContrBasis complexContrBasis) (complexContr.ρ M)) (i, j) k)
* coContrToMatrix v k.1 k.2) = _
erw [Finset.sum_product]
simp_rw [kroneckerMap_apply, Matrix.mul_apply, Matrix.transpose_apply]
have h1 : ∑ x, (∑ x1, (LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ x1 i *
coContrToMatrix v x1 x) * (LorentzGroup.toComplex (SL2C.toLorentzGroup M)) j x
= ∑ x, ∑ x1, ((LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ x1 i
* coContrToMatrix v x1 x) * (LorentzGroup.toComplex (SL2C.toLorentzGroup M)) j x := by
congr
funext x
rw [Finset.sum_mul]
erw [h1]
rw [Finset.sum_comm]
congr
funext x
congr
funext x1
simp only [complexCoBasis_ρ_apply, complexContrBasis_ρ_apply, transpose_apply,
Action.instMonoidalCategory_tensorObj_V]
ring
/-!
## The symm version of the group actions.
-/
lemma contrContrToMatrix_ρ_symm (v : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ) (M : SL(2,)) :
TensorProduct.map (complexContr.ρ M) (complexContr.ρ M) (contrContrToMatrix.symm v) =
contrContrToMatrix.symm ((LorentzGroup.toComplex (SL2C.toLorentzGroup M)) * v *
(LorentzGroup.toComplex (SL2C.toLorentzGroup M))ᵀ) := by
have h1 := contrContrToMatrix_ρ (contrContrToMatrix.symm v) M
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.apply_symm_apply] at h1
rw [← h1]
simp
lemma coCoToMatrix_ρ_symm (v : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ) (M : SL(2,)) :
TensorProduct.map (complexCo.ρ M) (complexCo.ρ M) (coCoToMatrix.symm v) =
coCoToMatrix.symm ((LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ᵀ * v *
(LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹) := by
have h1 := coCoToMatrix_ρ (coCoToMatrix.symm v) M
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.apply_symm_apply] at h1
rw [← h1]
simp
lemma contrCoToMatrix_ρ_symm (v : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ) (M : SL(2,)) :
TensorProduct.map (complexContr.ρ M) (complexCo.ρ M) (contrCoToMatrix.symm v) =
contrCoToMatrix.symm ((LorentzGroup.toComplex (SL2C.toLorentzGroup M)) * v *
(LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹) := by
have h1 := contrCoToMatrix_ρ (contrCoToMatrix.symm v) M
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.apply_symm_apply] at h1
rw [← h1]
simp
lemma coContrToMatrix_ρ_symm (v : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ) (M : SL(2,)) :
TensorProduct.map (complexCo.ρ M) (complexContr.ρ M) (coContrToMatrix.symm v) =
coContrToMatrix.symm ((LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ᵀ * v *
(LorentzGroup.toComplex (SL2C.toLorentzGroup M))ᵀ) := by
have h1 := coContrToMatrix_ρ (coContrToMatrix.symm v) M
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.apply_symm_apply] at h1
rw [← h1]
simp
end Lorentz
end

View file

@ -0,0 +1,89 @@
/-
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.SpaceTime.LorentzVector.Complex.Two
/-!
# Unit for complex Lorentz vectors
-/
noncomputable section
open Matrix
open MatrixGroups
open Complex
open TensorProduct
open SpaceTime
open CategoryTheory.MonoidalCategory
namespace Lorentz
/-- The contra-co unit for complex lorentz vectors. Usually denoted `δⁱᵢ`. -/
def contrCoUnitVal : (complexContr ⊗ complexCo).V :=
contrCoToMatrix.symm 1
/-- The contra-co unit for complex lorentz vectors as a morphism
`𝟙_ (Rep SL(2,)) ⟶ complexContr ⊗ complexCo`, manifesting the invaraince under
the `SL(2, )` action. -/
def contrCoUnit : 𝟙_ (Rep SL(2,)) ⟶ complexContr ⊗ complexCo where
hom := {
toFun := fun a =>
let a' : := a
a' • contrCoUnitVal,
map_add' := fun x y => by
simp only [add_smul],
map_smul' := fun m x => by
simp only [smul_smul]
rfl}
comm M := by
ext x : 2
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • contrCoUnitVal =
(TensorProduct.map (complexContr.ρ M) (complexCo.ρ M)) (x' • contrCoUnitVal)
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
simp only [Action.instMonoidalCategory_tensorObj_V, contrCoUnitVal]
erw [contrCoToMatrix_ρ_symm]
apply congrArg
simp
/-- The co-contra unit for complex lorentz vectors. Usually denoted `δᵢⁱ`. -/
def coContrUnitVal : (complexCo ⊗ complexContr).V :=
coContrToMatrix.symm 1
/-- The co-contra unit for complex lorentz vectors as a morphism
`𝟙_ (Rep SL(2,)) ⟶ complexCo ⊗ complexContr`, manifesting the invaraince under
the `SL(2, )` action. -/
def coContrUnit : 𝟙_ (Rep SL(2,)) ⟶ complexCo ⊗ complexContr where
hom := {
toFun := fun a =>
let a' : := a
a' • coContrUnitVal,
map_add' := fun x y => by
simp only [add_smul],
map_smul' := fun m x => by
simp only [smul_smul]
rfl}
comm M := by
ext x : 2
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • coContrUnitVal =
(TensorProduct.map (complexCo.ρ M) (complexContr.ρ M)) (x' • coContrUnitVal)
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
simp only [Action.instMonoidalCategory_tensorObj_V, coContrUnitVal]
erw [coContrToMatrix_ρ_symm]
apply congrArg
symm
refine transpose_eq_one.mp ?h.h.h.a
simp
end Lorentz
end

View file

@ -51,6 +51,19 @@ instance : AddCommGroup ContrModule := Equiv.addCommGroup toFin13Fun
with `Fin 1 ⊕ Fin 3 → `. -/
instance : Module ContrModule := Equiv.module toFin13Fun
@[ext]
lemma ext (ψ ψ' : ContrModule) (h : ψ.val = ψ'.val) : ψ = ψ' := by
cases ψ
cases ψ'
subst h
simp_all only
@[simp]
lemma val_add (ψ ψ' : ContrModule) : (ψ + ψ').val = ψ.val + ψ'.val := rfl
@[simp]
lemma val_smul (r : ) (ψ : ContrModule) : (r • ψ).val = r • ψ.val := rfl
/-- The linear equivalence between `ContrModule` and `(Fin 1 ⊕ Fin 3 → )`. -/
@[simps!]
def toFin13Equiv : ContrModule ≃ₗ[] (Fin 1 ⊕ Fin 3 → ) where
@ -68,7 +81,7 @@ abbrev toFin13 (ψ : ContrModule) := toFin13Equiv ψ
/-- The representation of the Lorentz group on `ContrModule`. -/
def lorentzGroupRep : Representation (LorentzGroup 3) ContrModule where
toFun M := {
toFun := fun v => toFin13Equiv.symm ((M.1.map ofReal) *ᵥ v.toFin13),
toFun := fun v => toFin13Equiv.symm (LorentzGroup.toComplex M *ᵥ v.toFin13),
map_add' := by
intro ψ ψ'
simp [mulVec_add]
@ -132,7 +145,7 @@ abbrev toFin13 (ψ : CoModule) := toFin13Equiv ψ
/-- The representation of the Lorentz group on `CoModule`. -/
def lorentzGroupRep : Representation (LorentzGroup 3) CoModule where
toFun M := {
toFun := fun v => toFin13Equiv.symm ((M.1.map ofReal)⁻¹ᵀ *ᵥ v.toFin13),
toFun := fun v => toFin13Equiv.symm ((LorentzGroup.toComplex M)⁻¹ᵀ *ᵥ v.toFin13),
map_add' := by
intro ψ ψ'
simp [mulVec_add]
@ -147,7 +160,7 @@ def lorentzGroupRep : Representation (LorentzGroup 3) CoModule where
simp only [SpecialLinearGroup.coe_mul, LinearMap.coe_mk, AddHom.coe_mk, LinearMap.mul_apply,
LinearEquiv.apply_symm_apply, mulVec_mulVec, EmbeddingLike.apply_eq_iff_eq]
refine (congrFun (congrArg _ ?_) _)
simp only [lorentzGroupIsGroup_mul_coe, Matrix.map_mul]
simp only [_root_.map_mul]
rw [Matrix.mul_inv_rev]
exact transpose_mul _ _

View file

@ -82,6 +82,14 @@ lemma off_diag_zero {μ ν : Fin 1 ⊕ Fin d} (h : μ ≠ ν) : η μ ν = 0 :=
simp only [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
exact diagonal_apply_ne _ h
lemma inl_0_inl_0 : @minkowskiMatrix d (Sum.inl 0) (Sum.inl 0) = 1 := by
simp only [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
rfl
lemma inr_i_inr_i (i : Fin d) : @minkowskiMatrix d (Sum.inr i) (Sum.inr i) = -1 := by
simp only [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
simp_all only [diagonal_apply_eq, Sum.elim_inr]
end minkowskiMatrix
/-!
@ -269,6 +277,16 @@ lemma det_dual : (dual Λ).det = Λ.det := by
norm_cast
simp
lemma dual_apply (μ ν : Fin 1 ⊕ Fin d) :
dual Λ μ ν = η μ μ * Λ ν μ * η ν ν := by
simp only [dual, minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, mul_diagonal,
diagonal_mul, transpose_apply, diagonal_apply_eq]
lemma dual_apply_minkowskiMatrix (μ ν : Fin 1 ⊕ Fin d) :
dual Λ μ ν * η ν ν = η μ μ * Λ ν μ := by
rw [dual_apply, mul_assoc]
simp
@[simp]
lemma dual_mulVec_right : ⟪x, (dual Λ) *ᵥ y⟫ₘ = ⟪Λ *ᵥ x, y⟫ₘ := by
simp only [minkowskiMetric, LinearMap.coe_mk, AddHom.coe_mk, dual, minkowskiLinearForm_apply,

View file

@ -0,0 +1,128 @@
/-
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.OverColor.Basic
import HepLean.Mathematics.PiTensorProduct
import HepLean.SpaceTime.LorentzVector.Complex.Basic
import HepLean.SpaceTime.WeylFermion.Two
import HepLean.SpaceTime.PauliMatrices.Basic
/-!
## Pauli matrices
-/
namespace PauliMatrix
open Complex
open Lorentz
open Fermion
open TensorProduct
open CategoryTheory.MonoidalCategory
noncomputable section
open Matrix
open MatrixGroups
open Complex
open TensorProduct
open SpaceTime
/-- The tensor `σ^μ^a^{dot a}` based on the Pauli-matrices as an element of
`complexContr ⊗ leftHanded ⊗ rightHanded`. -/
def asTensor : (complexContr ⊗ leftHanded ⊗ rightHanded).V :=
∑ i, complexContrBasis i ⊗ₜ leftRightToMatrix.symm (σSA i)
/-- The tensor `σ^μ^a^{dot a}` based on the Pauli-matrices as a morphism,
`𝟙_ (Rep SL(2,)) ⟶ complexContr ⊗ leftHanded ⊗ rightHanded` manifesting
the invariance under the `SL(2,)` action. -/
def asConsTensor : 𝟙_ (Rep SL(2,)) ⟶ complexContr ⊗ leftHanded ⊗ rightHanded where
hom := {
toFun := fun a =>
let a' : := a
a' • asTensor,
map_add' := fun x y => by
simp only [add_smul],
map_smul' := fun m x => by
simp only [smul_smul]
rfl}
comm M := by
ext x : 2
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • asTensor =
(TensorProduct.map (complexContr.ρ M) (
TensorProduct.map (leftHanded.ρ M) (rightHanded.ρ M))) (x' • asTensor)
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
nth_rewrite 2 [asTensor]
simp only [Action.instMonoidalCategory_tensorObj_V, CategoryTheory.Equivalence.symm_inverse,
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
map_sum, map_tmul]
symm
calc _ = ∑ x, ((complexContr.ρ M) (complexContrBasis x) ⊗ₜ[]
leftRightToMatrix.symm (SL2C.repSelfAdjointMatrix M (σSA x))) := by
refine Finset.sum_congr rfl (fun x _ => ?_)
rw [← leftRightToMatrix_ρ_symm_selfAdjoint]
rfl
_ = ∑ x, ((∑ i, (SL2C.toLorentzGroup M).1 i x • (complexContrBasis i)) ⊗ₜ[]
∑ j, leftRightToMatrix.symm ((SL2C.toLorentzGroup M⁻¹).1 x j • (σSA j))) := by
refine Finset.sum_congr rfl (fun x _ => ?_)
rw [SL2CRep_ρ_basis, SL2C.repSelfAdjointMatrix_σSA]
simp only [Action.instMonoidalCategory_tensorObj_V, SL2C.toLorentzGroup_apply_coe,
Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton, map_inv, lorentzGroupIsGroup_inv, AddSubgroup.coe_add,
selfAdjoint.val_smul, AddSubgroup.val_finset_sum, map_add, map_sum]
_ = ∑ x, ∑ i, ∑ j, ((SL2C.toLorentzGroup M).1 i x • (complexContrBasis i)) ⊗ₜ[]
leftRightToMatrix.symm.toLinearMap ((SL2C.toLorentzGroup M⁻¹).1 x j • (σSA j)) := by
refine Finset.sum_congr rfl (fun x _ => ?_)
rw [sum_tmul]
refine Finset.sum_congr rfl (fun i _ => ?_)
rw [tmul_sum]
rfl
_ = ∑ x, ∑ i, ∑ j, ((SL2C.toLorentzGroup M).1 i x • (complexContrBasis i)) ⊗ₜ[]
((SL2C.toLorentzGroup M⁻¹).1 x j • leftRightToMatrix.symm ((σSA j))) := by
refine Finset.sum_congr rfl (fun x _ => (Finset.sum_congr rfl (fun i _ =>
(Finset.sum_congr rfl (fun j _ => ?_)))))
simp only [Action.instMonoidalCategory_tensorObj_V, SL2C.toLorentzGroup_apply_coe,
map_inv, lorentzGroupIsGroup_inv, LinearMap.map_smul_of_tower, LinearEquiv.coe_coe,
tmul_smul]
_ = ∑ x, ∑ i, ∑ j, ((SL2C.toLorentzGroup M).1 i x * (SL2C.toLorentzGroup M⁻¹).1 x j)
• ((complexContrBasis i)) ⊗ₜ[] leftRightToMatrix.symm ((σSA j)) := by
refine Finset.sum_congr rfl (fun x _ => (Finset.sum_congr rfl (fun i _ =>
(Finset.sum_congr rfl (fun j _ => ?_)))))
rw [smul_tmul, smul_smul, tmul_smul]
_ = ∑ i, ∑ x, ∑ j, ((SL2C.toLorentzGroup M).1 i x * (SL2C.toLorentzGroup M⁻¹).1 x j)
• ((complexContrBasis i)) ⊗ₜ[] leftRightToMatrix.symm ((σSA j)) := Finset.sum_comm
_ = ∑ i, ∑ j, ∑ x, ((SL2C.toLorentzGroup M).1 i x * (SL2C.toLorentzGroup M⁻¹).1 x j)
• ((complexContrBasis i)) ⊗ₜ[] leftRightToMatrix.symm ((σSA j)) :=
Finset.sum_congr rfl (fun x _ => Finset.sum_comm)
_ = ∑ i, ∑ j, (∑ x, (SL2C.toLorentzGroup M).1 i x * (SL2C.toLorentzGroup M⁻¹).1 x j)
• ((complexContrBasis i)) ⊗ₜ[] leftRightToMatrix.symm ((σSA j)) := by
refine Finset.sum_congr rfl (fun i _ => (Finset.sum_congr rfl (fun j _ => ?_)))
rw [Finset.sum_smul]
_ = ∑ i, ∑ j, ((1 : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ) i j)
• ((complexContrBasis i)) ⊗ₜ[] leftRightToMatrix.symm ((σSA j)) := by
refine Finset.sum_congr rfl (fun i _ => (Finset.sum_congr rfl (fun j _ => ?_)))
congr
change ((SL2C.toLorentzGroup M) * (SL2C.toLorentzGroup M⁻¹)).1 i j = _
rw [← SL2C.toLorentzGroup.map_mul]
simp only [mul_inv_cancel, _root_.map_one, lorentzGroupIsGroup_one_coe]
_ = ∑ i, ((1 : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ) i i)
• ((complexContrBasis i)) ⊗ₜ[] leftRightToMatrix.symm ((σSA i)) := by
refine Finset.sum_congr rfl (fun i _ => ?_)
refine Finset.sum_eq_single i (fun b _ hb => ?_) (fun hb => ?_)
· simp [one_apply_ne' hb]
· simp only [Finset.mem_univ, not_true_eq_false] at hb
_ = asTensor := by
refine Finset.sum_congr rfl (fun i _ => ?_)
simp only [Action.instMonoidalCategory_tensorObj_V, one_apply_eq, one_smul,
CategoryTheory.Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj]
end
end PauliMatrix

View file

@ -0,0 +1,136 @@
/-
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.Mathematics.PiTensorProduct
import Mathlib.RepresentationTheory.Rep
import HepLean.Tensors.Basic
import Mathlib.Logic.Equiv.TransferInstance
import HepLean.SpaceTime.LorentzGroup.Basic
/-!
## Pauli matrices
-/
namespace PauliMatrix
open Complex
open Matrix
/-- The zeroth Pauli-matrix as a `2 x 2` complex matrix. -/
def σ0 : Matrix (Fin 2) (Fin 2) := !![1, 0; 0, 1]
/-- The first Pauli-matrix as a `2 x 2` complex matrix. -/
def σ1 : Matrix (Fin 2) (Fin 2) := !![0, 1; 1, 0]
/-- The second Pauli-matrix as a `2 x 2` complex matrix. -/
def σ2 : Matrix (Fin 2) (Fin 2) := !![0, -I; I, 0]
/-- The third Pauli-matrix as a `2 x 2` complex matrix. -/
def σ3 : Matrix (Fin 2) (Fin 2) := !![1, 0; 0, -1]
@[simp]
lemma σ0_selfAdjoint : σ0ᴴ = σ0 := by
rw [eta_fin_two σ0ᴴ]
simp [σ0]
@[simp]
lemma σ1_selfAdjoint : σ1ᴴ = σ1 := by
rw [eta_fin_two σ1ᴴ]
simp [σ1]
@[simp]
lemma σ2_selfAdjoint : σ2ᴴ = σ2 := by
rw [eta_fin_two σ2ᴴ]
simp [σ2]
@[simp]
lemma σ3_selfAdjoint : σ3ᴴ = σ3 := by
rw [eta_fin_two σ3ᴴ]
simp [σ3]
/-!
## Traces
-/
@[simp]
lemma σ0_σ0_trace : Matrix.trace (σ0 * σ0) = 2 := by
simp only [σ0, cons_mul, Nat.succ_eq_add_one, Nat.reduceAdd, vecMul_cons, head_cons, one_smul,
tail_cons, zero_smul, empty_vecMul, add_zero, zero_add, empty_mul, Equiv.symm_apply_apply,
trace_fin_two_of]
norm_num
@[simp]
lemma σ0_σ1_trace : Matrix.trace (σ0 * σ1) = 0 := by
simp [σ0, σ1]
@[simp]
lemma σ0_σ2_trace : Matrix.trace (σ0 * σ2) = 0 := by
simp [σ0, σ2]
@[simp]
lemma σ0_σ3_trace : Matrix.trace (σ0 * σ3) = 0 := by
simp [σ0, σ3]
@[simp]
lemma σ1_σ0_trace : Matrix.trace (σ1 * σ0) = 0 := by
simp [σ1, σ0]
@[simp]
lemma σ1_σ1_trace : Matrix.trace (σ1 * σ1) = 2 := by
simp only [σ1, cons_mul, Nat.succ_eq_add_one, Nat.reduceAdd, vecMul_cons, head_cons, one_smul,
tail_cons, zero_smul, empty_vecMul, add_zero, zero_add, empty_mul, Equiv.symm_apply_apply,
trace_fin_two_of]
norm_num
@[simp]
lemma σ1_σ2_trace : Matrix.trace (σ1 * σ2) = 0 := by
simp [σ1, σ2]
@[simp]
lemma σ1_σ3_trace : Matrix.trace (σ1 * σ3) = 0 := by
simp [σ1, σ3]
@[simp]
lemma σ2_σ0_trace : Matrix.trace (σ2 * σ0) = 0 := by
simp [σ2, σ0]
@[simp]
lemma σ2_σ1_trace : Matrix.trace (σ2 * σ1) = 0 := by
simp [σ2, σ1]
@[simp]
lemma σ2_σ2_trace : Matrix.trace (σ2 * σ2) = 2 := by
simp only [σ2, cons_mul, Nat.succ_eq_add_one, Nat.reduceAdd, vecMul_cons, head_cons, one_smul,
tail_cons, zero_smul, empty_vecMul, add_zero, zero_add, empty_mul, Equiv.symm_apply_apply,
trace_fin_two_of]
norm_num
@[simp]
lemma σ2_σ3_trace : Matrix.trace (σ2 * σ3) = 0 := by
simp [σ2, σ3]
@[simp]
lemma σ3_σ0_trace : Matrix.trace (σ3 * σ0) = 0 := by
simp [σ3, σ0]
@[simp]
lemma σ3_σ1_trace : Matrix.trace (σ3 * σ1) = 0 := by
simp [σ3, σ1]
@[simp]
lemma σ3_σ2_trace : Matrix.trace (σ3 * σ2) = 0 := by
simp [σ3, σ2]
@[simp]
lemma σ3_σ3_trace : Matrix.trace (σ3 * σ3) = 2 := by
simp only [σ3, cons_mul, Nat.succ_eq_add_one, Nat.reduceAdd, vecMul_cons, head_cons, one_smul,
tail_cons, zero_smul, empty_vecMul, add_zero, zero_add, empty_mul, Equiv.symm_apply_apply,
trace_fin_two_of]
norm_num
end PauliMatrix

View file

@ -0,0 +1,424 @@
/-
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.Mathematics.PiTensorProduct
import Mathlib.RepresentationTheory.Rep
import HepLean.Tensors.Basic
import Mathlib.Logic.Equiv.TransferInstance
import HepLean.SpaceTime.LorentzGroup.Basic
import HepLean.SpaceTime.PauliMatrices.Basic
/-!
## Interaction of Pauli matrices with self-adjoint matrices
-/
namespace PauliMatrix
open Matrix
lemma selfAdjoint_trace_σ0_real (A : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
(Matrix.trace (σ0 * A.1)).re = Matrix.trace (σ0 * A.1) := by
rw [eta_fin_two A.1]
simp only [σ0, Fin.isValue, cons_mul, Nat.succ_eq_add_one, Nat.reduceAdd, vecMul_cons, head_cons,
one_smul, tail_cons, zero_smul, empty_vecMul, add_zero, zero_add, empty_mul,
Equiv.symm_apply_apply, trace_fin_two_of, Complex.add_re, Complex.ofReal_add]
have hA : IsSelfAdjoint A.1 := A.2
rw [isSelfAdjoint_iff, star_eq_conjTranspose] at hA
rw [eta_fin_two (A.1)ᴴ] at hA
simp only [Fin.isValue, conjTranspose_apply, RCLike.star_def, of_apply, cons_val', cons_val_zero,
empty_val', cons_val_fin_one, cons_val_one, head_cons, head_fin_const] at hA
have h00 := congrArg (fun f => f 0 0) hA
simp only [Fin.isValue, of_apply, cons_val', cons_val_zero, empty_val', cons_val_fin_one] at h00
have hA00 : A.1 0 0 = (A.1 0 0).re := by
exact Eq.symm ((fun {z} => Complex.conj_eq_iff_re.mp) h00)
rw [hA00]
simp only [Fin.isValue, Complex.ofReal_re, add_right_inj]
have h11 := congrArg (fun f => f 1 1) hA
simp only [Fin.isValue, of_apply, cons_val', cons_val_one, head_cons, empty_val',
cons_val_fin_one, head_fin_const] at h11
exact Complex.conj_eq_iff_re.mp h11
lemma selfAdjoint_trace_σ1_real (A : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
(Matrix.trace (σ1 * A.1)).re = Matrix.trace (σ1 * A.1) := by
rw [eta_fin_two A.1]
simp only [σ1, Fin.isValue, cons_mul, Nat.succ_eq_add_one, Nat.reduceAdd, vecMul_cons, head_cons,
one_smul, tail_cons, zero_smul, empty_vecMul, add_zero, zero_add, empty_mul,
Equiv.symm_apply_apply, trace_fin_two_of, Complex.add_re, Complex.ofReal_add]
have hA : IsSelfAdjoint A.1 := A.2
rw [isSelfAdjoint_iff, star_eq_conjTranspose] at hA
rw [eta_fin_two (A.1)ᴴ] at hA
simp only [Fin.isValue, conjTranspose_apply, RCLike.star_def] at hA
have h01 := congrArg (fun f => f 0 1) hA
simp only [Fin.isValue, of_apply, cons_val', cons_val_one, head_cons, empty_val',
cons_val_fin_one, cons_val_zero] at h01
rw [← h01]
simp only [Fin.isValue, Complex.conj_re]
rw [Complex.add_conj]
simp only [Fin.isValue, Complex.ofReal_mul, Complex.ofReal_ofNat]
ring
lemma selfAdjoint_trace_σ2_real (A : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
(Matrix.trace (σ2 * A.1)).re = Matrix.trace (σ2 * A.1) := by
rw [eta_fin_two A.1]
simp only [σ2, Fin.isValue, cons_mul, Nat.succ_eq_add_one, Nat.reduceAdd, vecMul_cons, head_cons,
one_smul, tail_cons, zero_smul, empty_vecMul, add_zero, zero_add, empty_mul,
Equiv.symm_apply_apply, trace_fin_two_of, Complex.add_re, Complex.ofReal_add]
have hA : IsSelfAdjoint A.1 := A.2
rw [isSelfAdjoint_iff, star_eq_conjTranspose, eta_fin_two (A.1)ᴴ] at hA
simp only [Fin.isValue, conjTranspose_apply, RCLike.star_def] at hA
have h10 := congrArg (fun f => f 1 0) hA
simp only [Fin.isValue, of_apply, cons_val', cons_val_zero, empty_val', cons_val_fin_one,
cons_val_one, head_fin_const] at h10
rw [← h10]
simp only [Fin.isValue, neg_smul, smul_cons, smul_eq_mul, smul_empty, neg_cons, neg_empty,
trace_fin_two_of, Complex.add_re, Complex.neg_re, Complex.mul_re, Complex.I_re, Complex.conj_re,
zero_mul, Complex.I_im, Complex.conj_im, mul_neg, one_mul, sub_neg_eq_add, zero_add, zero_sub,
Complex.ofReal_add, Complex.ofReal_neg]
trans Complex.I * (A.1 0 1 - (starRingEnd ) (A.1 0 1))
· rw [Complex.sub_conj]
ring_nf
simp
· ring
lemma selfAdjoint_trace_σ3_real (A : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
(Matrix.trace (σ3 * A.1)).re = Matrix.trace (σ3 * A.1) := by
rw [eta_fin_two A.1]
simp only [σ3, Fin.isValue, cons_mul, Nat.succ_eq_add_one, Nat.reduceAdd, vecMul_cons, head_cons,
one_smul, tail_cons, zero_smul, empty_vecMul, add_zero, zero_add, empty_mul,
Equiv.symm_apply_apply, trace_fin_two_of, Complex.add_re, Complex.ofReal_add]
have hA : IsSelfAdjoint A.1 := A.2
rw [isSelfAdjoint_iff, star_eq_conjTranspose, eta_fin_two (A.1)ᴴ] at hA
simp only [Fin.isValue, conjTranspose_apply, RCLike.star_def] at hA
have h00 := congrArg (fun f => f 0 0) hA
have h11 := congrArg (fun f => f 1 1) hA
have hA00 : A.1 0 0 = (A.1 0 0).re := Eq.symm ((fun {z} => Complex.conj_eq_iff_re.mp) h00)
have hA11 : A.1 1 1 = (A.1 1 1).re := Eq.symm ((fun {z} => Complex.conj_eq_iff_re.mp) h11)
simp only [Fin.isValue, neg_smul, one_smul, neg_cons, neg_empty, trace_fin_two_of, Complex.add_re,
Complex.neg_re, Complex.ofReal_add, Complex.ofReal_neg]
rw [hA00, hA11]
simp
open Complex
lemma selfAdjoint_ext_complex {A B : selfAdjoint (Matrix (Fin 2) (Fin 2) )}
(h0 : ((Matrix.trace (PauliMatrix.σ0 * A.1)) = (Matrix.trace (PauliMatrix.σ0 * B.1))))
(h1 : Matrix.trace (PauliMatrix.σ1 * A.1) = Matrix.trace (PauliMatrix.σ1 * B.1))
(h2 : Matrix.trace (PauliMatrix.σ2 * A.1) = Matrix.trace (PauliMatrix.σ2 * B.1))
(h3 : Matrix.trace (PauliMatrix.σ3 * A.1) = Matrix.trace (PauliMatrix.σ3 * B.1)) : A = B := by
ext i j
rw [eta_fin_two A.1, eta_fin_two B.1] at h0 h1 h2 h3
simp only [PauliMatrix.σ0, Fin.isValue, cons_mul, Nat.succ_eq_add_one, Nat.reduceAdd, vecMul_cons,
head_cons, one_smul, tail_cons, zero_smul, empty_vecMul, add_zero, zero_add, empty_mul,
Equiv.symm_apply_apply, trace_fin_two_of] at h0
simp only [σ1, Fin.isValue, cons_mul, Nat.succ_eq_add_one, Nat.reduceAdd, vecMul_cons, head_cons,
zero_smul, tail_cons, one_smul, empty_vecMul, add_zero, zero_add, empty_mul,
Equiv.symm_apply_apply, trace_fin_two_of] at h1
simp only [σ2, Fin.isValue, cons_mul, Nat.succ_eq_add_one, Nat.reduceAdd, vecMul_cons, head_cons,
zero_smul, tail_cons, neg_smul, smul_cons, smul_eq_mul, smul_empty, neg_cons, neg_empty,
empty_vecMul, add_zero, zero_add, empty_mul, Equiv.symm_apply_apply, trace_fin_two_of] at h2
simp only [σ3, Fin.isValue, cons_mul, Nat.succ_eq_add_one, Nat.reduceAdd, vecMul_cons, head_cons,
one_smul, tail_cons, zero_smul, empty_vecMul, add_zero, neg_smul, neg_cons, neg_empty, zero_add,
empty_mul, Equiv.symm_apply_apply, trace_fin_two_of] at h3
match i, j with
| 0, 0 =>
linear_combination (norm := ring_nf) (h0 + h3) / 2
| 0, 1 =>
linear_combination (norm := ring_nf) (h1 - I * h2) / 2
simp
| 1, 0 =>
linear_combination (norm := ring_nf) (h1 + I * h2) / 2
simp
| 1, 1 =>
linear_combination (norm := ring_nf) (h0 - h3) / 2
lemma selfAdjoint_ext {A B : selfAdjoint (Matrix (Fin 2) (Fin 2) )}
(h0 : ((Matrix.trace (PauliMatrix.σ0 * A.1))).re = ((Matrix.trace (PauliMatrix.σ0 * B.1))).re)
(h1 : ((Matrix.trace (PauliMatrix.σ1 * A.1))).re = ((Matrix.trace (PauliMatrix.σ1 * B.1))).re)
(h2 : ((Matrix.trace (PauliMatrix.σ2 * A.1))).re = ((Matrix.trace (PauliMatrix.σ2 * B.1))).re)
(h3 : ((Matrix.trace (PauliMatrix.σ3 * A.1))).re = ((Matrix.trace (PauliMatrix.σ3 * B.1))).re) :
A = B := by
have h0' := congrArg ofReal h0
have h1' := congrArg ofReal h1
have h2' := congrArg ofReal h2
have h3' := congrArg ofReal h3
rw [ofReal_eq_coe, ofReal_eq_coe] at h0' h1' h2' h3'
rw [selfAdjoint_trace_σ0_real A, selfAdjoint_trace_σ0_real B] at h0'
rw [selfAdjoint_trace_σ1_real A, selfAdjoint_trace_σ1_real B] at h1'
rw [selfAdjoint_trace_σ2_real A, selfAdjoint_trace_σ2_real B] at h2'
rw [selfAdjoint_trace_σ3_real A, selfAdjoint_trace_σ3_real B] at h3'
exact selfAdjoint_ext_complex h0' h1' h2' h3'
noncomputable section
/-- An auxillary function which on `i : Fin 1 ⊕ Fin 3` returns the corresponding
Pauli-matrix as a self-adjoint matrix. -/
def σSA' (i : Fin 1 ⊕ Fin 3) : selfAdjoint (Matrix (Fin 2) (Fin 2) ) :=
match i with
| Sum.inl 0 => ⟨σ0, σ0_selfAdjoint⟩
| Sum.inr 0 => ⟨σ1, σ1_selfAdjoint⟩
| Sum.inr 1 => ⟨σ2, σ2_selfAdjoint⟩
| Sum.inr 2 => ⟨σ3, σ3_selfAdjoint⟩
lemma σSA_linearly_independent : LinearIndependent σSA' := by
apply Fintype.linearIndependent_iff.mpr
intro g hg
simp only [Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton] at hg
rw [Fin.sum_univ_three] at hg
simp only [Fin.isValue, σSA'] at hg
intro i
match i with
| Sum.inl 0 =>
simpa [mul_sub, mul_add] using congrArg (fun A => (Matrix.trace (PauliMatrix.σ0 * A.1))) hg
| Sum.inr 0 =>
simpa [mul_sub, mul_add] using congrArg (fun A => (Matrix.trace (PauliMatrix.σ1 * A.1))) hg
| Sum.inr 1 =>
simpa [mul_sub, mul_add] using congrArg (fun A => (Matrix.trace (PauliMatrix.σ2 * A.1))) hg
| Sum.inr 2 =>
simpa [mul_sub, mul_add] using congrArg (fun A => (Matrix.trace (PauliMatrix.σ3 * A.1))) hg
lemma σSA_span : ≤ Submodule.span (Set.range σSA') := by
refine (top_le_span_range_iff_forall_exists_fun ).mpr ?_
intro A
let c : Fin 1 ⊕ Fin 3 → := fun i =>
match i with
| Sum.inl 0 => 1/2 * (Matrix.trace (σ0 * A.1)).re
| Sum.inr 0 => 1/2 * (Matrix.trace (σ1 * A.1)).re
| Sum.inr 1 => 1/2 * (Matrix.trace (σ2 * A.1)).re
| Sum.inr 2 => 1/2 * (Matrix.trace (σ3 * A.1)).re
use c
simp only [one_div, Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton, Fin.sum_univ_three, c]
apply selfAdjoint_ext
· simp only [σSA', AddSubgroup.coe_add, selfAdjoint.val_smul, mul_add, Algebra.mul_smul_comm,
trace_add, trace_smul, σ0_σ0_trace, real_smul, ofReal_mul, ofReal_inv, ofReal_ofNat,
σ0_σ1_trace, smul_zero, σ0_σ2_trace, add_zero, σ0_σ3_trace, mul_re, inv_re, re_ofNat,
normSq_ofNat, div_self_mul_self', ofReal_re, inv_im, im_ofNat, neg_zero, zero_div, ofReal_im,
mul_zero, sub_zero, mul_im, zero_mul]
ring
· simp only [σSA', AddSubgroup.coe_add, selfAdjoint.val_smul, mul_add, Algebra.mul_smul_comm,
trace_add, trace_smul, σ1_σ0_trace, smul_zero, σ1_σ1_trace, real_smul, ofReal_mul, ofReal_inv,
ofReal_ofNat, σ1_σ2_trace, add_zero, σ1_σ3_trace, zero_add, mul_re, inv_re, re_ofNat,
normSq_ofNat, div_self_mul_self', ofReal_re, inv_im, im_ofNat, neg_zero, zero_div, ofReal_im,
mul_zero, sub_zero, mul_im, zero_mul]
ring
· simp only [σSA', AddSubgroup.coe_add, selfAdjoint.val_smul, mul_add, Algebra.mul_smul_comm,
trace_add, trace_smul, σ2_σ0_trace, smul_zero, σ2_σ1_trace, σ2_σ2_trace, real_smul, ofReal_mul,
ofReal_inv, ofReal_ofNat, zero_add, σ2_σ3_trace, add_zero, mul_re, inv_re, re_ofNat,
normSq_ofNat, div_self_mul_self', ofReal_re, inv_im, im_ofNat, neg_zero, zero_div, ofReal_im,
mul_zero, sub_zero, mul_im, zero_mul]
ring
· simp only [σSA', AddSubgroup.coe_add, selfAdjoint.val_smul, mul_add, Algebra.mul_smul_comm,
trace_add, trace_smul, σ3_σ0_trace, smul_zero, σ3_σ1_trace, σ3_σ2_trace, add_zero, σ3_σ3_trace,
real_smul, ofReal_mul, ofReal_inv, ofReal_ofNat, zero_add, mul_re, inv_re, re_ofNat,
normSq_ofNat, div_self_mul_self', ofReal_re, inv_im, im_ofNat, neg_zero, zero_div, ofReal_im,
mul_zero, sub_zero, mul_im, zero_mul]
ring
/-- The basis of `selfAdjoint (Matrix (Fin 2) (Fin 2) )` formed by Pauli matrices. -/
def σSA : Basis (Fin 1 ⊕ Fin 3) (selfAdjoint (Matrix (Fin 2) (Fin 2) )) :=
Basis.mk σSA_linearly_independent σSA_span
/-- An auxillary function which on `i : Fin 1 ⊕ Fin 3` returns the corresponding
Pauli-matrix as a self-adjoint matrix with a minus sign for `Sum.inr _`. -/
def σSAL' (i : Fin 1 ⊕ Fin 3) : selfAdjoint (Matrix (Fin 2) (Fin 2) ) :=
match i with
| Sum.inl 0 => ⟨σ0, σ0_selfAdjoint⟩
| Sum.inr 0 => ⟨-σ1, by rw [neg_mem_iff]; exact σ1_selfAdjoint⟩
| Sum.inr 1 => ⟨-σ2, by rw [neg_mem_iff]; exact σ2_selfAdjoint⟩
| Sum.inr 2 => ⟨-σ3, by rw [neg_mem_iff]; exact σ3_selfAdjoint⟩
lemma σSAL_linearly_independent : LinearIndependent σSAL' := by
apply Fintype.linearIndependent_iff.mpr
intro g hg
simp only [Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton] at hg
rw [Fin.sum_univ_three] at hg
simp only [Fin.isValue, σSAL'] at hg
intro i
match i with
| Sum.inl 0 =>
simpa [mul_sub, mul_add] using congrArg (fun A => (Matrix.trace (PauliMatrix.σ0 * A.1))) hg
| Sum.inr 0 =>
simpa [mul_sub, mul_add] using congrArg (fun A => (Matrix.trace (PauliMatrix.σ1 * A.1))) hg
| Sum.inr 1 =>
simpa [mul_sub, mul_add] using congrArg (fun A => (Matrix.trace (PauliMatrix.σ2 * A.1))) hg
| Sum.inr 2 =>
simpa [mul_sub, mul_add] using congrArg (fun A => (Matrix.trace (PauliMatrix.σ3 * A.1))) hg
lemma σSAL_span : ≤ Submodule.span (Set.range σSAL') := by
refine (top_le_span_range_iff_forall_exists_fun ).mpr ?_
intro A
let c : Fin 1 ⊕ Fin 3 → := fun i =>
match i with
| Sum.inl 0 => 1/2 * (Matrix.trace (σ0 * A.1)).re
| Sum.inr 0 => - 1/2 * (Matrix.trace (σ1 * A.1)).re
| Sum.inr 1 => - 1/2 * (Matrix.trace (σ2 * A.1)).re
| Sum.inr 2 => - 1/2 * (Matrix.trace (σ3 * A.1)).re
use c
simp only [one_div, Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton, Fin.sum_univ_three, c]
apply selfAdjoint_ext
· simp only [σSAL', AddSubgroup.coe_add, selfAdjoint.val_smul, smul_neg, mul_add,
Algebra.mul_smul_comm, mul_neg, trace_add, trace_smul, σ0_σ0_trace, real_smul, ofReal_mul,
ofReal_inv, ofReal_ofNat, trace_neg, σ0_σ1_trace, smul_zero, neg_zero, σ0_σ2_trace, add_zero,
σ0_σ3_trace, mul_re, inv_re, re_ofNat, normSq_ofNat, div_self_mul_self', ofReal_re, inv_im,
im_ofNat, zero_div, ofReal_im, mul_zero, sub_zero, mul_im, zero_mul]
ring
· simp only [σSAL', AddSubgroup.coe_add, selfAdjoint.val_smul, smul_neg, mul_add,
Algebra.mul_smul_comm, mul_neg, trace_add, trace_smul, σ1_σ0_trace, smul_zero, trace_neg,
σ1_σ1_trace, real_smul, ofReal_mul, ofReal_div, ofReal_neg, ofReal_one, ofReal_ofNat,
σ1_σ2_trace, neg_zero, add_zero, σ1_σ3_trace, zero_add, neg_re, mul_re, div_ofNat_re, one_re,
ofReal_re, div_ofNat_im, neg_im, one_im, zero_div, ofReal_im, mul_zero, sub_zero, re_ofNat,
mul_im, zero_mul, im_ofNat]
ring
· simp only [σSAL', AddSubgroup.coe_add, selfAdjoint.val_smul, smul_neg, mul_add,
Algebra.mul_smul_comm, mul_neg, trace_add, trace_smul, σ2_σ0_trace, smul_zero, trace_neg,
σ2_σ1_trace, neg_zero, σ2_σ2_trace, real_smul, ofReal_mul, ofReal_div, ofReal_neg, ofReal_one,
ofReal_ofNat, zero_add, σ2_σ3_trace, add_zero, neg_re, mul_re, div_ofNat_re, one_re, ofReal_re,
div_ofNat_im, neg_im, one_im, zero_div, ofReal_im, mul_zero, sub_zero, re_ofNat, mul_im,
zero_mul, im_ofNat]
ring
· simp only [σSAL', AddSubgroup.coe_add, selfAdjoint.val_smul, smul_neg, mul_add,
Algebra.mul_smul_comm, mul_neg, trace_add, trace_smul, σ3_σ0_trace, smul_zero, trace_neg,
σ3_σ1_trace, neg_zero, σ3_σ2_trace, add_zero, σ3_σ3_trace, real_smul, ofReal_mul, ofReal_div,
ofReal_neg, ofReal_one, ofReal_ofNat, zero_add, neg_re, mul_re, div_ofNat_re, one_re, ofReal_re,
div_ofNat_im, neg_im, one_im, zero_div, ofReal_im, mul_zero, sub_zero, re_ofNat, mul_im,
zero_mul, im_ofNat]
ring
/-- The basis of `selfAdjoint (Matrix (Fin 2) (Fin 2) )` formed by Pauli matrices
where the `1, 2, 3` pauli matrices are negated. -/
def σSAL : Basis (Fin 1 ⊕ Fin 3) (selfAdjoint (Matrix (Fin 2) (Fin 2) )) :=
Basis.mk σSAL_linearly_independent σSAL_span
lemma σSAL_decomp (M : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
M = (1/2 * (Matrix.trace (σ0 * M.1)).re) • σSAL (Sum.inl 0)
+ (-1/2 * (Matrix.trace (σ1 * M.1)).re) • σSAL (Sum.inr 0)
+ (-1/2 * (Matrix.trace (σ2 * M.1)).re) • σSAL (Sum.inr 1)
+ (-1/2 * (Matrix.trace (σ3 * M.1)).re) • σSAL (Sum.inr 2) := by
apply selfAdjoint_ext
· simp only [one_div, σSAL, Fin.isValue, Basis.coe_mk, σSAL', AddSubgroup.coe_add,
selfAdjoint.val_smul, smul_neg, mul_add, Algebra.mul_smul_comm, mul_neg, trace_add, trace_smul,
σ0_σ0_trace, real_smul, ofReal_mul, ofReal_inv, ofReal_ofNat, trace_neg, σ0_σ1_trace, smul_zero,
neg_zero, add_zero, σ0_σ2_trace, σ0_σ3_trace, mul_re, inv_re, re_ofNat, normSq_ofNat,
div_self_mul_self', ofReal_re, inv_im, im_ofNat, zero_div, ofReal_im, mul_zero, sub_zero,
mul_im, zero_mul]
ring
· simp only [one_div, σSAL, Fin.isValue, Basis.coe_mk, σSAL', AddSubgroup.coe_add,
selfAdjoint.val_smul, smul_neg, mul_add, Algebra.mul_smul_comm, mul_neg, trace_add, trace_smul,
σ1_σ0_trace, smul_zero, trace_neg, σ1_σ1_trace, real_smul, ofReal_mul, ofReal_div, ofReal_neg,
ofReal_one, ofReal_ofNat, zero_add, σ1_σ2_trace, neg_zero, add_zero, σ1_σ3_trace, neg_re,
mul_re, div_ofNat_re, one_re, ofReal_re, div_ofNat_im, neg_im, one_im, zero_div, ofReal_im,
mul_zero, sub_zero, re_ofNat, mul_im, zero_mul, im_ofNat]
ring
· simp only [one_div, σSAL, Fin.isValue, Basis.coe_mk, σSAL', AddSubgroup.coe_add,
selfAdjoint.val_smul, smul_neg, mul_add, Algebra.mul_smul_comm, mul_neg, trace_add, trace_smul,
σ2_σ0_trace, smul_zero, trace_neg, σ2_σ1_trace, neg_zero, add_zero, σ2_σ2_trace, real_smul,
ofReal_mul, ofReal_div, ofReal_neg, ofReal_one, ofReal_ofNat, zero_add, σ2_σ3_trace, neg_re,
mul_re, div_ofNat_re, one_re, ofReal_re, div_ofNat_im, neg_im, one_im, zero_div, ofReal_im,
mul_zero, sub_zero, re_ofNat, mul_im, zero_mul, im_ofNat]
ring
· simp only [one_div, σSAL, Fin.isValue, Basis.coe_mk, σSAL', AddSubgroup.coe_add,
selfAdjoint.val_smul, smul_neg, mul_add, Algebra.mul_smul_comm, mul_neg, trace_add, trace_smul,
σ3_σ0_trace, smul_zero, trace_neg, σ3_σ1_trace, neg_zero, add_zero, σ3_σ2_trace, σ3_σ3_trace,
real_smul, ofReal_mul, ofReal_div, ofReal_neg, ofReal_one, ofReal_ofNat, zero_add, neg_re,
mul_re, div_ofNat_re, one_re, ofReal_re, div_ofNat_im, neg_im, one_im, zero_div, ofReal_im,
mul_zero, sub_zero, re_ofNat, mul_im, zero_mul, im_ofNat]
ring
@[simp]
lemma σSAL_repr_inl_0 (M : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
σSAL.repr M (Sum.inl 0) = 1 / 2 * Matrix.trace (σ0 * M.1) := by
have hM : M = ∑ i, σSAL.repr M i • σSAL i := (Basis.sum_repr σSAL M).symm
simp only [Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton, Fin.sum_univ_three] at hM
have h0 := congrArg (fun A => Matrix.trace (σ0 * A.1)/ 2) hM
simp only [σSAL, Basis.mk_repr, Fin.isValue, Basis.coe_mk, σSAL', AddSubgroup.coe_add,
selfAdjoint.val_smul, smul_neg, mul_add, Algebra.mul_smul_comm, mul_neg, trace_add, trace_smul,
σ0_σ0_trace, real_smul, trace_neg, σ0_σ1_trace, smul_zero, neg_zero, σ0_σ2_trace, add_zero,
σ0_σ3_trace, isUnit_iff_ne_zero, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true,
IsUnit.mul_div_cancel_right] at h0
linear_combination (norm := ring_nf) -h0
simp [σSAL]
@[simp]
lemma σSAL_repr_inr_0 (M : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
σSAL.repr M (Sum.inr 0) = - 1 / 2 * Matrix.trace (σ1 * M.1) := by
have hM : M = ∑ i, σSAL.repr M i • σSAL i := (Basis.sum_repr σSAL M).symm
simp only [Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton, Fin.sum_univ_three] at hM
have h0 := congrArg (fun A => - Matrix.trace (σ1 * A.1)/ 2) hM
simp only [σSAL, Basis.mk_repr, Fin.isValue, Basis.coe_mk, σSAL', AddSubgroup.coe_add,
selfAdjoint.val_smul, smul_neg, mul_add, Algebra.mul_smul_comm, mul_neg, trace_add, trace_smul,
σ1_σ0_trace, smul_zero, trace_neg, σ1_σ1_trace, real_smul, σ1_σ2_trace, neg_zero, add_zero,
σ1_σ3_trace, zero_add, neg_neg, isUnit_iff_ne_zero, ne_eq, OfNat.ofNat_ne_zero,
not_false_eq_true, IsUnit.mul_div_cancel_right] at h0
linear_combination (norm := ring_nf) -h0
simp [σSAL]
@[simp]
lemma σSAL_repr_inr_1 (M : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
σSAL.repr M (Sum.inr 1) = - 1 / 2 * Matrix.trace (σ2 * M.1) := by
have hM : M = ∑ i, σSAL.repr M i • σSAL i := (Basis.sum_repr σSAL M).symm
simp only [Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton, Fin.sum_univ_three] at hM
have h0 := congrArg (fun A => - Matrix.trace (σ2 * A.1)/ 2) hM
simp only [σSAL, Basis.mk_repr, Fin.isValue, Basis.coe_mk, σSAL', AddSubgroup.coe_add,
selfAdjoint.val_smul, smul_neg, mul_add, Algebra.mul_smul_comm, mul_neg, trace_add, trace_smul,
σ2_σ0_trace, smul_zero, trace_neg, σ2_σ1_trace, neg_zero, σ2_σ2_trace, real_smul, zero_add,
σ2_σ3_trace, add_zero, neg_neg, isUnit_iff_ne_zero, ne_eq, OfNat.ofNat_ne_zero,
not_false_eq_true, IsUnit.mul_div_cancel_right] at h0
linear_combination (norm := ring_nf) -h0
simp [σSAL]
@[simp]
lemma σSAL_repr_inr_2 (M : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
σSAL.repr M (Sum.inr 2) = - 1 / 2 * Matrix.trace (σ3 * M.1) := by
have hM : M = ∑ i, σSAL.repr M i • σSAL i := (Basis.sum_repr σSAL M).symm
simp only [Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton, Fin.sum_univ_three] at hM
have h0 := congrArg (fun A => - Matrix.trace (σ3 * A.1)/ 2) hM
simp [σSAL, σSAL', mul_add] at h0
linear_combination (norm := ring_nf) -h0
simp [σSAL]
lemma σSA_minkowskiMetric_σSAL (i : Fin 1 ⊕ Fin 3) :
(σSA i) = minkowskiMatrix i i • (σSAL i) := by
match i with
| Sum.inl 0 =>
simp [σSA, σSAL, σSA', σSAL', minkowskiMatrix.inl_0_inl_0]
| Sum.inr 0 =>
simp only [σSA, Fin.isValue, Basis.coe_mk, σSA', minkowskiMatrix.inr_i_inr_i, σSAL, σSAL',
neg_smul, one_smul]
cases i with
| inl val =>
ext i j : 2
simp_all only [NegMemClass.coe_neg, neg_apply, neg_neg]
| inr val_1 =>
ext i j : 2
simp_all only [NegMemClass.coe_neg, neg_apply, neg_neg]
| Sum.inr 1 =>
simp only [σSA, Fin.isValue, Basis.coe_mk, σSA', minkowskiMatrix.inr_i_inr_i, σSAL, σSAL',
neg_smul, one_smul]
cases i with
| inl val =>
ext i j : 2
simp_all only [NegMemClass.coe_neg, neg_apply, neg_neg]
| inr val_1 =>
ext i j : 2
simp_all only [NegMemClass.coe_neg, neg_apply, neg_neg]
| Sum.inr 2 =>
simp only [σSA, Fin.isValue, Basis.coe_mk, σSA', minkowskiMatrix.inr_i_inr_i, σSAL, σSAL',
neg_smul, one_smul]
cases i with
| inl val =>
ext i j : 2
simp_all only [NegMemClass.coe_neg, neg_apply, neg_neg]
| inr val_1 =>
ext i j : 2
simp_all only [NegMemClass.coe_neg, neg_apply, neg_neg]
end
end PauliMatrix

View file

@ -43,6 +43,7 @@ lemma inverse_coe (M : SL(2, )) : M.1⁻¹ = (M⁻¹).1 := by
· simp
· simp
lemma transpose_coe (M : SL(2, )) : M.1ᵀ = (M.transpose).1 := rfl
/-!
## Representation of SL(2, ) on spacetime
@ -133,6 +134,64 @@ def toLorentzGroup : SL(2, ) →* LorentzGroup 3 where
simp only [toLorentzGroupElem, _root_.map_mul, LinearMap.toMatrix_mul,
lorentzGroupIsGroup_mul_coe]
lemma toLorentzGroup_eq_σSAL (M : SL(2, )) :
toLorentzGroup M = LinearMap.toMatrix
PauliMatrix.σSAL PauliMatrix.σSAL (repSelfAdjointMatrix M) := by
rfl
lemma toLorentzGroup_eq_stdBasis (M : SL(2, )) :
toLorentzGroup M = LinearMap.toMatrix LorentzVector.stdBasis LorentzVector.stdBasis
(repLorentzVector M) := by rfl
lemma repLorentzVector_apply_eq_mulVec (v : LorentzVector 3) :
SL2C.repLorentzVector M v = (SL2C.toLorentzGroup M).1 *ᵥ v := by
simp only [toLorentzGroup, MonoidHom.coe_mk, OneHom.coe_mk, toLorentzGroupElem_coe]
have hv : v = (Finsupp.linearEquivFunOnFinite (Fin 1 ⊕ Fin 3))
(LorentzVector.stdBasis.repr v) := by rfl
nth_rewrite 2 [hv]
change _ = toLorentzGroup M *ᵥ (LorentzVector.stdBasis.repr v)
rw [toLorentzGroup_eq_stdBasis, LinearMap.toMatrix_mulVec_repr]
rfl
lemma repSelfAdjointMatrix_basis (i : Fin 1 ⊕ Fin 3) :
SL2C.repSelfAdjointMatrix M (PauliMatrix.σSAL i) =
∑ j, (toLorentzGroup M).1 j i •
PauliMatrix.σSAL j := by
rw [toLorentzGroup_eq_σSAL]
simp only [LinearMap.toMatrix_apply, Finset.univ_unique,
Fin.default_eq_zero, Fin.isValue, Finset.sum_singleton]
nth_rewrite 1 [← (Basis.sum_repr PauliMatrix.σSAL
((repSelfAdjointMatrix M) (PauliMatrix.σSAL i)))]
congr
lemma repSelfAdjointMatrix_σSA (i : Fin 1 ⊕ Fin 3) :
SL2C.repSelfAdjointMatrix M (PauliMatrix.σSA i) =
∑ j, (toLorentzGroup M⁻¹).1 i j • PauliMatrix.σSA j := by
have h1 : (toLorentzGroup M⁻¹).1 = minkowskiMetric.dual (toLorentzGroup M).1 := by
simp
simp only [h1]
rw [PauliMatrix.σSA_minkowskiMetric_σSAL, _root_.map_smul]
rw [repSelfAdjointMatrix_basis]
rw [Finset.smul_sum]
apply congrArg
funext j
rw [smul_smul, PauliMatrix.σSA_minkowskiMetric_σSAL, smul_smul]
apply congrFun
apply congrArg
exact Eq.symm (minkowskiMetric.dual_apply_minkowskiMatrix ((toLorentzGroup M).1) i j)
lemma repLorentzVector_stdBasis (i : Fin 1 ⊕ Fin 3) :
SL2C.repLorentzVector M (LorentzVector.stdBasis i) =
∑ j, (toLorentzGroup M).1 j i • LorentzVector.stdBasis j := by
simp only [repLorentzVector, MonoidHom.coe_mk, OneHom.coe_mk, LinearMap.coe_comp,
LinearEquiv.coe_coe, Function.comp_apply]
rw [toSelfAdjointMatrix_stdBasis]
rw [repSelfAdjointMatrix_basis]
rw [map_sum]
apply congrArg
funext j
simp
/-!
## Homomorphism to the restricted Lorentz group

View file

@ -9,7 +9,6 @@ import Mathlib.RepresentationTheory.Rep
import HepLean.Tensors.Basic
import HepLean.SpaceTime.WeylFermion.Modules
import Mathlib.Logic.Equiv.TransferInstance
import LLMlean
/-!
# Weyl fermions
@ -28,7 +27,7 @@ open Complex
open TensorProduct
/-- The vector space ^2 carrying the fundamental representation of SL(2,C).
In index notation corresponds to a Weyl fermion with indices ψ_a. -/
In index notation corresponds to a Weyl fermion with indices ψ^a. -/
def leftHanded : Rep SL(2,) := Rep.of {
toFun := fun M => {
toFun := fun (ψ : LeftHandedModule) =>
@ -56,12 +55,12 @@ def leftBasis : Basis (Fin 2) leftHanded := Basis.ofEquivFun
lemma leftBasis_ρ_apply (M : SL(2,)) (i j : Fin 2) :
(LinearMap.toMatrix leftBasis leftBasis) (leftHanded.ρ M) i j = M.1 i j := by
rw [LinearMap.toMatrix_apply]
simp [leftBasis]
simp only [leftBasis, Basis.coe_ofEquivFun, Basis.ofEquivFun_repr_apply]
change (M.1 *ᵥ (Pi.single j 1)) i = _
simp only [mulVec_single, mul_one]
/-- The vector space ^2 carrying the representation of SL(2,C) given by
M → (M⁻¹)ᵀ. In index notation corresponds to a Weyl fermion with indices ψ^a. -/
M → (M⁻¹)ᵀ. In index notation corresponds to a Weyl fermion with indices ψ_a. -/
def altLeftHanded : Rep SL(2,) := Rep.of {
toFun := fun M => {
toFun := fun (ψ : AltLeftHandedModule) =>
@ -96,7 +95,7 @@ lemma altLeftBasis_ρ_apply (M : SL(2,)) (i j : Fin 2) :
simp only [mulVec_single, transpose_apply, mul_one]
/-- The vector space ^2 carrying the conjugate representation of SL(2,C).
In index notation corresponds to a Weyl fermion with indices ψ_{dot a}. -/
In index notation corresponds to a Weyl fermion with indices ψ^{dot a}. -/
def rightHanded : Rep SL(2,) := Rep.of {
toFun := fun M => {
toFun := fun (ψ : RightHandedModule) =>
@ -129,7 +128,7 @@ lemma rightBasis_ρ_apply (M : SL(2,)) (i j : Fin 2) :
/-- The vector space ^2 carrying the representation of SL(2,C) given by
M → (M⁻¹)^†.
In index notation this corresponds to a Weyl fermion with index `ψ^{dot a}`. -/
In index notation this corresponds to a Weyl fermion with index `ψ_{dot a}`. -/
def altRightHanded : Rep SL(2,) := Rep.of {
toFun := fun M => {
toFun := fun (ψ : AltRightHandedModule) =>

View file

@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.WeylFermion.Basic
import LLMlean
/-!
# Contraction of Weyl fermions
@ -21,7 +20,6 @@ open MatrixGroups
open Complex
open TensorProduct
/-!
## Contraction of Weyl fermions.
@ -76,7 +74,6 @@ def altLeftBi : altLeftHanded →ₗ[] leftHanded →ₗ[] where
simp only [_root_.map_smul, smul_dotProduct, vec2_dotProduct, Fin.isValue, smul_eq_mul,
LinearMap.coe_mk, AddHom.coe_mk, RingHom.id_apply, LinearMap.smul_apply]
/-- The bi-linear map corresponding to contraction of a right-handed Weyl fermion with a
alt-right-handed Weyl fermion. -/
def rightAltBi : rightHanded →ₗ[] altRightHanded →ₗ[] where
@ -128,7 +125,7 @@ def altRightBi : altRightHanded →ₗ[] rightHanded →ₗ[] where
summing over components of leftHandedWeyl and altLeftHandedWeyl in the
standard basis (i.e. the dot product).
Physically, the contraction of a left-handed Weyl fermion with a alt-left-handed Weyl fermion.
In index notation this is ψ_a φ^a. -/
In index notation this is ψ^a φ_a. -/
def leftAltContraction : leftHanded ⊗ altLeftHanded ⟶ 𝟙_ (Rep SL(2,)) where
hom := TensorProduct.lift leftAltBi
comm M := TensorProduct.ext' fun ψ φ => by
@ -146,7 +143,7 @@ lemma leftAltContraction_hom_tmul (ψ : leftHanded) (φ : altLeftHanded) :
summing over components of altLeftHandedWeyl and leftHandedWeyl in the
standard basis (i.e. the dot product).
Physically, the contraction of a alt-left-handed Weyl fermion with a left-handed Weyl fermion.
In index notation this is φ^a ψ_a. -/
In index notation this is φ_a ψ^a. -/
def altLeftContraction : altLeftHanded ⊗ leftHanded ⟶ 𝟙_ (Rep SL(2,)) where
hom := TensorProduct.lift altLeftBi
comm M := TensorProduct.ext' fun φ ψ => by
@ -165,12 +162,13 @@ The linear map from rightHandedWeyl ⊗ altRightHandedWeyl to given by
summing over components of rightHandedWeyl and altRightHandedWeyl in the
standard basis (i.e. the dot product).
The contraction of a right-handed Weyl fermion with a left-handed Weyl fermion.
In index notation this is ψ_{dot a} φ^{dot a}.
In index notation this is ψ^{dot a} φ_{dot a}.
-/
def rightAltContraction : rightHanded ⊗ altRightHanded ⟶ 𝟙_ (Rep SL(2,)) where
hom := TensorProduct.lift rightAltBi
comm M := TensorProduct.ext' fun ψ φ => by
change (M.1.map star *ᵥ ψ.toFin2) ⬝ᵥ (M.1⁻¹.conjTranspose *ᵥ φ.toFin2) = ψ.toFin2 ⬝ᵥ φ.toFin2
change (M.1.map star *ᵥ ψ.toFin2) ⬝ᵥ (M.1⁻¹.conjTranspose *ᵥ φ.toFin2) =
ψ.toFin2 ⬝ᵥ φ.toFin2
have h1 : (M.1)⁻¹ᴴ = ((M.1)⁻¹.map star)ᵀ := by rfl
rw [dotProduct_mulVec, h1, vecMul_transpose, mulVec_mulVec]
have h2 : ((M.1)⁻¹.map star * (M.1).map star) = 1 := by
@ -189,12 +187,13 @@ def rightAltContraction : rightHanded ⊗ altRightHanded ⟶ 𝟙_ (Rep SL(2
summing over components of altRightHandedWeyl and rightHandedWeyl in the
standard basis (i.e. the dot product).
The contraction of a right-handed Weyl fermion with a left-handed Weyl fermion.
In index notation this is φ^{dot a} ψ_{dot a}.
In index notation this is φ_{dot a} ψ^{dot a}.
-/
def altRightContraction : altRightHanded ⊗ rightHanded ⟶ 𝟙_ (Rep SL(2,)) where
hom := TensorProduct.lift altRightBi
comm M := TensorProduct.ext' fun φ ψ => by
change (M.1⁻¹.conjTranspose *ᵥ φ.toFin2) ⬝ᵥ (M.1.map star *ᵥ ψ.toFin2) = φ.toFin2 ⬝ᵥ ψ.toFin2
comm M := TensorProduct.ext' fun φ ψ => by
change (M.1⁻¹.conjTranspose *ᵥ φ.toFin2) ⬝ᵥ (M.1.map star *ᵥ ψ.toFin2) =
φ.toFin2 ⬝ᵥ ψ.toFin2
have h1 : (M.1)⁻¹ᴴ = ((M.1)⁻¹.map star)ᵀ := by rfl
rw [dotProduct_mulVec, h1, mulVec_transpose, vecMul_vecMul]
have h2 : ((M.1)⁻¹.map star * (M.1).map star) = 1 := by
@ -251,7 +250,6 @@ informal_lemma altLeftWeylContraction_invariant where
the action of SL(2,C) on leftHandedWeyl and altLeftHandedWeyl."
deps :≈ [``altLeftContraction]
informal_lemma rightAltWeylContraction_invariant where
math :≈ "The contraction rightAltWeylContraction is invariant with respect to
the action of SL(2,C) on rightHandedWeyl and altRightHandedWeyl."

View file

@ -13,7 +13,7 @@ import HepLean.SpaceTime.WeylFermion.Two
We define the metrics for Weyl fermions, often denoted `ε` in the literature.
These allow us to go from left-handed to alt-left-handed Weyl fermions and back,
and from right-handed to alt-right-handed Weyl fermions and back.
and from right-handed to alt-right-handed Weyl fermions and back.
-/
@ -26,6 +26,7 @@ open Complex
open TensorProduct
open CategoryTheory.MonoidalCategory
/-- The raw `2x2` matrix corresponding to the metric for fermions. -/
def metricRaw : Matrix (Fin 2) (Fin 2) := !![0, 1; -1, 0]
lemma comm_metricRaw (M : SL(2,)) : M.1 * metricRaw = metricRaw * (M.1⁻¹)ᵀ := by
@ -49,7 +50,7 @@ lemma metricRaw_comm (M : SL(2,)) : metricRaw * M.1 = (M.1⁻¹)ᵀ * metricR
mul_one, smul_empty, tail_cons, neg_smul, mul_neg, neg_cons, neg_neg, neg_zero, neg_empty,
empty_vecMul, add_cons, empty_add_empty, empty_mul, Equiv.symm_apply_apply]
lemma star_comm_metricRaw (M : SL(2,)) : M.1.map star * metricRaw = metricRaw * ((M.1)⁻¹)ᴴ := by
lemma star_comm_metricRaw (M : SL(2,)) : M.1.map star * metricRaw = metricRaw * ((M.1)⁻¹)ᴴ := by
rw [metricRaw]
rw [SpaceTime.SL2C.inverse_coe, eta_fin_two M.1]
rw [SpecialLinearGroup.coe_inv, Matrix.adjugate_fin_two,
@ -83,13 +84,15 @@ def leftMetric : 𝟙_ (Rep SL(2,)) ⟶ leftHanded ⊗ leftHanded where
rfl}
comm M := by
ext x : 2
simp
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • leftMetricVal =
(TensorProduct.map (leftHanded.ρ M) (leftHanded.ρ M)) (x' • leftMetricVal)
simp
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
simp [leftMetricVal]
simp only [Action.instMonoidalCategory_tensorObj_V, leftMetricVal, map_neg, neg_inj]
erw [leftLeftToMatrix_ρ_symm]
apply congrArg
rw [comm_metricRaw, mul_assoc, ← @transpose_mul]
@ -114,20 +117,21 @@ def altLeftMetric : 𝟙_ (Rep SL(2,)) ⟶ altLeftHanded ⊗ altLeftHande
rfl}
comm M := by
ext x : 2
simp
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • altLeftMetricVal =
(TensorProduct.map (altLeftHanded.ρ M) (altLeftHanded.ρ M)) (x' • altLeftMetricVal)
simp
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
simp [altLeftMetricVal]
simp only [Action.instMonoidalCategory_tensorObj_V, altLeftMetricVal]
erw [altLeftaltLeftToMatrix_ρ_symm]
apply congrArg
rw [← metricRaw_comm, mul_assoc]
simp only [SpecialLinearGroup.det_coe, isUnit_iff_ne_zero, ne_eq, one_ne_zero,
not_false_eq_true, mul_nonsing_inv, mul_one]
/-- The metric `ε_{dot a}_{dot a}` as an element of `(rightHanded ⊗ rightHanded).V`. -/
def rightMetricVal : (rightHanded ⊗ rightHanded).V :=
rightRightToMatrix.symm (- metricRaw)
@ -146,7 +150,9 @@ def rightMetric : 𝟙_ (Rep SL(2,)) ⟶ rightHanded ⊗ rightHanded wher
rfl}
comm M := by
ext x : 2
simp
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • rightMetricVal =
(TensorProduct.map (rightHanded.ρ M) (rightHanded.ρ M)) (x' • rightMetricVal)
@ -186,7 +192,9 @@ def altRightMetric : 𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ altRightHa
rfl}
comm M := by
ext x : 2
simp
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • altRightMetricVal =
(TensorProduct.map (altRightHanded.ρ M) (altRightHanded.ρ M)) (x' • altRightMetricVal)
@ -201,7 +209,7 @@ def altRightMetric : 𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ altRightHa
have h1 : ((M.1).map star * (M.1)⁻¹ᴴᵀ) = 1 := by
refine transpose_eq_one.mp ?_
rw [@transpose_mul]
simp
simp only [transpose_transpose, RCLike.star_def]
change (M.1)⁻¹ᴴ * (M.1)ᴴ = 1
rw [← @conjTranspose_mul]
simp

View file

@ -10,7 +10,6 @@ import Mathlib.LinearAlgebra.TensorProduct.Matrix
# Tensor product of two Weyl fermion
-/
namespace Fermion
@ -76,6 +75,18 @@ def altRightRightToMatrix : (altRightHanded ⊗ rightHanded).V ≃ₗ[] Matri
Finsupp.linearEquivFunOnFinite (Fin 2 × Fin 2) ≪≫ₗ
LinearEquiv.curry (Fin 2) (Fin 2)
/-- Equivalence of `altLeftHanded ⊗ altRightHanded` to `2 x 2` complex matrices. -/
def altLeftAltRightToMatrix : (altLeftHanded ⊗ altRightHanded).V ≃ₗ[] Matrix (Fin 2) (Fin 2) :=
(Basis.tensorProduct altLeftBasis altRightBasis).repr ≪≫ₗ
Finsupp.linearEquivFunOnFinite (Fin 2 × Fin 2) ≪≫ₗ
LinearEquiv.curry (Fin 2) (Fin 2)
/-- Equivalence of `leftHanded ⊗ rightHanded` to `2 x 2` complex matrices. -/
def leftRightToMatrix : (leftHanded ⊗ rightHanded).V ≃ₗ[] Matrix (Fin 2) (Fin 2) :=
(Basis.tensorProduct leftBasis rightBasis).repr ≪≫ₗ
Finsupp.linearEquivFunOnFinite (Fin 2 × Fin 2) ≪≫ₗ
LinearEquiv.curry (Fin 2) (Fin 2)
/-!
## Group actions
@ -96,7 +107,7 @@ lemma leftLeftToMatrix_ρ (v : (leftHanded ⊗ leftHanded).V) (M : SL(2,)) :
((leftBasis.tensorProduct leftBasis).repr (v))))
· apply congrArg
have h1 := (LinearMap.toMatrix_mulVec_repr (leftBasis.tensorProduct leftBasis)
(leftBasis.tensorProduct leftBasis) (TensorProduct.map (leftHanded.ρ M) (leftHanded.ρ M)) v)
(leftBasis.tensorProduct leftBasis) (TensorProduct.map (leftHanded.ρ M) (leftHanded.ρ M)) v)
erw [h1]
rfl
rw [TensorProduct.toMatrix_map]
@ -108,7 +119,7 @@ lemma leftLeftToMatrix_ρ (v : (leftHanded ⊗ leftHanded).V) (M : SL(2,)) :
erw [Finset.sum_product]
simp_rw [kroneckerMap_apply, Matrix.mul_apply, Matrix.transpose_apply]
have h1 : ∑ x : Fin 2, (∑ j : Fin 2, M.1 i j * leftLeftToMatrix v j x) * M.1 j x
= ∑ x : Fin 2, ∑ x1 : Fin 2, (M.1 i x1 * leftLeftToMatrix v x1 x) * M.1 j x := by
= ∑ x : Fin 2, ∑ x1 : Fin 2, (M.1 i x1 * leftLeftToMatrix v x1 x) * M.1 j x := by
congr
funext x
rw [Finset.sum_mul]
@ -138,8 +149,8 @@ lemma altLeftaltLeftToMatrix_ρ (v : (altLeftHanded ⊗ altLeftHanded).V) (M : S
((altLeftBasis.tensorProduct altLeftBasis).repr v)))
· apply congrArg
have h1 := (LinearMap.toMatrix_mulVec_repr (altLeftBasis.tensorProduct altLeftBasis)
(altLeftBasis.tensorProduct altLeftBasis)
(TensorProduct.map (altLeftHanded.ρ M) (altLeftHanded.ρ M)) v)
(altLeftBasis.tensorProduct altLeftBasis)
(TensorProduct.map (altLeftHanded.ρ M) (altLeftHanded.ρ M)) v)
erw [h1]
rfl
rw [TensorProduct.toMatrix_map]
@ -178,8 +189,8 @@ lemma leftAltLeftToMatrix_ρ (v : (leftHanded ⊗ altLeftHanded).V) (M : SL(2,
((leftBasis.tensorProduct altLeftBasis).repr (v))))
· apply congrArg
have h1 := (LinearMap.toMatrix_mulVec_repr (leftBasis.tensorProduct altLeftBasis)
(leftBasis.tensorProduct altLeftBasis)
(TensorProduct.map (leftHanded.ρ M) (altLeftHanded.ρ M)) v)
(leftBasis.tensorProduct altLeftBasis)
(TensorProduct.map (leftHanded.ρ M) (altLeftHanded.ρ M)) v)
erw [h1]
rfl
rw [TensorProduct.toMatrix_map]
@ -219,8 +230,8 @@ lemma altLeftLeftToMatrix_ρ (v : (altLeftHanded ⊗ leftHanded).V) (M : SL(2,
((altLeftBasis.tensorProduct leftBasis).repr (v))))
· apply congrArg
have h1 := (LinearMap.toMatrix_mulVec_repr (altLeftBasis.tensorProduct leftBasis)
(altLeftBasis.tensorProduct leftBasis)
(TensorProduct.map (altLeftHanded.ρ M) (leftHanded.ρ M)) v)
(altLeftBasis.tensorProduct leftBasis)
(TensorProduct.map (altLeftHanded.ρ M) (leftHanded.ρ M)) v)
erw [h1]
rfl
rw [TensorProduct.toMatrix_map]
@ -260,7 +271,8 @@ lemma rightRightToMatrix_ρ (v : (rightHanded ⊗ rightHanded).V) (M : SL(2,)
((rightBasis.tensorProduct rightBasis).repr (v))))
· apply congrArg
have h1 := (LinearMap.toMatrix_mulVec_repr (rightBasis.tensorProduct rightBasis)
(rightBasis.tensorProduct rightBasis) (TensorProduct.map (rightHanded.ρ M) (rightHanded.ρ M)) v)
(rightBasis.tensorProduct rightBasis)
(TensorProduct.map (rightHanded.ρ M) (rightHanded.ρ M)) v)
erw [h1]
rfl
rw [TensorProduct.toMatrix_map]
@ -271,8 +283,9 @@ lemma rightRightToMatrix_ρ (v : (rightHanded ⊗ rightHanded).V) (M : SL(2,)
* rightRightToMatrix v k.1 k.2) = _
erw [Finset.sum_product]
simp_rw [kroneckerMap_apply, Matrix.mul_apply, Matrix.transpose_apply]
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, (M.1.map star) i x1 * rightRightToMatrix v x1 x) * (M.1.map star) j x
= ∑ x : Fin 2, ∑ x1 : Fin 2, ((M.1.map star) i x1 * rightRightToMatrix v x1 x) * (M.1.map star) j x:= by
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, (M.1.map star) i x1 * rightRightToMatrix v x1 x) *
(M.1.map star) j x = ∑ x : Fin 2, ∑ x1 : Fin 2,
((M.1.map star) i x1 * rightRightToMatrix v x1 x) * (M.1.map star) j x:= by
congr
funext x
rw [Finset.sum_mul]
@ -300,8 +313,8 @@ lemma altRightAltRightToMatrix_ρ (v : (altRightHanded ⊗ altRightHanded).V) (M
((altRightBasis.tensorProduct altRightBasis).repr (v))))
· apply congrArg
have h1 := (LinearMap.toMatrix_mulVec_repr (altRightBasis.tensorProduct altRightBasis)
(altRightBasis.tensorProduct altRightBasis)
(TensorProduct.map (altRightHanded.ρ M) (altRightHanded.ρ M)) v)
(altRightBasis.tensorProduct altRightBasis)
(TensorProduct.map (altRightHanded.ρ M) (altRightHanded.ρ M)) v)
erw [h1]
rfl
rw [TensorProduct.toMatrix_map]
@ -312,8 +325,9 @@ lemma altRightAltRightToMatrix_ρ (v : (altRightHanded ⊗ altRightHanded).V) (M
* altRightAltRightToMatrix v k.1 k.2) = _
erw [Finset.sum_product]
simp_rw [kroneckerMap_apply, Matrix.mul_apply, Matrix.transpose_apply]
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, (↑M)⁻¹ᴴ i x1 * altRightAltRightToMatrix v x1 x) * (↑M)⁻¹ᴴ j x
= ∑ x : Fin 2, ∑ x1 : Fin 2, ((↑M)⁻¹ᴴ i x1 * altRightAltRightToMatrix v x1 x) * (↑M)⁻¹ᴴ j x := by
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, (↑M)⁻¹ᴴ i x1 * altRightAltRightToMatrix v x1 x) *
(↑M)⁻¹ᴴ j x = ∑ x : Fin 2, ∑ x1 : Fin 2,
((↑M)⁻¹ᴴ i x1 * altRightAltRightToMatrix v x1 x) * (↑M)⁻¹ᴴ j x := by
congr
funext x
rw [Finset.sum_mul]
@ -340,8 +354,8 @@ lemma rightAltRightToMatrix_ρ (v : (rightHanded ⊗ altRightHanded).V) (M : SL(
((rightBasis.tensorProduct altRightBasis).repr (v))))
· apply congrArg
have h1 := (LinearMap.toMatrix_mulVec_repr (rightBasis.tensorProduct altRightBasis)
(rightBasis.tensorProduct altRightBasis)
(TensorProduct.map (rightHanded.ρ M) (altRightHanded.ρ M)) v)
(rightBasis.tensorProduct altRightBasis)
(TensorProduct.map (rightHanded.ρ M) (altRightHanded.ρ M)) v)
erw [h1]
rfl
rw [TensorProduct.toMatrix_map]
@ -352,8 +366,9 @@ lemma rightAltRightToMatrix_ρ (v : (rightHanded ⊗ altRightHanded).V) (M : SL(
* rightAltRightToMatrix v k.1 k.2) = _
erw [Finset.sum_product]
simp_rw [kroneckerMap_apply, Matrix.mul_apply, Matrix.transpose_apply]
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, (M.1.map star) i x1 * rightAltRightToMatrix v x1 x) * (↑M)⁻¹ᴴ j x
= ∑ x : Fin 2, ∑ x1 : Fin 2, ((M.1.map star) i x1 * rightAltRightToMatrix v x1 x) * (↑M)⁻¹ᴴ j x := by
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, (M.1.map star) i x1 * rightAltRightToMatrix v x1 x)
* (↑M)⁻¹ᴴ j x = ∑ x : Fin 2, ∑ x1 : Fin 2,
((M.1.map star) i x1 * rightAltRightToMatrix v x1 x) * (↑M)⁻¹ᴴ j x := by
congr
funext x
rw [Finset.sum_mul]
@ -381,8 +396,8 @@ lemma altRightRightToMatrix_ρ (v : (altRightHanded ⊗ rightHanded).V) (M : SL(
((altRightBasis.tensorProduct rightBasis).repr (v))))
· apply congrArg
have h1 := (LinearMap.toMatrix_mulVec_repr (altRightBasis.tensorProduct rightBasis)
(altRightBasis.tensorProduct rightBasis)
(TensorProduct.map (altRightHanded.ρ M) (rightHanded.ρ M)) v)
(altRightBasis.tensorProduct rightBasis)
(TensorProduct.map (altRightHanded.ρ M) (rightHanded.ρ M)) v)
erw [h1]
rfl
rw [TensorProduct.toMatrix_map]
@ -393,8 +408,10 @@ lemma altRightRightToMatrix_ρ (v : (altRightHanded ⊗ rightHanded).V) (M : SL(
* altRightRightToMatrix v k.1 k.2) = _
erw [Finset.sum_product]
simp_rw [kroneckerMap_apply, Matrix.mul_apply, Matrix.transpose_apply]
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, (↑M)⁻¹ᴴ i x1 * altRightRightToMatrix v x1 x) * (M.1.map star) j x
= ∑ x : Fin 2, ∑ x1 : Fin 2, ((↑M)⁻¹ᴴ i x1 * altRightRightToMatrix v x1 x) * (M.1.map star) j x := by
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2,
(↑M)⁻¹ᴴ i x1 * altRightRightToMatrix v x1 x) * (M.1.map star) j x
= ∑ x : Fin 2, ∑ x1 : Fin 2, ((↑M)⁻¹ᴴ i x1 * altRightRightToMatrix v x1 x) *
(M.1.map star) j x := by
congr
funext x
rw [Finset.sum_mul]
@ -408,6 +425,87 @@ lemma altRightRightToMatrix_ρ (v : (altRightHanded ⊗ rightHanded).V) (M : SL(
Action.instMonoidalCategory_tensorObj_V]
ring
lemma altLeftAltRightToMatrix_ρ (v : (altLeftHanded ⊗ altRightHanded).V) (M : SL(2,)) :
altLeftAltRightToMatrix (TensorProduct.map (altLeftHanded.ρ M) (altRightHanded.ρ M) v) =
(M.1⁻¹)ᵀ * altLeftAltRightToMatrix v * ((M.1⁻¹).conjTranspose)ᵀ := by
nth_rewrite 1 [altLeftAltRightToMatrix]
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.trans_apply]
trans (LinearEquiv.curry (Fin 2) (Fin 2)) ((LinearMap.toMatrix
(altLeftBasis.tensorProduct altRightBasis) (altLeftBasis.tensorProduct altRightBasis)
(TensorProduct.map (altLeftHanded.ρ M) (altRightHanded.ρ M)))
*ᵥ ((Finsupp.linearEquivFunOnFinite (Fin 2 × Fin 2))
((altLeftBasis.tensorProduct altRightBasis).repr (v))))
· apply congrArg
have h1 := (LinearMap.toMatrix_mulVec_repr (altLeftBasis.tensorProduct altRightBasis)
(altLeftBasis.tensorProduct altRightBasis)
(TensorProduct.map (altLeftHanded.ρ M) (altRightHanded.ρ M)) v)
erw [h1]
rfl
rw [TensorProduct.toMatrix_map]
funext i j
change ∑ k, ((kroneckerMap (fun x1 x2 => x1 * x2)
((LinearMap.toMatrix altLeftBasis altLeftBasis) (altLeftHanded.ρ M))
((LinearMap.toMatrix altRightBasis altRightBasis) (altRightHanded.ρ M)) (i, j) k)
* altLeftAltRightToMatrix v k.1 k.2) = _
erw [Finset.sum_product]
simp_rw [kroneckerMap_apply, Matrix.mul_apply, Matrix.transpose_apply]
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, (M.1)⁻¹ x1 i * altLeftAltRightToMatrix v x1 x) *
(M.1)⁻¹ᴴ j x = ∑ x : Fin 2, ∑ x1 : Fin 2,
((M.1)⁻¹ x1 i * altLeftAltRightToMatrix v x1 x) * (M.1)⁻¹ᴴ j x:= by
congr
funext x
rw [Finset.sum_mul]
erw [h1]
rw [Finset.sum_comm]
congr
funext x
congr
funext x1
simp only [altLeftBasis_ρ_apply, altRightBasis_ρ_apply, transpose_apply,
Action.instMonoidalCategory_tensorObj_V]
ring
lemma leftRightToMatrix_ρ (v : (leftHanded ⊗ rightHanded).V) (M : SL(2,)) :
leftRightToMatrix (TensorProduct.map (leftHanded.ρ M) (rightHanded.ρ M) v) =
M.1 * leftRightToMatrix v * (M.1)ᴴ := by
nth_rewrite 1 [leftRightToMatrix]
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.trans_apply]
trans (LinearEquiv.curry (Fin 2) (Fin 2)) ((LinearMap.toMatrix
(leftBasis.tensorProduct rightBasis) (leftBasis.tensorProduct rightBasis)
(TensorProduct.map (leftHanded.ρ M) (rightHanded.ρ M)))
*ᵥ ((Finsupp.linearEquivFunOnFinite (Fin 2 × Fin 2))
((leftBasis.tensorProduct rightBasis).repr (v))))
· apply congrArg
have h1 := (LinearMap.toMatrix_mulVec_repr (leftBasis.tensorProduct rightBasis)
(leftBasis.tensorProduct rightBasis)
(TensorProduct.map (leftHanded.ρ M) (rightHanded.ρ M)) v)
erw [h1]
rfl
rw [TensorProduct.toMatrix_map]
funext i j
change ∑ k, ((kroneckerMap (fun x1 x2 => x1 * x2)
((LinearMap.toMatrix leftBasis leftBasis) (leftHanded.ρ M))
((LinearMap.toMatrix rightBasis rightBasis) (rightHanded.ρ M)) (i, j) k)
* leftRightToMatrix v k.1 k.2) = _
erw [Finset.sum_product]
simp_rw [kroneckerMap_apply, Matrix.mul_apply]
have h1 : ∑ x : Fin 2, (∑ x1 : Fin 2, M.1 i x1 * leftRightToMatrix v x1 x) * (M.1)ᴴ x j
= ∑ x : Fin 2, ∑ x1 : Fin 2, (M.1 i x1 * leftRightToMatrix v x1 x) * (M.1)ᴴ x j := by
congr
funext x
rw [Finset.sum_mul]
erw [h1]
rw [Finset.sum_comm]
congr
funext x
congr
funext x1
simp only [leftBasis_ρ_apply, rightBasis_ρ_apply, transpose_apply,
Action.instMonoidalCategory_tensorObj_V]
rw [Matrix.conjTranspose]
simp only [RCLike.star_def, map_apply, transpose_apply]
ring
/-!
## The symm version of the group actions.
@ -464,7 +562,7 @@ lemma altRightAltRightToMatrix_ρ_symm (v : Matrix (Fin 2) (Fin 2) ) (M : SL(
lemma rightAltRightToMatrix_ρ_symm (v : Matrix (Fin 2) (Fin 2) ) (M : SL(2,)) :
TensorProduct.map (rightHanded.ρ M) (altRightHanded.ρ M) (rightAltRightToMatrix.symm v) =
rightAltRightToMatrix.symm ((M.1.map star) * v * (((M.1⁻¹).conjTranspose)ᵀ) ) := by
rightAltRightToMatrix.symm ((M.1.map star) * v * (((M.1⁻¹).conjTranspose)ᵀ)) := by
have h1 := rightAltRightToMatrix_ρ (rightAltRightToMatrix.symm v) M
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.apply_symm_apply] at h1
rw [← h1]
@ -478,7 +576,51 @@ lemma altRightRightToMatrix_ρ_symm (v : Matrix (Fin 2) (Fin 2) ) (M : SL(2,
rw [← h1]
simp
lemma altLeftAltRightToMatrix_ρ_symm (v : Matrix (Fin 2) (Fin 2) ) (M : SL(2,)) :
TensorProduct.map (altLeftHanded.ρ M) (altRightHanded.ρ M) (altLeftAltRightToMatrix.symm v) =
altLeftAltRightToMatrix.symm ((M.1⁻¹)ᵀ * v * ((M.1⁻¹).conjTranspose)ᵀ) := by
have h1 := altLeftAltRightToMatrix_ρ (altLeftAltRightToMatrix.symm v) M
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.apply_symm_apply] at h1
rw [← h1]
simp
lemma leftRightToMatrix_ρ_symm (v : Matrix (Fin 2) (Fin 2) ) (M : SL(2,)) :
TensorProduct.map (leftHanded.ρ M) (rightHanded.ρ M) (leftRightToMatrix.symm v) =
leftRightToMatrix.symm (M.1 * v * (M.1)ᴴ) := by
have h1 := leftRightToMatrix_ρ (leftRightToMatrix.symm v) M
simp only [Action.instMonoidalCategory_tensorObj_V, LinearEquiv.apply_symm_apply] at h1
rw [← h1]
simp
open SpaceTime
lemma altLeftAltRightToMatrix_ρ_symm_selfAdjoint (v : Matrix (Fin 2) (Fin 2) )
(hv : IsSelfAdjoint v) (M : SL(2,)) :
TensorProduct.map (altLeftHanded.ρ M) (altRightHanded.ρ M) (altLeftAltRightToMatrix.symm v) =
altLeftAltRightToMatrix.symm
(SL2C.repSelfAdjointMatrix (M.transpose⁻¹) ⟨v, hv⟩) := by
rw [altLeftAltRightToMatrix_ρ_symm]
apply congrArg
simp only [SL2C.repSelfAdjointMatrix, MonoidHom.coe_mk, OneHom.coe_mk,
SL2C.toLinearMapSelfAdjointMatrix_apply_coe, SpecialLinearGroup.coe_inv,
SpecialLinearGroup.coe_transpose]
congr
· rw [SL2C.inverse_coe]
simp only [SpecialLinearGroup.coe_inv]
rw [@adjugate_transpose]
· rw [SL2C.inverse_coe]
simp only [SpecialLinearGroup.coe_inv]
rw [← @adjugate_transpose]
rfl
lemma leftRightToMatrix_ρ_symm_selfAdjoint (v : Matrix (Fin 2) (Fin 2) )
(hv : IsSelfAdjoint v) (M : SL(2,)) :
TensorProduct.map (leftHanded.ρ M) (rightHanded.ρ M) (leftRightToMatrix.symm v) =
leftRightToMatrix.symm
(SL2C.repSelfAdjointMatrix M ⟨v, hv⟩) := by
rw [leftRightToMatrix_ρ_symm]
apply congrArg
simp [SpaceTime.SL2C.repSelfAdjointMatrix]
end
end Fermion

View file

@ -24,11 +24,11 @@ open Complex
open TensorProduct
open CategoryTheory.MonoidalCategory
/-- The left-alt-left unit `δᵃ` as an element of `(leftHanded ⊗ altLeftHanded).V`. -/
/-- The left-alt-left unit `δᵃ` as an element of `(leftHanded ⊗ altLeftHanded).V`. -/
def leftAltLeftUnitVal : (leftHanded ⊗ altLeftHanded).V :=
leftAltLeftToMatrix.symm 1
/-- The left-alt-left unit `δᵃ` as a morphism `𝟙_ (Rep SL(2,)) ⟶ leftHanded ⊗ altLeftHanded `,
/-- The left-alt-left unit `δᵃ` as a morphism `𝟙_ (Rep SL(2,)) ⟶ leftHanded ⊗ altLeftHanded `,
manifesting the invariance under the `SL(2,)` action. -/
def leftAltLeftUnit : 𝟙_ (Rep SL(2,)) ⟶ leftHanded ⊗ altLeftHanded where
hom := {
@ -42,22 +42,24 @@ def leftAltLeftUnit : 𝟙_ (Rep SL(2,)) ⟶ leftHanded ⊗ altLeftHanded
rfl}
comm M := by
ext x : 2
simp
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • leftAltLeftUnitVal =
(TensorProduct.map (leftHanded.ρ M) (altLeftHanded.ρ M)) (x' • leftAltLeftUnitVal)
simp
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
simp [leftAltLeftUnitVal]
simp only [Action.instMonoidalCategory_tensorObj_V, leftAltLeftUnitVal]
erw [leftAltLeftToMatrix_ρ_symm]
apply congrArg
simp
/-- The alt-left-left unit `δₐ` as an element of `(altLeftHanded ⊗ leftHanded).V`. -/
/-- The alt-left-left unit `δₐ` as an element of `(altLeftHanded ⊗ leftHanded).V`. -/
def altLeftLeftUnitVal : (altLeftHanded ⊗ leftHanded).V :=
altLeftLeftToMatrix.symm 1
/-- The alt-left-left unit `δₐ` as a morphism `𝟙_ (Rep SL(2,)) ⟶ altLeftHanded ⊗ leftHanded `,
/-- The alt-left-left unit `δₐ` as a morphism `𝟙_ (Rep SL(2,)) ⟶ altLeftHanded ⊗ leftHanded `,
manifesting the invariance under the `SL(2,)` action. -/
def altLeftLeftUnit : 𝟙_ (Rep SL(2,)) ⟶ altLeftHanded ⊗ leftHanded where
hom := {
@ -71,24 +73,26 @@ def altLeftLeftUnit : 𝟙_ (Rep SL(2,)) ⟶ altLeftHanded ⊗ leftHanded
rfl}
comm M := by
ext x : 2
simp
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • altLeftLeftUnitVal =
(TensorProduct.map (altLeftHanded.ρ M) (leftHanded.ρ M)) (x' • altLeftLeftUnitVal)
simp
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
simp [altLeftLeftUnitVal]
simp only [Action.instMonoidalCategory_tensorObj_V, altLeftLeftUnitVal]
erw [altLeftLeftToMatrix_ρ_symm]
apply congrArg
simp only [mul_one, ← transpose_mul, SpecialLinearGroup.det_coe, isUnit_iff_ne_zero, ne_eq,
one_ne_zero, not_false_eq_true, mul_nonsing_inv, transpose_one]
/-- The right-alt-right unit `δ_{dot a}^{dot a}` as an element of
/-- The right-alt-right unit `δ^{dot a}_{dot a}` as an element of
`(rightHanded ⊗ altRightHanded).V`. -/
def rightAltRightUnitVal : (rightHanded ⊗ altRightHanded).V :=
rightAltRightToMatrix.symm 1
/-- The right-alt-right unit `δ_{dot a}^{dot a}` as a morphism
/-- The right-alt-right unit `δ^{dot a}_{dot a}` as a morphism
`𝟙_ (Rep SL(2,)) ⟶ rightHanded ⊗ altRightHanded`, manifesting
the invariance under the `SL(2,)` action. -/
def rightAltRightUnit : 𝟙_ (Rep SL(2,)) ⟶ rightHanded ⊗ altRightHanded where
@ -103,29 +107,31 @@ def rightAltRightUnit : 𝟙_ (Rep SL(2,)) ⟶ rightHanded ⊗ altRightHa
rfl}
comm M := by
ext x : 2
simp
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • rightAltRightUnitVal =
(TensorProduct.map (rightHanded.ρ M) (altRightHanded.ρ M)) (x' • rightAltRightUnitVal)
simp
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
simp [rightAltRightUnitVal]
simp only [Action.instMonoidalCategory_tensorObj_V, rightAltRightUnitVal]
erw [rightAltRightToMatrix_ρ_symm]
apply congrArg
simp
simp only [RCLike.star_def, mul_one]
symm
refine transpose_eq_one.mp ?h.h.h.a
simp
change (M.1)⁻¹ᴴ * (M.1)ᴴ = 1
simp only [transpose_mul, transpose_transpose]
change (M.1)⁻¹ᴴ * (M.1)ᴴ = 1
rw [@conjTranspose_nonsing_inv]
simp
/-- The alt-right-right unit `δ^{dot a}_{dot a}` as an element of
/-- The alt-right-right unit `δ_{dot a}^{dot a}` as an element of
`(rightHanded ⊗ altRightHanded).V`. -/
def altRightRightUnitVal : (altRightHanded ⊗ rightHanded).V :=
altRightRightToMatrix.symm 1
/-- The alt-right-right unit `δ^{dot a}_{dot a}` as a morphism
/-- The alt-right-right unit `δ_{dot a}^{dot a}` as a morphism
`𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ rightHanded`, manifesting
the invariance under the `SL(2,)` action. -/
def altRightRightUnit : 𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ rightHanded where
@ -140,18 +146,20 @@ def altRightRightUnit : 𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ rightHa
rfl}
comm M := by
ext x : 2
simp
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • altRightRightUnitVal =
(TensorProduct.map (altRightHanded.ρ M) (rightHanded.ρ M)) (x' • altRightRightUnitVal)
simp
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
simp [altRightRightUnitVal]
simp only [Action.instMonoidalCategory_tensorObj_V, altRightRightUnitVal]
erw [altRightRightToMatrix_ρ_symm]
apply congrArg
simp
simp only [mul_one, RCLike.star_def]
symm
change (M.1)⁻¹ᴴ * (M.1)ᴴ = 1
change (M.1)⁻¹ᴴ * (M.1)ᴴ = 1
rw [@conjTranspose_nonsing_inv]
simp

View file

@ -10,7 +10,7 @@ import Mathlib.CategoryTheory.Comma.Over
import Mathlib.CategoryTheory.Core
import Mathlib.CategoryTheory.Monoidal.Braided.Basic
import HepLean.SpaceTime.WeylFermion.Basic
import HepLean.SpaceTime.LorentzVector.Complex
import HepLean.SpaceTime.LorentzVector.Complex.Basic
/-!
# Over color category.

View file

@ -47,14 +47,14 @@ def discreteFunctorMapIso {c1 c2 : Discrete C} (h : c1 ⟶ c2) :
lemma discreteFun_hom_trans {c1 c2 c3 : Discrete C} (h1 : c1 = c2) (h2 : c2 = c3)
(v : F.obj c1) : (F.map (eqToHom h2)).hom ((F.map (eqToHom h1)).hom v)
= (F.map (eqToHom ( h1.trans h2))).hom v := by
= (F.map (eqToHom (h1.trans h2))).hom v := by
subst h2 h1
simp_all only [eqToHom_refl, Discrete.functor_map_id, Action.id_hom, ModuleCat.id_apply]
/-- The linear equivalence between `(F.obj c1).V ≃ₗ[k] (F.obj c2).V` induced by an equality of
`c1` and `c2`. -/
def discreteFunctorMapEqIso {c1 c2 : Discrete C} (h : c1.as = c2.as) :
(F.obj c1).V ≃ₗ[k] (F.obj c2).V := LinearEquiv.ofLinear
(F.obj c1).V ≃ₗ[k] (F.obj c2).V := LinearEquiv.ofLinear
(F.mapIso (Discrete.eqToIso h)).hom.hom (F.mapIso (Discrete.eqToIso h)).inv.hom
(by
ext x : 2
@ -64,7 +64,7 @@ def discreteFunctorMapEqIso {c1 c2 : Discrete C} (h : c1.as = c2.as) :
rw [ModuleCat.ext_iff] at h1
have h1x := h1 x
simp only [CategoryStruct.comp] at h1x
simpa using h1x )
simpa using h1x)
(by
ext x : 2
simp only [Functor.mapIso_inv, eqToIso.inv, Functor.mapIso_hom, eqToIso.hom, LinearMap.coe_comp,
@ -123,7 +123,6 @@ lemma objObj'_ρ_empty (g : G) : (objObj' F (𝟙_ (OverColor C))).ρ g = Linear
funext i
exact Empty.elim i
open TensorProduct in
@[simp]
lemma objObj'_V_carrier (f : OverColor C) :
@ -150,7 +149,6 @@ lemma mapToLinearEquiv'_tprod {f g : OverColor C} (m : f ⟶ g)
rw [PiTensorProduct.reindex_tprod, PiTensorProduct.congr_tprod]
rfl
/-- Given a morphism in `OverColor C` the corresopnding map of representations induced by
reindexing. -/
def objMap' {f g : OverColor C} (m : f ⟶ g) : objObj' F f ⟶ objObj' F g where
@ -232,7 +230,6 @@ lemma μModEquiv_tmul_tprod {X Y : OverColor C}
rw [PiTensorProduct.congr_tprod]
rfl
/-- The natural isomorphism corresponding to the tensorate. -/
def μ (X Y : OverColor C) : objObj' F X ⊗ objObj' F Y ≅ objObj' F (X ⊗ Y) :=
Action.mkIso (μModEquiv F X Y).toModuleIso
@ -264,7 +261,6 @@ lemma μ_tmul_tprod {X Y : OverColor C} (p : (i : X.left) → F.obj (Discrete.mk
discreteSumEquiv F i (HepLean.PiTensorProduct.elimPureTensor p q i) := by
exact μModEquiv_tmul_tprod F p q
lemma μ_natural_left {X Y : OverColor C} (f : X ⟶ Y) (Z : OverColor C) :
MonoidalCategory.whiskerRight (objMap' F f) (objObj' F Z) ≫ (μ F Y Z).hom =
(μ F X Z).hom ≫ objMap' F (MonoidalCategory.whiskerRight f Z) := by
@ -302,7 +298,6 @@ lemma μ_natural_left {X Y : OverColor C} (f : X ⟶ Y) (Z : OverColor C) :
Functor.mapIso_refl, Iso.refl_hom, Action.id_hom, Iso.refl_inv]
rfl
lemma μ_natural_right {X Y : OverColor C} (X' : OverColor C) (f : X ⟶ Y) :
MonoidalCategory.whiskerLeft (objObj' F X') (objMap' F f) ≫ (μ F X' Y).hom =
(μ F X' X).hom ≫ objMap' F (MonoidalCategory.whiskerLeft X' f) := by
@ -338,7 +333,6 @@ lemma μ_natural_right {X Y : OverColor C} (X' : OverColor C) (f : X ⟶ Y) :
rfl
| Sum.inr i => rfl
lemma associativity (X Y Z : OverColor C) :
whiskerRight (μ F X Y).hom (objObj' F Z) ≫
(μ F (X ⊗ Y) Z).hom ≫ objMap' F (associator X Y Z).hom =
@ -534,7 +528,7 @@ def mapApp' (X : OverColor C) : (obj' F).obj X ⟶ (obj' F').obj X where
funext i
simpa using LinearMap.congr_fun ((η.app (Discrete.mk (X.hom i))).comm M) (x i)
lemma mapApp'_tprod (X : OverColor C) (p : (i : X.left) → F.obj (Discrete.mk (X.hom i))) :
lemma mapApp'_tprod (X : OverColor C) (p : (i : X.left) → F.obj (Discrete.mk (X.hom i))) :
(mapApp' η X).hom (PiTensorProduct.tprod k p) =
PiTensorProduct.tprod k fun i => (η.app (Discrete.mk (X.hom i))).hom (p i) := by
change (mapApp' η X).hom (PiTensorProduct.tprod k p) = _
@ -619,7 +613,7 @@ end lift
noncomputable def lift : (Discrete C ⥤ Rep k G) ⥤ MonoidalFunctor (OverColor C) (Rep k G) where
obj F := lift.obj' F
map η := lift.map' η
map_id F := by
map_id F := by
simp only [lift.map']
refine MonoidalNatTrans.ext' (fun X => ?_)
ext x : 2

View file

@ -36,7 +36,10 @@ into Lean 4.
| [Video](https://www.youtube.com/watch?v=IjodVUawTjM) | Index notation in Lean 4 |
### Papers referencing HepLean
- Hu, Jiewen, Thomas Zhu, and Sean Welleck. "miniCTX: Neural Theorem Proving with (Long-) Contexts." arXiv preprint arXiv:2408.03350 (2024).
- Hu, Jiewen, Thomas Zhu, and Sean Welleck. "miniCTX: Neural Theorem Proving with (Long-) Contexts." arXiv preprint [arXiv:2408.03350](https://www.arxiv.org/abs/2408.03350) (2024). [Project page]( https://cmu-l3.github.io/minictx/)
How HepLean was used: *Theorems from the space-time files of HepLean were included in a data set used to evaluate the ability of models to prove theorems from real-world repositories, which requires working with definitions, theorems, and other context not seen in training.*
## Contributing
We follow here roughly the same contribution policies as MathLib4 (which can be found [here](https://leanprover-community.github.io/contribute/index.html)).
@ -61,5 +64,5 @@ The installation instructions for Lean 4 are given here: https://leanprover-comm
### Optional extras
- [Lean Copilot](https://github.com/lean-dojo/LeanCopilot) allows the use of large language models in Lean.
- [Lean Copilot](https://github.com/lean-dojo/LeanCopilot) and [LLMLean](https://github.com/cmu-l3/llmlean) allow for the use of large language models in Lean
- [tryAtEachStep](https://github.com/dwrensha/tryAtEachStep) allows one to apply a tactic, e.g. `exact?` at each step of a lemma in a file to see if it completes the goal. This is useful for golfing proofs.