feat: Add metric invariance for Real Lorentz Tensors

This commit is contained in:
jstoobysmith 2024-07-31 08:52:09 -04:00
parent 8b2c853fd8
commit 7b0b979d51
7 changed files with 175 additions and 24 deletions

View file

@ -167,6 +167,27 @@ lemma subtype_mul_inv : (Subtype.val Λ) * (Subtype.val Λ)⁻¹ = 1 := by
rw [mul_inv_self Λ]
simp only [lorentzGroupIsGroup_one_coe]
@[simp]
lemma mul_minkowskiMatrix_mul_transpose :
(Subtype.val Λ) * minkowskiMatrix * (Subtype.val Λ).transpose = minkowskiMatrix := by
have h2 := Λ.prop
rw [LorentzGroup.mem_iff_self_mul_dual] at h2
simp [minkowskiMetric.dual] at h2
symm
refine right_inv_eq_left_inv minkowskiMatrix.sq ?_
rw [← h2]
noncomm_ring
@[simp]
lemma transpose_mul_minkowskiMatrix_mul_self :
(Subtype.val Λ).transpose * minkowskiMatrix * (Subtype.val Λ) = minkowskiMatrix := by
have h2 := Λ.prop
rw [LorentzGroup.mem_iff_dual_mul_self] at h2
simp [minkowskiMetric.dual] at h2
refine right_inv_eq_left_inv ?_ minkowskiMatrix.sq
rw [← h2]
noncomm_ring
/-- The transpose of a matrix in the Lorentz group is an element of the Lorentz group. -/
def transpose (Λ : LorentzGroup d) : LorentzGroup d :=
⟨Λ.1ᵀ, mem_iff_transpose.mp Λ.2⟩

View file

@ -81,12 +81,12 @@ def realLorentzTensor (d : ) : TensorStructure where
| .down => LorentzVector.unitDown_rid
metric μ :=
match μ with
| realTensor.ColorType.up => asProdLorentzVector
| realTensor.ColorType.down => asProdCovariantLorentzVector
| realTensor.ColorType.up => asTenProd
| realTensor.ColorType.down => asCoTenProd
metric_dual μ :=
match μ with
| realTensor.ColorType.up => asProdLorentzVector_contr_asCovariantProdLorentzVector
| realTensor.ColorType.down => asProdCovariantLorentzVector_contr_asProdLorentzVector
| realTensor.ColorType.up => asTenProd_contr_asCoTenProd
| realTensor.ColorType.down => asCoTenProd_contr_asTenProd
/-- The action of the Lorentz group on real Lorentz tensors. -/
instance : MulActionTensor (LorentzGroup d) (realLorentzTensor d) where
@ -98,5 +98,8 @@ instance : MulActionTensor (LorentzGroup d) (realLorentzTensor d) where
match μ with
| .up => TensorProduct.ext' (fun _ _ => LorentzVector.contrUpDown_invariant_lorentzAction)
| .down => TensorProduct.ext' (fun _ _ => LorentzVector.contrDownUp_invariant_lorentzAction)
metric_inv μ g :=
match μ with
| .up => asTenProd_invariant g
| .down => asCoTenProd_invariant g
end

View file

@ -176,7 +176,7 @@ lemma dualizeAll_equivariant (g : G) : (𝓣.dualizeAll.toLinearMap) ∘ₗ (@re
simp [map_add, add_tmul, hx]
simp_all only [Function.comp_apply, LinearMap.coe_comp, LinearMap.id_coe, id_eq])
intro rx fx
simp
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, LinearMapClass.map_smul, rep_tprod]
apply congrArg
change (PiTensorProduct.map _) ((PiTensorProduct.tprod R) _) =
(𝓣.rep g) ((PiTensorProduct.map _) ((PiTensorProduct.tprod R) fx))

View file

@ -175,18 +175,41 @@ open scoped minkowskiMatrix
variable {d : }
/-- The metric tensor as an element of `LorentzVector d ⊗[] LorentzVector d`. -/
def asProdLorentzVector : LorentzVector d ⊗[] LorentzVector d :=
∑ μ, η μ μ • (LorentzVector.stdBasis μ ⊗ₜ[] LorentzVector.stdBasis μ)
def asTenProd : LorentzVector d ⊗[] LorentzVector d :=
∑ μ, ∑ ν, η μ ν • (LorentzVector.stdBasis μ ⊗ₜ[] LorentzVector.stdBasis ν)
lemma asTenProd_diag :
@asTenProd d = ∑ μ, η μ μ • (LorentzVector.stdBasis μ ⊗ₜ[] LorentzVector.stdBasis μ) := by
simp only [asTenProd]
refine Finset.sum_congr rfl (fun μ _ => ?_)
rw [Finset.sum_eq_single μ]
intro ν _ hμν
rw [minkowskiMatrix.off_diag_zero hμν.symm]
simp only [zero_smul]
intro a
simp_all only
/-- The metric tensor as an element of `CovariantLorentzVector d ⊗[] CovariantLorentzVector d`. -/
def asProdCovariantLorentzVector : CovariantLorentzVector d ⊗[] CovariantLorentzVector d :=
∑ μ, η μ μ • (CovariantLorentzVector.stdBasis μ ⊗ₜ[] CovariantLorentzVector.stdBasis μ)
def asCoTenProd : CovariantLorentzVector d ⊗[] CovariantLorentzVector d :=
∑ μ, ∑ ν, η μ ν • (CovariantLorentzVector.stdBasis μ ⊗ₜ[] CovariantLorentzVector.stdBasis ν)
lemma asCoTenProd_diag :
@asCoTenProd d = ∑ μ, η μ μ • (CovariantLorentzVector.stdBasis μ ⊗ₜ[]
CovariantLorentzVector.stdBasis μ) := by
simp only [asCoTenProd]
refine Finset.sum_congr rfl (fun μ _ => ?_)
rw [Finset.sum_eq_single μ]
intro ν _ hμν
rw [minkowskiMatrix.off_diag_zero hμν.symm]
simp only [zero_smul]
intro a
simp_all only
@[simp]
lemma contrLeft_asProdLorentzVector (x : CovariantLorentzVector d) :
contrLeftAux contrDownUp (x ⊗ₜ[] asProdLorentzVector) =
lemma contrLeft_asTenProd (x : CovariantLorentzVector d) :
contrLeftAux contrDownUp (x ⊗ₜ[] asTenProd) =
∑ μ, ((η μ μ * x μ) • LorentzVector.stdBasis μ) := by
simp only [asProdLorentzVector]
simp only [asTenProd_diag]
rw [tmul_sum]
rw [map_sum]
refine Finset.sum_congr rfl (fun μ _ => ?_)
@ -196,10 +219,10 @@ lemma contrLeft_asProdLorentzVector (x : CovariantLorentzVector d) :
exact smul_smul (η μ μ) (x μ) (e μ)
@[simp]
lemma contrLeft_asProdCovariantLorentzVector (x : LorentzVector d) :
contrLeftAux contrUpDown (x ⊗ₜ[] asProdCovariantLorentzVector) =
lemma contrLeft_asCoTenProd (x : LorentzVector d) :
contrLeftAux contrUpDown (x ⊗ₜ[] asCoTenProd) =
∑ μ, ((η μ μ * x μ) • CovariantLorentzVector.stdBasis μ) := by
simp only [asProdCovariantLorentzVector]
simp only [asCoTenProd_diag]
rw [tmul_sum]
rw [map_sum]
refine Finset.sum_congr rfl (fun μ _ => ?_)
@ -209,17 +232,17 @@ lemma contrLeft_asProdCovariantLorentzVector (x : LorentzVector d) :
exact smul_smul (η μ μ) (x μ) (CovariantLorentzVector.stdBasis μ)
@[simp]
lemma asProdLorentzVector_contr_asCovariantProdLorentzVector :
(contrMidAux (contrUpDown) (asProdLorentzVector ⊗ₜ[] asProdCovariantLorentzVector))
lemma asTenProd_contr_asCoTenProd :
(contrMidAux (contrUpDown) (asTenProd ⊗ₜ[] asCoTenProd))
= TensorProduct.comm _ _ (@unitUp d) := by
simp only [contrMidAux, LinearEquiv.refl_toLinearMap, asProdLorentzVector, LinearMap.coe_comp,
simp only [contrMidAux, LinearEquiv.refl_toLinearMap, asTenProd_diag, LinearMap.coe_comp,
LinearEquiv.coe_coe, Function.comp_apply]
rw [sum_tmul, map_sum, map_sum, unitUp]
simp only [Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton, map_add, comm_tmul, map_sum]
refine Finset.sum_congr rfl (fun μ _ => ?_)
rw [← tmul_smul, TensorProduct.assoc_tmul]
simp only [map_tmul, LinearMap.id_coe, id_eq, contrLeft_asProdCovariantLorentzVector]
simp only [map_tmul, LinearMap.id_coe, id_eq, contrLeft_asCoTenProd]
rw [tmul_sum, Finset.sum_eq_single_of_mem μ, tmul_smul]
change (η μ μ * (η μ μ * e μ μ)) • e μ ⊗ₜ[] CovariantLorentzVector.stdBasis μ = _
rw [LorentzVector.stdBasis]
@ -236,10 +259,10 @@ lemma asProdLorentzVector_contr_asCovariantProdLorentzVector :
exact fun a => False.elim (hμν (id (Eq.symm a)))
@[simp]
lemma asProdCovariantLorentzVector_contr_asProdLorentzVector :
(contrMidAux (contrDownUp) (asProdCovariantLorentzVector ⊗ₜ[] asProdLorentzVector))
lemma asCoTenProd_contr_asTenProd :
(contrMidAux (contrDownUp) (asCoTenProd ⊗ₜ[] asTenProd))
= TensorProduct.comm _ _ (@unitDown d) := by
simp only [contrMidAux, LinearEquiv.refl_toLinearMap, asProdCovariantLorentzVector,
simp only [contrMidAux, LinearEquiv.refl_toLinearMap, asCoTenProd_diag,
LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply]
rw [sum_tmul, map_sum, map_sum, unitDown]
simp only [Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
@ -247,7 +270,7 @@ lemma asProdCovariantLorentzVector_contr_asProdLorentzVector :
refine Finset.sum_congr rfl (fun μ _ => ?_)
rw [smul_tmul, TensorProduct.assoc_tmul]
simp only [tmul_smul, LinearMapClass.map_smul, map_tmul, LinearMap.id_coe, id_eq,
contrLeft_asProdLorentzVector]
contrLeft_asTenProd]
rw [tmul_sum, Finset.sum_eq_single_of_mem μ, tmul_smul, smul_smul, LorentzVector.stdBasis]
erw [Pi.basisFun_apply]
simp only [LinearMap.stdBasis_same, mul_one, η_apply_mul_η_apply_diag, one_smul]
@ -260,6 +283,86 @@ lemma asProdCovariantLorentzVector_contr_asProdLorentzVector :
ite_eq_right_iff, smul_eq_zero, mul_eq_zero]
exact fun a => False.elim (hμν (id (Eq.symm a)))
lemma asTenProd_invariant (g : LorentzGroup d) :
TensorProduct.map (LorentzVector.rep g) (LorentzVector.rep g) asTenProd = asTenProd := by
simp only [asTenProd, map_sum, LinearMapClass.map_smul, map_tmul, rep_apply_stdBasis,
Matrix.transpose_apply]
trans ∑ μ : Fin 1 ⊕ Fin d, ∑ ν : Fin 1 ⊕ Fin d, ∑ φ : Fin 1 ⊕ Fin d,
η μ ν • (g.1 φ μ • e φ) ⊗ₜ[] ∑ ρ : Fin 1 ⊕ Fin d, g.1 ρ ν • e ρ
refine Finset.sum_congr rfl (fun μ _ => Finset.sum_congr rfl (fun ν _ => ?_))
rw [sum_tmul, Finset.smul_sum]
trans ∑ μ : Fin 1 ⊕ Fin d, ∑ ν : Fin 1 ⊕ Fin d, ∑ φ : Fin 1 ⊕ Fin d, ∑ ρ : Fin 1 ⊕ Fin d,
(η μ ν • (g.1 φ μ • e φ) ⊗ₜ[] (g.1 ρ ν • e ρ))
· refine Finset.sum_congr rfl (fun μ _ => Finset.sum_congr rfl (fun ν _ =>
Finset.sum_congr rfl (fun φ _ => ?_)))
rw [tmul_sum, Finset.smul_sum]
rw [Finset.sum_congr rfl (fun μ _ => Finset.sum_comm)]
rw [Finset.sum_congr rfl (fun μ _ => Finset.sum_congr rfl (fun ν _ => Finset.sum_comm))]
rw [Finset.sum_comm]
rw [Finset.sum_congr rfl (fun φ _ => Finset.sum_comm)]
trans ∑ φ : Fin 1 ⊕ Fin d, ∑ ρ : Fin 1 ⊕ Fin d, ∑ μ : Fin 1 ⊕ Fin d, ∑ ν : Fin 1 ⊕ Fin d,
((g.1 φ μ * η μ ν * g.1 ρ ν) • (e φ) ⊗ₜ[] (e ρ))
· refine Finset.sum_congr rfl (fun φ _ => Finset.sum_congr rfl (fun ρ _ =>
Finset.sum_congr rfl (fun μ _ => Finset.sum_congr rfl (fun ν _ => ?_))))
rw [smul_tmul, tmul_smul, tmul_smul, smul_smul, smul_smul]
ring_nf
rw [Finset.sum_congr rfl (fun φ _ => Finset.sum_congr rfl (fun ρ _ =>
Finset.sum_congr rfl (fun μ _ => Finset.sum_smul.symm)))]
rw [Finset.sum_congr rfl (fun φ _ => Finset.sum_congr rfl (fun ρ _ => Finset.sum_smul.symm))]
trans ∑ φ : Fin 1 ⊕ Fin d, ∑ ρ : Fin 1 ⊕ Fin d,
(((g.1 * minkowskiMatrix * g.1.transpose : Matrix _ _ _) φ ρ) • (e φ) ⊗ₜ[] (e ρ))
· refine Finset.sum_congr rfl (fun φ _ => Finset.sum_congr rfl (fun ρ _ => ?_))
apply congrFun
apply congrArg
simp only [Matrix.mul_apply, Matrix.transpose_apply]
rw [Finset.sum_comm]
refine Finset.sum_congr rfl (fun μ _ => ?_)
rw [Finset.sum_mul]
simp
open CovariantLorentzVector in
lemma asCoTenProd_invariant (g : LorentzGroup d) :
TensorProduct.map (CovariantLorentzVector.rep g) (CovariantLorentzVector.rep g)
asCoTenProd = asCoTenProd := by
simp only [asCoTenProd, map_sum, LinearMapClass.map_smul, map_tmul,
CovariantLorentzVector.rep_apply_stdBasis, Matrix.transpose_apply]
trans ∑ μ : Fin 1 ⊕ Fin d, ∑ ν : Fin 1 ⊕ Fin d, ∑ φ : Fin 1 ⊕ Fin d,
η μ ν • (g⁻¹.1 μ φ • CovariantLorentzVector.stdBasis φ) ⊗ₜ[]
ρ : Fin 1 ⊕ Fin d, g⁻¹.1 ν ρ • CovariantLorentzVector.stdBasis ρ
· refine Finset.sum_congr rfl (fun μ _ => Finset.sum_congr rfl (fun ν _ => ?_))
rw [sum_tmul, Finset.smul_sum]
trans ∑ μ : Fin 1 ⊕ Fin d, ∑ ν : Fin 1 ⊕ Fin d, ∑ φ : Fin 1 ⊕ Fin d, ∑ ρ : Fin 1 ⊕ Fin d,
(η μ ν • (g⁻¹.1 μ φ • CovariantLorentzVector.stdBasis φ) ⊗ₜ[]
(g⁻¹.1 ν ρ • CovariantLorentzVector.stdBasis ρ))
· refine Finset.sum_congr rfl (fun μ _ => Finset.sum_congr rfl (fun ν _ =>
Finset.sum_congr rfl (fun φ _ => ?_)))
rw [tmul_sum, Finset.smul_sum]
rw [Finset.sum_congr rfl (fun μ _ => Finset.sum_comm)]
rw [Finset.sum_congr rfl (fun μ _ => Finset.sum_congr rfl (fun ν _ => Finset.sum_comm))]
rw [Finset.sum_comm]
rw [Finset.sum_congr rfl (fun φ _ => Finset.sum_comm)]
trans ∑ φ : Fin 1 ⊕ Fin d, ∑ ρ : Fin 1 ⊕ Fin d, ∑ μ : Fin 1 ⊕ Fin d, ∑ ν : Fin 1 ⊕ Fin d,
((g⁻¹.1 μ φ * η μ ν * g⁻¹.1 ν ρ) • (CovariantLorentzVector.stdBasis φ) ⊗ₜ[]
(CovariantLorentzVector.stdBasis ρ))
· refine Finset.sum_congr rfl (fun φ _ => Finset.sum_congr rfl (fun ρ _ =>
Finset.sum_congr rfl (fun μ _ => Finset.sum_congr rfl (fun ν _ => ?_))))
rw [smul_tmul, tmul_smul, tmul_smul, smul_smul, smul_smul]
ring_nf
rw [Finset.sum_congr rfl (fun φ _ => Finset.sum_congr rfl (fun ρ _ =>
Finset.sum_congr rfl (fun μ _ => Finset.sum_smul.symm)))]
rw [Finset.sum_congr rfl (fun φ _ => Finset.sum_congr rfl (fun ρ _ => Finset.sum_smul.symm))]
trans ∑ φ : Fin 1 ⊕ Fin d, ∑ ρ : Fin 1 ⊕ Fin d,
(((g⁻¹.1.transpose * minkowskiMatrix * g⁻¹.1 : Matrix _ _ _) φ ρ) •
(CovariantLorentzVector.stdBasis φ) ⊗ₜ[] (CovariantLorentzVector.stdBasis ρ))
· refine Finset.sum_congr rfl (fun φ _ => Finset.sum_congr rfl (fun ρ _ => ?_))
apply congrFun
apply congrArg
simp only [Matrix.mul_apply, Matrix.transpose_apply]
rw [Finset.sum_comm]
refine Finset.sum_congr rfl (fun μ _ => ?_)
rw [Finset.sum_mul]
rw [LorentzGroup.transpose_mul_minkowskiMatrix_mul_self]
end minkowskiMatrix
end

View file

@ -94,6 +94,16 @@ lemma rep_apply (g : LorentzGroup d) : rep g v = (g.1⁻¹)ᵀ *ᵥ v := by
rw [← LorentzGroup.coe_inv]
rfl
lemma rep_apply_stdBasis (g : LorentzGroup d) (μ : Fin 1 ⊕ Fin d) :
rep g (stdBasis μ) = ∑ ν, g⁻¹.1 μ ν • stdBasis ν := by
simp only [rep_apply, Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton, decomp_stdBasis']
funext ν
simp [LorentzVector.stdBasis, Pi.basisFun_apply]
erw [Pi.basisFun_apply, Matrix.mulVec_stdBasis]
rw [← LorentzGroup.coe_inv]
simp
end CovariantLorentzVector
end

View file

@ -28,6 +28,14 @@ def rep : Representation (LorentzGroup d) (LorentzVector d) where
open Matrix in
lemma rep_apply (g : LorentzGroup d) : rep g v = g *ᵥ v := rfl
lemma rep_apply_stdBasis (g : LorentzGroup d) (μ : Fin 1 ⊕ Fin d) :
rep g (stdBasis μ) = ∑ ν, g.1.transpose μ ν • stdBasis ν := by
simp only [rep_apply, Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton, decomp_stdBasis']
funext ν
simp [LorentzVector.stdBasis, Pi.basisFun_apply]
erw [Pi.basisFun_apply, Matrix.mulVec_stdBasis]
end LorentzVector
end

View file

@ -79,6 +79,12 @@ lemma as_block : @minkowskiMatrix d = (
rw [← diagonal_neg]
rfl
@[simp]
lemma off_diag_zero {μ ν : Fin 1 ⊕ Fin d} (h : μ ≠ ν) : η μ ν = 0 := by
simp only [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
apply diagonal_apply_ne
exact h
end minkowskiMatrix
/-!