2024-05-17 11:52:16 -04:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license.
|
|
|
|
|
Authors: Joseph Tooby-Smith
|
|
|
|
|
-/
|
2024-07-01 16:56:15 -04:00
|
|
|
|
import HepLean.SpaceTime.MinkowskiMetric
|
2024-07-02 10:13:52 -04:00
|
|
|
|
import HepLean.SpaceTime.LorentzVector.AsSelfAdjointMatrix
|
2024-07-01 16:56:15 -04:00
|
|
|
|
import HepLean.SpaceTime.LorentzVector.NormOne
|
2024-05-17 11:52:16 -04:00
|
|
|
|
/-!
|
|
|
|
|
# The Lorentz Group
|
|
|
|
|
|
|
|
|
|
We define the Lorentz group.
|
|
|
|
|
|
|
|
|
|
## TODO
|
|
|
|
|
|
|
|
|
|
- Show that the Lorentz is a Lie group.
|
|
|
|
|
- Prove that the restricted Lorentz group is equivalent to the connected component of the
|
|
|
|
|
identity.
|
|
|
|
|
- Define the continuous maps from `ℝ³` to `restrictedLorentzGroup` defining boosts.
|
|
|
|
|
|
|
|
|
|
## References
|
|
|
|
|
|
|
|
|
|
- http://home.ku.edu.tr/~amostafazadeh/phys517_518/phys517_2016f/Handouts/A_Jaffi_Lorentz_Group.pdf
|
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
noncomputable section
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
open Manifold
|
|
|
|
|
open Matrix
|
|
|
|
|
open Complex
|
|
|
|
|
open ComplexConjugate
|
|
|
|
|
|
2024-06-15 17:08:08 -04:00
|
|
|
|
/-!
|
2024-07-01 16:56:15 -04:00
|
|
|
|
## Matrices which preserves the Minkowski metric
|
2024-06-15 17:08:08 -04:00
|
|
|
|
|
|
|
|
|
We start studying the properties of matrices which preserve `ηLin`.
|
|
|
|
|
These matrices form the Lorentz group, which we will define in the next section at `lorentzGroup`.
|
|
|
|
|
|
|
|
|
|
-/
|
2024-07-01 16:56:15 -04:00
|
|
|
|
variable {d : ℕ}
|
2024-06-15 17:08:08 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
open minkowskiMetric in
|
|
|
|
|
/-- The Lorentz group is the subset of matrices which preserve the minkowski metric. -/
|
|
|
|
|
def LorentzGroup (d : ℕ) : Set (Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ) :=
|
|
|
|
|
{Λ : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ |
|
|
|
|
|
∀ (x y : LorentzVector d), ⟪Λ *ᵥ x, Λ *ᵥ y⟫ₘ = ⟪x, y⟫ₘ}
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
namespace LorentzGroup
|
|
|
|
|
/-- Notation for the Lorentz group. -/
|
|
|
|
|
scoped[LorentzGroup] notation (name := lorentzGroup_notation) "𝓛" => LorentzGroup
|
|
|
|
|
|
|
|
|
|
open minkowskiMetric
|
|
|
|
|
|
|
|
|
|
variable {Λ Λ' : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ}
|
|
|
|
|
|
|
|
|
|
/-!
|
|
|
|
|
|
|
|
|
|
# Membership conditions
|
|
|
|
|
|
|
|
|
|
-/
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma mem_iff_norm : Λ ∈ LorentzGroup d ↔
|
|
|
|
|
∀ (x : LorentzVector d), ⟪Λ *ᵥ x, Λ *ᵥ x⟫ₘ = ⟪x, x⟫ₘ := by
|
2024-06-13 10:57:25 -04:00
|
|
|
|
refine Iff.intro (fun h x => h x x) (fun h x y => ?_)
|
|
|
|
|
have hp := h (x + y)
|
|
|
|
|
have hn := h (x - y)
|
|
|
|
|
rw [mulVec_add] at hp
|
|
|
|
|
rw [mulVec_sub] at hn
|
|
|
|
|
simp only [map_add, LinearMap.add_apply, map_sub, LinearMap.sub_apply] at hp hn
|
2024-07-01 16:56:15 -04:00
|
|
|
|
rw [symm (Λ *ᵥ y) (Λ *ᵥ x), symm y x] at hp hn
|
2024-06-13 10:57:25 -04:00
|
|
|
|
linear_combination hp / 4 + -1 * hn / 4
|
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma mem_iff_on_right : Λ ∈ LorentzGroup d ↔
|
|
|
|
|
∀ (x y : LorentzVector d), ⟪x, (dual Λ * Λ) *ᵥ y⟫ₘ = ⟪x, y⟫ₘ := by
|
2024-05-17 11:52:16 -04:00
|
|
|
|
apply Iff.intro
|
2024-06-13 08:10:08 -04:00
|
|
|
|
intro h x y
|
2024-05-17 11:52:16 -04:00
|
|
|
|
have h1 := h x y
|
2024-07-01 16:56:15 -04:00
|
|
|
|
rw [← dual_mulVec_right, mulVec_mulVec] at h1
|
2024-05-17 11:52:16 -04:00
|
|
|
|
exact h1
|
2024-06-13 08:10:08 -04:00
|
|
|
|
intro h x y
|
2024-07-01 16:56:15 -04:00
|
|
|
|
rw [← dual_mulVec_right, mulVec_mulVec]
|
2024-05-17 11:52:16 -04:00
|
|
|
|
exact h x y
|
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma mem_iff_dual_mul_self : Λ ∈ LorentzGroup d ↔ dual Λ * Λ = 1 := by
|
|
|
|
|
rw [mem_iff_on_right, matrix_eq_id_iff]
|
|
|
|
|
exact forall_comm
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma mem_iff_self_mul_dual : Λ ∈ LorentzGroup d ↔ Λ * dual Λ = 1 := by
|
|
|
|
|
rw [mem_iff_dual_mul_self]
|
2024-06-13 08:10:08 -04:00
|
|
|
|
exact mul_eq_one_comm
|
|
|
|
|
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma mem_iff_transpose : Λ ∈ LorentzGroup d ↔ Λᵀ ∈ LorentzGroup d := by
|
2024-05-17 11:52:16 -04:00
|
|
|
|
apply Iff.intro
|
2024-07-01 16:56:15 -04:00
|
|
|
|
· intro h
|
|
|
|
|
have h1 := congrArg transpose ((mem_iff_dual_mul_self).mp h)
|
|
|
|
|
rw [dual, transpose_mul, transpose_mul, transpose_mul, minkowskiMatrix.eq_transpose,
|
|
|
|
|
← mul_assoc, transpose_one] at h1
|
|
|
|
|
rw [mem_iff_self_mul_dual, ← h1, dual]
|
|
|
|
|
noncomm_ring
|
|
|
|
|
· intro h
|
|
|
|
|
have h1 := congrArg transpose ((mem_iff_dual_mul_self).mp h)
|
|
|
|
|
rw [dual, transpose_mul, transpose_mul, transpose_mul, minkowskiMatrix.eq_transpose,
|
|
|
|
|
← mul_assoc, transpose_one, transpose_transpose] at h1
|
|
|
|
|
rw [mem_iff_self_mul_dual, ← h1, dual]
|
|
|
|
|
noncomm_ring
|
|
|
|
|
|
|
|
|
|
lemma mem_mul (hΛ : Λ ∈ LorentzGroup d) (hΛ' : Λ' ∈ LorentzGroup d) : Λ * Λ' ∈ LorentzGroup d := by
|
|
|
|
|
rw [mem_iff_dual_mul_self, dual_mul]
|
|
|
|
|
trans dual Λ' * (dual Λ * Λ) * Λ'
|
2024-06-13 10:57:25 -04:00
|
|
|
|
noncomm_ring
|
2024-07-01 16:56:15 -04:00
|
|
|
|
rw [(mem_iff_dual_mul_self).mp hΛ]
|
|
|
|
|
simp [(mem_iff_dual_mul_self).mp hΛ']
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma one_mem : 1 ∈ LorentzGroup d := by
|
|
|
|
|
rw [mem_iff_dual_mul_self]
|
2024-05-22 13:34:53 -04:00
|
|
|
|
simp
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma dual_mem (h : Λ ∈ LorentzGroup d) : dual Λ ∈ LorentzGroup d := by
|
|
|
|
|
rw [mem_iff_dual_mul_self, dual_dual]
|
|
|
|
|
exact mem_iff_self_mul_dual.mp h
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
end LorentzGroup
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-06-15 17:08:08 -04:00
|
|
|
|
/-!
|
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
# The Lorentz group as a group
|
2024-06-15 17:08:08 -04:00
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
2024-05-22 13:34:53 -04:00
|
|
|
|
@[simps mul_coe one_coe inv div]
|
2024-07-01 16:56:15 -04:00
|
|
|
|
instance lorentzGroupIsGroup : Group (LorentzGroup d) where
|
|
|
|
|
mul A B := ⟨A.1 * B.1, LorentzGroup.mem_mul A.2 B.2⟩
|
2024-05-22 13:34:53 -04:00
|
|
|
|
mul_assoc A B C := by
|
|
|
|
|
apply Subtype.eq
|
|
|
|
|
exact Matrix.mul_assoc A.1 B.1 C.1
|
2024-07-01 16:56:15 -04:00
|
|
|
|
one := ⟨1, LorentzGroup.one_mem⟩
|
2024-05-22 13:34:53 -04:00
|
|
|
|
one_mul A := by
|
|
|
|
|
apply Subtype.eq
|
|
|
|
|
exact Matrix.one_mul A.1
|
|
|
|
|
mul_one A := by
|
|
|
|
|
apply Subtype.eq
|
|
|
|
|
exact Matrix.mul_one A.1
|
2024-07-01 16:56:15 -04:00
|
|
|
|
inv A := ⟨minkowskiMetric.dual A.1, LorentzGroup.dual_mem A.2⟩
|
2024-05-22 13:34:53 -04:00
|
|
|
|
mul_left_inv A := by
|
|
|
|
|
apply Subtype.eq
|
2024-07-01 16:56:15 -04:00
|
|
|
|
exact LorentzGroup.mem_iff_dual_mul_self.mp A.2
|
2024-05-22 13:34:53 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
/-- `LorentzGroup` has the subtype topology. -/
|
|
|
|
|
instance : TopologicalSpace (LorentzGroup d) := instTopologicalSpaceSubtype
|
2024-05-22 13:34:53 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
namespace LorentzGroup
|
2024-05-22 13:34:53 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
open minkowskiMetric
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
variable {Λ Λ' : LorentzGroup d}
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
@[simp]
|
|
|
|
|
lemma coe_inv : (Λ⁻¹).1 = Λ.1⁻¹:= by
|
2024-05-22 13:34:53 -04:00
|
|
|
|
refine (inv_eq_left_inv ?h).symm
|
2024-07-01 16:56:15 -04:00
|
|
|
|
exact mem_iff_dual_mul_self.mp Λ.2
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-06-15 17:08:08 -04:00
|
|
|
|
/-- The transpose of an matrix in the Lorentz group is an element of the Lorentz group. -/
|
2024-07-01 16:56:15 -04:00
|
|
|
|
def transpose (Λ : LorentzGroup d) : LorentzGroup d :=
|
|
|
|
|
⟨Λ.1ᵀ, mem_iff_transpose.mp Λ.2⟩
|
2024-06-15 17:08:08 -04:00
|
|
|
|
|
|
|
|
|
/-!
|
|
|
|
|
|
|
|
|
|
## Lorentz group as a topological group
|
|
|
|
|
|
|
|
|
|
We now show that the Lorentz group is a topological group.
|
|
|
|
|
We do this by showing that the natrual map from the Lorentz group to `GL (Fin 4) ℝ` is an
|
|
|
|
|
embedding.
|
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
2024-05-22 13:34:53 -04:00
|
|
|
|
/-- The homomorphism of the Lorentz group into `GL (Fin 4) ℝ`. -/
|
2024-07-01 16:56:15 -04:00
|
|
|
|
def toGL : LorentzGroup d →* GL (Fin 1 ⊕ Fin d) ℝ where
|
|
|
|
|
toFun A := ⟨A.1, (A⁻¹).1, mul_eq_one_comm.mpr $ mem_iff_dual_mul_self.mp A.2,
|
|
|
|
|
mem_iff_dual_mul_self.mp A.2⟩
|
2024-05-22 13:34:53 -04:00
|
|
|
|
map_one' := by
|
|
|
|
|
simp
|
|
|
|
|
rfl
|
|
|
|
|
map_mul' x y := by
|
|
|
|
|
simp only [lorentzGroupIsGroup, _root_.mul_inv_rev, coe_inv]
|
|
|
|
|
ext
|
|
|
|
|
rfl
|
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma toGL_injective : Function.Injective (@toGL d) := by
|
2024-05-22 13:34:53 -04:00
|
|
|
|
intro A B h
|
|
|
|
|
apply Subtype.eq
|
|
|
|
|
rw [@Units.ext_iff] at h
|
2024-06-13 08:10:08 -04:00
|
|
|
|
exact h
|
2024-05-22 13:34:53 -04:00
|
|
|
|
|
|
|
|
|
/-- The homomorphism from the Lorentz Group into the monoid of matrices times the opposite of
|
|
|
|
|
the monoid of matrices. -/
|
|
|
|
|
@[simps!]
|
2024-07-01 16:56:15 -04:00
|
|
|
|
def toProd : LorentzGroup d →* (Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ) ×
|
|
|
|
|
(Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ)ᵐᵒᵖ :=
|
2024-05-22 13:34:53 -04:00
|
|
|
|
MonoidHom.comp (Units.embedProduct _) toGL
|
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma toProd_eq_transpose_η : toProd Λ = (Λ.1, MulOpposite.op $ minkowskiMetric.dual Λ.1) := rfl
|
2024-05-22 13:34:53 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma toProd_injective : Function.Injective (@toProd d) := by
|
2024-05-22 13:34:53 -04:00
|
|
|
|
intro A B h
|
|
|
|
|
rw [toProd_eq_transpose_η, toProd_eq_transpose_η] at h
|
|
|
|
|
rw [@Prod.mk.inj_iff] at h
|
|
|
|
|
apply Subtype.eq
|
|
|
|
|
exact h.1
|
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma toProd_continuous : Continuous (@toProd d) := by
|
|
|
|
|
change Continuous (fun A => (A.1, ⟨dual A.1⟩))
|
2024-05-22 13:34:53 -04:00
|
|
|
|
refine continuous_prod_mk.mpr (And.intro ?_ ?_)
|
|
|
|
|
exact continuous_iff_le_induced.mpr fun U a => a
|
|
|
|
|
refine Continuous.comp' ?_ ?_
|
|
|
|
|
exact MulOpposite.continuous_op
|
|
|
|
|
refine Continuous.matrix_mul (Continuous.matrix_mul continuous_const ?_) continuous_const
|
|
|
|
|
refine Continuous.matrix_transpose ?_
|
|
|
|
|
exact continuous_iff_le_induced.mpr fun U a => a
|
|
|
|
|
|
|
|
|
|
/-- The embedding from the Lorentz Group into the monoid of matrices times the opposite of
|
|
|
|
|
the monoid of matrices. -/
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma toProd_embedding : Embedding (@toProd d) where
|
2024-05-22 13:34:53 -04:00
|
|
|
|
inj := toProd_injective
|
|
|
|
|
induced := by
|
|
|
|
|
refine (inducing_iff ⇑toProd).mp ?_
|
|
|
|
|
refine inducing_of_inducing_compose toProd_continuous continuous_fst ?hgf
|
|
|
|
|
exact (inducing_iff (Prod.fst ∘ ⇑toProd)).mpr rfl
|
|
|
|
|
|
|
|
|
|
/-- The embedding from the Lorentz Group into `GL (Fin 4) ℝ`. -/
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma toGL_embedding : Embedding (@toGL d).toFun where
|
2024-05-22 13:34:53 -04:00
|
|
|
|
inj := toGL_injective
|
|
|
|
|
induced := by
|
|
|
|
|
refine ((fun {X} {t t'} => TopologicalSpace.ext_iff.mpr) ?_).symm
|
|
|
|
|
intro s
|
|
|
|
|
rw [TopologicalSpace.ext_iff.mp toProd_embedding.induced s ]
|
|
|
|
|
rw [isOpen_induced_iff, isOpen_induced_iff]
|
2024-06-13 08:10:08 -04:00
|
|
|
|
exact exists_exists_and_eq_and
|
|
|
|
|
|
2024-05-22 13:34:53 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
instance : TopologicalGroup (LorentzGroup d) :=
|
|
|
|
|
Inducing.topologicalGroup toGL toGL_embedding.toInducing
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
section
|
|
|
|
|
open LorentzVector
|
|
|
|
|
/-!
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
# To a norm one Lorentz vector
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
-/
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
/-- The first column of a lorentz matrix as a `NormOneLorentzVector`. -/
|
|
|
|
|
@[simps!]
|
|
|
|
|
def toNormOneLorentzVector (Λ : LorentzGroup d) : NormOneLorentzVector d :=
|
|
|
|
|
⟨Λ.1 *ᵥ timeVec, by rw [NormOneLorentzVector.mem_iff, Λ.2, minkowskiMetric.on_timeVec]⟩
|
2024-05-22 13:34:53 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
/-!
|
|
|
|
|
|
|
|
|
|
# The time like element
|
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
/-- The time like element of a Lorentz matrix. -/
|
|
|
|
|
@[simp]
|
|
|
|
|
def timeComp (Λ : LorentzGroup d) : ℝ := Λ.1 (Sum.inl 0) (Sum.inl 0)
|
|
|
|
|
|
|
|
|
|
lemma timeComp_eq_toNormOneLorentzVector : timeComp Λ = (toNormOneLorentzVector Λ).1.time := by
|
|
|
|
|
simp only [time, toNormOneLorentzVector, timeVec, Fin.isValue, timeComp]
|
|
|
|
|
erw [Pi.basisFun_apply, mulVec_stdBasis]
|
|
|
|
|
|
|
|
|
|
lemma timeComp_mul (Λ Λ' : LorentzGroup d) : timeComp (Λ * Λ') =
|
|
|
|
|
⟪toNormOneLorentzVector (transpose Λ), (toNormOneLorentzVector Λ').1.spaceReflection⟫ₘ := by
|
|
|
|
|
simp only [timeComp, Fin.isValue, lorentzGroupIsGroup_mul_coe, mul_apply, Fintype.sum_sum_type,
|
|
|
|
|
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton, toNormOneLorentzVector,
|
|
|
|
|
transpose, timeVec, right_spaceReflection, time, space, PiLp.inner_apply, Function.comp_apply,
|
|
|
|
|
RCLike.inner_apply, conj_trivial]
|
|
|
|
|
erw [Pi.basisFun_apply, mulVec_stdBasis]
|
|
|
|
|
simp
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
end
|
|
|
|
|
end LorentzGroup
|