refactor: Move Lorentz Group
This commit is contained in:
parent
5cc188146f
commit
573ea890cd
13 changed files with 20 additions and 20 deletions
372
HepLean/Lorentz/Group/Basic.lean
Normal file
372
HepLean/Lorentz/Group/Basic.lean
Normal file
|
@ -0,0 +1,372 @@
|
|||
/-
|
||||
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.MinkowskiMatrix
|
||||
/-!
|
||||
# The Lorentz Group
|
||||
|
||||
We define the Lorentz group.
|
||||
|
||||
## References
|
||||
|
||||
- http://home.ku.edu.tr/~amostafazadeh/phys517_518/phys517_2016f/Handouts/A_Jaffi_Lorentz_Group.pdf
|
||||
|
||||
-/
|
||||
/-! TODO: Show that the Lorentz is a Lie group. -/
|
||||
|
||||
noncomputable section
|
||||
|
||||
open Matrix
|
||||
open Complex
|
||||
open ComplexConjugate
|
||||
|
||||
/-!
|
||||
## Matrices which preserves the Minkowski metric
|
||||
|
||||
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`.
|
||||
|
||||
-/
|
||||
variable {d : ℕ}
|
||||
|
||||
open minkowskiMatrix in
|
||||
/-- The Lorentz group is the subset of matrices for which
|
||||
`Λ * dual Λ = 1`. -/
|
||||
def LorentzGroup (d : ℕ) : Set (Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ) :=
|
||||
{Λ : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ | Λ * dual Λ = 1}
|
||||
|
||||
namespace LorentzGroup
|
||||
/-- Notation for the Lorentz group. -/
|
||||
scoped[LorentzGroup] notation (name := lorentzGroup_notation) "𝓛" => LorentzGroup
|
||||
|
||||
open minkowskiMatrix
|
||||
variable {Λ Λ' : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ}
|
||||
|
||||
/-!
|
||||
|
||||
# Membership conditions
|
||||
|
||||
-/
|
||||
|
||||
lemma mem_iff_self_mul_dual : Λ ∈ LorentzGroup d ↔ Λ * dual Λ = 1 := by
|
||||
rfl
|
||||
|
||||
lemma mem_iff_dual_mul_self : Λ ∈ LorentzGroup d ↔ dual Λ * Λ = 1 := by
|
||||
rw [mem_iff_self_mul_dual]
|
||||
exact mul_eq_one_comm
|
||||
|
||||
lemma mem_iff_transpose : Λ ∈ LorentzGroup d ↔ Λᵀ ∈ LorentzGroup d := by
|
||||
refine Iff.intro (fun h ↦ ?_) (fun 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
|
||||
· 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 Λ * Λ) * Λ'
|
||||
· noncomm_ring
|
||||
· rw [(mem_iff_dual_mul_self).mp hΛ]
|
||||
simp [(mem_iff_dual_mul_self).mp hΛ']
|
||||
|
||||
lemma one_mem : 1 ∈ LorentzGroup d := by
|
||||
rw [mem_iff_dual_mul_self]
|
||||
simp
|
||||
|
||||
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
|
||||
|
||||
|
||||
end LorentzGroup
|
||||
|
||||
/-!
|
||||
|
||||
# The Lorentz group as a group
|
||||
|
||||
-/
|
||||
|
||||
@[simps! mul_coe one_coe inv div]
|
||||
instance lorentzGroupIsGroup : Group (LorentzGroup d) where
|
||||
mul A B := ⟨A.1 * B.1, LorentzGroup.mem_mul A.2 B.2⟩
|
||||
mul_assoc A B C := Subtype.eq (Matrix.mul_assoc A.1 B.1 C.1)
|
||||
one := ⟨1, LorentzGroup.one_mem⟩
|
||||
one_mul A := Subtype.eq (Matrix.one_mul A.1)
|
||||
mul_one A := Subtype.eq (Matrix.mul_one A.1)
|
||||
inv A := ⟨minkowskiMatrix.dual A.1, LorentzGroup.dual_mem A.2⟩
|
||||
inv_mul_cancel A := Subtype.eq (LorentzGroup.mem_iff_dual_mul_self.mp A.2)
|
||||
|
||||
/-- `LorentzGroup` has the subtype topology. -/
|
||||
instance : TopologicalSpace (LorentzGroup d) := instTopologicalSpaceSubtype
|
||||
|
||||
namespace LorentzGroup
|
||||
|
||||
open minkowskiMatrix
|
||||
|
||||
variable {Λ Λ' : LorentzGroup d}
|
||||
|
||||
lemma coe_inv : (Λ⁻¹).1 = Λ.1⁻¹:= (inv_eq_left_inv (mem_iff_dual_mul_self.mp Λ.2)).symm
|
||||
|
||||
instance (M : LorentzGroup d) : Invertible M.1 where
|
||||
invOf := M⁻¹
|
||||
invOf_mul_self := by
|
||||
rw [← coe_inv]
|
||||
exact (mem_iff_dual_mul_self.mp M.2)
|
||||
mul_invOf_self := by
|
||||
rw [← coe_inv]
|
||||
exact (mem_iff_self_mul_dual.mp M.2)
|
||||
|
||||
@[simp]
|
||||
lemma subtype_inv_mul : (Subtype.val Λ)⁻¹ * (Subtype.val Λ) = 1 := by
|
||||
trans Subtype.val (Λ⁻¹ * Λ)
|
||||
· rw [← coe_inv]
|
||||
rfl
|
||||
· rw [inv_mul_cancel Λ]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma subtype_mul_inv : (Subtype.val Λ) * (Subtype.val Λ)⁻¹ = 1 := by
|
||||
trans Subtype.val (Λ * Λ⁻¹)
|
||||
· rw [← coe_inv]
|
||||
rfl
|
||||
· rw [mul_inv_cancel Λ]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma mul_minkowskiMatrix_mul_transpose :
|
||||
(Subtype.val Λ) * minkowskiMatrix * (Subtype.val Λ).transpose = minkowskiMatrix := by
|
||||
have h2 := Λ.prop
|
||||
rw [LorentzGroup.mem_iff_self_mul_dual] at h2
|
||||
simp only [dual] at h2
|
||||
refine (right_inv_eq_left_inv minkowskiMatrix.sq ?_).symm
|
||||
rw [← h2]
|
||||
noncomm_ring
|
||||
|
||||
@[simp]
|
||||
lemma transpose_mul_minkowskiMatrix_mul_self :
|
||||
(Subtype.val Λ).transpose * minkowskiMatrix * (Subtype.val Λ) = minkowskiMatrix := by
|
||||
have h2 := Λ.prop
|
||||
rw [LorentzGroup.mem_iff_dual_mul_self] at h2
|
||||
simp only [dual] at h2
|
||||
refine right_inv_eq_left_inv ?_ minkowskiMatrix.sq
|
||||
rw [← h2]
|
||||
noncomm_ring
|
||||
|
||||
/-- The transpose of a matrix in the Lorentz group is an element of the Lorentz group. -/
|
||||
def transpose (Λ : LorentzGroup d) : LorentzGroup d :=
|
||||
⟨Λ.1ᵀ, mem_iff_transpose.mp Λ.2⟩
|
||||
|
||||
@[simp]
|
||||
lemma transpose_one : @transpose d 1 = 1 := Subtype.eq Matrix.transpose_one
|
||||
|
||||
@[simp]
|
||||
lemma transpose_mul : transpose (Λ * Λ') = transpose Λ' * transpose Λ :=
|
||||
Subtype.eq (Matrix.transpose_mul Λ.1 Λ'.1)
|
||||
|
||||
lemma transpose_val : (transpose Λ).1 = Λ.1ᵀ := rfl
|
||||
|
||||
lemma transpose_inv : (transpose Λ)⁻¹ = transpose Λ⁻¹ := by
|
||||
refine Subtype.eq ?_
|
||||
rw [transpose_val, coe_inv, transpose_val, coe_inv, Matrix.transpose_nonsing_inv]
|
||||
|
||||
lemma comm_minkowskiMatrix : Λ.1 * minkowskiMatrix = minkowskiMatrix * (transpose Λ⁻¹).1 := by
|
||||
conv_rhs => rw [← @mul_minkowskiMatrix_mul_transpose d Λ]
|
||||
rw [← transpose_inv, coe_inv, transpose_val]
|
||||
rw [mul_assoc]
|
||||
have h1 : ((Λ.1)ᵀ * (Λ.1)ᵀ⁻¹) = 1 := by
|
||||
rw [← transpose_val]
|
||||
simp only [subtype_mul_inv]
|
||||
rw [h1]
|
||||
simp
|
||||
|
||||
lemma minkowskiMatrix_comm : minkowskiMatrix * Λ.1 = (transpose Λ⁻¹).1 * minkowskiMatrix := by
|
||||
conv_rhs => rw [← @transpose_mul_minkowskiMatrix_mul_self d Λ]
|
||||
rw [← transpose_inv, coe_inv, transpose_val]
|
||||
rw [← mul_assoc, ← mul_assoc]
|
||||
have h1 : ((Λ.1)ᵀ⁻¹ * (Λ.1)ᵀ) = 1 := by
|
||||
rw [← transpose_val]
|
||||
simp only [subtype_inv_mul]
|
||||
rw [h1]
|
||||
simp
|
||||
|
||||
/-!
|
||||
|
||||
## 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.
|
||||
|
||||
-/
|
||||
|
||||
/-- The homomorphism of the Lorentz group into `GL (Fin 4) ℝ`. -/
|
||||
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⟩
|
||||
map_one' :=
|
||||
(GeneralLinearGroup.ext_iff _ 1).mpr fun _ => congrFun rfl
|
||||
map_mul' _ _ :=
|
||||
(GeneralLinearGroup.ext_iff _ _).mpr fun _ => congrFun rfl
|
||||
|
||||
lemma toGL_injective : Function.Injective (@toGL d) := by
|
||||
refine fun A B h => Subtype.eq ?_
|
||||
rw [@Units.ext_iff] at h
|
||||
exact h
|
||||
|
||||
/-- The homomorphism from the Lorentz Group into the monoid of matrices times the opposite of
|
||||
the monoid of matrices. -/
|
||||
@[simps!]
|
||||
def toProd : LorentzGroup d →* (Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ) ×
|
||||
(Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ)ᵐᵒᵖ :=
|
||||
MonoidHom.comp (Units.embedProduct _) toGL
|
||||
|
||||
lemma toProd_eq_transpose_η : toProd Λ = (Λ.1, MulOpposite.op $ minkowskiMatrix.dual Λ.1) := rfl
|
||||
|
||||
lemma toProd_injective : Function.Injective (@toProd d) := by
|
||||
intro A B h
|
||||
rw [toProd_eq_transpose_η, toProd_eq_transpose_η] at h
|
||||
rw [@Prod.mk.inj_iff] at h
|
||||
exact Subtype.eq h.1
|
||||
|
||||
lemma toProd_continuous : Continuous (@toProd d) := by
|
||||
change Continuous (fun A => (A.1, ⟨dual A.1⟩))
|
||||
refine continuous_prod_mk.mpr ⟨continuous_iff_le_induced.mpr fun U a ↦ a,
|
||||
MulOpposite.continuous_op.comp' ((continuous_const.matrix_mul (continuous_iff_le_induced.mpr
|
||||
fun U a => a).matrix_transpose).matrix_mul continuous_const)⟩
|
||||
|
||||
/-- The embedding from the Lorentz Group into the monoid of matrices times the opposite of
|
||||
the monoid of matrices. -/
|
||||
lemma toProd_embedding : IsEmbedding (@toProd d) where
|
||||
inj := toProd_injective
|
||||
eq_induced :=
|
||||
(isInducing_iff ⇑toProd).mp (IsInducing.of_comp toProd_continuous continuous_fst
|
||||
((isInducing_iff (Prod.fst ∘ ⇑toProd)).mpr rfl))
|
||||
|
||||
/-- The embedding from the Lorentz Group into `GL (Fin 4) ℝ`. -/
|
||||
lemma toGL_embedding : IsEmbedding (@toGL d).toFun where
|
||||
inj := toGL_injective
|
||||
eq_induced := by
|
||||
refine ((fun {X} {t t'} => TopologicalSpace.ext_iff.mpr) fun _ ↦ ?_).symm
|
||||
rw [TopologicalSpace.ext_iff.mp toProd_embedding.eq_induced _, isOpen_induced_iff,
|
||||
isOpen_induced_iff]
|
||||
exact exists_exists_and_eq_and
|
||||
|
||||
instance : TopologicalGroup (LorentzGroup d) :=
|
||||
IsInducing.topologicalGroup toGL toGL_embedding.toIsInducing
|
||||
|
||||
/-!
|
||||
|
||||
## To Complex matrices
|
||||
|
||||
-/
|
||||
|
||||
/-- The monoid homomorphisms taking the lorentz group to complex matrices. -/
|
||||
def toComplex : LorentzGroup d →* Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℂ where
|
||||
toFun Λ := Λ.1.map ofRealHom
|
||||
map_one' := by
|
||||
ext i j
|
||||
simp only [lorentzGroupIsGroup_one_coe, map_apply, ofRealHom_eq_coe]
|
||||
simp only [Matrix.one_apply, ofReal_one, ofReal_zero]
|
||||
split_ifs
|
||||
· rfl
|
||||
· rfl
|
||||
map_mul' Λ Λ' := by
|
||||
ext i j
|
||||
simp only [lorentzGroupIsGroup_mul_coe, map_apply, ofRealHom_eq_coe]
|
||||
simp only [← Matrix.map_mul, RingHom.map_matrix_mul]
|
||||
rfl
|
||||
|
||||
instance (M : LorentzGroup d) : Invertible (toComplex M) where
|
||||
invOf := toComplex M⁻¹
|
||||
invOf_mul_self := by
|
||||
rw [← toComplex.map_mul, Group.inv_mul_cancel]
|
||||
simp
|
||||
mul_invOf_self := by
|
||||
rw [← toComplex.map_mul]
|
||||
rw [@mul_inv_cancel]
|
||||
simp
|
||||
|
||||
lemma toComplex_inv (Λ : LorentzGroup d) : (toComplex Λ)⁻¹ = toComplex Λ⁻¹ := by
|
||||
refine inv_eq_right_inv ?h
|
||||
rw [← toComplex.map_mul, mul_inv_cancel]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma toComplex_mul_minkowskiMatrix_mul_transpose (Λ : LorentzGroup d) :
|
||||
toComplex Λ * minkowskiMatrix.map ofRealHom * (toComplex Λ)ᵀ =
|
||||
minkowskiMatrix.map ofRealHom := by
|
||||
simp only [toComplex, MonoidHom.coe_mk, OneHom.coe_mk]
|
||||
have h1 : ((Λ.1).map ⇑ofRealHom)ᵀ = (Λ.1ᵀ).map ofRealHom := rfl
|
||||
rw [h1]
|
||||
trans (Λ.1 * minkowskiMatrix * Λ.1ᵀ).map ofRealHom
|
||||
· simp only [Matrix.map_mul]
|
||||
simp only [mul_minkowskiMatrix_mul_transpose]
|
||||
|
||||
@[simp]
|
||||
lemma toComplex_transpose_mul_minkowskiMatrix_mul_self (Λ : LorentzGroup d) :
|
||||
(toComplex Λ)ᵀ * minkowskiMatrix.map ofRealHom * toComplex Λ =
|
||||
minkowskiMatrix.map ofRealHom := by
|
||||
simp only [toComplex, MonoidHom.coe_mk, OneHom.coe_mk]
|
||||
have h1 : ((Λ.1).map ofRealHom)ᵀ = (Λ.1ᵀ).map ofRealHom := rfl
|
||||
rw [h1]
|
||||
trans (Λ.1ᵀ * minkowskiMatrix * Λ.1).map ofRealHom
|
||||
· simp only [Matrix.map_mul]
|
||||
simp only [transpose_mul_minkowskiMatrix_mul_self]
|
||||
|
||||
lemma toComplex_mulVec_ofReal (v : Fin 1 ⊕ Fin d → ℝ) (Λ : LorentzGroup d) :
|
||||
toComplex Λ *ᵥ (ofRealHom ∘ v) = ofRealHom ∘ (Λ *ᵥ v) := by
|
||||
simp only [toComplex, MonoidHom.coe_mk, OneHom.coe_mk]
|
||||
funext i
|
||||
rw [← RingHom.map_mulVec]
|
||||
rfl
|
||||
/-!
|
||||
/-!
|
||||
|
||||
# To a norm one Lorentz vector
|
||||
|
||||
-/
|
||||
|
||||
/-- 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]⟩
|
||||
|
||||
/-!
|
||||
|
||||
# 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, Matrix.mulVec_single_one]
|
||||
rfl
|
||||
|
||||
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, Matrix.mulVec_single_one]
|
||||
simp
|
||||
-/
|
||||
|
||||
/-- The parity transformation. -/
|
||||
def parity : LorentzGroup d := ⟨minkowskiMatrix, by
|
||||
rw [mem_iff_dual_mul_self]
|
||||
simp only [dual_eta, minkowskiMatrix.sq]⟩
|
||||
|
||||
end LorentzGroup
|
236
HepLean/Lorentz/Group/Boosts.lean
Normal file
236
HepLean/Lorentz/Group/Boosts.lean
Normal file
|
@ -0,0 +1,236 @@
|
|||
/-
|
||||
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.Lorentz.Group.Proper
|
||||
import Mathlib.Topology.Constructions
|
||||
import HepLean.SpaceTime.LorentzVector.Real.NormOne
|
||||
/-!
|
||||
# Boosts
|
||||
|
||||
This file defines those Lorentz which are boosts.
|
||||
|
||||
We first define generalised boosts, which are restricted lorentz transformations taking
|
||||
a four velocity `u` to a four velocity `v`.
|
||||
|
||||
A boost is the special case of a generalised boost when `u = stdBasis 0`.
|
||||
|
||||
## References
|
||||
|
||||
- The main argument follows: Guillem Cobos, The Lorentz Group, 2015:
|
||||
https://diposit.ub.edu/dspace/bitstream/2445/68763/2/memoria.pdf
|
||||
|
||||
-/
|
||||
/-! TODO: Show that generalised boosts are in the restricted Lorentz group. -/
|
||||
noncomputable section
|
||||
|
||||
namespace LorentzGroup
|
||||
|
||||
open Lorentz.Contr.NormOne
|
||||
open Lorentz
|
||||
open TensorProduct
|
||||
|
||||
variable {d : ℕ}
|
||||
|
||||
/-- An auxillary linear map used in the definition of a generalised boost. -/
|
||||
def genBoostAux₁ (u v : FuturePointing d) : ContrMod d →ₗ[ℝ] ContrMod d where
|
||||
toFun x := (2 * ⟪x, u.val.val⟫ₘ) • v.1.1
|
||||
map_add' x y := by
|
||||
simp [map_add, LinearMap.add_apply, tmul_add, add_tmul, mul_add, add_smul]
|
||||
map_smul' c x := by
|
||||
simp only [ Action.instMonoidalCategory_tensorObj_V,
|
||||
Action.instMonoidalCategory_tensorUnit_V, CategoryTheory.Equivalence.symm_inverse,
|
||||
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
|
||||
smul_tmul, tmul_smul, map_smul, smul_eq_mul, RingHom.id_apply]
|
||||
rw [← mul_assoc, mul_comm 2 c, mul_assoc, mul_smul]
|
||||
|
||||
/-- An auxillary linear map used in the definition of a genearlised boost. -/
|
||||
def genBoostAux₂ (u v : FuturePointing d) : ContrMod d →ₗ[ℝ] ContrMod d where
|
||||
toFun x := - (⟪x, u.1.1 + v.1.1⟫ₘ / (1 + ⟪u.1.1, v.1.1⟫ₘ)) • (u.1.1 + v.1.1)
|
||||
map_add' x y := by
|
||||
simp only
|
||||
rw [← add_smul]
|
||||
apply congrFun (congrArg _ _)
|
||||
field_simp [add_tmul]
|
||||
apply congrFun (congrArg _ _)
|
||||
ring
|
||||
map_smul' c x := by
|
||||
simp only [smul_tmul, tmul_smul]
|
||||
rw [map_smul]
|
||||
conv =>
|
||||
lhs; lhs; rhs; lhs
|
||||
change (c * (contrContrContractField (x ⊗ₜ[ℝ] (u.val.val + v.val.val))))
|
||||
rw [mul_div_assoc, neg_mul_eq_mul_neg, smul_smul]
|
||||
rfl
|
||||
|
||||
lemma genBoostAux₂_self (u : FuturePointing d) : genBoostAux₂ u u = - genBoostAux₁ u u := by
|
||||
ext1 x
|
||||
simp only [genBoostAux₂, LinearMap.coe_mk, AddHom.coe_mk, genBoostAux₁, LinearMap.neg_apply]
|
||||
rw [neg_smul]
|
||||
apply congrArg
|
||||
conv => lhs; rhs; rw [← (two_smul ℝ u.val.val)]
|
||||
rw [smul_smul]
|
||||
congr 1
|
||||
rw [u.1.2]
|
||||
conv => lhs; lhs; rhs; rhs; change 1
|
||||
rw [show 1 + (1 : ℝ) = (2 : ℝ ) by ring]
|
||||
simp only [isUnit_iff_ne_zero, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true,
|
||||
IsUnit.div_mul_cancel]
|
||||
rw [← (two_smul ℝ u.val.val)]
|
||||
simp only [tmul_smul, map_smul, smul_eq_mul]
|
||||
|
||||
/-- An generalised boost. This is a Lorentz transformation which takes the four velocity `u`
|
||||
to `v`. -/
|
||||
def genBoost (u v : FuturePointing d) : ContrMod d →ₗ[ℝ] ContrMod d :=
|
||||
LinearMap.id + genBoostAux₁ u v + genBoostAux₂ u v
|
||||
|
||||
lemma genBoost_mul_one_plus_contr (u v : FuturePointing d) (x : Contr d) :
|
||||
(1 + ⟪u.1.1, v.1.1⟫ₘ) • genBoost u v x =
|
||||
(1 + ⟪u.1.1, v.1.1⟫ₘ) • x +
|
||||
(2 * ⟪x, u.val.val⟫ₘ * (1 + ⟪u.1.1, v.1.1⟫ₘ)) • v.1.1
|
||||
- (⟪x, u.1.1 + v.1.1⟫ₘ) • (u.1.1 + v.1.1) := by
|
||||
simp only [genBoost, LinearMap.add_apply, LinearMap.id_apply, id_eq]
|
||||
rw [smul_add, smul_add]
|
||||
trans (1 + ⟪u.1.1, v.1.1⟫ₘ) • x +
|
||||
(2 * ⟪x, u.val.val⟫ₘ * (1 + ⟪u.1.1, v.1.1⟫ₘ)) • v.1.1
|
||||
+ (- ⟪x, u.1.1 + v.1.1⟫ₘ) • (u.1.1 + v.1.1)
|
||||
· congr 1
|
||||
· congr 1
|
||||
rw [genBoostAux₁]
|
||||
simp
|
||||
rw [smul_smul]
|
||||
congr 1
|
||||
ring
|
||||
· rw [genBoostAux₂]
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
CategoryTheory.Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
|
||||
Action.FunctorCategoryEquivalence.functor_obj_obj, neg_smul, LinearMap.coe_mk,
|
||||
AddHom.coe_mk, smul_neg]
|
||||
rw [smul_smul]
|
||||
congr
|
||||
have h1 := FuturePointing.one_add_metric_non_zero u v
|
||||
field_simp
|
||||
· rw [neg_smul]
|
||||
rfl
|
||||
|
||||
namespace genBoost
|
||||
|
||||
/--
|
||||
This lemma states that for a given four-velocity `u`, the general boost
|
||||
transformation `genBoost u u` is equal to the identity linear map `LinearMap.id`.
|
||||
-/
|
||||
lemma self (u : FuturePointing d) : genBoost u u = LinearMap.id := by
|
||||
ext x
|
||||
simp only [genBoost, LinearMap.add_apply, LinearMap.id_coe, id_eq]
|
||||
rw [genBoostAux₂_self]
|
||||
simp only [LinearMap.neg_apply, add_neg_cancel_right]
|
||||
|
||||
/-- A generalised boost as a matrix. -/
|
||||
def toMatrix (u v : FuturePointing d) : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ :=
|
||||
LinearMap.toMatrix ContrMod.stdBasis ContrMod.stdBasis (genBoost u v)
|
||||
|
||||
lemma toMatrix_mulVec (u v : FuturePointing d) (x : Contr d) :
|
||||
(toMatrix u v) *ᵥ x = genBoost u v x :=
|
||||
ContrMod.ext (LinearMap.toMatrix_mulVec_repr ContrMod.stdBasis ContrMod.stdBasis (genBoost u v) x)
|
||||
|
||||
open minkowskiMatrix in
|
||||
@[simp]
|
||||
lemma toMatrix_apply (u v : FuturePointing d) (μ ν : Fin 1 ⊕ Fin d) :
|
||||
(toMatrix u v) μ ν = η μ μ * (toField d ⟪ContrMod.stdBasis μ, ContrMod.stdBasis ν⟫ₘ + 2 *
|
||||
toField d ⟪ContrMod.stdBasis ν, u.val.val⟫ₘ * toField d ⟪ContrMod.stdBasis μ, v.val.val⟫ₘ
|
||||
- toField d ⟪ContrMod.stdBasis μ, u.val.val + v.val.val⟫ₘ *
|
||||
toField d ⟪ContrMod.stdBasis ν, u.val.val + v.val.val⟫ₘ /
|
||||
(1 + toField d ⟪u.val.val, v.val.val⟫ₘ)) := by
|
||||
rw [contrContrContractField.matrix_apply_stdBasis (Λ := toMatrix u v) μ ν, toMatrix_mulVec]
|
||||
simp only [genBoost, genBoostAux₁, genBoostAux₂, map_add, smul_add, neg_smul, LinearMap.add_apply,
|
||||
LinearMap.id_apply, LinearMap.coe_mk, AddHom.coe_mk, contrContrContractField.basis_left, map_smul, smul_eq_mul, map_neg,
|
||||
mul_eq_mul_left_iff, toField_apply]
|
||||
ring_nf
|
||||
simp only [Pi.add_apply, Action.instMonoidalCategory_tensorObj_V,
|
||||
Action.instMonoidalCategory_tensorUnit_V, CategoryTheory.Equivalence.symm_inverse,
|
||||
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
|
||||
Pi.smul_apply, smul_eq_mul, Pi.sub_apply, Pi.neg_apply]
|
||||
apply Or.inl
|
||||
ring
|
||||
|
||||
open minkowskiMatrix in
|
||||
lemma toMatrix_continuous (u : FuturePointing d) : Continuous (toMatrix u) := by
|
||||
refine continuous_matrix ?_
|
||||
intro i j
|
||||
simp only [toMatrix_apply]
|
||||
refine (continuous_mul_left (η i i)).comp' ?_
|
||||
refine Continuous.sub ?_ ?_
|
||||
· refine .comp' (continuous_add_left _) ?_
|
||||
refine .comp' (continuous_mul_left _) ?_
|
||||
exact FuturePointing.metric_continuous _
|
||||
refine .mul ?_ ?_
|
||||
· refine .mul ?_ ?_
|
||||
· simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
CategoryTheory.Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
|
||||
Action.FunctorCategoryEquivalence.functor_obj_obj, tmul_add, map_add]
|
||||
refine .comp' ?_ ?_
|
||||
· exact continuous_add_left _
|
||||
· exact FuturePointing.metric_continuous _
|
||||
· simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
CategoryTheory.Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
|
||||
Action.FunctorCategoryEquivalence.functor_obj_obj, tmul_add, map_add]
|
||||
refine .comp' ?_ ?_
|
||||
· exact continuous_add_left _
|
||||
· exact FuturePointing.metric_continuous _
|
||||
· refine .inv₀ ?_ ?_
|
||||
· refine .comp' (continuous_add_left 1) (FuturePointing.metric_continuous _)
|
||||
exact fun x => FuturePointing.one_add_metric_non_zero u x
|
||||
|
||||
lemma toMatrix_in_lorentzGroup (u v : FuturePointing d) : (toMatrix u v) ∈ LorentzGroup d := by
|
||||
rw [LorentzGroup.mem_iff_invariant]
|
||||
intro x y
|
||||
rw [toMatrix_mulVec, toMatrix_mulVec]
|
||||
have h1 : (((1 + ⟪u.1.1, v.1.1⟫ₘ)) * (1 + ⟪u.1.1, v.1.1⟫ₘ)) •
|
||||
(contrContrContractField ((genBoost u v) y ⊗ₜ[ℝ] (genBoost u v) x))
|
||||
= (((1 + ⟪u.1.1, v.1.1⟫ₘ)) * (1 + ⟪u.1.1, v.1.1⟫ₘ)) • ⟪y, x⟫ₘ := by
|
||||
conv_lhs =>
|
||||
erw [← map_smul]
|
||||
rw [← smul_smul]
|
||||
rw [← tmul_smul, ← smul_tmul, ← tmul_smul, genBoost_mul_one_plus_contr,
|
||||
genBoost_mul_one_plus_contr]
|
||||
simp only [add_smul, one_smul, tmul_add, map_add, smul_add, tmul_sub, sub_tmul, add_tmul,
|
||||
smul_tmul, tmul_smul, map_sub, map_smul, smul_eq_mul, contr_self, mul_one]
|
||||
rw [contrContrContractField.symm v.1.1 u, contrContrContractField.symm v.1.1 x,
|
||||
contrContrContractField.symm u.1.1 x]
|
||||
simp
|
||||
ring
|
||||
have hn (a : ℝ ) {b c : ℝ} (h : a ≠ 0) (hab : a * b = a * c ) : b = c := by
|
||||
simp_all only [smul_eq_mul, ne_eq, mul_eq_mul_left_iff, or_false]
|
||||
refine hn _ ?_ h1
|
||||
simpa using (FuturePointing.one_add_metric_non_zero u v)
|
||||
|
||||
/-- A generalised boost as an element of the Lorentz Group. -/
|
||||
def toLorentz (u v : FuturePointing d) : LorentzGroup d :=
|
||||
⟨toMatrix u v, toMatrix_in_lorentzGroup u v⟩
|
||||
|
||||
lemma toLorentz_continuous (u : FuturePointing d) : Continuous (toLorentz u) := by
|
||||
refine Continuous.subtype_mk ?_ (fun x => toMatrix_in_lorentzGroup u x)
|
||||
exact toMatrix_continuous u
|
||||
|
||||
lemma toLorentz_joined_to_1 (u v : FuturePointing d) : Joined 1 (toLorentz u v) := by
|
||||
obtain ⟨f, _⟩ := FuturePointing.isPathConnected.joinedIn u trivial v trivial
|
||||
use ContinuousMap.comp ⟨toLorentz u, toLorentz_continuous u⟩ f
|
||||
· simp only [ContinuousMap.toFun_eq_coe, ContinuousMap.comp_apply, ContinuousMap.coe_coe,
|
||||
Path.source, ContinuousMap.coe_mk]
|
||||
rw [@Subtype.ext_iff, toLorentz]
|
||||
simp [toMatrix, self u]
|
||||
· simp
|
||||
|
||||
lemma toLorentz_in_connected_component_1 (u v : FuturePointing d) :
|
||||
toLorentz u v ∈ connectedComponent 1 :=
|
||||
pathComponent_subset_component _ (toLorentz_joined_to_1 u v)
|
||||
|
||||
lemma isProper (u v : FuturePointing d) : IsProper (toLorentz u v) :=
|
||||
(isProper_on_connected_component (toLorentz_in_connected_component_1 u v)).mp id_IsProper
|
||||
|
||||
end genBoost
|
||||
|
||||
end LorentzGroup
|
||||
|
||||
end
|
170
HepLean/Lorentz/Group/Orthochronous.lean
Normal file
170
HepLean/Lorentz/Group/Orthochronous.lean
Normal file
|
@ -0,0 +1,170 @@
|
|||
/-
|
||||
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.NormOne
|
||||
import HepLean.Lorentz.Group.Proper
|
||||
/-!
|
||||
# The Orthochronous Lorentz Group
|
||||
|
||||
We define the give a series of lemmas related to the orthochronous property of lorentz
|
||||
matrices.
|
||||
|
||||
-/
|
||||
/-! TODO: Prove topological properties of the Orthochronous Lorentz Group. -/
|
||||
|
||||
noncomputable section
|
||||
|
||||
open Matrix
|
||||
open Complex
|
||||
open ComplexConjugate
|
||||
|
||||
namespace LorentzGroup
|
||||
|
||||
variable {d : ℕ}
|
||||
variable (Λ : LorentzGroup d)
|
||||
open Lorentz.Contr
|
||||
|
||||
/-- A Lorentz transformation is `orthochronous` if its `0 0` element is non-negative. -/
|
||||
def IsOrthochronous : Prop := 0 ≤ Λ.1 (Sum.inl 0) (Sum.inl 0)
|
||||
|
||||
lemma IsOrthochronous_iff_futurePointing :
|
||||
IsOrthochronous Λ ↔ toNormOne Λ ∈ NormOne.FuturePointing d := by
|
||||
simp only [IsOrthochronous]
|
||||
rw [NormOne.FuturePointing.mem_iff_inl_nonneg, toNormOne_inl]
|
||||
|
||||
lemma IsOrthochronous_iff_transpose :
|
||||
IsOrthochronous Λ ↔ IsOrthochronous (transpose Λ) := by rfl
|
||||
|
||||
lemma IsOrthochronous_iff_ge_one :
|
||||
IsOrthochronous Λ ↔ 1 ≤ Λ.1 (Sum.inl 0) (Sum.inl 0) := by
|
||||
rw [IsOrthochronous_iff_futurePointing, NormOne.FuturePointing.mem_iff_inl_one_le_inl,
|
||||
toNormOne_inl]
|
||||
|
||||
lemma not_orthochronous_iff_le_neg_one :
|
||||
¬ IsOrthochronous Λ ↔ Λ.1 (Sum.inl 0) (Sum.inl 0) ≤ -1 := by
|
||||
rw [IsOrthochronous_iff_futurePointing, NormOne.FuturePointing.not_mem_iff_inl_le_neg_one,
|
||||
toNormOne_inl]
|
||||
|
||||
lemma not_orthochronous_iff_le_zero : ¬ IsOrthochronous Λ ↔ Λ.1 (Sum.inl 0) (Sum.inl 0) ≤ 0 := by
|
||||
rw [IsOrthochronous_iff_futurePointing, NormOne.FuturePointing.not_mem_iff_inl_le_zero,
|
||||
toNormOne_inl]
|
||||
|
||||
/-- The continuous map taking a Lorentz transformation to its `0 0` element. -/
|
||||
def timeCompCont : C(LorentzGroup d, ℝ) := ⟨fun Λ => Λ.1 (Sum.inl 0) (Sum.inl 0),
|
||||
Continuous.matrix_elem (continuous_iff_le_induced.mpr fun _ a => a) (Sum.inl 0) (Sum.inl 0)⟩
|
||||
|
||||
/-- An auxillary function used in the definition of `orthchroMapReal`. -/
|
||||
def stepFunction : ℝ → ℝ := fun t =>
|
||||
if t ≤ -1 then -1 else
|
||||
if 1 ≤ t then 1 else t
|
||||
|
||||
lemma stepFunction_continuous : Continuous stepFunction := by
|
||||
apply Continuous.if ?_ continuous_const (Continuous.if ?_ continuous_const continuous_id)
|
||||
<;> intro a ha
|
||||
· rw [@Set.Iic_def, @frontier_Iic, @Set.mem_singleton_iff] at ha
|
||||
rw [ha]
|
||||
simp only [le_neg_self_iff, id_eq]
|
||||
have h1 : ¬ (1 : ℝ) ≤ 0 := by simp
|
||||
exact Eq.symm (if_neg h1)
|
||||
· rw [Set.Ici_def, @frontier_Ici, @Set.mem_singleton_iff] at ha
|
||||
exact id (Eq.symm ha)
|
||||
|
||||
/-- The continuous map from `lorentzGroup` to `ℝ` wh
|
||||
taking Orthochronous elements to `1` and non-orthochronous to `-1`. -/
|
||||
def orthchroMapReal : C(LorentzGroup d, ℝ) := ContinuousMap.comp
|
||||
⟨stepFunction, stepFunction_continuous⟩ timeCompCont
|
||||
|
||||
lemma orthchroMapReal_on_IsOrthochronous {Λ : LorentzGroup d} (h : IsOrthochronous Λ) :
|
||||
orthchroMapReal Λ = 1 := by
|
||||
rw [IsOrthochronous_iff_ge_one] at h
|
||||
change stepFunction (Λ.1 _ _) = 1
|
||||
rw [stepFunction, if_pos h, if_neg]
|
||||
linarith
|
||||
|
||||
lemma orthchroMapReal_on_not_IsOrthochronous {Λ : LorentzGroup d} (h : ¬ IsOrthochronous Λ) :
|
||||
orthchroMapReal Λ = - 1 := by
|
||||
rw [not_orthochronous_iff_le_neg_one] at h
|
||||
change stepFunction (_)= - 1
|
||||
rw [stepFunction, if_pos]
|
||||
exact h
|
||||
|
||||
lemma orthchroMapReal_minus_one_or_one (Λ : LorentzGroup d) :
|
||||
orthchroMapReal Λ = -1 ∨ orthchroMapReal Λ = 1 := by
|
||||
by_cases h : IsOrthochronous Λ
|
||||
· exact Or.inr $ orthchroMapReal_on_IsOrthochronous h
|
||||
· exact Or.inl $ orthchroMapReal_on_not_IsOrthochronous h
|
||||
|
||||
local notation "ℤ₂" => Multiplicative (ZMod 2)
|
||||
|
||||
/-- A continuous map from `lorentzGroup` to `ℤ₂` whose kernal are the Orthochronous elements. -/
|
||||
def orthchroMap : C(LorentzGroup d, ℤ₂) :=
|
||||
ContinuousMap.comp coeForℤ₂ {
|
||||
toFun := fun Λ => ⟨orthchroMapReal Λ, orthchroMapReal_minus_one_or_one Λ⟩,
|
||||
continuous_toFun := Continuous.subtype_mk (ContinuousMap.continuous orthchroMapReal) _}
|
||||
|
||||
lemma orthchroMap_IsOrthochronous {Λ : LorentzGroup d} (h : IsOrthochronous Λ) :
|
||||
orthchroMap Λ = 1 := by
|
||||
simp [orthchroMap, orthchroMapReal_on_IsOrthochronous h]
|
||||
|
||||
lemma orthchroMap_not_IsOrthochronous {Λ : LorentzGroup d} (h : ¬ IsOrthochronous Λ) :
|
||||
orthchroMap Λ = Additive.toMul (1 : ZMod 2) := by
|
||||
simp only [orthchroMap, ContinuousMap.comp_apply, ContinuousMap.coe_mk,
|
||||
orthchroMapReal_on_not_IsOrthochronous h, coeForℤ₂_apply, Subtype.mk.injEq, Nat.reduceAdd]
|
||||
rw [if_neg]
|
||||
· rfl
|
||||
· linarith
|
||||
|
||||
lemma mul_othchron_of_othchron_othchron {Λ Λ' : LorentzGroup d} (h : IsOrthochronous Λ)
|
||||
(h' : IsOrthochronous Λ') : IsOrthochronous (Λ * Λ') := by
|
||||
rw [IsOrthochronous_iff_transpose] at h
|
||||
rw [IsOrthochronous_iff_futurePointing] at h h'
|
||||
rw [IsOrthochronous, LorentzGroup.inl_inl_mul]
|
||||
exact NormOne.FuturePointing.metric_reflect_mem_mem h h'
|
||||
|
||||
lemma mul_othchron_of_not_othchron_not_othchron {Λ Λ' : LorentzGroup d} (h : ¬ IsOrthochronous Λ)
|
||||
(h' : ¬ IsOrthochronous Λ') : IsOrthochronous (Λ * Λ') := by
|
||||
rw [IsOrthochronous_iff_transpose] at h
|
||||
rw [IsOrthochronous_iff_futurePointing] at h h'
|
||||
rw [IsOrthochronous, LorentzGroup.inl_inl_mul]
|
||||
exact NormOne.FuturePointing.metric_reflect_not_mem_not_mem h h'
|
||||
|
||||
lemma mul_not_othchron_of_othchron_not_othchron {Λ Λ' : LorentzGroup d} (h : IsOrthochronous Λ)
|
||||
(h' : ¬ IsOrthochronous Λ') : ¬ IsOrthochronous (Λ * Λ') := by
|
||||
rw [not_orthochronous_iff_le_zero, LorentzGroup.inl_inl_mul]
|
||||
rw [IsOrthochronous_iff_transpose] at h
|
||||
rw [IsOrthochronous_iff_futurePointing] at h h'
|
||||
exact NormOne.FuturePointing.metric_reflect_mem_not_mem h h'
|
||||
|
||||
lemma mul_not_othchron_of_not_othchron_othchron {Λ Λ' : LorentzGroup d} (h : ¬ IsOrthochronous Λ)
|
||||
(h' : IsOrthochronous Λ') : ¬ IsOrthochronous (Λ * Λ') := by
|
||||
rw [not_orthochronous_iff_le_zero, LorentzGroup.inl_inl_mul]
|
||||
rw [IsOrthochronous_iff_transpose] at h
|
||||
rw [IsOrthochronous_iff_futurePointing] at h h'
|
||||
exact NormOne.FuturePointing.metric_reflect_not_mem_mem h h'
|
||||
|
||||
/-- The homomorphism from `LorentzGroup` to `ℤ₂` whose kernel are the Orthochronous elements. -/
|
||||
def orthchroRep : LorentzGroup d →* ℤ₂ where
|
||||
toFun := orthchroMap
|
||||
map_one' := orthchroMap_IsOrthochronous (by simp [IsOrthochronous])
|
||||
map_mul' Λ Λ' := by
|
||||
simp only
|
||||
by_cases h : IsOrthochronous Λ
|
||||
<;> by_cases h' : IsOrthochronous Λ'
|
||||
· rw [orthchroMap_IsOrthochronous h, orthchroMap_IsOrthochronous h',
|
||||
orthchroMap_IsOrthochronous (mul_othchron_of_othchron_othchron h h')]
|
||||
rfl
|
||||
· rw [orthchroMap_IsOrthochronous h, orthchroMap_not_IsOrthochronous h',
|
||||
orthchroMap_not_IsOrthochronous (mul_not_othchron_of_othchron_not_othchron h h')]
|
||||
rfl
|
||||
· rw [orthchroMap_not_IsOrthochronous h, orthchroMap_IsOrthochronous h',
|
||||
orthchroMap_not_IsOrthochronous (mul_not_othchron_of_not_othchron_othchron h h')]
|
||||
rfl
|
||||
· rw [orthchroMap_not_IsOrthochronous h, orthchroMap_not_IsOrthochronous h',
|
||||
orthchroMap_IsOrthochronous (mul_othchron_of_not_othchron_not_othchron h h')]
|
||||
rfl
|
||||
|
||||
end LorentzGroup
|
||||
|
||||
end
|
184
HepLean/Lorentz/Group/Proper.lean
Normal file
184
HepLean/Lorentz/Group/Proper.lean
Normal file
|
@ -0,0 +1,184 @@
|
|||
/-
|
||||
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.Lorentz.Group.Basic
|
||||
/-!
|
||||
# The Proper Lorentz Group
|
||||
|
||||
The proper Lorentz group is the subgroup of the Lorentz group with determinant `1`.
|
||||
|
||||
We define the give a series of lemmas related to the determinant of the Lorentz group.
|
||||
|
||||
-/
|
||||
|
||||
noncomputable section
|
||||
|
||||
open Matrix
|
||||
open Complex
|
||||
open ComplexConjugate
|
||||
|
||||
namespace LorentzGroup
|
||||
|
||||
open minkowskiMatrix
|
||||
|
||||
variable {d : ℕ}
|
||||
|
||||
/-- The determinant of a member of the Lorentz group is `1` or `-1`. -/
|
||||
lemma det_eq_one_or_neg_one (Λ : 𝓛 d) : Λ.1.det = 1 ∨ Λ.1.det = -1 := by
|
||||
refine mul_self_eq_one_iff.mp ?_
|
||||
simpa only [det_mul, det_dual, det_one] using congrArg det ((mem_iff_self_mul_dual).mp Λ.2)
|
||||
|
||||
local notation "ℤ₂" => Multiplicative (ZMod 2)
|
||||
|
||||
instance : TopologicalSpace ℤ₂ := instTopologicalSpaceFin
|
||||
|
||||
instance : DiscreteTopology ℤ₂ := by
|
||||
exact forall_open_iff_discrete.mp fun _ => trivial
|
||||
|
||||
instance : TopologicalGroup ℤ₂ := TopologicalGroup.mk
|
||||
|
||||
/-- A continuous function from `({-1, 1} : Set ℝ)` to `ℤ₂`. -/
|
||||
@[simps!]
|
||||
def coeForℤ₂ : C(({-1, 1} : Set ℝ), ℤ₂) where
|
||||
toFun x := if x = ⟨1, Set.mem_insert_of_mem (-1) rfl⟩
|
||||
then Additive.toMul 0 else Additive.toMul (1 : ZMod 2)
|
||||
continuous_toFun := continuous_of_discreteTopology
|
||||
|
||||
/-- The continuous map taking a Lorentz matrix to its determinant. -/
|
||||
def detContinuous : C(𝓛 d, ℤ₂) :=
|
||||
ContinuousMap.comp coeForℤ₂ {
|
||||
toFun := fun Λ => ⟨Λ.1.det, Or.symm (LorentzGroup.det_eq_one_or_neg_one _)⟩,
|
||||
continuous_toFun := by
|
||||
refine Continuous.subtype_mk ?_ _
|
||||
exact Continuous.matrix_det $
|
||||
Continuous.comp' (continuous_iff_le_induced.mpr fun U a => a) continuous_id'
|
||||
}
|
||||
|
||||
lemma detContinuous_eq_one (Λ : LorentzGroup d) :
|
||||
detContinuous Λ = Additive.toMul 0 ↔ Λ.1.det = 1 := by
|
||||
simp only [detContinuous, ContinuousMap.comp_apply, ContinuousMap.coe_mk, coeForℤ₂_apply,
|
||||
Subtype.mk.injEq, ite_eq_left_iff, toMul_eq_one]
|
||||
simp only [toMul_zero, ite_eq_left_iff, toMul_eq_one]
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· by_contra hn
|
||||
have h' := h hn
|
||||
change (1 : Fin 2) = (0 : Fin 2) at h'
|
||||
simp only [Fin.isValue, one_ne_zero] at h'
|
||||
· intro h'
|
||||
exact False.elim (h' h)
|
||||
|
||||
lemma detContinuous_eq_zero (Λ : LorentzGroup d) :
|
||||
detContinuous Λ = Additive.toMul (1 : ZMod 2) ↔ Λ.1.det = - 1 := by
|
||||
simp only [detContinuous, ContinuousMap.comp_apply, ContinuousMap.coe_mk, coeForℤ₂_apply,
|
||||
Subtype.mk.injEq, Nat.reduceAdd]
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· by_contra hn
|
||||
rw [if_pos] at h
|
||||
· change (0 : Fin 2) = (1 : Fin 2) at h
|
||||
simp only [Fin.isValue, zero_ne_one] at h
|
||||
· cases' det_eq_one_or_neg_one Λ with h2 h2
|
||||
· simp_all only [ite_true]
|
||||
· simp_all only [not_true_eq_false]
|
||||
· rw [if_neg]
|
||||
· rfl
|
||||
· cases' det_eq_one_or_neg_one Λ with h2 h2
|
||||
· rw [h]
|
||||
linarith
|
||||
· linarith
|
||||
|
||||
lemma detContinuous_eq_iff_det_eq (Λ Λ' : LorentzGroup d) :
|
||||
detContinuous Λ = detContinuous Λ' ↔ Λ.1.det = Λ'.1.det := by
|
||||
cases' det_eq_one_or_neg_one Λ with h1 h1
|
||||
· rw [h1, (detContinuous_eq_one Λ).mpr h1]
|
||||
cases' det_eq_one_or_neg_one Λ' with h2 h2
|
||||
· rw [h2, (detContinuous_eq_one Λ').mpr h2]
|
||||
simp only [toMul_zero]
|
||||
· rw [h2, (detContinuous_eq_zero Λ').mpr h2]
|
||||
erw [Additive.toMul.apply_eq_iff_eq]
|
||||
change (0 : Fin 2) = (1 : Fin 2) ↔ _
|
||||
simp only [Fin.isValue, zero_ne_one, false_iff]
|
||||
linarith
|
||||
· rw [h1, (detContinuous_eq_zero Λ).mpr h1]
|
||||
cases' det_eq_one_or_neg_one Λ' with h2 h2
|
||||
· rw [h2, (detContinuous_eq_one Λ').mpr h2]
|
||||
erw [Additive.toMul.apply_eq_iff_eq]
|
||||
change (1 : Fin 2) = (0 : Fin 2) ↔ _
|
||||
simp only [Fin.isValue, one_ne_zero, false_iff]
|
||||
linarith
|
||||
· rw [h2, (detContinuous_eq_zero Λ').mpr h2]
|
||||
simp only [Nat.reduceAdd]
|
||||
|
||||
/-- The representation taking a Lorentz matrix to its determinant. -/
|
||||
@[simps!]
|
||||
def detRep : 𝓛 d →* ℤ₂ where
|
||||
toFun Λ := detContinuous Λ
|
||||
map_one' := by
|
||||
simp only [detContinuous, ContinuousMap.comp_apply, ContinuousMap.coe_mk,
|
||||
lorentzGroupIsGroup_one_coe, det_one, coeForℤ₂_apply, ↓reduceIte]
|
||||
map_mul' Λ1 Λ2 := by
|
||||
simp only [Submonoid.coe_mul, Subgroup.coe_toSubmonoid, Units.val_mul, det_mul, toMul_zero,
|
||||
mul_ite, mul_one, ite_mul, one_mul]
|
||||
cases' det_eq_one_or_neg_one Λ1 with h1 h1
|
||||
· rw [(detContinuous_eq_one Λ1).mpr h1]
|
||||
cases' det_eq_one_or_neg_one Λ2 with h2 h2
|
||||
· rw [(detContinuous_eq_one Λ2).mpr h2]
|
||||
apply (detContinuous_eq_one _).mpr
|
||||
simp only [lorentzGroupIsGroup_mul_coe, det_mul, h1, h2, mul_one]
|
||||
· rw [(detContinuous_eq_zero Λ2).mpr h2]
|
||||
apply (detContinuous_eq_zero _).mpr
|
||||
simp only [lorentzGroupIsGroup_mul_coe, det_mul, h1, h2, mul_neg, mul_one]
|
||||
· rw [(detContinuous_eq_zero Λ1).mpr h1]
|
||||
cases' det_eq_one_or_neg_one Λ2 with h2 h2
|
||||
· rw [(detContinuous_eq_one Λ2).mpr h2]
|
||||
apply (detContinuous_eq_zero _).mpr
|
||||
simp only [lorentzGroupIsGroup_mul_coe, det_mul, h1, h2, mul_one]
|
||||
· rw [(detContinuous_eq_zero Λ2).mpr h2]
|
||||
apply (detContinuous_eq_one _).mpr
|
||||
simp only [lorentzGroupIsGroup_mul_coe, det_mul, h1, h2, mul_neg, mul_one, neg_neg]
|
||||
|
||||
lemma detRep_continuous : Continuous (@detRep d) := detContinuous.2
|
||||
|
||||
lemma det_on_connected_component {Λ Λ' : LorentzGroup d} (h : Λ' ∈ connectedComponent Λ) :
|
||||
Λ.1.det = Λ'.1.det := by
|
||||
obtain ⟨s, hs, hΛ'⟩ := h
|
||||
let f : ContinuousMap s ℤ₂ := ContinuousMap.restrict s detContinuous
|
||||
haveI : PreconnectedSpace s := isPreconnected_iff_preconnectedSpace.mp hs.1
|
||||
simpa [f, detContinuous_eq_iff_det_eq] using
|
||||
(@IsPreconnected.subsingleton ℤ₂ _ _ _ (isPreconnected_range f.2))
|
||||
(Set.mem_range_self ⟨Λ, hs.2⟩) (Set.mem_range_self ⟨Λ', hΛ'⟩)
|
||||
|
||||
lemma detRep_on_connected_component {Λ Λ' : LorentzGroup d} (h : Λ' ∈ connectedComponent Λ) :
|
||||
detRep Λ = detRep Λ' := by
|
||||
simp only [detRep_apply, detContinuous, ContinuousMap.comp_apply, ContinuousMap.coe_mk,
|
||||
coeForℤ₂_apply, Subtype.mk.injEq]
|
||||
rw [det_on_connected_component h]
|
||||
|
||||
lemma det_of_joined {Λ Λ' : LorentzGroup d} (h : Joined Λ Λ') : Λ.1.det = Λ'.1.det :=
|
||||
det_on_connected_component $ pathComponent_subset_component _ h
|
||||
|
||||
/-- A Lorentz Matrix is proper if its determinant is 1. -/
|
||||
@[simp]
|
||||
def IsProper (Λ : LorentzGroup d) : Prop := Λ.1.det = 1
|
||||
|
||||
instance : DecidablePred (@IsProper d) := by
|
||||
intro Λ
|
||||
apply Real.decidableEq
|
||||
|
||||
lemma IsProper_iff (Λ : LorentzGroup d) : IsProper Λ ↔ detRep Λ = 1 := by
|
||||
rw [show 1 = detRep 1 from Eq.symm (MonoidHom.map_one detRep), detRep_apply, detRep_apply,
|
||||
detContinuous_eq_iff_det_eq]
|
||||
simp only [IsProper, lorentzGroupIsGroup_one_coe, det_one]
|
||||
|
||||
lemma id_IsProper : @IsProper d 1 := by
|
||||
simp [IsProper]
|
||||
|
||||
lemma isProper_on_connected_component {Λ Λ' : LorentzGroup d} (h : Λ' ∈ connectedComponent Λ) :
|
||||
IsProper Λ ↔ IsProper Λ' := by
|
||||
simp only [IsProper]
|
||||
rw [det_on_connected_component h]
|
||||
|
||||
end LorentzGroup
|
||||
|
||||
end
|
33
HepLean/Lorentz/Group/Restricted.lean
Normal file
33
HepLean/Lorentz/Group/Restricted.lean
Normal file
|
@ -0,0 +1,33 @@
|
|||
/-
|
||||
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.Lorentz.Group.Basic
|
||||
import HepLean.Lorentz.Group.Proper
|
||||
import HepLean.Lorentz.Group.Orthochronous
|
||||
import HepLean.Meta.Informal
|
||||
/-!
|
||||
# The Restricted Lorentz Group
|
||||
|
||||
This file is currently a stub.
|
||||
|
||||
-/
|
||||
/-! TODO: Add definition of the restricted Lorentz group. -/
|
||||
/-! TODO: Prove member of the restricted Lorentz group is combo of boost and rotation. -/
|
||||
/-! TODO: Prove restricted Lorentz group equivalent to connected component of identity. -/
|
||||
|
||||
noncomputable section
|
||||
|
||||
open Matrix
|
||||
open Complex
|
||||
open ComplexConjugate
|
||||
|
||||
namespace LorentzGroup
|
||||
|
||||
informal_definition Restricted where
|
||||
math :≈ "The subgroup of the Lorentz group consisting of elements which
|
||||
are proper and orthochronous."
|
||||
deps :≈ [``LorentzGroup, ``IsProper, ``IsOrthochronous]
|
||||
|
||||
end LorentzGroup
|
55
HepLean/Lorentz/Group/Rotations.lean
Normal file
55
HepLean/Lorentz/Group/Rotations.lean
Normal file
|
@ -0,0 +1,55 @@
|
|||
/-
|
||||
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.Lorentz.Group.Basic
|
||||
import HepLean.Mathematics.SO3.Basic
|
||||
import Mathlib.Topology.Constructions
|
||||
/-!
|
||||
# Rotations
|
||||
|
||||
This file describes the embedding of `SO(3)` into `LorentzGroup 3`.
|
||||
|
||||
-/
|
||||
/-! TODO: Generalize the inclusion of rotations into LorentzGroup to abitary dimension. -/
|
||||
noncomputable section
|
||||
|
||||
namespace LorentzGroup
|
||||
open GroupTheory
|
||||
|
||||
/-- Given a element of `SO(3)` the matrix corresponding to a space-time rotation. -/
|
||||
@[simp]
|
||||
def SO3ToMatrix (A : SO(3)) : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ :=
|
||||
(Matrix.fromBlocks 1 0 0 A.1)
|
||||
|
||||
lemma SO3ToMatrix_in_LorentzGroup (A : SO(3)) : SO3ToMatrix A ∈ LorentzGroup 3 := by
|
||||
rw [LorentzGroup.mem_iff_dual_mul_self]
|
||||
simp only [minkowskiMatrix.dual, minkowskiMatrix.as_block, SO3ToMatrix,
|
||||
Matrix.fromBlocks_transpose, Matrix.transpose_one, Matrix.transpose_zero,
|
||||
Matrix.fromBlocks_multiply, mul_one, Matrix.mul_zero, add_zero, Matrix.zero_mul, Matrix.mul_one,
|
||||
neg_mul, one_mul, zero_add, Matrix.mul_neg, neg_zero, mul_neg, neg_neg,
|
||||
Matrix.mul_eq_one_comm.mpr A.2.2, Matrix.fromBlocks_one]
|
||||
|
||||
lemma SO3ToMatrix_injective : Function.Injective SO3ToMatrix := by
|
||||
intro A B h
|
||||
erw [Equiv.apply_eq_iff_eq] at h
|
||||
have h1 := congrArg Matrix.toBlocks₂₂ h
|
||||
erw [Matrix.toBlocks_fromBlocks₂₂, Matrix.toBlocks_fromBlocks₂₂] at h1
|
||||
apply Subtype.eq
|
||||
exact h1
|
||||
|
||||
/-- Given a element of `SO(3)` the element of the Lorentz group corresponding to a
|
||||
space-time rotation. -/
|
||||
def SO3ToLorentz : SO(3) →* LorentzGroup 3 where
|
||||
toFun A := ⟨SO3ToMatrix A, SO3ToMatrix_in_LorentzGroup A⟩
|
||||
map_one' := Subtype.eq <| by
|
||||
simp only [SO3ToMatrix, SO3Group_one_coe, Matrix.fromBlocks_one, lorentzGroupIsGroup_one_coe]
|
||||
map_mul' A B := Subtype.eq <| by
|
||||
simp only [SO3ToMatrix, SO3Group_mul_coe, lorentzGroupIsGroup_mul_coe,
|
||||
Matrix.fromBlocks_multiply, mul_one, Matrix.mul_zero, add_zero, Matrix.zero_mul,
|
||||
Matrix.mul_one, zero_add]
|
||||
|
||||
end LorentzGroup
|
||||
|
||||
end
|
Loading…
Add table
Add a link
Reference in a new issue