2024-05-17 11:52:16 -04:00
|
|
|
|
/-
|
|
|
|
|
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.
|
2024-05-17 11:52:16 -04:00
|
|
|
|
Authors: Joseph Tooby-Smith
|
|
|
|
|
-/
|
2024-11-09 17:37:12 +00:00
|
|
|
|
import HepLean.Lorentz.RealVector.NormOne
|
2024-11-09 17:29:43 +00:00
|
|
|
|
import HepLean.Lorentz.Group.Proper
|
2024-05-17 11:52:16 -04:00
|
|
|
|
/-!
|
|
|
|
|
# The Orthochronous Lorentz Group
|
|
|
|
|
|
|
|
|
|
We define the give a series of lemmas related to the orthochronous property of lorentz
|
|
|
|
|
matrices.
|
|
|
|
|
|
|
|
|
|
-/
|
2024-07-09 19:22:16 -04:00
|
|
|
|
/-! TODO: Prove topological properties of the Orthochronous Lorentz Group. -/
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
|
|
|
|
noncomputable section
|
|
|
|
|
|
|
|
|
|
open Matrix
|
|
|
|
|
open Complex
|
|
|
|
|
open ComplexConjugate
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
namespace LorentzGroup
|
2024-07-01 16:56:15 -04:00
|
|
|
|
|
|
|
|
|
variable {d : ℕ}
|
|
|
|
|
variable (Λ : LorentzGroup d)
|
2024-11-08 16:24:58 +00:00
|
|
|
|
open Lorentz.Contr
|
2024-05-17 15:10:35 -04:00
|
|
|
|
|
2024-07-12 11:23:02 -04:00
|
|
|
|
/-- A Lorentz transformation is `orthochronous` if its `0 0` element is non-negative. -/
|
2024-11-08 16:24:58 +00:00
|
|
|
|
def IsOrthochronous : Prop := 0 ≤ Λ.1 (Sum.inl 0) (Sum.inl 0)
|
2024-07-01 16:56:15 -04:00
|
|
|
|
|
|
|
|
|
lemma IsOrthochronous_iff_futurePointing :
|
2024-11-08 16:24:58 +00:00
|
|
|
|
IsOrthochronous Λ ↔ toNormOne Λ ∈ NormOne.FuturePointing d := by
|
|
|
|
|
simp only [IsOrthochronous]
|
|
|
|
|
rw [NormOne.FuturePointing.mem_iff_inl_nonneg, toNormOne_inl]
|
2024-05-17 15:10:35 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma IsOrthochronous_iff_transpose :
|
2024-06-13 08:10:08 -04:00
|
|
|
|
IsOrthochronous Λ ↔ IsOrthochronous (transpose Λ) := by rfl
|
2024-05-17 15:10:35 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma IsOrthochronous_iff_ge_one :
|
2024-11-08 16:24:58 +00:00
|
|
|
|
IsOrthochronous Λ ↔ 1 ≤ Λ.1 (Sum.inl 0) (Sum.inl 0) := by
|
|
|
|
|
rw [IsOrthochronous_iff_futurePointing, NormOne.FuturePointing.mem_iff_inl_one_le_inl,
|
|
|
|
|
toNormOne_inl]
|
2024-07-01 16:56:15 -04:00
|
|
|
|
|
|
|
|
|
lemma not_orthochronous_iff_le_neg_one :
|
2024-11-08 16:24:58 +00:00
|
|
|
|
¬ 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]
|
2024-07-01 16:56:15 -04:00
|
|
|
|
|
2024-05-17 15:28:05 -04:00
|
|
|
|
/-- The continuous map taking a Lorentz transformation to its `0 0` element. -/
|
2024-11-08 16:24:58 +00:00
|
|
|
|
def timeCompCont : C(LorentzGroup d, ℝ) := ⟨fun Λ => Λ.1 (Sum.inl 0) (Sum.inl 0),
|
2024-07-19 17:00:32 -04:00
|
|
|
|
Continuous.matrix_elem (continuous_iff_le_induced.mpr fun _ a => a) (Sum.inl 0) (Sum.inl 0)⟩
|
2024-05-17 15:10:35 -04:00
|
|
|
|
|
2024-05-17 15:28:05 -04:00
|
|
|
|
/-- An auxillary function used in the definition of `orthchroMapReal`. -/
|
2024-05-17 15:10:35 -04:00
|
|
|
|
def stepFunction : ℝ → ℝ := fun t =>
|
|
|
|
|
if t ≤ -1 then -1 else
|
2024-07-12 10:36:39 -04:00
|
|
|
|
if 1 ≤ t then 1 else t
|
2024-05-17 15:10:35 -04:00
|
|
|
|
|
|
|
|
|
lemma stepFunction_continuous : Continuous stepFunction := by
|
|
|
|
|
apply Continuous.if ?_ continuous_const (Continuous.if ?_ continuous_const continuous_id)
|
2024-07-19 17:00:32 -04:00
|
|
|
|
<;> intro a ha
|
2024-08-21 06:40:58 -04:00
|
|
|
|
· rw [@Set.Iic_def, @frontier_Iic, @Set.mem_singleton_iff] at ha
|
|
|
|
|
rw [ha]
|
2024-09-04 07:31:34 -04:00
|
|
|
|
simp only [le_neg_self_iff, id_eq]
|
2024-08-21 06:40:58 -04:00
|
|
|
|
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)
|
2024-05-17 15:10:35 -04:00
|
|
|
|
|
2024-05-17 15:28:05 -04:00
|
|
|
|
/-- The continuous map from `lorentzGroup` to `ℝ` wh
|
|
|
|
|
taking Orthochronous elements to `1` and non-orthochronous to `-1`. -/
|
2024-07-01 16:56:15 -04:00
|
|
|
|
def orthchroMapReal : C(LorentzGroup d, ℝ) := ContinuousMap.comp
|
|
|
|
|
⟨stepFunction, stepFunction_continuous⟩ timeCompCont
|
2024-05-17 15:10:35 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma orthchroMapReal_on_IsOrthochronous {Λ : LorentzGroup d} (h : IsOrthochronous Λ) :
|
2024-05-17 15:10:35 -04:00
|
|
|
|
orthchroMapReal Λ = 1 := by
|
2024-11-08 16:24:58 +00:00
|
|
|
|
rw [IsOrthochronous_iff_ge_one] at h
|
2024-07-01 16:56:15 -04:00
|
|
|
|
change stepFunction (Λ.1 _ _) = 1
|
|
|
|
|
rw [stepFunction, if_pos h, if_neg]
|
|
|
|
|
linarith
|
2024-05-17 15:10:35 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma orthchroMapReal_on_not_IsOrthochronous {Λ : LorentzGroup d} (h : ¬ IsOrthochronous Λ) :
|
2024-05-17 15:10:35 -04:00
|
|
|
|
orthchroMapReal Λ = - 1 := by
|
2024-07-01 16:56:15 -04:00
|
|
|
|
rw [not_orthochronous_iff_le_neg_one] at h
|
2024-11-08 16:24:58 +00:00
|
|
|
|
change stepFunction (_)= - 1
|
|
|
|
|
rw [stepFunction, if_pos]
|
|
|
|
|
exact h
|
2024-05-17 15:10:35 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma orthchroMapReal_minus_one_or_one (Λ : LorentzGroup d) :
|
2024-05-17 15:10:35 -04:00
|
|
|
|
orthchroMapReal Λ = -1 ∨ orthchroMapReal Λ = 1 := by
|
|
|
|
|
by_cases h : IsOrthochronous Λ
|
2024-08-21 06:40:58 -04:00
|
|
|
|
· exact Or.inr $ orthchroMapReal_on_IsOrthochronous h
|
|
|
|
|
· exact Or.inl $ orthchroMapReal_on_not_IsOrthochronous h
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-07-12 10:36:39 -04:00
|
|
|
|
local notation "ℤ₂" => Multiplicative (ZMod 2)
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-05-17 15:28:05 -04:00
|
|
|
|
/-- A continuous map from `lorentzGroup` to `ℤ₂` whose kernal are the Orthochronous elements. -/
|
2024-07-01 16:56:15 -04:00
|
|
|
|
def orthchroMap : C(LorentzGroup d, ℤ₂) :=
|
2024-05-17 15:10:35 -04:00
|
|
|
|
ContinuousMap.comp coeForℤ₂ {
|
|
|
|
|
toFun := fun Λ => ⟨orthchroMapReal Λ, orthchroMapReal_minus_one_or_one Λ⟩,
|
2024-07-12 10:36:39 -04:00
|
|
|
|
continuous_toFun := Continuous.subtype_mk (ContinuousMap.continuous orthchroMapReal) _}
|
2024-05-17 15:10:35 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma orthchroMap_IsOrthochronous {Λ : LorentzGroup d} (h : IsOrthochronous Λ) :
|
2024-05-17 15:10:35 -04:00
|
|
|
|
orthchroMap Λ = 1 := by
|
|
|
|
|
simp [orthchroMap, orthchroMapReal_on_IsOrthochronous h]
|
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma orthchroMap_not_IsOrthochronous {Λ : LorentzGroup d} (h : ¬ IsOrthochronous Λ) :
|
2024-07-12 10:36:39 -04:00
|
|
|
|
orthchroMap Λ = Additive.toMul (1 : ZMod 2) := by
|
2024-09-04 06:28:46 -04:00
|
|
|
|
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
|
2024-05-17 15:10:35 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma mul_othchron_of_othchron_othchron {Λ Λ' : LorentzGroup d} (h : IsOrthochronous Λ)
|
2024-05-17 15:10:35 -04:00
|
|
|
|
(h' : IsOrthochronous Λ') : IsOrthochronous (Λ * Λ') := by
|
|
|
|
|
rw [IsOrthochronous_iff_transpose] at h
|
2024-07-01 16:56:15 -04:00
|
|
|
|
rw [IsOrthochronous_iff_futurePointing] at h h'
|
2024-11-08 16:24:58 +00:00
|
|
|
|
rw [IsOrthochronous, LorentzGroup.inl_inl_mul]
|
|
|
|
|
exact NormOne.FuturePointing.metric_reflect_mem_mem h h'
|
2024-05-17 15:10:35 -04:00
|
|
|
|
|
2024-07-01 16:56:15 -04:00
|
|
|
|
lemma mul_othchron_of_not_othchron_not_othchron {Λ Λ' : LorentzGroup d} (h : ¬ IsOrthochronous Λ)
|
2024-05-17 15:10:35 -04:00
|
|
|
|
(h' : ¬ IsOrthochronous Λ') : IsOrthochronous (Λ * Λ') := by
|
|
|
|
|
rw [IsOrthochronous_iff_transpose] at h
|
2024-07-01 16:56:15 -04:00
|
|
|
|
rw [IsOrthochronous_iff_futurePointing] at h h'
|
2024-11-08 16:24:58 +00:00
|
|
|
|
rw [IsOrthochronous, LorentzGroup.inl_inl_mul]
|
|
|
|
|
exact NormOne.FuturePointing.metric_reflect_not_mem_not_mem h h'
|
2024-05-17 15:10:35 -04:00
|
|
|
|
|
2024-07-12 10:36:39 -04:00
|
|
|
|
lemma mul_not_othchron_of_othchron_not_othchron {Λ Λ' : LorentzGroup d} (h : IsOrthochronous Λ)
|
2024-05-17 15:10:35 -04:00
|
|
|
|
(h' : ¬ IsOrthochronous Λ') : ¬ IsOrthochronous (Λ * Λ') := by
|
2024-11-08 16:24:58 +00:00
|
|
|
|
rw [not_orthochronous_iff_le_zero, LorentzGroup.inl_inl_mul]
|
2024-05-17 15:10:35 -04:00
|
|
|
|
rw [IsOrthochronous_iff_transpose] at h
|
2024-07-01 16:56:15 -04:00
|
|
|
|
rw [IsOrthochronous_iff_futurePointing] at h h'
|
2024-11-08 16:24:58 +00:00
|
|
|
|
exact NormOne.FuturePointing.metric_reflect_mem_not_mem h h'
|
2024-07-01 16:56:15 -04:00
|
|
|
|
|
|
|
|
|
lemma mul_not_othchron_of_not_othchron_othchron {Λ Λ' : LorentzGroup d} (h : ¬ IsOrthochronous Λ)
|
2024-05-17 15:10:35 -04:00
|
|
|
|
(h' : IsOrthochronous Λ') : ¬ IsOrthochronous (Λ * Λ') := by
|
2024-11-08 16:24:58 +00:00
|
|
|
|
rw [not_orthochronous_iff_le_zero, LorentzGroup.inl_inl_mul]
|
2024-05-17 15:10:35 -04:00
|
|
|
|
rw [IsOrthochronous_iff_transpose] at h
|
2024-07-01 16:56:15 -04:00
|
|
|
|
rw [IsOrthochronous_iff_futurePointing] at h h'
|
2024-11-08 16:24:58 +00:00
|
|
|
|
exact NormOne.FuturePointing.metric_reflect_not_mem_mem h h'
|
2024-05-17 15:10:35 -04:00
|
|
|
|
|
2024-07-12 16:22:06 -04:00
|
|
|
|
/-- The homomorphism from `LorentzGroup` to `ℤ₂` whose kernel are the Orthochronous elements. -/
|
2024-07-01 16:56:15 -04:00
|
|
|
|
def orthchroRep : LorentzGroup d →* ℤ₂ where
|
2024-05-17 15:10:35 -04:00
|
|
|
|
toFun := orthchroMap
|
2024-06-13 08:10:08 -04:00
|
|
|
|
map_one' := orthchroMap_IsOrthochronous (by simp [IsOrthochronous])
|
2024-05-17 15:10:35 -04:00
|
|
|
|
map_mul' Λ Λ' := by
|
2024-05-17 15:28:05 -04:00
|
|
|
|
simp only
|
2024-05-17 15:10:35 -04:00
|
|
|
|
by_cases h : IsOrthochronous Λ
|
2024-07-19 17:00:32 -04:00
|
|
|
|
<;> by_cases h' : IsOrthochronous Λ'
|
2024-08-21 06:40:58 -04:00
|
|
|
|
· 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
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
end LorentzGroup
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
|
|
|
|
end
|