PhysLean/HepLean/SpaceTime/LorentzVector/Basic.lean

136 lines
3.9 KiB
Text
Raw Normal View History

/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
2024-07-12 16:39:44 -04:00
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
/-!
# Lorentz vectors
(aka 4-vectors)
In this file we define a Lorentz vector (in 4d, this is more often called a 4-vector).
One of the most important example of a Lorentz vector is SpaceTime.
2024-07-29 16:54:59 -04:00
We will define the group action of the Lorentz group on Lorentz vectors in
`HepLean.SpaceTime.LorentzVector.LorentzAction` in such a way that `LorentzVector`
corresponds to contravariant Lorentz tensors.
-/
/- The number of space dimensions . -/
variable (d : )
2024-07-29 16:54:59 -04:00
/-- The type of (contravariant) Lorentz Vectors in `d`-space dimensions. -/
def LorentzVector : Type := (Fin 1 ⊕ Fin d) →
2024-07-15 14:52:50 -04:00
/-- An instance of an additive commutative monoid on `LorentzVector`. -/
instance : AddCommMonoid (LorentzVector d) := Pi.addCommMonoid
/-- An instance of a module on `LorentzVector`. -/
noncomputable instance : Module (LorentzVector d) := Pi.module _ _ _
instance : AddCommGroup (LorentzVector d) := Pi.addCommGroup
2024-07-02 10:13:52 -04:00
/-- The structure of a topological space `LorentzVector d`. -/
instance : TopologicalSpace (LorentzVector d) :=
haveI : NormedAddCommGroup (LorentzVector d) := Pi.normedAddCommGroup
UniformSpace.toTopologicalSpace
namespace LorentzVector
2024-07-18 08:36:56 -04:00
variable {d : } (v : LorentzVector d)
/-- The space components. -/
@[simp]
def space : EuclideanSpace (Fin d) := v ∘ Sum.inr
/-- The time component. -/
@[simp]
def time : := v (Sum.inl 0)
/-!
# The standard basis
-/
2024-07-02 10:13:52 -04:00
/-- The standard basis of `LorentzVector` indexed by `Fin 1 ⊕ Fin (d)`. -/
@[simps!]
noncomputable def stdBasis : Basis (Fin 1 ⊕ Fin (d)) (LorentzVector d) := Pi.basisFun _
2024-07-02 10:26:21 -04:00
/-- Notation for `stdBasis`. -/
2024-07-02 10:13:52 -04:00
scoped[LorentzVector] notation "e" => stdBasis
lemma stdBasis_apply (μ ν : Fin 1 ⊕ Fin d) : e μ ν = if μ = ν then 1 else 0 := by
rw [stdBasis]
erw [Pi.basisFun_apply]
2024-07-02 10:40:35 -04:00
exact LinearMap.stdBasis_apply' μ ν
2024-07-02 10:13:52 -04:00
2024-07-29 16:54:59 -04:00
lemma decomp_stdBasis (v : LorentzVector d) : ∑ i, v i • e i = v := by
funext ν
rw [Finset.sum_apply]
rw [Finset.sum_eq_single_of_mem ν]
simp [HSMul.hSMul, SMul.smul, stdBasis, Pi.basisFun_apply]
erw [Pi.basisFun_apply]
simp only [LinearMap.stdBasis_same, mul_one]
exact Finset.mem_univ ν
intros b _ hbi
simp [HSMul.hSMul, SMul.smul, stdBasis, Pi.basisFun_apply]
erw [Pi.basisFun_apply]
simp [LinearMap.stdBasis_apply]
exact Or.inr hbi
@[simp]
lemma decomp_stdBasis' (v : LorentzVector d) :
2024-07-30 07:51:07 -04:00
v (Sum.inl 0) • e (Sum.inl 0) + ∑ a₂ : Fin d, v (Sum.inr a₂) • e (Sum.inr a₂) = v := by
2024-07-29 16:54:59 -04:00
trans ∑ i, v i • e i
simp only [Fin.isValue, Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero,
Finset.sum_singleton]
exact decomp_stdBasis v
/-- The standard unit time vector. -/
2024-07-02 10:13:52 -04:00
noncomputable abbrev timeVec : (LorentzVector d) := e (Sum.inl 0)
lemma timeVec_space : (@timeVec d).space = 0 := by
funext i
2024-07-02 10:13:52 -04:00
simp only [space, Function.comp_apply, stdBasis_apply, Fin.isValue, ↓reduceIte, PiLp.zero_apply]
lemma timeVec_time: (@timeVec d).time = 1 := by
2024-07-02 10:13:52 -04:00
simp only [time, Fin.isValue, stdBasis_apply, ↓reduceIte]
/-!
# Reflection of space
-/
/-- The reflection of space as a linear map. -/
@[simps!]
def spaceReflectionLin : LorentzVector d →ₗ[] LorentzVector d where
toFun x := Sum.elim (x ∘ Sum.inl) (- x ∘ Sum.inr)
map_add' x y := by
funext i
rcases i with i | i
2024-07-02 10:40:35 -04:00
· rfl
· simp only [Sum.elim_inr, Pi.neg_apply]
apply neg_add
map_smul' c x := by
funext i
rcases i with i | i
2024-07-02 10:40:35 -04:00
· rfl
· simp [HSMul.hSMul, SMul.smul]
/-- The reflection of space. -/
@[simp]
def spaceReflection : LorentzVector d := spaceReflectionLin v
2024-07-18 08:36:56 -04:00
lemma spaceReflection_space : v.spaceReflection.space = - v.space := rfl
2024-07-18 08:36:56 -04:00
lemma spaceReflection_time : v.spaceReflection.time = v.time := rfl
end LorentzVector