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
|
|
|
|
|
-/
|
|
|
|
|
import HepLean.SpaceTime.LorentzGroup.Basic
|
|
|
|
|
/-!
|
|
|
|
|
# The Proper Lorentz Group
|
|
|
|
|
|
|
|
|
|
We define the give a series of lemmas related to the determinant of the lorentz group.
|
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
noncomputable section
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
namespace SpaceTime
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
|
|
|
|
open Manifold
|
|
|
|
|
open Matrix
|
|
|
|
|
open Complex
|
|
|
|
|
open ComplexConjugate
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
namespace LorentzGroup
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
|
|
|
|
/-- The determinant of a member of the lorentz group is `1` or `-1`. -/
|
2024-05-22 13:34:53 -04:00
|
|
|
|
lemma det_eq_one_or_neg_one (Λ : 𝓛) : Λ.1.det = 1 ∨ Λ.1.det = -1 := by
|
2024-05-17 11:52:16 -04:00
|
|
|
|
simpa [← sq, det_one, det_mul, det_mul, det_mul, det_transpose, det_η] using
|
2024-05-22 13:34:53 -04:00
|
|
|
|
(congrArg det ((PreservesηLin.iff_matrix' Λ.1).mp Λ.2))
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
|
|
|
|
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
|
2024-06-13 08:10:08 -04:00
|
|
|
|
toFun x := if x = ⟨1, Set.mem_insert_of_mem (-1) rfl⟩
|
|
|
|
|
then (Additive.toMul 0) else (Additive.toMul (1 : ZMod 2))
|
2024-05-17 11:52:16 -04:00
|
|
|
|
continuous_toFun := by
|
|
|
|
|
haveI : DiscreteTopology ({-1, 1} : Set ℝ) := discrete_of_t1_of_finite
|
|
|
|
|
exact continuous_of_discreteTopology
|
|
|
|
|
|
|
|
|
|
/-- The continuous map taking a lorentz matrix to its determinant. -/
|
2024-05-22 13:34:53 -04:00
|
|
|
|
def detContinuous : C(𝓛, ℤ₂) :=
|
2024-05-17 11:52:16 -04:00
|
|
|
|
ContinuousMap.comp coeForℤ₂ {
|
2024-06-26 11:54:02 -04:00
|
|
|
|
toFun := fun Λ => ⟨Λ.1.det, Or.symm (LorentzGroup.det_eq_one_or_neg_one _)⟩,
|
2024-05-17 11:52:16 -04:00
|
|
|
|
continuous_toFun := by
|
|
|
|
|
refine Continuous.subtype_mk ?_ _
|
2024-05-22 13:34:53 -04:00
|
|
|
|
apply Continuous.matrix_det $
|
|
|
|
|
Continuous.comp' (continuous_iff_le_induced.mpr fun U a => a) continuous_id'
|
|
|
|
|
}
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma detContinuous_eq_iff_det_eq (Λ Λ' : LorentzGroup) :
|
2024-05-22 13:34:53 -04:00
|
|
|
|
detContinuous Λ = detContinuous Λ' ↔ Λ.1.det = Λ'.1.det := by
|
2024-05-17 11:52:16 -04:00
|
|
|
|
apply Iff.intro
|
|
|
|
|
intro h
|
|
|
|
|
simp [detContinuous] at h
|
|
|
|
|
cases' det_eq_one_or_neg_one Λ with h1 h1
|
|
|
|
|
<;> cases' det_eq_one_or_neg_one Λ' with h2 h2
|
|
|
|
|
<;> simp_all [h1, h2, h]
|
|
|
|
|
rw [← toMul_zero, @Equiv.apply_eq_iff_eq] at h
|
2024-06-25 09:18:43 -04:00
|
|
|
|
· change (0 : Fin 2) = (1 : Fin 2) at h
|
|
|
|
|
simp only [Fin.isValue, zero_ne_one] at h
|
|
|
|
|
· change (1 : Fin 2) = (0 : Fin 2) at h
|
|
|
|
|
simp only [Fin.isValue, one_ne_zero] at h
|
|
|
|
|
· intro h
|
|
|
|
|
simp [detContinuous, h]
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
|
|
|
|
|
|
|
|
|
/-- The representation taking a lorentz matrix to its determinant. -/
|
|
|
|
|
@[simps!]
|
2024-05-22 13:34:53 -04:00
|
|
|
|
def detRep : 𝓛 →* ℤ₂ where
|
2024-05-17 11:52:16 -04:00
|
|
|
|
toFun Λ := detContinuous Λ
|
|
|
|
|
map_one' := by
|
2024-05-22 13:34:53 -04:00
|
|
|
|
simp [detContinuous, lorentzGroupIsGroup]
|
2024-05-17 11:52:16 -04:00
|
|
|
|
map_mul' := by
|
|
|
|
|
intro Λ1 Λ2
|
|
|
|
|
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
|
|
|
|
|
<;> cases' (det_eq_one_or_neg_one Λ2) with h2 h2
|
|
|
|
|
<;> simp [h1, h2, detContinuous]
|
|
|
|
|
rfl
|
|
|
|
|
|
|
|
|
|
lemma detRep_continuous : Continuous detRep := detContinuous.2
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma det_on_connected_component {Λ Λ' : LorentzGroup} (h : Λ' ∈ connectedComponent Λ) :
|
2024-05-22 13:34:53 -04:00
|
|
|
|
Λ.1.det = Λ'.1.det := by
|
2024-05-17 11:52:16 -04:00
|
|
|
|
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Λ'⟩)
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma detRep_on_connected_component {Λ Λ' : LorentzGroup} (h : Λ' ∈ connectedComponent Λ) :
|
2024-05-20 16:20:26 -04:00
|
|
|
|
detRep Λ = detRep Λ' := by
|
|
|
|
|
simp [detRep_apply, detRep_apply, detContinuous]
|
|
|
|
|
rw [det_on_connected_component h]
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma det_of_joined {Λ Λ' : LorentzGroup} (h : Joined Λ Λ') : Λ.1.det = Λ'.1.det :=
|
2024-05-17 11:52:16 -04:00
|
|
|
|
det_on_connected_component $ pathComponent_subset_component _ h
|
|
|
|
|
|
|
|
|
|
/-- A Lorentz Matrix is proper if its determinant is 1. -/
|
|
|
|
|
@[simp]
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def IsProper (Λ : LorentzGroup) : Prop := Λ.1.det = 1
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
|
|
|
|
instance : DecidablePred IsProper := by
|
|
|
|
|
intro Λ
|
|
|
|
|
apply Real.decidableEq
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma IsProper_iff (Λ : LorentzGroup) : IsProper Λ ↔ detRep Λ = 1 := by
|
2024-06-13 08:10:08 -04:00
|
|
|
|
rw [show 1 = detRep 1 from Eq.symm (MonoidHom.map_one detRep)]
|
2024-05-17 11:52:16 -04:00
|
|
|
|
rw [detRep_apply, detRep_apply, detContinuous_eq_iff_det_eq]
|
2024-05-22 13:34:53 -04:00
|
|
|
|
simp only [IsProper, lorentzGroupIsGroup_one_coe, det_one]
|
2024-05-17 11:52:16 -04:00
|
|
|
|
|
2024-05-20 16:20:26 -04:00
|
|
|
|
lemma id_IsProper : IsProper 1 := by
|
|
|
|
|
simp [IsProper]
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isProper_on_connected_component {Λ Λ' : LorentzGroup} (h : Λ' ∈ connectedComponent Λ) :
|
2024-05-20 16:20:26 -04:00
|
|
|
|
IsProper Λ ↔ IsProper Λ' := by
|
|
|
|
|
simp [detRep_apply, detRep_apply, detContinuous]
|
|
|
|
|
rw [det_on_connected_component h]
|
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
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
end SpaceTime
|
2024-05-17 11:52:16 -04:00
|
|
|
|
end
|