feat: Properties of the group SO(3)

This commit is contained in:
jstoobysmith 2024-05-20 16:20:26 -04:00
parent a40d21378b
commit 8024aa92a1
4 changed files with 172 additions and 1 deletions

View file

@ -163,6 +163,12 @@ lemma toLorentz_joined_to_1 (u v : FourVelocity) : Joined 1 (toLorentz u v) := b
simp [PreservesηLin.liftGL, toMatrix, self u]
· simp
lemma toLorentz_in_connected_component_1 (u v : FourVelocity) :
toLorentz u v ∈ connectedComponent 1 :=
pathComponent_subset_component _ (toLorentz_joined_to_1 u v)
lemma isProper (u v : FourVelocity) : IsProper (toLorentz u v) :=
(isProper_on_connected_component (toLorentz_in_connected_component_1 u v)).mp id_IsProper
end genBoost

View file

@ -100,6 +100,11 @@ lemma det_on_connected_component {Λ Λ' : lorentzGroup} (h : Λ' ∈ connected
(@IsPreconnected.subsingleton ℤ₂ _ _ _ (isPreconnected_range f.2))
(Set.mem_range_self ⟨Λ, hs.2⟩) (Set.mem_range_self ⟨Λ', hΛ'⟩)
lemma detRep_on_connected_component {Λ Λ' : lorentzGroup} (h : Λ' ∈ connectedComponent Λ) :
detRep Λ = detRep Λ' := by
simp [detRep_apply, detRep_apply, detContinuous]
rw [det_on_connected_component h]
lemma det_of_joined {Λ Λ' : lorentzGroup} (h : Joined Λ Λ') : Λ.1.1.det = Λ'.1.1.det :=
det_on_connected_component $ pathComponent_subset_component _ h
@ -116,6 +121,13 @@ lemma IsProper_iff (Λ : lorentzGroup) : IsProper Λ ↔ detRep Λ = 1 := by
rw [detRep_apply, detRep_apply, detContinuous_eq_iff_det_eq]
simp only [IsProper, OneMemClass.coe_one, Units.val_one, det_one]
lemma id_IsProper : IsProper 1 := by
simp [IsProper]
lemma isProper_on_connected_component {Λ Λ' : lorentzGroup} (h : Λ' ∈ connectedComponent Λ) :
IsProper Λ ↔ IsProper Λ' := by
simp [detRep_apply, detRep_apply, detContinuous]
rw [det_on_connected_component h]
end lorentzGroup