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

@ -63,7 +63,7 @@ def repCharges {n : } : Representation (PermGroup n) (SMCharges n).Charge
erw [toSMSpecies_toSpecies_inv]
rfl
/-- The species chages of a set of charges acted on by a family permutation is the permutation
/-- The species charges of a set of charges acted on by a family permutation is the permutation
of those species charges with the corresponding part of the family permutation. -/
lemma repCharges_toSpecies (f : PermGroup n) (S : (SMCharges n).Charges) (j : Fin 5) :
toSpecies j (repCharges f S) = toSpecies j S ∘ f⁻¹ j := by
@ -78,7 +78,7 @@ lemma toSpecies_sum_invariant (m : ) (f : PermGroup n) (S : (SMCharges n).Cha
exact Fintype.sum_equiv (f⁻¹ j) (fun x => ((fun a => a ^ m) ∘ (toSpecies j) S ∘ ⇑(f⁻¹ j)) x)
(fun x => ((fun a => a ^ m) ∘ (toSpecies j) S) x) (congrFun rfl)
/-- The gravitional anomaly equations is invariant under family permutations. -/
/-- The gravitational anomaly equations is invariant under family permutations. -/
lemma accGrav_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
accGrav (repCharges f S) = accGrav S := accGrav_ext
(by simpa using toSpecies_sum_invariant 1 f S)

View file

@ -38,7 +38,7 @@ informal_lemma inclSM_ker where
deps := [``inclSM, ``StandardModel.gaugeGroup₆SubGroup]
/-- The group embedding from `StandardModel.GaugeGroup₆` to `GaugeGroupI` induced by `inclSM` by
quotienting by the kernal `inclSM_ker`.
quotienting by the kernel `inclSM_ker`.
-/
informal_definition embedSM₆ where
deps := [``inclSM, ``StandardModel.GaugeGroup₆, ``GaugeGroupI, ``inclSM_ker]

View file

@ -43,7 +43,7 @@ informal_lemma inclSM_ker where
deps := [``inclSM, ``StandardModel.gaugeGroup₃SubGroup]
/-- The group embedding from `StandardModel.GaugeGroup₃` to `GaugeGroupI` induced by `inclSM` by
quotienting by the kernal `inclSM_ker`.
quotienting by the kernel `inclSM_ker`.
-/
informal_definition embedSM₃ where
deps := [``inclSM, ``StandardModel.GaugeGroup₃, ``GaugeGroupI, ``inclSM_ker]

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

View file

@ -23,7 +23,7 @@ lemma insertionSortMin_lt_length_succ {α : Type} (r : αα → Prop) [Deci
rw [eraseIdx_length']
simp
/-- Given a list `i :: l` the left-most minimial position `a` of `i :: l` wrt `r`
/-- Given a list `i :: l` the left-most minimal position `a` of `i :: l` wrt `r`
as an element of `Fin (insertionSortDropMinPos r i l).length.succ`. -/
def insertionSortMinPosFin {α : Type} (r : αα → Prop) [DecidableRel r] (i : α) (l : List α) :
Fin (insertionSortDropMinPos r i l).length.succ :=

View file

@ -25,7 +25,7 @@ def RemarkInfo.toFullName (r : RemarkInfo) : Name :=
else
r.name
/-- A Bool which is true if a name correponds to a remark. -/
/-- A Bool which is true if a name corresponds to a remark. -/
def RemarkInfo.IsRemark (n : Name) : m Bool := do
let allRemarks ← allRemarkInfo
let r := allRemarks.find? fun r => r.toFullName == n

View file

@ -12,7 +12,7 @@ import Mathlib.Algebra.BigOperators.Group.Finset
/-- The type `CreateAnnihilate` is the type containing two elements `create` and `annihilate`.
This type is used to specify if an operator is a creation, or annihilation, operator
or the sum thereof or intergral thereover etc. -/
or the sum thereof or integral thereover etc. -/
inductive CreateAnnihilate where
| create : CreateAnnihilate
| annihilate : CreateAnnihilate

View file

@ -83,17 +83,17 @@ def euclidInner : F.HalfEdgeMomenta →ₗ[] F.HalfEdgeMomenta →ₗ[]
Corresponding to that spanned by its total outflowing momentum. -/
def EdgeMomenta : Type := F.𝓔
/-- The edge momenta form an additive commuative group. -/
/-- The edge momenta form an additive commutative group. -/
instance : AddCommGroup F.EdgeMomenta := Pi.addCommGroup
/-- The edge momenta form a module over ``. -/
instance : Module F.EdgeMomenta := Pi.module _ _ _
/-- The type which associates to each ege a `1`-dimensional vector space.
/-- The type which associates to each edge a `1`-dimensional vector space.
Corresponding to that spanned by its total inflowing momentum. -/
def VertexMomenta : Type := F.𝓥
/-- The vertex momenta carries the structure of an additive commuative group. -/
/-- The vertex momenta carries the structure of an additive commutative group. -/
instance : AddCommGroup F.VertexMomenta := Pi.addCommGroup
/-- The vertex momenta carries the structure of a module over ``. -/
@ -106,7 +106,7 @@ def EdgeVertexMomentaMap : Fin 2 → Type := fun i =>
| 1 => F.VertexMomenta
/-- The target of the map `EdgeVertexMomentaMap` is either the type of edge momenta
or vertex momenta and thus carries the structure of an additive commuative group. -/
or vertex momenta and thus carries the structure of an additive commutative group. -/
instance (i : Fin 2) : AddCommGroup (EdgeVertexMomentaMap F i) :=
match i with
| 0 => instAddCommGroupEdgeMomenta F

View file

@ -40,7 +40,7 @@ def fieldOpIdealSet : Set (FieldOpFreeAlgebra 𝓕) :=
This corresponds to the condition that two annihilation operators always super-commute.
- `[ofCrAnOpF φ, ofCrAnOpF φ']ₛca` for `φ` and `φ'` operators with different statistics.
This corresponds to the condition that two operators with different statistics always
super-commute. In otherwords, fermions and bosons always super-commute.
super-commute. In other words, fermions and bosons always super-commute.
- `[ofCrAnOpF φ1, [ofCrAnOpF φ2, ofCrAnOpF φ3]ₛca]ₛca`. This corresponds to the condition,
when combined with the conditions above, that the super-commutor is in the center of the
of the algebra.
@ -218,7 +218,7 @@ lemma ι_superCommuteF_ofCrAnOpF_ofCrAnOpF_mem_center (φ ψ : 𝓕.CrAnFieldOp)
/-!
## The kernal of ι
## The kernel of ι
-/
lemma ι_eq_zero_iff_mem_ideal (x : FieldOpFreeAlgebra 𝓕) :

View file

@ -282,7 +282,7 @@ lemma ofCrAnOp_superCommute_normalOrder_ofFieldOpList_sum (φ : 𝓕.CrAnFieldOp
/--
The commutor of the annihilation part of a field operator with a normal ordered list of field
operators can be decomponsed into the sum of the commutators of the annihilation part with each
operators can be decomposed into the sum of the commutators of the annihilation part with each
element of the list of field operators, i.e.
`[anPart φ, 𝓝(φ₀…φₙ)]ₛ= ∑ i, 𝓢(φ, φ₀…φᵢ₋₁) • [anPart φ, φᵢ]ₛ * 𝓝(φ₀…φᵢ₋₁φᵢ₊₁…φₙ)`.
-/

View file

@ -104,7 +104,7 @@ is equal the product of
- `φsΛ.timeContract`
- `s • [anPart φ, ofFieldOp φs[k]]ₛ` where `s` is the sign associated with moving `φ` through
uncontracted fields in `φ₀…φₖ₋₁`
- the normal ordering `[φsΛ]ᵘᶜ` with the field corresonding to `k` removed.
- the normal ordering `[φsΛ]ᵘᶜ` with the field corresponding to `k` removed.
The proof of this result relies on
- `timeContract_insert_some_of_not_lt`

View file

@ -34,7 +34,7 @@ This result follows from
those `𝓣(φsΛ.staticWickTerm)` for which `φsΛ` has a contracted pair which are not
equal time to zero.
- `staticContract_eq_timeContract_of_eqTimeOnly` to rewrite the static contract
in the reminaing `𝓣(φsΛ.staticWickTerm)` as a time contract.
in the remaining `𝓣(φsΛ.staticWickTerm)` as a time contract.
- `timeOrder_timeContract_mul_of_eqTimeOnly_left` to move the time contracts out of the time
ordering.
-/

View file

@ -88,7 +88,7 @@ As some intuition, if `f` corresponds to a Weyl-fermion field, then
- `position f e x`, `e` would correspond to a Lorentz index `α`, and `position f e x` would,
once represented in the operator algebra, be proportional to the operator
`∑ s, ∫ d^3p/(…) (x_α(p,s) a(p, s) e^{-i p x} + y_α(p,s) a^†(p, s) e^{-i p x})`.
- `outAsymp f e p`, `e` would corresond to a spin `s`, and `outAsymp f e p` would,
- `outAsymp f e p`, `e` would correspond to a spin `s`, and `outAsymp f e p` would,
once represented in the operator algebra, be proportional to the
annihilation operator `a^†(p, s)`.

View file

@ -83,7 +83,7 @@ As some intuition, if `f` corresponds to a Weyl-fermion field, it would contribu
- an element corresponding to the creation parts of position operators for each each Lorentz
index `α`:
`∑ s, ∫ d^3p/(…) (x_α(p,s) a(p, s) e^{-i p x})`.
- an element corresponding to anihilation parts of position operator,
- an element corresponding to annihilation parts of position operator,
for each each Lorentz index `α`:
`∑ s, ∫ d^3p/(…) (y_α(p,s) a^†(p, s) e^{-i p x})`.
- an element corresponding to outgoing asymptotic operators for each spin `s`: `a^†(p, s)`.
@ -98,7 +98,7 @@ def crAnFieldOpToFieldOp : 𝓕.CrAnFieldOp → 𝓕.FieldOp := Sigma.fst
lemma crAnFieldOpToFieldOp_prod (s : 𝓕.FieldOp) (t : 𝓕.fieldOpToCrAnType s) :
𝓕.crAnFieldOpToFieldOp ⟨s, t⟩ = s := rfl
/-- For a field specficiation `𝓕`, `𝓕.crAnFieldOpToCreateAnnihilate` is the map from
/-- For a field specification `𝓕`, `𝓕.crAnFieldOpToCreateAnnihilate` is the map from
`𝓕.CrAnFieldOp` to `CreateAnnihilate` taking `φ` to `create` if
- `φ` corresponds to an incoming asymptotic field operator or the creation part of a position based
field operator.

View file

@ -126,7 +126,7 @@ def singletonEquiv {φ : 𝓕.FieldOp} : CrAnSection [φ] ≃
simp only [head]
rfl
/-- An equivalence seperating the head of a creation and annihilation section
/-- An equivalence separating the head of a creation and annihilation section
from the tail. -/
def consEquiv {φ : 𝓕.FieldOp} {φs : List 𝓕.FieldOp} : CrAnSection (φ :: φs) ≃
𝓕.fieldOpToCrAnType φ × CrAnSection φs where

View file

@ -32,7 +32,7 @@ def timeOrderRel : 𝓕.FieldOp → 𝓕.FieldOp → Prop
| FieldOp.inAsymp _, FieldOp.position _ => False
| FieldOp.inAsymp _, FieldOp.inAsymp _ => True
/-- The relation `timeOrderRel` is decidable, but not computablly so due to
/-- The relation `timeOrderRel` is decidable, but not computable so due to
`Real.decidableLE`. -/
noncomputable instance : (φ φ' : 𝓕.FieldOp) → Decidable (timeOrderRel φ φ')
| FieldOp.outAsymp _, _ => isTrue True.intro
@ -206,7 +206,7 @@ it is needed that the operator with the greatest time is to the left.
-/
def crAnTimeOrderRel (a b : 𝓕.CrAnFieldOp) : Prop := 𝓕.timeOrderRel a.1 b.1
/-- The relation `crAnTimeOrderRel` is decidable, but not computablly so due to
/-- The relation `crAnTimeOrderRel` is decidable, but not computable so due to
`Real.decidableLE`. -/
noncomputable instance (φ φ' : 𝓕.CrAnFieldOp) : Decidable (crAnTimeOrderRel φ φ') :=
inferInstanceAs (Decidable (𝓕.timeOrderRel φ.1 φ'.1))
@ -508,7 +508,7 @@ lemma sum_crAnSections_timeOrder {φs : List 𝓕.FieldOp} [AddCommMonoid M]
def normTimeOrderRel (a b : 𝓕.CrAnFieldOp) : Prop :=
crAnTimeOrderRel a b ∧ (crAnTimeOrderRel b a → normalOrderRel a b)
/-- The relation `normTimeOrderRel` is decidable, but not computablly so due to
/-- The relation `normTimeOrderRel` is decidable, but not computable so due to
`Real.decidableLE`. -/
noncomputable instance (φ φ' : 𝓕.CrAnFieldOp) : Decidable (normTimeOrderRel φ φ') :=
instDecidableAnd

View file

@ -26,7 +26,7 @@ namespace FieldStatistic
variable {𝓕 : Type}
/-- The type `FieldStatistic` carries an instance of a commuative group in which
/-- The type `FieldStatistic` carries an instance of a commutative group in which
- `bosonic * bosonic = bosonic`
- `bosonic * fermionic = fermionic`
- `fermionic * bosonic = fermionic`

View file

@ -21,7 +21,7 @@ open FieldStatistic
/-- Given a Wick contraction `c : WickContraction n` and `i1 i2 : Fin n` the finite set
of elements of `Fin n` between `i1` and `i2` which are either uncontracted
or are contracted but are contracted with an element occuring after `i1`.
or are contracted but are contracted with an element occurring after `i1`.
I.e. the elements of `Fin n` between `i1` and `i2` which are not contracted with before `i1`.
One should assume `i1 < i2` otherwise this finite set is empty. -/
def signFinset (c : WickContraction n) (i1 i2 : Fin n) : Finset (Fin n) :=

View file

@ -506,7 +506,7 @@ lemma hasEqTimeEquiv_ext_sigma {φs : List 𝓕.FieldOp} {x1 x2 :
simp only [ne_eq, congr_refl] at h2
simp [h2]
/-- The equivalence which seperates a Wick contraction which has an equal time contraction
/-- The equivalence which separates a Wick contraction which has an equal time contraction
into a non-empty contraction only between equal-time fields and a Wick contraction which
does not have equal time contractions. -/
def hasEqTimeEquiv (φs : List 𝓕.FieldOp) :

View file

@ -43,7 +43,7 @@ noncomputable def higgsRepUnitary : GaugeGroupI →* unitaryGroup (Fin 2) wh
repeat rw [mul_assoc]
map_one' := by simp
/-- Using the orthonormal basis of `HiggsVec`, turns a `2×2`-matrix intoa a linear map
/-- Using the orthonormal basis of `HiggsVec`, turns a `2×2`-matrix into a linear map
of `HiggsVec`. -/
noncomputable def matrixToLin : Matrix (Fin 2) (Fin 2) →* (HiggsVec →L[] HiggsVec) where
toFun g := LinearMap.toContinuousLinearMap
@ -220,7 +220,7 @@ informal_lemma guage_orbit where
deps := [``rotate_fst_zero_snd_real]
/-- The Higgs boson breaks electroweak symmetry down to the electromagnetic force, i.e., the
stablity group of the action of `rep` on `![0, Complex.ofReal ‖φ‖]`, for non-zero `‖φ‖`, is the
stability group of the action of `rep` on `![0, Complex.ofReal ‖φ‖]`, for non-zero `‖φ‖`, is the
`SU(3) × U(1)` subgroup of `gaugeGroup := SU(3) × SU(2) × U(1)` with the embedding given by
`(g, e^{i θ}) ↦ (g, diag (e ^ {3 * i θ}, e ^ {- 3 * i θ}), e^{i θ})`.
-/

View file

@ -77,7 +77,7 @@ syntax num : indexExpr
/-- Notation to describe the jiggle of a tensor index. -/
syntax "τ(" ident ")" : indexExpr
/-- Bool which is ture if an index is a num. -/
/-- Bool which is true if an index is a num. -/
def indexExprIsNum (stx : Syntax) : Bool :=
match stx with
| `(indexExpr|$_:num) => true