refactor: Large refactor of Lorentz vecs

This commit is contained in:
jstoobysmith 2024-11-08 16:24:58 +00:00
parent a69cf91919
commit 3eb5da875f
8 changed files with 402 additions and 88 deletions

View file

@ -30,10 +30,18 @@ open minkowskiMatrix
Lorentz vectors. In index notation these have an up index `ψⁱ`. -/
def Contr (d : ) : Rep (LorentzGroup d) := Rep.of ContrMod.rep
instance : TopologicalSpace (Contr d) :=
haveI : NormedAddCommGroup (Contr d) := ContrMod.norm
UniformSpace.toTopologicalSpace
/-- The representation of `LorentzGroup d` on real vectors corresponding to covariant
Lorentz vectors. In index notation these have an up index `ψⁱ`. -/
def Co (d : ) : Rep (LorentzGroup d) := Rep.of CoMod.rep
open CategoryTheory.MonoidalCategory
def toField (d : ) : (𝟙_ (Rep ↑(LorentzGroup d))) →ₗ[] := LinearMap.id
/-!
## Isomorphism between contravariant and covariant Lorentz vectors

View file

@ -183,6 +183,17 @@ namespace contrContrContract
variable (x y : Contr d)
@[simp]
lemma action_tmul (g : LorentzGroup d) : ⟪(Contr d).ρ g x, (Contr d).ρ g y⟫ₘ = ⟪x, y⟫ₘ := by
conv =>
lhs
change (CategoryTheory.CategoryStruct.comp
((CategoryTheory.MonoidalCategory.tensorObj (Contr d) (Contr d)).ρ g)
contrContrContract.hom) (x ⊗ₜ[] y)
arg 1
apply contrContrContract.comm g
rfl
lemma as_sum : ⟪x, y⟫ₘ = x.val (Sum.inl 0) * y.val (Sum.inl 0) -
∑ i, x.val (Sum.inr i) * y.val (Sum.inr i) := by
rw [contrContrContract_hom_tmul]
@ -191,6 +202,26 @@ lemma as_sum : ⟪x, y⟫ₘ = x.val (Sum.inl 0) * y.val (Sum.inl 0) -
one_mul, Finset.sum_singleton, Sum.elim_inr, neg_mul, mul_neg, Finset.sum_neg_distrib]
rfl
open InnerProductSpace
lemma as_sum_toSpace : ⟪x, y⟫ₘ = x.val (Sum.inl 0) * y.val (Sum.inl 0) -
⟪x.toSpace, y.toSpace⟫_ := by
rw [as_sum]
rfl
@[simp]
lemma stdBasis_inl {d : } :
⟪@ContrMod.stdBasis d (Sum.inl 0), ContrMod.stdBasis (Sum.inl 0)⟫ₘ = (1 : ) := by
rw [as_sum]
trans (1 : ) - (0 : )
congr
· rw [ContrMod.stdBasis_apply_same]
simp
· rw [Fintype.sum_eq_zero]
intro a
simp
· ring
lemma symm : ⟪x, y⟫ₘ = ⟪y, x⟫ₘ := by
rw [as_sum, as_sum]
congr 1
@ -316,6 +347,39 @@ lemma _root_.LorentzGroup.mem_iff_norm : Λ ∈ LorentzGroup d ↔
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, add_sub_cancel, neg_add_cancel, e]
/-!
## Some equalities and inequalities
-/
lemma inl_sq_eq (v : Contr d) : v.val (Sum.inl 0) ^ 2 =
(toField d ⟪v, v⟫ₘ) + ∑ i, v.val (Sum.inr i) ^ 2:= by
rw [as_sum]
apply sub_eq_iff_eq_add.mp
congr
· exact pow_two (v.val (Sum.inl 0))
· funext i
exact pow_two (v.val (Sum.inr i))
lemma le_inl_sq (v : Contr d) : toField d ⟪v, v⟫ₘ ≤ v.val (Sum.inl 0) ^ 2 := by
rw [inl_sq_eq]
apply (le_add_iff_nonneg_right _).mpr
refine Fintype.sum_nonneg ?hf
exact fun i => pow_two_nonneg (v.val (Sum.inr i))
lemma ge_abs_inner_product (v w : Contr d) : v.val (Sum.inl 0) * w.val (Sum.inl 0) -
‖⟪v.toSpace, w.toSpace⟫_‖ ≤ ⟪v, w⟫ₘ := by
rw [as_sum_toSpace, sub_le_sub_iff_left]
exact Real.le_norm_self ⟪v.toSpace, w.toSpace⟫_
lemma ge_sub_norm (v w : Contr d) : v.val (Sum.inl 0) * w.val (Sum.inl 0) -
‖v.toSpace‖ * ‖w.toSpace‖ ≤ ⟪v, w⟫ₘ := by
apply le_trans _ (ge_abs_inner_product v w)
rw [sub_le_sub_iff_left]
exact norm_inner_le_norm v.toSpace w.toSpace
end contrContrContract
end Lorentz

View file

@ -93,6 +93,10 @@ lemma stdBasis_toFin1dEquiv_apply_same (μ : Fin 1 ⊕ Fin d) :
rw [@LinearEquiv.apply_symm_apply]
exact Pi.single_eq_same μ 1
@[simp]
lemma stdBasis_apply_same (μ : Fin 1 ⊕ Fin d) : (stdBasis μ).val μ = 1 :=
stdBasis_toFin1dEquiv_apply_same μ
lemma stdBasis_toFin1dEquiv_apply_ne {μ ν : Fin 1 ⊕ Fin d} (h : μ ≠ ν) :
toFin1dEquiv (stdBasis μ) ν = 0 := by
simp only [stdBasis, Basis.ofEquivFun, Basis.coe_ofRepr, LinearEquiv.trans_symm,
@ -100,6 +104,11 @@ lemma stdBasis_toFin1dEquiv_apply_ne {μ ν : Fin 1 ⊕ Fin d} (h : μ ≠ ν
rw [@LinearEquiv.apply_symm_apply]
exact Pi.single_eq_of_ne' h 1
@[simp]
lemma stdBasis_inl_apply_inr (i : Fin d) : (stdBasis (Sum.inl 0)).val (Sum.inr i) = 0 := by
refine stdBasis_toFin1dEquiv_apply_ne ?_
simp
/-- 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 toFin1dEquiv.injective
@ -151,6 +160,22 @@ lemma mulVec_mulVec (M N : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ) (v :
/-!
## The norm
(Not the Minkowski norm, but the norm of a vector in `ContrModule d`.)
-/
def norm : NormedAddCommGroup (ContrMod d) where
norm v := ‖v.val‖₊
dist_self x := Pi.normedAddCommGroup.dist_self x.val
dist_triangle x y z := Pi.normedAddCommGroup.dist_triangle x.val y.val z.val
dist_comm x y := Pi.normedAddCommGroup.dist_comm x.val y.val
eq_of_dist_eq_zero {x y} := fun h => ext (MetricSpace.eq_of_dist_eq_zero h)
def toSpace (v : ContrMod d) : EuclideanSpace (Fin d) := v.val ∘ Sum.inr
/-!
## The representation.
-/
@ -230,6 +255,10 @@ lemma stdBasis_toFin1dEquiv_apply_same (μ : Fin 1 ⊕ Fin d) :
rw [@LinearEquiv.apply_symm_apply]
exact Pi.single_eq_same μ 1
@[simp]
lemma stdBasis_apply_same (μ : Fin 1 ⊕ Fin d) : (stdBasis μ).val μ = 1 :=
stdBasis_toFin1dEquiv_apply_same μ
lemma stdBasis_toFin1dEquiv_apply_ne {μ ν : Fin 1 ⊕ Fin d} (h : μ ≠ ν) :
toFin1dEquiv (stdBasis μ) ν = 0 := by
simp only [stdBasis, Basis.ofEquivFun, Basis.coe_ofRepr, LinearEquiv.trans_symm,

View file

@ -0,0 +1,218 @@
/-
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.Real.Contraction
import Mathlib.GroupTheory.GroupAction.Blocks
/-!
# Lorentz vectors with norm one
-/
open TensorProduct
namespace Lorentz
namespace Contr
variable {d : }
/-- The set of Lorentz vectors with norm 1. -/
def NormOne (d : ) : Set (Contr d) := fun v => ⟪v, v⟫ₘ = (1 : )
noncomputable section
namespace NormOne
lemma mem_iff {v : Contr d} : v ∈ NormOne d ↔ ⟪v, v⟫ₘ = (1 : ) := by
rfl
lemma mem_mulAction (g : LorentzGroup d) (v : Contr d) :
v ∈ NormOne d ↔ (Contr d).ρ g v ∈ NormOne d := by
rw [mem_iff, mem_iff, contrContrContract.action_tmul]
instance : TopologicalSpace (NormOne d) := instTopologicalSpaceSubtype
variable (v w : NormOne d)
/-- The negative of a `NormOne` as a `NormOne`. -/
def neg : NormOne d := ⟨- v, by
rw [mem_iff]
simp only [Action.instMonoidalCategory_tensorUnit_V, Action.instMonoidalCategory_tensorObj_V,
CategoryTheory.Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, tmul_neg, neg_tmul, neg_neg]
exact v.2⟩
/-- The first column of a Lorentz matrix as a `NormOneLorentzVector`. -/
@[simps!]
def _root_.LorentzGroup.toNormOne (Λ : LorentzGroup d) : NormOne d :=
⟨(Contr d).ρ Λ (ContrMod.stdBasis (Sum.inl 0)), by
rw [mem_iff, contrContrContract.action_tmul, contrContrContract.stdBasis_inl]⟩
lemma _root_.LorentzGroup.toNormOne_inl (Λ : LorentzGroup d) :
(LorentzGroup.toNormOne Λ).val.val (Sum.inl 0) = Λ.1 (Sum.inl 0) (Sum.inl 0) := by
simp only [Fin.isValue, LorentzGroup.toNormOne_coe_val, Finsupp.single, one_ne_zero, ↓reduceIte,
Finsupp.coe_mk, Matrix.mulVec_single, mul_one]
lemma _root_.LorentzGroup.toNormOne_inr (Λ : LorentzGroup d) (i : Fin d) :
(LorentzGroup.toNormOne Λ).val.val (Sum.inr i) = Λ.1 (Sum.inr i) (Sum.inl 0) := by
simp only [LorentzGroup.toNormOne_coe_val, Finsupp.single, one_ne_zero, ↓reduceIte, Fin.isValue,
Finsupp.coe_mk, Matrix.mulVec_single, mul_one]
lemma _root_.LorentzGroup.inl_inl_mul (Λ Λ' : LorentzGroup d) : (Λ * Λ').1 (Sum.inl 0) (Sum.inl 0) =
⟪(LorentzGroup.toNormOne (LorentzGroup.transpose Λ)).1,
(Contr d).ρ LorentzGroup.parity (LorentzGroup.toNormOne Λ').1⟫ₘ := by
rw [contrContrContract.right_parity]
simp only [Fin.isValue, lorentzGroupIsGroup_mul_coe, Matrix.mul_apply, Fintype.sum_sum_type,
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton,
LorentzGroup.transpose, PiLp.inner_apply, Function.comp_apply,
RCLike.inner_apply, conj_trivial]
congr
· rw [LorentzGroup.toNormOne_inl]
rfl
· rw [LorentzGroup.toNormOne_inl]
· funext x
rw [LorentzGroup.toNormOne_inr, LorentzGroup.toNormOne_inr]
rfl
lemma inl_sq : v.val.val (Sum.inl 0) ^ 2 = 1 + ‖ContrMod.toSpace v.val‖ ^ 2 := by
rw [contrContrContract.inl_sq_eq, v.2]
congr
rw [← real_inner_self_eq_norm_sq]
simp only [PiLp.inner_apply, RCLike.inner_apply, conj_trivial]
congr
funext x
exact pow_two ((v.val).val (Sum.inr x))
lemma one_le_abs_inl : 1 ≤ |v.val.val (Sum.inl 0)| := by
have h1 := contrContrContract.le_inl_sq v.val
rw [v.2] at h1
exact (one_le_sq_iff_one_le_abs _).mp h1
lemma inl_le_neg_one_or_one_le_inl : v.val.val (Sum.inl 0) ≤ -1 1 ≤ v.val.val (Sum.inl 0) :=
le_abs'.mp (one_le_abs_inl v)
lemma norm_space_le_abs_inl : ‖v.1.toSpace‖ < |v.val.val (Sum.inl 0)| := by
rw [(abs_norm _).symm, ← @sq_lt_sq, inl_sq]
change ‖ContrMod.toSpace v.val‖ ^ 2 < 1 + ‖ContrMod.toSpace v.val‖ ^ 2
exact lt_one_add (‖(v.1).toSpace‖ ^ 2)
lemma norm_space_leq_abs_inl : ‖v.1.toSpace‖ ≤ |v.val.val (Sum.inl 0)| :=
le_of_lt (norm_space_le_abs_inl v)
lemma inl_abs_sub_space_norm :
0 ≤ |v.val.val (Sum.inl 0)| * |w.val.val (Sum.inl 0)| - ‖v.1.toSpace‖ * ‖w.1.toSpace‖ := by
apply sub_nonneg.mpr
apply mul_le_mul (norm_space_leq_abs_inl v) (norm_space_leq_abs_inl w) ?_ ?_
· exact norm_nonneg _
· exact abs_nonneg _
/-!
# Future pointing norm one Lorentz vectors
-/
/-- The future pointing Lorentz vectors with Norm one. -/
def FuturePointing (d : ) : Set (NormOne d) :=
fun x => 0 < x.val.val (Sum.inl 0)
namespace FuturePointing
lemma mem_iff : v ∈ FuturePointing d ↔ 0 < v.val.val (Sum.inl 0) := by
rfl
lemma mem_iff_inl_nonneg : v ∈ FuturePointing d ↔ 0 ≤ v.val.val (Sum.inl 0) := by
refine Iff.intro (fun h => le_of_lt h) (fun h => ?_)
rw [mem_iff]
rcases inl_le_neg_one_or_one_le_inl v with (h | h)
· linarith
· linarith
lemma mem_iff_inl_one_le_inl : v ∈ FuturePointing d ↔ 1 ≤ v.val.val (Sum.inl 0) := by
rw [mem_iff_inl_nonneg]
refine Iff.intro (fun h => ?_) (fun h => ?_)
· rcases inl_le_neg_one_or_one_le_inl v with (h | h)
· linarith
· linarith
· linarith
lemma mem_iff_parity_mem : v ∈ FuturePointing d ↔ ⟨(Contr d).ρ LorentzGroup.parity v.1,
(NormOne.mem_mulAction _ _).mp v.2⟩ ∈ FuturePointing d := by
rw [mem_iff, mem_iff]
change _ ↔ 0 < (minkowskiMatrix.mulVec v.val.val) (Sum.inl 0)
simp only [Fin.isValue, minkowskiMatrix.mulVec_inl_0]
lemma not_mem_iff_inl_le_zero : v ∉ FuturePointing d ↔ v.val.val (Sum.inl 0) ≤ 0 := by
rw [mem_iff]
simp
lemma not_mem_iff_inl_lt_zero : v ∉ FuturePointing d ↔ v.val.val (Sum.inl 0) < 0 := by
rw [mem_iff_inl_nonneg]
simp
lemma not_mem_iff_inl_le_neg_one : v ∉ FuturePointing d ↔ v.val.val (Sum.inl 0) ≤ -1 := by
rw [not_mem_iff_inl_le_zero]
refine Iff.intro (fun h => ?_) (fun h => ?_)
· rcases inl_le_neg_one_or_one_le_inl v with (h | h)
· linarith
· linarith
· linarith
lemma not_mem_iff_neg : v ∉ FuturePointing d ↔ neg v ∈ FuturePointing d := by
rw [not_mem_iff_inl_le_zero, mem_iff_inl_nonneg]
simp only [Fin.isValue, neg]
change (v).val.val (Sum.inl 0) ≤ 0 ↔ 0 ≤ - (v.val).val (Sum.inl 0)
simp
variable (f f' : FuturePointing d)
lemma inl_nonneg : 0 ≤ f.val.val.val (Sum.inl 0):= le_of_lt f.2
lemma abs_inl : |f.val.val.val (Sum.inl 0)| = f.val.val.val (Sum.inl 0) :=
abs_of_nonneg (inl_nonneg f)
open InnerProductSpace
lemma metric_nonneg : 0 ≤ toField d ⟪f.1.1, f'.1.1⟫ₘ := by
apply le_trans (inl_abs_sub_space_norm f f'.1)
rw [abs_inl f, abs_inl f']
exact contrContrContract.ge_sub_norm f.1.1 f'.1.1
variable {v w : NormOne d}
lemma metric_reflect_mem_mem (h : v ∈ FuturePointing d) (hw : w ∈ FuturePointing d) :
0 ≤ toField d ⟪v.val, (Contr d).ρ LorentzGroup.parity w.1⟫ₘ :=
metric_nonneg ⟨v, h⟩ ⟨⟨(Contr d).ρ LorentzGroup.parity w.1,
(NormOne.mem_mulAction _ _).mp w.2⟩, (mem_iff_parity_mem w).mp hw⟩
lemma metric_reflect_not_mem_not_mem (h : v ∉ FuturePointing d) (hw : w ∉ FuturePointing d) :
0 ≤ toField d ⟪v.val, (Contr d).ρ LorentzGroup.parity w.1⟫ₘ := by
have h1 := metric_reflect_mem_mem ((not_mem_iff_neg v).mp h) ((not_mem_iff_neg w).mp hw)
apply le_of_le_of_eq h1 ?_
simp [neg, neg_tmul, tmul_neg]
lemma metric_reflect_mem_not_mem (h : v ∈ FuturePointing d) (hw : w ∉ FuturePointing d) :
toField d ⟪v.val, (Contr d).ρ LorentzGroup.parity w.1⟫ₘ ≤ 0 := by
rw [show (0 : ) = - 0 from zero_eq_neg.mpr rfl, le_neg]
have h1 := metric_reflect_mem_mem h ((not_mem_iff_neg w).mp hw)
apply le_of_le_of_eq h1 ?_
simp [neg, neg_tmul, tmul_neg]
lemma metric_reflect_not_mem_mem (h : v ∉ FuturePointing d) (hw : w ∈ FuturePointing d) :
toField d ⟪v.val, (Contr d).ρ LorentzGroup.parity w.1⟫ₘ ≤ 0 := by
rw [show (0 : ) = - 0 from zero_eq_neg.mpr rfl, le_neg]
have h1 := metric_reflect_mem_mem ((not_mem_iff_neg v).mp h) hw
apply le_of_le_of_eq h1 ?_
simp [neg, neg_tmul, tmul_neg]
end FuturePointing
end NormOne
end
end Contr
end Lorentz