refactor: Major refactor of Lorentz vecs
This commit is contained in:
parent
3eb5da875f
commit
5cc188146f
20 changed files with 494 additions and 1005 deletions
|
@ -7,6 +7,7 @@ import HepLean.Meta.Informal
|
|||
import HepLean.SpaceTime.LorentzGroup.Basic
|
||||
import Mathlib.RepresentationTheory.Rep
|
||||
import Mathlib.Logic.Equiv.TransferInstance
|
||||
import HepLean.SpaceTime.PauliMatrices.SelfAdjoint
|
||||
/-!
|
||||
|
||||
## Modules associated with Real Lorentz vectors
|
||||
|
@ -74,6 +75,7 @@ def toFin1dℝEquiv : ContrMod d ≃ₗ[ℝ] (Fin 1 ⊕ Fin d → ℝ) :=
|
|||
through the linear equivalence `toFin1dℝEquiv`. -/
|
||||
abbrev toFin1dℝ (ψ : ContrMod d) := toFin1dℝEquiv ψ
|
||||
|
||||
lemma toFin1dℝ_eq_val (ψ : ContrMod d) : ψ.toFin1dℝ = ψ.val := by rfl
|
||||
/-!
|
||||
|
||||
## The standard basis.
|
||||
|
@ -109,6 +111,13 @@ lemma stdBasis_inl_apply_inr (i : Fin d) : (stdBasis (Sum.inl 0)).val (Sum.inr i
|
|||
refine stdBasis_toFin1dℝEquiv_apply_ne ?_
|
||||
simp
|
||||
|
||||
lemma stdBasis_apply (μ ν : Fin 1 ⊕ Fin d) : (stdBasis μ).val ν = if μ = ν then 1 else 0 := by
|
||||
simp [stdBasis, Pi.basisFun_apply, Pi.single_apply]
|
||||
change Pi.single μ 1 ν = _
|
||||
simp [Pi.single_apply]
|
||||
refine ite_congr ?h₁ (congrFun rfl) (congrFun rfl)
|
||||
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)
|
||||
|
||||
/-- Decomposition of a contrvariant Lorentz vector into the standard basis. -/
|
||||
lemma stdBasis_decomp (v : ContrMod d) : v = ∑ i, v.toFin1dℝ i • stdBasis i := by
|
||||
apply toFin1dℝEquiv.injective
|
||||
|
@ -192,6 +201,76 @@ lemma rep_apply_toFin1dℝ (g : LorentzGroup d) (ψ : ContrMod d) :
|
|||
(rep g ψ).toFin1dℝ = g.1 *ᵥ ψ.toFin1dℝ := by
|
||||
rfl
|
||||
|
||||
/-!
|
||||
|
||||
## To Self-Adjoint Matrix
|
||||
-/
|
||||
|
||||
/-- The linear equivalence between the vector-space `ContrMod 3` and self-adjoint
|
||||
`2×2`-complex matrices. -/
|
||||
def toSelfAdjoint : ContrMod 3 ≃ₗ[ℝ] selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ) :=
|
||||
toFin1dℝEquiv ≪≫ₗ (Finsupp.linearEquivFunOnFinite ℝ ℝ (Fin 1 ⊕ Fin 3)).symm ≪≫ₗ
|
||||
PauliMatrix.σSAL.repr.symm
|
||||
|
||||
lemma toSelfAdjoint_apply (x : ContrMod 3) : toSelfAdjoint x =
|
||||
x.toFin1dℝ (Sum.inl 0) • ⟨PauliMatrix.σ0, PauliMatrix.σ0_selfAdjoint⟩
|
||||
- x.toFin1dℝ (Sum.inr 0) • ⟨PauliMatrix.σ1, PauliMatrix.σ1_selfAdjoint⟩
|
||||
- x.toFin1dℝ (Sum.inr 1) • ⟨PauliMatrix.σ2, PauliMatrix.σ2_selfAdjoint⟩
|
||||
- x.toFin1dℝ (Sum.inr 2) • ⟨PauliMatrix.σ3, PauliMatrix.σ3_selfAdjoint⟩ := by
|
||||
simp only [toSelfAdjoint, 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.toFin1dℝ 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 toSelfAdjoint_apply_coe (x : ContrMod 3) : (toSelfAdjoint x).1 =
|
||||
x.toFin1dℝ (Sum.inl 0) • PauliMatrix.σ0
|
||||
- x.toFin1dℝ (Sum.inr 0) • PauliMatrix.σ1
|
||||
- x.toFin1dℝ (Sum.inr 1) • PauliMatrix.σ2
|
||||
- x.toFin1dℝ (Sum.inr 2) • PauliMatrix.σ3 := by
|
||||
rw [toSelfAdjoint_apply]
|
||||
rfl
|
||||
|
||||
|
||||
lemma toSelfAdjoint_stdBasis (i : Fin 1 ⊕ Fin 3) :
|
||||
toSelfAdjoint (stdBasis i) = PauliMatrix.σSAL i := by
|
||||
rw [toSelfAdjoint_apply]
|
||||
match i with
|
||||
| Sum.inl 0 =>
|
||||
simp only [stdBasis, Fin.isValue, Basis.coe_ofEquivFun, LinearEquiv.apply_symm_apply,
|
||||
Pi.single_eq_same, one_smul, ne_eq, reduceCtorEq, not_false_eq_true, Pi.single_eq_of_ne,
|
||||
zero_smul, sub_zero, PauliMatrix.σSAL, Basis.coe_mk, PauliMatrix.σSAL']
|
||||
| Sum.inr 0 =>
|
||||
simp only [stdBasis, Fin.isValue, Basis.coe_ofEquivFun, LinearEquiv.apply_symm_apply, 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']
|
||||
rfl
|
||||
| Sum.inr 1 =>
|
||||
simp only [stdBasis, Fin.isValue, Basis.coe_ofEquivFun, LinearEquiv.apply_symm_apply, 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']
|
||||
rfl
|
||||
| Sum.inr 2 =>
|
||||
simp only [stdBasis, Fin.isValue, Basis.coe_ofEquivFun, LinearEquiv.apply_symm_apply, 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']
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma toSelfAdjoint_symm_basis (i : Fin 1 ⊕ Fin 3) :
|
||||
toSelfAdjoint.symm (PauliMatrix.σSAL i) = (stdBasis i) := by
|
||||
refine (LinearEquiv.symm_apply_eq toSelfAdjoint).mpr ?_
|
||||
rw [toSelfAdjoint_stdBasis]
|
||||
|
||||
end ContrMod
|
||||
|
||||
/-- The module for covariant (up-index) complex Lorentz vectors. -/
|
||||
|
@ -266,6 +345,14 @@ lemma stdBasis_toFin1dℝEquiv_apply_ne {μ ν : Fin 1 ⊕ Fin d} (h : μ ≠ ν
|
|||
rw [@LinearEquiv.apply_symm_apply]
|
||||
exact Pi.single_eq_of_ne' h 1
|
||||
|
||||
lemma stdBasis_apply (μ ν : Fin 1 ⊕ Fin d) : (stdBasis μ).val ν = if μ = ν then 1 else 0 := by
|
||||
simp [stdBasis, Pi.basisFun_apply, Pi.single_apply]
|
||||
change Pi.single μ 1 ν = _
|
||||
simp [Pi.single_apply]
|
||||
refine ite_congr ?h₁ (congrFun rfl) (congrFun rfl)
|
||||
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)
|
||||
|
||||
|
||||
/-- Decomposition of a covariant Lorentz vector into the standard basis. -/
|
||||
lemma stdBasis_decomp (v : CoMod d) : v = ∑ i, v.toFin1dℝ i • stdBasis i := by
|
||||
apply toFin1dℝEquiv.injective
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue