PhysLean/HepLean/SpaceTime/LorentzVector/Basic.lean
2024-07-02 10:40:35 -04:00

113 lines
2.9 KiB
Text
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 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.
-/
/- The number of space dimensions . -/
variable (d : )
/-- The type of Lorentz Vectors in `d`-space dimensions. -/
def LorentzVector : Type := (Fin 1 ⊕ Fin d) →
/-- An instance of a 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
/-- The structure of a topological space `LorentzVector d`. -/
instance : TopologicalSpace (LorentzVector d) :=
haveI : NormedAddCommGroup (LorentzVector d) := Pi.normedAddCommGroup
UniformSpace.toTopologicalSpace
namespace LorentzVector
variable {d : }
variable (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
-/
/-- The standard basis of `LorentzVector` indexed by `Fin 1 ⊕ Fin (d)`. -/
@[simps!]
noncomputable def stdBasis : Basis (Fin 1 ⊕ Fin (d)) (LorentzVector d) := Pi.basisFun _
/-- Notation for `stdBasis`. -/
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]
exact LinearMap.stdBasis_apply' μ ν
/-- The standard unit time vector. -/
noncomputable abbrev timeVec : (LorentzVector d) := e (Sum.inl 0)
lemma timeVec_space : (@timeVec d).space = 0 := by
funext i
simp only [space, Function.comp_apply, stdBasis_apply, Fin.isValue, ↓reduceIte, PiLp.zero_apply]
lemma timeVec_time: (@timeVec d).time = 1 := by
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
· rfl
· simp only [Sum.elim_inr, Pi.neg_apply]
apply neg_add
map_smul' c x := by
funext i
rcases i with i | i
· rfl
· simp [HSMul.hSMul, SMul.smul]
/-- The reflection of space. -/
@[simp]
def spaceReflection : LorentzVector d := spaceReflectionLin v
lemma spaceReflection_space : v.spaceReflection.space = - v.space := by
rfl
lemma spaceReflection_time : v.spaceReflection.time = v.time := by
rfl
end LorentzVector