refactor: Golfing
This commit is contained in:
parent
fbda420da9
commit
d7b6cf7246
13 changed files with 73 additions and 117 deletions
|
@ -48,13 +48,11 @@ variable (Λ : Matrix (Fin 4) (Fin 4) ℝ)
|
|||
lemma iff_on_right : PreservesηLin Λ ↔
|
||||
∀ (x y : spaceTime), ηLin x ((η * Λᵀ * η * Λ) *ᵥ y) = ηLin x y := by
|
||||
apply Iff.intro
|
||||
intro h
|
||||
intro x y
|
||||
intro h x y
|
||||
have h1 := h x y
|
||||
rw [ηLin_mulVec_left, mulVec_mulVec] at h1
|
||||
exact h1
|
||||
intro h
|
||||
intro x y
|
||||
intro h x y
|
||||
rw [ηLin_mulVec_left, mulVec_mulVec]
|
||||
exact h x y
|
||||
|
||||
|
@ -62,16 +60,12 @@ lemma iff_matrix : PreservesηLin Λ ↔ η * Λᵀ * η * Λ = 1 := by
|
|||
rw [iff_on_right, ηLin_matrix_eq_identity_iff (η * Λᵀ * η * Λ)]
|
||||
apply Iff.intro
|
||||
· simp_all [ηLin, implies_true, iff_true, one_mulVec]
|
||||
· simp_all only [ηLin, LinearMap.coe_mk, AddHom.coe_mk, linearMapForSpaceTime_apply,
|
||||
mulVec_mulVec, implies_true]
|
||||
· exact fun a x y => Eq.symm (Real.ext_cauchy (congrArg Real.cauchy (a x y)))
|
||||
|
||||
lemma iff_matrix' : PreservesηLin Λ ↔ Λ * (η * Λᵀ * η) = 1 := by
|
||||
rw [iff_matrix]
|
||||
apply Iff.intro
|
||||
intro h
|
||||
exact mul_eq_one_comm.mp h
|
||||
intro h
|
||||
exact mul_eq_one_comm.mp h
|
||||
exact mul_eq_one_comm
|
||||
|
||||
|
||||
lemma iff_transpose : PreservesηLin Λ ↔ PreservesηLin Λᵀ := by
|
||||
apply Iff.intro
|
||||
|
@ -80,7 +74,8 @@ lemma iff_transpose : PreservesηLin Λ ↔ PreservesηLin Λᵀ := by
|
|||
rw [transpose_mul, transpose_mul, transpose_mul, η_transpose,
|
||||
← mul_assoc, transpose_one] at h1
|
||||
rw [iff_matrix' Λ.transpose, ← h1]
|
||||
repeat rw [← mul_assoc]
|
||||
rw [← mul_assoc, ← mul_assoc]
|
||||
exact Matrix.mul_assoc (Λᵀ * η) Λᵀᵀ η
|
||||
intro h
|
||||
have h1 := congrArg transpose ((iff_matrix Λ.transpose).mp h)
|
||||
rw [transpose_mul, transpose_mul, transpose_mul, η_transpose,
|
||||
|
@ -157,7 +152,7 @@ lemma toGL_injective : Function.Injective toGL := by
|
|||
intro A B h
|
||||
apply Subtype.eq
|
||||
rw [@Units.ext_iff] at h
|
||||
simpa using h
|
||||
exact h
|
||||
|
||||
/-- The homomorphism from the Lorentz Group into the monoid of matrices times the opposite of
|
||||
the monoid of matrices. -/
|
||||
|
@ -201,21 +196,8 @@ lemma toGL_embedding : Embedding toGL.toFun where
|
|||
intro s
|
||||
rw [TopologicalSpace.ext_iff.mp toProd_embedding.induced s ]
|
||||
rw [isOpen_induced_iff, isOpen_induced_iff]
|
||||
apply Iff.intro ?_ ?_
|
||||
· intro h
|
||||
obtain ⟨U, hU1, hU2⟩ := h
|
||||
rw [isOpen_induced_iff] at hU1
|
||||
obtain ⟨V, hV1, hV2⟩ := hU1
|
||||
use V
|
||||
simp [hV1]
|
||||
rw [← hU2, ← hV2]
|
||||
rfl
|
||||
· intro h
|
||||
obtain ⟨U, hU1, hU2⟩ := h
|
||||
let t := (Units.embedProduct _) ⁻¹' U
|
||||
use t
|
||||
apply And.intro (isOpen_induced hU1)
|
||||
exact hU2
|
||||
exact exists_exists_and_eq_and
|
||||
|
||||
|
||||
instance : TopologicalGroup 𝓛 := Inducing.topologicalGroup toGL toGL_embedding.toInducing
|
||||
|
||||
|
|
|
@ -68,6 +68,10 @@ def genBoost (u v : FourVelocity) : spaceTime →ₗ[ℝ] spaceTime :=
|
|||
|
||||
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 : FourVelocity) : genBoost u u = LinearMap.id := by
|
||||
ext x
|
||||
simp [genBoost]
|
||||
|
|
|
@ -44,16 +44,13 @@ def fstCol (Λ : lorentzGroup) : PreFourVelocity := ⟨Λ.1 *ᵥ stdBasis 0, by
|
|||
cons_val_fin_one, vecCons_const, one_mul, mul_one, cons_val_one, head_cons, mul_neg, neg_mul,
|
||||
cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons, cons_val_three,
|
||||
head_fin_const] at h00
|
||||
rw [← h00]
|
||||
ring⟩
|
||||
exact h00⟩
|
||||
|
||||
/-- A Lorentz transformation is `orthochronous` if its `0 0` element is non-negative. -/
|
||||
def IsOrthochronous (Λ : lorentzGroup) : Prop := 0 ≤ Λ.1 0 0
|
||||
|
||||
lemma IsOrthochronous_iff_transpose (Λ : lorentzGroup) :
|
||||
IsOrthochronous Λ ↔ IsOrthochronous (transpose Λ) := by
|
||||
simp only [IsOrthochronous, Fin.isValue, transpose, PreservesηLin.liftGL,
|
||||
transpose_transpose, transpose_apply]
|
||||
IsOrthochronous Λ ↔ IsOrthochronous (transpose Λ) := by rfl
|
||||
|
||||
lemma IsOrthochronous_iff_fstCol_IsFourVelocity (Λ : lorentzGroup) :
|
||||
IsOrthochronous Λ ↔ IsFourVelocity (fstCol Λ) := by
|
||||
|
@ -61,9 +58,8 @@ lemma IsOrthochronous_iff_fstCol_IsFourVelocity (Λ : lorentzGroup) :
|
|||
rw [stdBasis_mulVec]
|
||||
|
||||
/-- The continuous map taking a Lorentz transformation to its `0 0` element. -/
|
||||
def mapZeroZeroComp : C(lorentzGroup, ℝ) := ⟨fun Λ => Λ.1 0 0, by
|
||||
refine Continuous.matrix_elem ?_ 0 0
|
||||
refine Continuous.comp' (continuous_iff_le_induced.mpr fun U a => a) continuous_id'⟩
|
||||
def mapZeroZeroComp : C(lorentzGroup, ℝ) := ⟨fun Λ => Λ.1 0 0,
|
||||
Continuous.matrix_elem (continuous_iff_le_induced.mpr fun _ a => a) 0 0⟩
|
||||
|
||||
/-- An auxillary function used in the definition of `orthchroMapReal`. -/
|
||||
def stepFunction : ℝ → ℝ := fun t =>
|
||||
|
@ -77,9 +73,9 @@ lemma stepFunction_continuous : Continuous stepFunction := by
|
|||
rw [ha]
|
||||
simp [neg_lt_self_iff, zero_lt_one, ↓reduceIte]
|
||||
have h1 : ¬ (1 : ℝ) ≤ 0 := by simp
|
||||
rw [if_neg h1]
|
||||
exact Eq.symm (if_neg h1)
|
||||
rw [Set.Ici_def, @frontier_Ici, @Set.mem_singleton_iff] at ha
|
||||
simp [ha]
|
||||
exact id (Eq.symm ha)
|
||||
|
||||
/-- The continuous map from `lorentzGroup` to `ℝ` wh
|
||||
taking Orthochronous elements to `1` and non-orthochronous to `-1`. -/
|
||||
|
@ -174,25 +170,22 @@ lemma mul_not_othchron_of_not_othchron_othchron {Λ Λ' : lorentzGroup} (h : ¬
|
|||
/-- The homomorphism from `lorentzGroup` to `ℤ₂` whose kernel are the Orthochronous elements. -/
|
||||
def orthchroRep : lorentzGroup →* ℤ₂ where
|
||||
toFun := orthchroMap
|
||||
map_one' := by
|
||||
have h1 : IsOrthochronous 1 := by simp [IsOrthochronous]
|
||||
rw [orthchroMap_IsOrthochronous h1]
|
||||
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')]
|
||||
simp only [mul_one]
|
||||
rfl
|
||||
rw [orthchroMap_IsOrthochronous h, orthchroMap_not_IsOrthochronous h',
|
||||
orthchroMap_not_IsOrthochronous (mul_not_othchron_of_othchron_not_othchron h h')]
|
||||
simp only [Nat.reduceAdd, one_mul]
|
||||
rfl
|
||||
rw [orthchroMap_not_IsOrthochronous h, orthchroMap_IsOrthochronous h',
|
||||
orthchroMap_not_IsOrthochronous (mul_not_othchron_of_not_othchron_othchron h h')]
|
||||
simp only [Nat.reduceAdd, mul_one]
|
||||
rfl
|
||||
rw [orthchroMap_not_IsOrthochronous h, orthchroMap_not_IsOrthochronous h',
|
||||
orthchroMap_IsOrthochronous (mul_othchron_of_not_othchron_not_othchron h h')]
|
||||
simp only [Nat.reduceAdd]
|
||||
rfl
|
||||
|
||||
end lorentzGroup
|
||||
|
|
|
@ -41,7 +41,8 @@ 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, by simp⟩ then (Additive.toMul 0) else (Additive.toMul (1 : ZMod 2))
|
||||
toFun x := if x = ⟨1, Set.mem_insert_of_mem (-1) rfl⟩
|
||||
then (Additive.toMul 0) else (Additive.toMul (1 : ZMod 2))
|
||||
continuous_toFun := by
|
||||
haveI : DiscreteTopology ({-1, 1} : Set ℝ) := discrete_of_t1_of_finite
|
||||
exact continuous_of_discreteTopology
|
||||
|
@ -118,7 +119,7 @@ instance : DecidablePred IsProper := by
|
|||
apply Real.decidableEq
|
||||
|
||||
lemma IsProper_iff (Λ : lorentzGroup) : IsProper Λ ↔ detRep Λ = 1 := by
|
||||
rw [show 1 = detRep 1 by simp]
|
||||
rw [show 1 = detRep 1 from Eq.symm (MonoidHom.map_one detRep)]
|
||||
rw [detRep_apply, detRep_apply, detContinuous_eq_iff_det_eq]
|
||||
simp only [IsProper, lorentzGroupIsGroup_one_coe, det_one]
|
||||
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue