refactor: Spelling and typos

This commit is contained in:
jstoobysmith 2025-02-10 10:51:44 +00:00
parent b30a49d7db
commit dc5b63c4a7
25 changed files with 37 additions and 37 deletions

View file

@ -216,7 +216,7 @@ lemma basis_contr_pauliMatrix_basis_tree_expand' {n : } {c : Fin n → comple
rfl
/-- The map to color which appears when contracting a basis vector with
puali matrices. -/
Pauli matrices. -/
def pauliMatrixBasisProdMap
{n : } {c : Fin n → complexLorentzTensor.C}
(b : Π k, Fin (complexLorentzTensor.repDim (c k))) (i1 i2 i3 : Fin 4) :

View file

@ -43,7 +43,7 @@ def genBoostAux₁ (u v : FuturePointing d) : ContrMod d →ₗ[] ContrMod d
smul_tmul, tmul_smul, map_smul, smul_eq_mul, RingHom.id_apply]
rw [← mul_assoc, mul_comm 2 c, mul_assoc, mul_smul]
/-- An auxiliary linear map used in the definition of a genearlised boost. -/
/-- An auxiliary linear map used in the definition of a generalised boost. -/
def genBoostAux₂ (u v : FuturePointing d) : ContrMod d →ₗ[] ContrMod d where
toFun x := - (⟪x, u.1.1 + v.1.1⟫ₘ / (1 + ⟪u.1.1, v.1.1⟫ₘ)) • (u.1.1 + v.1.1)
map_add' x y := by

View file

@ -113,7 +113,7 @@ lemma orthchroMapReal_minus_one_or_one (Λ : LorentzGroup d) :
local notation "ℤ₂" => Multiplicative (ZMod 2)
/-- A continuous map from `lorentzGroup` to `ℤ₂` whose kernal are the Orthochronous elements. -/
/-- A continuous map from `lorentzGroup` to `ℤ₂` whose kernel are the Orthochronous elements. -/
def orthchroMap : C(LorentzGroup d, ℤ₂) :=
ContinuousMap.comp coeFor₂ {
toFun := fun Λ => ⟨orthchroMapReal Λ, orthchroMapReal_minus_one_or_one Λ⟩,

View file

@ -163,7 +163,7 @@ lemma dual_apply (μ ν : Fin 1 ⊕ Fin d) :
diagonal_mul, transpose_apply, diagonal_apply_eq]
/-- The components of the Minkowski dual of a matrix multiplied by the Minkowski matrix
in tems of the original matrix. -/
in terms of the original matrix. -/
lemma dual_apply_minkowskiMatrix (μ ν : Fin 1 ⊕ Fin d) :
dual Λ μ ν * η ν ν = η μ μ * Λ ν μ := by
rw [dual_apply, mul_assoc]

View file

@ -338,7 +338,7 @@ lemma σSAL_decomp (M : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
ring
/-- The component of a self-adjoint matrix in the direction `σ0` under
the basis formed by the covaraiant Pauli matrices. -/
the basis formed by the covariant Pauli matrices. -/
@[simp]
lemma σSAL_repr_inl_0 (M : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
σSAL.repr M (Sum.inl 0) = 1 / 2 * Matrix.trace (σ0 * M.1) := by
@ -355,7 +355,7 @@ lemma σSAL_repr_inl_0 (M : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
simp [σSAL]
/-- The component of a self-adjoint matrix in the direction `-σ1` under
the basis formed by the covaraiant Pauli matrices. -/
the basis formed by the covariant Pauli matrices. -/
@[simp]
lemma σSAL_repr_inr_0 (M : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
σSAL.repr M (Sum.inr 0) = - 1 / 2 * Matrix.trace (σ1 * M.1) := by
@ -372,7 +372,7 @@ lemma σSAL_repr_inr_0 (M : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
simp [σSAL]
/-- The component of a self-adjoint matrix in the direction `-σ2` under
the basis formed by the covaraiant Pauli matrices. -/
the basis formed by the covariant Pauli matrices. -/
@[simp]
lemma σSAL_repr_inr_1 (M : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
σSAL.repr M (Sum.inr 1) = - 1 / 2 * Matrix.trace (σ2 * M.1) := by
@ -389,7 +389,7 @@ lemma σSAL_repr_inr_1 (M : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
simp [σSAL]
/-- The component of a self-adjoint matrix in the direction `-σ3` under
the basis formed by the covaraiant Pauli matrices. -/
the basis formed by the covariant Pauli matrices. -/
@[simp]
lemma σSAL_repr_inr_2 (M : selfAdjoint (Matrix (Fin 2) (Fin 2) )) :
σSAL.repr M (Sum.inr 2) = - 1 / 2 * Matrix.trace (σ3 * M.1) := by