Merge pull request #31 from pitmonticone/master
Fix typos in docstrings
This commit is contained in:
commit
bdb960b7c3
12 changed files with 24 additions and 24 deletions
|
@ -77,7 +77,7 @@ namespace ACCSystemLinear
|
|||
structure LinSols (χ : ACCSystemLinear) where
|
||||
/-- The underlying charge. -/
|
||||
val : χ.1.charges
|
||||
/-- The condition that the charge satifies the linear ACCs. -/
|
||||
/-- The condition that the charge satisfies the linear ACCs. -/
|
||||
linearSol : ∀ i : Fin χ.numberLinear, χ.linearACCs i val = 0
|
||||
|
||||
/-- Two solutions are equal if the underlying charges are equal. -/
|
||||
|
@ -172,7 +172,7 @@ namespace ACCSystemQuad
|
|||
|
||||
/-- The type of solutions to the linear and quadratic ACCs. -/
|
||||
structure QuadSols (χ : ACCSystemQuad) extends χ.LinSols where
|
||||
/-- The condition that the charge satifies the quadratic ACCs. -/
|
||||
/-- The condition that the charge satisfies the quadratic ACCs. -/
|
||||
quadSol : ∀ i : Fin χ.numberQuadratic, (χ.quadraticACCs i) val = 0
|
||||
|
||||
/-- Two `QuadSols` are equal if the underlying charges are equal. -/
|
||||
|
@ -225,7 +225,7 @@ namespace ACCSystem
|
|||
|
||||
/-- The type of solutions to the anomaly cancellation conditions. -/
|
||||
structure Sols (χ : ACCSystem) extends χ.QuadSols where
|
||||
/-- The condition that the charge satifies the cubic ACC. -/
|
||||
/-- The condition that the charge satisfies the cubic ACC. -/
|
||||
cubicSol : χ.cubicACC val = 0
|
||||
|
||||
/-- Two solutions are equal if the underlying charges are equal. -/
|
||||
|
|
|
@ -206,7 +206,7 @@ lemma accSU2_ext {S T : MSSMCharges.charges}
|
|||
rw [hd, hu]
|
||||
rfl
|
||||
|
||||
/-- The anomaly cancelation condition for SU(3) anomaly. -/
|
||||
/-- The anomaly cancellation condition for SU(3) anomaly. -/
|
||||
@[simp]
|
||||
def accSU3 : MSSMCharges.charges →ₗ[ℚ] ℚ where
|
||||
toFun S := ∑ i, (2 * (Q S i) + (U S i) + (D S i))
|
||||
|
|
|
@ -8,7 +8,7 @@ import Mathlib.Tactic.Polyrith
|
|||
/-!
|
||||
# Hypercharge in MSSM.
|
||||
|
||||
Relavent definitions for the MSSM hypercharge.
|
||||
Relevant definitions for the MSSM hypercharge.
|
||||
|
||||
-/
|
||||
|
||||
|
|
|
@ -67,7 +67,7 @@ lemma line_in_cubic_P_P_P! {S : (PureU1 (2 * n.succ)).LinSols} (h : lineInCubic
|
|||
linear_combination 2 / 3 * (lineInCubic_expand h g f hS 1 1) -
|
||||
(lineInCubic_expand h g f hS 1 2) / 6
|
||||
|
||||
/-- We say a `LinSol` satifies `lineInCubicPerm` if all its permutations satsify `lineInCubic`. -/
|
||||
/-- We say a `LinSol` satisfies `lineInCubicPerm` if all its permutations satisfy `lineInCubic`. -/
|
||||
def lineInCubicPerm (S : (PureU1 (2 * n.succ)).LinSols) : Prop :=
|
||||
∀ (M : (FamilyPermutations (2 * n.succ)).group ),
|
||||
lineInCubic ((FamilyPermutations (2 * n.succ)).linSolRep M S)
|
||||
|
|
|
@ -12,7 +12,7 @@ import Mathlib.RepresentationTheory.Basic
|
|||
/-!
|
||||
# Line in plane condition
|
||||
|
||||
We say a `LinSol` satifies the `line in plane` condition if for all distinct `i1`, `i2`, `i3` in
|
||||
We say a `LinSol` satisfies the `line in plane` condition if for all distinct `i1`, `i2`, `i3` in
|
||||
`Fin n`, we have
|
||||
`S i1 = S i2` or `S i1 = - S i2` or `2 S i3 + S i1 + S i2 = 0`.
|
||||
|
||||
|
@ -21,7 +21,7 @@ The main reference for this material is
|
|||
|
||||
- https://arxiv.org/pdf/1912.04804.pdf
|
||||
|
||||
We will show that `n ≥ 4` the `line in plane` condition on solutions inplies the
|
||||
We will show that `n ≥ 4` the `line in plane` condition on solutions implies the
|
||||
`constAbs` condition.
|
||||
|
||||
-/
|
||||
|
|
|
@ -10,7 +10,7 @@ import Mathlib.Logic.Equiv.Fin
|
|||
/-!
|
||||
# Basis of `LinSols` in the odd case
|
||||
|
||||
We give a basis of `LinSols` in the odd case. This basis has the special propoerty
|
||||
We give a basis of `LinSols` in the odd case. This basis has the special property
|
||||
that splits into two planes on which every point is a solution to the ACCs.
|
||||
-/
|
||||
universe v u
|
||||
|
|
|
@ -16,10 +16,10 @@ import Mathlib.RepresentationTheory.Basic
|
|||
# Line In Cubic Odd case
|
||||
|
||||
We say that a linear solution satisfies the `lineInCubic` property
|
||||
if the line through that point and through the two different planes formed by the baisis of
|
||||
if the line through that point and through the two different planes formed by the basis of
|
||||
`LinSols` lies in the cubic.
|
||||
|
||||
We show that for a solution all its permutations satsfiy this property,
|
||||
We show that for a solution all its permutations satisfy this property,
|
||||
then the charge must be zero.
|
||||
|
||||
The main reference for this file is:
|
||||
|
@ -34,7 +34,7 @@ open BigOperators
|
|||
variable {n : ℕ}
|
||||
open VectorLikeOddPlane
|
||||
|
||||
/-- A property on `LinSols`, statified if every point on the line between the two planes
|
||||
/-- A property on `LinSols`, satisfied if every point on the line between the two planes
|
||||
in the basis through that point is in the cubic. -/
|
||||
def lineInCubic (S : (PureU1 (2 * n + 1)).LinSols) : Prop :=
|
||||
∀ (g f : Fin n → ℚ) (_ : S.val = Pa g f) (a b : ℚ) ,
|
||||
|
@ -64,7 +64,7 @@ lemma line_in_cubic_P_P_P! {S : (PureU1 (2 * n + 1)).LinSols} (h : lineInCubic S
|
|||
|
||||
|
||||
|
||||
/-- We say a `LinSol` satifies `lineInCubicPerm` if all its permutations satsify `lineInCubic`. -/
|
||||
/-- We say a `LinSol` satisfies `lineInCubicPerm` if all its permutations satisfy `lineInCubic`. -/
|
||||
def lineInCubicPerm (S : (PureU1 (2 * n + 1)).LinSols) : Prop :=
|
||||
∀ (M : (FamilyPermutations (2 * n + 1)).group ),
|
||||
lineInCubic ((FamilyPermutations (2 * n + 1)).linSolRep M S)
|
||||
|
|
|
@ -13,7 +13,7 @@ import Mathlib.Logic.Equiv.Fin
|
|||
/-!
|
||||
# Anomaly cancellation conditions for n family SM.
|
||||
|
||||
We define the ACC system for the Standard Model with`n`-famlies and no RHN.
|
||||
We define the ACC system for the Standard Model with`n`-families and no RHN.
|
||||
|
||||
-/
|
||||
|
||||
|
@ -21,7 +21,7 @@ universe v u
|
|||
open Nat
|
||||
open BigOperators
|
||||
|
||||
/-- Aassociate to each (including RHN) SM fermion a set of charges-/
|
||||
/-- Associate to each (including RHN) SM fermion a set of charges-/
|
||||
@[simps!]
|
||||
def SMCharges (n : ℕ) : ACCSystemCharges := ACCSystemChargesMk (5 * n)
|
||||
|
||||
|
|
|
@ -7,7 +7,7 @@ import HepLean.AnomalyCancellation.SM.Basic
|
|||
/-!
|
||||
# Family maps for the Standard Model ACCs
|
||||
|
||||
We define the a series of maps between the charges for different numbers of famlies.
|
||||
We define the a series of maps between the charges for different numbers of families.
|
||||
|
||||
-/
|
||||
|
||||
|
@ -84,7 +84,7 @@ other charges zero. -/
|
|||
def familyEmbedding (m n : ℕ) : (SMCharges m).charges →ₗ[ℚ] (SMCharges n).charges :=
|
||||
chargesMapOfSpeciesMap (speciesEmbed m n)
|
||||
|
||||
/-- For species, the embeddding of the `1`-family charges into the `n`-family charges in
|
||||
/-- For species, the embedding of the `1`-family charges into the `n`-family charges in
|
||||
a universal manor. -/
|
||||
@[simps!]
|
||||
def speciesFamilyUniversial (n : ℕ) :
|
||||
|
@ -98,7 +98,7 @@ def speciesFamilyUniversial (n : ℕ) :
|
|||
simp [HSMul.hSMul]
|
||||
--rw [show Rat.cast a = a from rfl]
|
||||
|
||||
/-- The embeddding of the `1`-family charges into the `n`-family charges in
|
||||
/-- The embedding of the `1`-family charges into the `n`-family charges in
|
||||
a universal manor. -/
|
||||
def familyUniversal ( n : ℕ) : (SMCharges 1).charges →ₗ[ℚ] (SMCharges n).charges :=
|
||||
chargesMapOfSpeciesMap (speciesFamilyUniversial n)
|
||||
|
|
|
@ -12,7 +12,7 @@ import HepLean.AnomalyCancellation.SM.NoGrav.One.LinearParameterization
|
|||
The main result of this file is the conclusion of this paper:
|
||||
https://arxiv.org/abs/1907.00514
|
||||
|
||||
That eveery solution to the ACCs without gravity satifies for free the gravitational anomaly.
|
||||
That eveery solution to the ACCs without gravity satisfies for free the gravitational anomaly.
|
||||
-/
|
||||
|
||||
universe v u
|
||||
|
@ -67,7 +67,7 @@ lemma accGrav_Q_neq_zero {S : (SMNoGrav 1).Sols} (hQ : Q S.val (0 : Fin 1) ≠ 0
|
|||
rw [← hS']
|
||||
exact S'.grav_of_cubic hC FLTThree
|
||||
|
||||
/-- Any solution to the ACCs without gravity satifies the gravitational ACC. -/
|
||||
/-- Any solution to the ACCs without gravity satisfies the gravitational ACC. -/
|
||||
theorem accGravSatisfied {S : (SMNoGrav 1).Sols} (FLTThree : FermatLastTheoremWith ℚ 3) :
|
||||
accGrav S.val = 0 := by
|
||||
by_cases hQ : Q S.val (0 : Fin 1)= 0
|
||||
|
|
|
@ -7,7 +7,7 @@ import HepLean.AnomalyCancellation.SMNu.Basic
|
|||
/-!
|
||||
# Family maps for the Standard Model for RHN ACCs
|
||||
|
||||
We define the a series of maps between the charges for different numbers of famlies.
|
||||
We define the a series of maps between the charges for different numbers of families.
|
||||
|
||||
-/
|
||||
|
||||
|
@ -89,7 +89,7 @@ other charges zero. -/
|
|||
def familyEmbedding (m n : ℕ) : (SMνCharges m).charges →ₗ[ℚ] (SMνCharges n).charges :=
|
||||
chargesMapOfSpeciesMap (speciesEmbed m n)
|
||||
|
||||
/-- For species, the embeddding of the `1`-family charges into the `n`-family charges in
|
||||
/-- For species, the embedding of the `1`-family charges into the `n`-family charges in
|
||||
a universal manor. -/
|
||||
@[simps!]
|
||||
def speciesFamilyUniversial (n : ℕ) :
|
||||
|
@ -103,7 +103,7 @@ def speciesFamilyUniversial (n : ℕ) :
|
|||
simp [HSMul.hSMul]
|
||||
-- rw [show Rat.cast a = a from rfl]
|
||||
|
||||
/-- The embeddding of the `1`-family charges into the `n`-family charges in
|
||||
/-- The embedding of the `1`-family charges into the `n`-family charges in
|
||||
a universal manor. -/
|
||||
def familyUniversal (n : ℕ) : (SMνCharges 1).charges →ₗ[ℚ] (SMνCharges n).charges :=
|
||||
chargesMapOfSpeciesMap (speciesFamilyUniversial n)
|
||||
|
|
|
@ -9,7 +9,7 @@ import Mathlib.Tactic.Polyrith
|
|||
/-!
|
||||
# The CKM Matrix
|
||||
|
||||
The definition of the type of CKM matries as unitary $3×3$-matries.
|
||||
The definition of the type of CKM matrices as unitary $3×3$-matrices.
|
||||
|
||||
An equivalence relation on CKM matrices is defined, where two matrices are equivalent if they are
|
||||
related by phase shifts.
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue