refactor: Spellings
This commit is contained in:
parent
4cd71f5ec6
commit
4a55351b72
33 changed files with 88 additions and 88 deletions
|
@ -79,7 +79,7 @@ lemma tensorNode_pauliContrDown : {pauliContrDown | μ α β}ᵀ.tensor =
|
|||
-/
|
||||
|
||||
set_option maxRecDepth 5000 in
|
||||
/-- A rearanging of `pauliCoDown` to place the pauli matrices on the right. -/
|
||||
/-- A rearranging of `pauliCoDown` to place the pauli matrices on the right. -/
|
||||
lemma pauliCoDown_eq_metric_mul_pauliCo :
|
||||
{pauliCoDown | μ α' β' = εL' | α α' ⊗ εR' | β β' ⊗ pauliCo | μ α β}ᵀ := by
|
||||
conv =>
|
||||
|
@ -128,7 +128,7 @@ lemma pauliCoDown_eq_metric_mul_pauliCo :
|
|||
decide
|
||||
|
||||
set_option maxRecDepth 5000 in
|
||||
/-- A rearanging of `pauliContrDown` to place the pauli matrices on the right. -/
|
||||
/-- A rearranging of `pauliContrDown` to place the pauli matrices on the right. -/
|
||||
lemma pauliContrDown_eq_metric_mul_pauliContr :
|
||||
{pauliContrDown | μ α' β' = εL' | α α' ⊗
|
||||
εR' | β β' ⊗ pauliContr | μ α β}ᵀ := by
|
||||
|
|
|
@ -38,7 +38,7 @@ lemma contrCoUnitVal_expand_tmul : contrCoUnitVal =
|
|||
rfl
|
||||
|
||||
/-- The contra-co unit for complex lorentz vectors as a morphism
|
||||
`𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexContr ⊗ complexCo`, manifesting the invaraince under
|
||||
`𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexContr ⊗ complexCo`, manifesting the invariance under
|
||||
the `SL(2, ℂ)` action. -/
|
||||
def contrCoUnit : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexContr ⊗ complexCo where
|
||||
hom := ModuleCat.ofHom {
|
||||
|
@ -88,7 +88,7 @@ lemma coContrUnitVal_expand_tmul : coContrUnitVal =
|
|||
rfl
|
||||
|
||||
/-- The co-contra unit for complex lorentz vectors as a morphism
|
||||
`𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexCo ⊗ complexContr`, manifesting the invaraince under
|
||||
`𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexCo ⊗ complexContr`, manifesting the invariance under
|
||||
the `SL(2, ℂ)` action. -/
|
||||
def coContrUnit : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexCo ⊗ complexContr where
|
||||
hom := ModuleCat.ofHom {
|
||||
|
|
|
@ -29,7 +29,7 @@ open Lorentz.Contr
|
|||
/-- A Lorentz transformation is `orthochronous` if its `0 0` element is non-negative. -/
|
||||
def IsOrthochronous : Prop := 0 ≤ Λ.1 (Sum.inl 0) (Sum.inl 0)
|
||||
|
||||
/-- A Lorentz transformation is `orthochronous` if and only if its fist column is
|
||||
/-- A Lorentz transformation is `orthochronous` if and only if its first column is
|
||||
future pointing. -/
|
||||
lemma IsOrthochronous_iff_futurePointing :
|
||||
IsOrthochronous Λ ↔ toNormOne Λ ∈ NormOne.FuturePointing d := by
|
||||
|
@ -66,7 +66,7 @@ def timeCompCont : C(LorentzGroup d, ℝ) := ⟨fun Λ => Λ.1 (Sum.inl 0) (Sum.
|
|||
|
||||
/-- An auxiliary function used in the definition of `orthchroMapReal`.
|
||||
This function takes all elements of `ℝ` less then `-1` to `-1`,
|
||||
all elements of `R` geater then `1` to `1` and peserves all other elements. -/
|
||||
all elements of `R` greater then `1` to `1` and preserves all other elements. -/
|
||||
def stepFunction : ℝ → ℝ := fun t =>
|
||||
if t ≤ -1 then -1 else
|
||||
if 1 ≤ t then 1 else t
|
||||
|
@ -152,7 +152,7 @@ lemma mul_othchron_of_not_othchron_not_othchron {Λ Λ' : LorentzGroup d} (h :
|
|||
exact NormOne.FuturePointing.metric_reflect_not_mem_not_mem h h'
|
||||
|
||||
/-- The product of an orthochronous Lorentz transformations with a
|
||||
non-orthochronous Loentz transformation is not orthochronous. -/
|
||||
non-orthochronous Lorentz transformation is not orthochronous. -/
|
||||
lemma mul_not_othchron_of_othchron_not_othchron {Λ Λ' : LorentzGroup d} (h : IsOrthochronous Λ)
|
||||
(h' : ¬ IsOrthochronous Λ') : ¬ IsOrthochronous (Λ * Λ') := by
|
||||
rw [not_orthochronous_iff_le_zero, LorentzGroup.inl_inl_mul]
|
||||
|
@ -161,7 +161,7 @@ lemma mul_not_othchron_of_othchron_not_othchron {Λ Λ' : LorentzGroup d} (h : I
|
|||
exact NormOne.FuturePointing.metric_reflect_mem_not_mem h h'
|
||||
|
||||
/-- The product of a non-orthochronous Lorentz transformations with an
|
||||
orthochronous Loentz transformation is not orthochronous. -/
|
||||
orthochronous Lorentz transformation is not orthochronous. -/
|
||||
lemma mul_not_othchron_of_not_othchron_othchron {Λ Λ' : LorentzGroup d} (h : ¬ IsOrthochronous Λ)
|
||||
(h' : IsOrthochronous Λ') : ¬ IsOrthochronous (Λ * Λ') := by
|
||||
rw [not_orthochronous_iff_le_zero, LorentzGroup.inl_inl_mul]
|
||||
|
|
|
@ -29,7 +29,7 @@ variable {d : ℕ}
|
|||
/-- Notation for `minkowskiMatrix`. -/
|
||||
scoped[minkowskiMatrix] notation "η" => minkowskiMatrix
|
||||
|
||||
/-- The Minkowski matrix is self-inveting. -/
|
||||
/-- The Minkowski matrix is self-inverting. -/
|
||||
@[simp]
|
||||
lemma sq : @minkowskiMatrix d * minkowskiMatrix = 1 := by
|
||||
simp only [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, diagonal_mul_diagonal]
|
||||
|
@ -55,8 +55,8 @@ lemma sq : @minkowskiMatrix d * minkowskiMatrix = 1 := by
|
|||
lemma eq_transpose : minkowskiMatrixᵀ = @minkowskiMatrix d := by
|
||||
simp only [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, diagonal_transpose]
|
||||
|
||||
/-- The deteminant of the Minkowski matrix is equal to `-1` to the power
|
||||
of the number of spactial dimensions. -/
|
||||
/-- The determinant of the Minkowski matrix is equal to `-1` to the power
|
||||
of the number of spatial dimensions. -/
|
||||
@[simp]
|
||||
lemma det_eq_neg_one_pow_d : (@minkowskiMatrix d).det = (- 1) ^ d := by
|
||||
simp [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
|
||||
|
@ -68,7 +68,7 @@ lemma η_apply_mul_η_apply_diag (μ : Fin 1 ⊕ Fin d) : η μ μ * η μ μ =
|
|||
| Sum.inl _ => simp [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
|
||||
| Sum.inr _ => simp [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
|
||||
|
||||
/-- The Minkowski matix as a block matrix. -/
|
||||
/-- The Minkowski matrix as a block matrix. -/
|
||||
lemma as_block : @minkowskiMatrix d =
|
||||
Matrix.fromBlocks (1 : Matrix (Fin 1) (Fin 1) ℝ) 0 0 (-1 : Matrix (Fin d) (Fin d) ℝ) := by
|
||||
rw [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, ← fromBlocks_diagonal]
|
||||
|
|
|
@ -301,7 +301,7 @@ lemma σSAL_span : ⊤ ≤ Submodule.span ℝ (Set.range σSAL') := by
|
|||
def σSAL : Basis (Fin 1 ⊕ Fin 3) ℝ (selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ)) :=
|
||||
Basis.mk σSAL_linearly_independent σSAL_span
|
||||
|
||||
/-- The decomposition of a self-adjint matrix into the Pauli matrices (where `σi` are negated). -/
|
||||
/-- The decomposition of a self-adjoint matrix into the Pauli matrices (where `σi` are negated). -/
|
||||
lemma σSAL_decomp (M : selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ)) :
|
||||
M = (1/2 * (Matrix.trace (σ0 * M.1)).re) • σSAL (Sum.inl 0)
|
||||
+ (-1/2 * (Matrix.trace (σ1 * M.1)).re) • σSAL (Sum.inr 0)
|
||||
|
|
|
@ -175,7 +175,7 @@ lemma toSelfAdjointMap_det_one' {M : ℂ²ˣ²} (hM : M.IsUpperTriangular) (detM
|
|||
_ = 1 := by rw [hE]; simp [detD_one, detA_one]
|
||||
|
||||
/-- This promotes `Lorentz.SL2C.toSelfAdjointMap M` and its definitional equivalence,
|
||||
`Lorentz.SL2C.toSelfAdjointMap' M`, to a linear equivalence by recognising the linear inverse to be
|
||||
`Lorentz.SL2C.toSelfAdjointMap' M`, to a linear equivalence by recognizing the linear inverse to be
|
||||
`Lorentz.SL2C.toSelfAdjointMap M⁻¹`, i.e., `Lorentz.SL2C.toSelfAdjointMap' M⁻¹`. -/
|
||||
noncomputable def toSelfAdjointEquiv (M : ℂ²ˣ²) [Invertible M] : ℍ₂ ≃ₗ[ℝ] ℍ₂ where
|
||||
toLinearMap := toSelfAdjointMap' M
|
||||
|
|
|
@ -174,11 +174,11 @@ lemma altLeftContraction_basis (i j : Fin 2) :
|
|||
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)
|
||||
|
||||
/--
|
||||
The linear map from rightHandedWeyl ⊗ altRightHandedWeyl to ℂ given by
|
||||
summing over components of rightHandedWeyl and altRightHandedWeyl in the
|
||||
The linear map from `rightHandedWeyl ⊗ altRightHandedWeyl` to `ℂ` given by
|
||||
summing over components of `rightHandedWeyl` and `altRightHandedWeyl` in the
|
||||
standard basis (i.e. the dot product).
|
||||
The contraction of a right-handed Weyl fermion with a left-handed Weyl fermion.
|
||||
In index notation this is ψ^{dot a} φ_{dot a}.
|
||||
In index notation this is `ψ^{dot a} φ_{dot a}`.
|
||||
-/
|
||||
def rightAltContraction : rightHanded ⊗ altRightHanded ⟶ 𝟙_ (Rep ℂ SL(2,ℂ)) where
|
||||
hom := ModuleCat.ofHom <| TensorProduct.lift rightAltBi
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue