reactor: Removal of double spaces

This commit is contained in:
jstoobysmith 2024-07-12 11:23:02 -04:00
parent ce92e1d649
commit 13f62a50eb
64 changed files with 550 additions and 546 deletions

View file

@ -28,7 +28,7 @@ def lorentzAlgebra : LieSubalgebra (Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin
namespace lorentzAlgebra
open minkowskiMatrix
lemma transpose_eta (A : lorentzAlgebra) : A.1ᵀ * η = - η * A.1 := by
lemma transpose_eta (A : lorentzAlgebra) : A.1ᵀ * η = - η * A.1 := by
have h1 := A.2
erw [mem_skewAdjointMatricesLieSubalgebra] at h1
simpa [LieAlgebra.Orthogonal.so', IsSkewAdjoint, IsAdjointPair] using h1
@ -39,7 +39,7 @@ lemma mem_of_transpose_eta_eq_eta_mul_self {A : Matrix (Fin 1 ⊕ Fin 3) (Fin 1
simpa [LieAlgebra.Orthogonal.so', IsSkewAdjoint, IsAdjointPair] using h
lemma mem_iff {A : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) } :
A ∈ lorentzAlgebra ↔ Aᵀ * η = - η * A :=
A ∈ lorentzAlgebra ↔ Aᵀ * η = - η * A :=
Iff.intro (fun h => transpose_eta ⟨A, h⟩) (fun h => mem_of_transpose_eta_eq_eta_mul_self h)
lemma mem_iff' (A : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ) :
@ -52,7 +52,7 @@ lemma mem_iff' (A : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ) :
rw [minkowskiMatrix.sq]
all_goals noncomm_ring
· nth_rewrite 2 [h]
trans (η * η) * Aᵀ * η
trans (η * η) * Aᵀ * η
rw [minkowskiMatrix.sq]
all_goals noncomm_ring
@ -80,7 +80,7 @@ end lorentzAlgebra
@[simps!]
instance lorentzVectorAsLieRingModule : LieRingModule lorentzAlgebra (LorentzVector 3) where
bracket Λ x := Λ.1.mulVec x
bracket Λ x := Λ.1.mulVec x
add_lie Λ1 Λ2 x := by
simp [add_mulVec]
lie_add Λ x1 x2 := by

View file

@ -30,7 +30,7 @@ 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 : }
variable {d : }
open minkowskiMetric in
/-- The Lorentz group is the subset of matrices which preserve the minkowski metric. -/
@ -74,7 +74,7 @@ lemma mem_iff_on_right : Λ ∈ LorentzGroup d ↔
rw [← dual_mulVec_right, mulVec_mulVec]
exact h x y
lemma mem_iff_dual_mul_self : Λ ∈ LorentzGroup d ↔ dual Λ * Λ = 1 := by
lemma mem_iff_dual_mul_self : Λ ∈ LorentzGroup d ↔ dual Λ * Λ = 1 := by
rw [mem_iff_on_right, matrix_eq_id_iff]
exact forall_comm
@ -145,7 +145,7 @@ namespace LorentzGroup
open minkowskiMetric
variable {Λ Λ' : LorentzGroup d}
variable {Λ Λ' : LorentzGroup d}
lemma coe_inv : (Λ⁻¹).1 = Λ.1⁻¹:= by
refine (inv_eq_left_inv ?h).symm
@ -172,7 +172,7 @@ def toGL : LorentzGroup d →* GL (Fin 1 ⊕ Fin d) where
map_one' := by
simp
rfl
map_mul' x y := by
map_mul' x y := by
simp only [lorentzGroupIsGroup, _root_.mul_inv_rev, coe_inv]
ext
rfl

View file

@ -99,7 +99,7 @@ lemma toMatrix_mulVec (u v : FuturePointing d) (x : LorentzVector d) :
open minkowskiMatrix LorentzVector in
@[simp]
lemma toMatrix_apply (u v : FuturePointing d) (μ ν : Fin 1 ⊕ Fin d) :
lemma toMatrix_apply (u v : FuturePointing d) (μ ν : Fin 1 ⊕ Fin d) :
(toMatrix u v) μ ν = η μ μ * (⟪e μ, e ν⟫ₘ + 2 * ⟪e ν, u⟫ₘ * ⟪e μ, v⟫ₘ
- ⟪e μ, u + v⟫ₘ * ⟪e ν, u + v⟫ₘ / (1 + ⟪u, v.1.1⟫ₘ)) := by
rw [matrix_apply_stdBasis (toMatrix u v) μ ν, toMatrix_mulVec]

View file

@ -27,7 +27,7 @@ variable (Λ : LorentzGroup d)
open LorentzVector
open minkowskiMetric
/-- A Lorentz transformation is `orthochronous` if its `0 0` element is non-negative. -/
/-- A Lorentz transformation is `orthochronous` if its `0 0` element is non-negative. -/
def IsOrthochronous : Prop := 0 ≤ timeComp Λ
lemma IsOrthochronous_iff_futurePointing :
@ -62,7 +62,7 @@ lemma not_orthochronous_iff_le_zero :
linarith
/-- The continuous map taking a Lorentz transformation to its `0 0` element. -/
def timeCompCont : C(LorentzGroup d, ) := ⟨fun Λ => timeComp Λ ,
def timeCompCont : C(LorentzGroup d, ) := ⟨fun Λ => timeComp Λ,
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`. -/
@ -75,7 +75,7 @@ lemma stepFunction_continuous : Continuous stepFunction := by
<;> intro a ha
rw [@Set.Iic_def, @frontier_Iic, @Set.mem_singleton_iff] at ha
rw [ha]
simp [neg_lt_self_iff, zero_lt_one, ↓reduceIte]
simp [neg_lt_self_iff, zero_lt_one, ↓reduceIte]
have h1 : ¬ (1 : ) ≤ 0 := by simp
exact Eq.symm (if_neg h1)
rw [Set.Ici_def, @frontier_Ici, @Set.mem_singleton_iff] at ha

View file

@ -31,7 +31,7 @@ lemma det_eq_one_or_neg_one (Λ : 𝓛 d) : Λ.1.det = 1 Λ.1.det = -1 := by
simp [det_mul, det_dual] at h1
exact mul_self_eq_one_iff.mp h1
local notation "ℤ₂" => Multiplicative (ZMod 2)
local notation "ℤ₂" => Multiplicative (ZMod 2)
instance : TopologicalSpace ℤ₂ := instTopologicalSpaceFin
@ -51,7 +51,7 @@ def coeFor₂ : C(({-1, 1} : Set ), ℤ₂) where
/-- The continuous map taking a Lorentz matrix to its determinant. -/
def detContinuous : C(𝓛 d, ℤ₂) :=
ContinuousMap.comp coeFor₂ {
ContinuousMap.comp coeFor₂ {
toFun := fun Λ => ⟨Λ.1.det, Or.symm (LorentzGroup.det_eq_one_or_neg_one _)⟩,
continuous_toFun := by
refine Continuous.subtype_mk ?_ _
@ -65,7 +65,7 @@ lemma detContinuous_eq_iff_det_eq (Λ Λ' : LorentzGroup d) :
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
<;> 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
· change (0 : Fin 2) = (1 : Fin 2) at h

View file

@ -47,7 +47,7 @@ instance (d : ) (μ : RealLorentzTensor.Colors) : Fintype (RealLorentzTensor.
| RealLorentzTensor.Colors.down => instFintypeSum (Fin 1) (Fin d)
/-- An `IndexValue` is a set of actual values an index can take. e.g. for a
3-tensor (0, 1, 2). -/
3-tensor (0, 1, 2). -/
@[simp]
def RealLorentzTensor.IndexValue {X : FintypeCat} (d : ) (c : X → RealLorentzTensor.Colors) :
Type 0 := (x : X) → RealLorentzTensor.ColorsIndex d (c x)
@ -107,9 +107,9 @@ def congrColorsDual {μ ν : Colors} (h : μ = τ ν) :
ColorsIndex d μ ≃ ColorsIndex d ν :=
(castColorsIndex h).trans dualColorsIndex.symm
lemma congrColorsDual_symm {μ ν : Colors} (h : μ = τ ν) :
lemma congrColorsDual_symm {μ ν : Colors} (h : μ = τ ν) :
(congrColorsDual h).symm =
@congrColorsDual d _ _ ((Function.Involutive.eq_iff τ_involutive).mp h.symm) := by
@congrColorsDual d _ _ ((Function.Involutive.eq_iff τ_involutive).mp h.symm) := by
match μ, ν with
| Colors.up, Colors.down => rfl
| Colors.down, Colors.up => rfl
@ -124,7 +124,7 @@ lemma color_eq_dual_symm {μ ν : Colors} (h : μ = τ ν) : ν = τ μ :=
-/
/-- An equivalence of Index values from an equality of color maps. -/
def castIndexValue {X : FintypeCat} {T S : X → Colors} (h : T = S) :
def castIndexValue {X : FintypeCat} {T S : X → Colors} (h : T = S) :
IndexValue d T ≃ IndexValue d S where
toFun i := (fun μ => castColorsIndex (congrFun h μ) (i μ))
invFun i := (fun μ => (castColorsIndex (congrFun h μ)).symm (i μ))
@ -133,7 +133,7 @@ def castIndexValue {X : FintypeCat} {T S : X → Colors} (h : T = S) :
right_inv i := by
simp
lemma indexValue_eq {T₁ T₂ : X → RealLorentzTensor.Colors} (d : ) (h : T₁ = T₂) :
lemma indexValue_eq {T₁ T₂ : X → RealLorentzTensor.Colors} (d : ) (h : T₁ = T₂) :
IndexValue d T₁ = IndexValue d T₂ :=
pi_congr fun a => congrArg (ColorsIndex d) (congrFun h a)
@ -176,7 +176,7 @@ lemma ext' {T₁ T₂ : RealLorentzTensor d X} (h : T₁.color = T₂.color)
/-- An equivalence between `X → Fin 1 ⊕ Fin d` and `Y → Fin 1 ⊕ Fin d` given an isomorphism
between `X` and `Y`. -/
def congrSetIndexValue (d : ) (f : X ≃ Y) (i : X → Colors) :
IndexValue d i ≃ IndexValue d (i ∘ f.symm) :=
IndexValue d i ≃ IndexValue d (i ∘ f.symm) :=
Equiv.piCongrLeft' _ f
/-- Given an equivalence of indexing sets, a map on Lorentz tensors. -/
@ -259,7 +259,7 @@ def inrIndexValue {Tc : X → Colors} {Sc : Y → Colors}
def sumCommIndexValue {X Y : FintypeCat} (Tc : X → Colors) (Sc : Y → Colors) :
IndexValue d (sumElimIndexColor Tc Sc) ≃ IndexValue d (sumElimIndexColor Sc Tc) :=
(congrSetIndexValue d (sumCommFintypeCat X Y) (sumElimIndexColor Tc Sc)).trans
(castIndexValue ((sumElimIndexColor_symm Sc Tc).symm))
(castIndexValue ((sumElimIndexColor_symm Sc Tc).symm))
lemma sumCommIndexValue_inlIndexValue {X Y : FintypeCat} {Tc : X → Colors} {Sc : Y → Colors}
(i : IndexValue d (sumElimIndexColor Tc Sc)) :
@ -300,7 +300,7 @@ def unmarkedColor (T : Marked d X n) : X → Colors :=
T.color ∘ Sum.inl
/-- The colors of marked indices. -/
def markedColor (T : Marked d X n) : FintypeCat.of (Σ _ : Fin n, PUnit) → Colors :=
def markedColor (T : Marked d X n) : FintypeCat.of (Σ _ : Fin n, PUnit) → Colors :=
T.color ∘ Sum.inr
/-- The index values restricted to unmarked indices. -/