Merge pull request #67 from HEPLean/Update-versions

refactor: Change case of Type and Prop
This commit is contained in:
Joseph Tooby-Smith 2024-06-26 12:40:47 -04:00 committed by GitHub
commit d2c89d8cc2
No known key found for this signature in database
GPG key ID: B5690EEEBB952194
59 changed files with 750 additions and 696 deletions

View file

@ -34,14 +34,14 @@ def ACCSystemChargesMk (n : ) : ACCSystemCharges where
namespace ACCSystemCharges namespace ACCSystemCharges
/-- The charges as functions from `Fin χ.numberCharges → `. -/ /-- The charges as functions from `Fin χ.numberCharges → `. -/
def charges (χ : ACCSystemCharges) : Type := Fin χ.numberCharges → def Charges (χ : ACCSystemCharges) : Type := Fin χ.numberCharges →
/-- /--
An instance to provide the necessary operations and properties for `charges` to form an additive An instance to provide the necessary operations and properties for `charges` to form an additive
commutative monoid. commutative monoid.
-/ -/
@[simps!] @[simps!]
instance chargesAddCommMonoid (χ : ACCSystemCharges) : AddCommMonoid χ.charges := instance chargesAddCommMonoid (χ : ACCSystemCharges) : AddCommMonoid χ.Charges :=
Pi.addCommMonoid Pi.addCommMonoid
/-- /--
@ -49,17 +49,17 @@ instance chargesAddCommMonoid (χ : ACCSystemCharges) : AddCommMonoid χ.charges
the field ``. the field ``.
-/ -/
@[simps!] @[simps!]
instance chargesModule (χ : ACCSystemCharges) : Module χ.charges := instance chargesModule (χ : ACCSystemCharges) : Module χ.Charges :=
Pi.module _ _ _ Pi.module _ _ _
/-- /--
An instance provides the necessary operations and properties for `charges` to form an additive An instance provides the necessary operations and properties for `charges` to form an additive
commutative group. commutative group.
-/ -/
instance ChargesAddCommGroup (χ : ACCSystemCharges) : AddCommGroup χ.charges := instance ChargesAddCommGroup (χ : ACCSystemCharges) : AddCommGroup χ.Charges :=
Module.addCommMonoidToAddCommGroup Module.addCommMonoidToAddCommGroup
instance (χ : ACCSystemCharges) : Module.Finite χ.charges := instance (χ : ACCSystemCharges) : Module.Finite χ.Charges :=
FiniteDimensional.finiteDimensional_pi FiniteDimensional.finiteDimensional_pi
end ACCSystemCharges end ACCSystemCharges
@ -69,14 +69,14 @@ structure ACCSystemLinear extends ACCSystemCharges where
/-- The number of linear ACCs. -/ /-- The number of linear ACCs. -/
numberLinear : numberLinear :
/-- The linear ACCs. -/ /-- The linear ACCs. -/
linearACCs : Fin numberLinear → (toACCSystemCharges.charges →ₗ[] ) linearACCs : Fin numberLinear → (toACCSystemCharges.Charges →ₗ[] )
namespace ACCSystemLinear namespace ACCSystemLinear
/-- The type of solutions to the linear ACCs. -/ /-- The type of solutions to the linear ACCs. -/
structure LinSols (χ : ACCSystemLinear) where structure LinSols (χ : ACCSystemLinear) where
/-- The underlying charge. -/ /-- The underlying charge. -/
val : χ.1.charges val : χ.1.Charges
/-- The condition that the charge satisfies the linear ACCs. -/ /-- The condition that the charge satisfies the linear ACCs. -/
linearSol : ∀ i : Fin χ.numberLinear, χ.linearACCs i val = 0 linearSol : ∀ i : Fin χ.numberLinear, χ.linearACCs i val = 0
@ -154,7 +154,7 @@ instance linSolsAddCommGroup (χ : ACCSystemLinear) : AddCommGroup χ.LinSols :=
Module.addCommMonoidToAddCommGroup Module.addCommMonoidToAddCommGroup
/-- The inclusion of `LinSols` into `charges`. -/ /-- The inclusion of `LinSols` into `charges`. -/
def linSolsIncl (χ : ACCSystemLinear) : χ.LinSols →ₗ[] χ.charges where def linSolsIncl (χ : ACCSystemLinear) : χ.LinSols →ₗ[] χ.Charges where
toFun S := S.val toFun S := S.val
map_add' _ _ := rfl map_add' _ _ := rfl
map_smul' _ _ := rfl map_smul' _ _ := rfl
@ -166,7 +166,7 @@ structure ACCSystemQuad extends ACCSystemLinear where
/-- The number of quadratic ACCs. -/ /-- The number of quadratic ACCs. -/
numberQuadratic : numberQuadratic :
/-- The quadratic ACCs. -/ /-- The quadratic ACCs. -/
quadraticACCs : Fin numberQuadratic → HomogeneousQuadratic toACCSystemCharges.charges quadraticACCs : Fin numberQuadratic → HomogeneousQuadratic toACCSystemCharges.Charges
namespace ACCSystemQuad namespace ACCSystemQuad
@ -211,7 +211,7 @@ def linSolsInclQuadSolsZero (χ : ACCSystemQuad) (h : χ.numberQuadratic = 0) :
map_smul' _ _ := rfl map_smul' _ _ := rfl
/-- The inclusion of quadratic solutions into all charges. -/ /-- The inclusion of quadratic solutions into all charges. -/
def quadSolsIncl (χ : ACCSystemQuad) : χ.QuadSols →[] χ.charges := def quadSolsIncl (χ : ACCSystemQuad) : χ.QuadSols →[] χ.Charges :=
MulActionHom.comp χ.linSolsIncl.toMulActionHom χ.quadSolsInclLinSols MulActionHom.comp χ.linSolsIncl.toMulActionHom χ.quadSolsInclLinSols
end ACCSystemQuad end ACCSystemQuad
@ -219,7 +219,7 @@ end ACCSystemQuad
/-- The type of charges plus the anomaly cancellation conditions. -/ /-- The type of charges plus the anomaly cancellation conditions. -/
structure ACCSystem extends ACCSystemQuad where structure ACCSystem extends ACCSystemQuad where
/-- The cubic ACC. -/ /-- The cubic ACC. -/
cubicACC : HomogeneousCubic toACCSystemCharges.charges cubicACC : HomogeneousCubic toACCSystemCharges.Charges
namespace ACCSystem namespace ACCSystem
@ -236,7 +236,7 @@ lemma Sols.ext {χ : ACCSystem} {S T : χ.Sols} (h : S.val = T.val) :
simp_all only simp_all only
/-- We say a charge S is a solution if it extends to a solution. -/ /-- We say a charge S is a solution if it extends to a solution. -/
def isSolution (χ : ACCSystem) (S : χ.charges) : Prop := def IsSolution (χ : ACCSystem) (S : χ.Charges) : Prop :=
∃ (sol : χ.Sols), sol.val = S ∃ (sol : χ.Sols), sol.val = S
/-- An instance giving the properties and structures to define an action of `` on `Sols`. -/ /-- An instance giving the properties and structures to define an action of `` on `Sols`. -/
@ -262,13 +262,13 @@ def solsInclLinSols (χ : ACCSystem) : χ.Sols →[] χ.LinSols :=
MulActionHom.comp χ.quadSolsInclLinSols χ.solsInclQuadSols MulActionHom.comp χ.quadSolsInclLinSols χ.solsInclQuadSols
/-- The inclusion of `Sols` into `LinSols`. -/ /-- The inclusion of `Sols` into `LinSols`. -/
def solsIncl (χ : ACCSystem) : χ.Sols →[] χ.charges := def solsIncl (χ : ACCSystem) : χ.Sols →[] χ.Charges :=
MulActionHom.comp χ.quadSolsIncl χ.solsInclQuadSols MulActionHom.comp χ.quadSolsIncl χ.solsInclQuadSols
/-- The structure of a map between two ACCSystems. -/ /-- The structure of a map between two ACCSystems. -/
structure Hom (χ η : ACCSystem) where structure Hom (χ η : ACCSystem) where
/-- The linear map between vector spaces of charges. -/ /-- The linear map between vector spaces of charges. -/
charges : χ.charges →ₗ[] η.charges charges : χ.Charges →ₗ[] η.Charges
/-- The map between solutions. -/ /-- The map between solutions. -/
anomalyFree : χ.Sols → η.Sols anomalyFree : χ.Sols → η.Sols
/-- The condition that the map commutes with the relevant inclusions. -/ /-- The condition that the map commutes with the relevant inclusions. -/

View file

@ -27,7 +27,7 @@ structure ACCSystemGroupAction (χ : ACCSystem) where
/-- An instance given group the structure of a group. -/ /-- An instance given group the structure of a group. -/
groupInst : Group group groupInst : Group group
/-- The representation of group acting on the vector space of charges. -/ /-- The representation of group acting on the vector space of charges. -/
rep : Representation group χ.charges rep : Representation group χ.Charges
/-- The invariance of the linear ACCs under the group action. -/ /-- The invariance of the linear ACCs under the group action. -/
linearInvariant : ∀ i g S, χ.linearACCs i (rep g S) = χ.linearACCs i S linearInvariant : ∀ i g S, χ.linearACCs i (rep g S) = χ.linearACCs i S
/-- The invariance of the quadratic ACCs under the group action. -/ /-- The invariance of the quadratic ACCs under the group action. -/

View file

@ -26,7 +26,7 @@ open BigOperators
/-- `B₃` is the charge which is $B-L$ in all families, but with the third /-- `B₃` is the charge which is $B-L$ in all families, but with the third
family of the opposite sign. -/ family of the opposite sign. -/
def B₃AsCharge : MSSMACC.charges := toSpecies.symm def B₃AsCharge : MSSMACC.Charges := toSpecies.symm
⟨fun s => fun i => ⟨fun s => fun i =>
match s, i with match s, i with
| 0, 0 => 1 | 0, 0 => 1

View file

@ -37,7 +37,7 @@ namespace MSSMCharges
`(Fin 18 ⊕ Fin 2 → )`. The first 18 factors corresponds to the SM fermions, while the last two `(Fin 18 ⊕ Fin 2 → )`. The first 18 factors corresponds to the SM fermions, while the last two
are the higgsions. -/ are the higgsions. -/
@[simps!] @[simps!]
def toSMPlusH : MSSMCharges.charges ≃ (Fin 18 ⊕ Fin 2 → ) := def toSMPlusH : MSSMCharges.Charges ≃ (Fin 18 ⊕ Fin 2 → ) :=
((@finSumFinEquiv 18 2).arrowCongr (Equiv.refl )).symm ((@finSumFinEquiv 18 2).arrowCongr (Equiv.refl )).symm
/-- An equivalence between `Fin 18 ⊕ Fin 2 → ` and `(Fin 18 → ) × (Fin 2 → )`. -/ /-- An equivalence between `Fin 18 ⊕ Fin 2 → ` and `(Fin 18 → ) × (Fin 2 → )`. -/
@ -51,7 +51,7 @@ def splitSMPlusH : (Fin 18 ⊕ Fin 2 → ) ≃ (Fin 18 → ) × (Fin 2 →
/-- An equivalence between `MSSMCharges.charges` and `(Fin 18 → ) × (Fin 2 → )`. This /-- An equivalence between `MSSMCharges.charges` and `(Fin 18 → ) × (Fin 2 → )`. This
splits the charges up into the SM and the additional ones for the MSSM. -/ splits the charges up into the SM and the additional ones for the MSSM. -/
@[simps!] @[simps!]
def toSplitSMPlusH : MSSMCharges.charges ≃ (Fin 18 → ) × (Fin 2 → ) := def toSplitSMPlusH : MSSMCharges.Charges ≃ (Fin 18 → ) × (Fin 2 → ) :=
toSMPlusH.trans splitSMPlusH toSMPlusH.trans splitSMPlusH
/-- An equivalence between `(Fin 18 → )` and `(Fin 6 → Fin 3 → )`. -/ /-- An equivalence between `(Fin 18 → )` and `(Fin 6 → Fin 3 → )`. -/
@ -64,13 +64,13 @@ def toSpeciesMaps' : (Fin 18 → ) ≃ (Fin 6 → Fin 3 → ) :=
This split charges up into the SM and additional fermions, and further splits the SM into This split charges up into the SM and additional fermions, and further splits the SM into
species. -/ species. -/
@[simps!] @[simps!]
def toSpecies : MSSMCharges.charges ≃ (Fin 6 → Fin 3 → ) × (Fin 2 → ) := def toSpecies : MSSMCharges.Charges ≃ (Fin 6 → Fin 3 → ) × (Fin 2 → ) :=
toSplitSMPlusH.trans (Equiv.prodCongr toSpeciesMaps' (Equiv.refl _)) toSplitSMPlusH.trans (Equiv.prodCongr toSpeciesMaps' (Equiv.refl _))
/-- For a given `i ∈ Fin 6` the projection of `MSSMCharges.charges` down to the /-- For a given `i ∈ Fin 6` the projection of `MSSMCharges.charges` down to the
corresponding SM species of charges. -/ corresponding SM species of charges. -/
@[simps!] @[simps!]
def toSMSpecies (i : Fin 6) : MSSMCharges.charges →ₗ[] MSSMSpecies.charges where def toSMSpecies (i : Fin 6) : MSSMCharges.Charges →ₗ[] MSSMSpecies.Charges where
toFun S := (Prod.fst ∘ toSpecies) S i toFun S := (Prod.fst ∘ toSpecies) S i
map_add' _ _ := by rfl map_add' _ _ := by rfl
map_smul' _ _ := by rfl map_smul' _ _ := by rfl
@ -95,19 +95,19 @@ abbrev N := toSMSpecies 5
/-- The charge `Hd`. -/ /-- The charge `Hd`. -/
@[simps!] @[simps!]
def Hd : MSSMCharges.charges →ₗ[] where def Hd : MSSMCharges.Charges →ₗ[] where
toFun S := S ⟨18, Nat.lt_of_sub_eq_succ rfl⟩ toFun S := S ⟨18, Nat.lt_of_sub_eq_succ rfl⟩
map_add' _ _ := by rfl map_add' _ _ := by rfl
map_smul' _ _ := by rfl map_smul' _ _ := by rfl
/-- The charge `Hu`. -/ /-- The charge `Hu`. -/
@[simps!] @[simps!]
def Hu : MSSMCharges.charges →ₗ[] where def Hu : MSSMCharges.Charges →ₗ[] where
toFun S := S ⟨19, Nat.lt_of_sub_eq_succ rfl⟩ toFun S := S ⟨19, Nat.lt_of_sub_eq_succ rfl⟩
map_add' _ _ := by rfl map_add' _ _ := by rfl
map_smul' _ _ := by rfl map_smul' _ _ := by rfl
lemma charges_eq_toSpecies_eq (S T : MSSMCharges.charges) : lemma charges_eq_toSpecies_eq (S T : MSSMCharges.Charges) :
S = T ↔ (∀ i, toSMSpecies i S = toSMSpecies i T) ∧ Hd S = Hd T ∧ Hu S = Hu T := by S = T ↔ (∀ i, toSMSpecies i S = toSMSpecies i T) ∧ Hd S = Hd T ∧ Hu S = Hu T := by
apply Iff.intro apply Iff.intro
intro h intro h
@ -139,7 +139,7 @@ open MSSMCharges
/-- The gravitational anomaly equation. -/ /-- The gravitational anomaly equation. -/
@[simp] @[simp]
def accGrav : MSSMCharges.charges →ₗ[] where def accGrav : MSSMCharges.Charges →ₗ[] where
toFun S := ∑ i, (6 * Q S i + 3 * U S i + 3 * D S i toFun S := ∑ i, (6 * Q S i + 3 * U S i + 3 * D S i
+ 2 * L S i + E S i + N S i) + 2 * (Hd S + Hu S) + 2 * L S i + E S i + N S i) + 2 * (Hd S + Hu S)
map_add' S T := by map_add' S T := by
@ -159,7 +159,7 @@ def accGrav : MSSMCharges.charges →ₗ[] where
ring ring
/-- Extensionality lemma for `accGrav`. -/ /-- Extensionality lemma for `accGrav`. -/
lemma accGrav_ext {S T : MSSMCharges.charges} lemma accGrav_ext {S T : MSSMCharges.Charges}
(hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i) (hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i)
(hd : Hd S = Hd T) (hu : Hu S = Hu T) : (hd : Hd S = Hd T) (hu : Hu S = Hu T) :
accGrav S = accGrav T := by accGrav S = accGrav T := by
@ -173,7 +173,7 @@ lemma accGrav_ext {S T : MSSMCharges.charges}
/-- The anomaly cancellation condition for SU(2) anomaly. -/ /-- The anomaly cancellation condition for SU(2) anomaly. -/
@[simp] @[simp]
def accSU2 : MSSMCharges.charges →ₗ[] where def accSU2 : MSSMCharges.Charges →ₗ[] where
toFun S := ∑ i, (3 * Q S i + L S i) + Hd S + Hu S toFun S := ∑ i, (3 * Q S i + L S i) + Hd S + Hu S
map_add' S T := by map_add' S T := by
simp only simp only
@ -192,7 +192,7 @@ def accSU2 : MSSMCharges.charges →ₗ[] where
ring ring
/-- Extensionality lemma for `accSU2`. -/ /-- Extensionality lemma for `accSU2`. -/
lemma accSU2_ext {S T : MSSMCharges.charges} lemma accSU2_ext {S T : MSSMCharges.Charges}
(hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i) (hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i)
(hd : Hd S = Hd T) (hu : Hu S = Hu T) : (hd : Hd S = Hd T) (hu : Hu S = Hu T) :
accSU2 S = accSU2 T := by accSU2 S = accSU2 T := by
@ -206,7 +206,7 @@ lemma accSU2_ext {S T : MSSMCharges.charges}
/-- The anomaly cancellation condition for SU(3) anomaly. -/ /-- The anomaly cancellation condition for SU(3) anomaly. -/
@[simp] @[simp]
def accSU3 : MSSMCharges.charges →ₗ[] where def accSU3 : MSSMCharges.Charges →ₗ[] where
toFun S := ∑ i, (2 * (Q S i) + (U S i) + (D S i)) toFun S := ∑ i, (2 * (Q S i) + (U S i) + (D S i))
map_add' S T := by map_add' S T := by
simp only simp only
@ -224,7 +224,7 @@ def accSU3 : MSSMCharges.charges →ₗ[] where
ring ring
/-- Extensionality lemma for `accSU3`. -/ /-- Extensionality lemma for `accSU3`. -/
lemma accSU3_ext {S T : MSSMCharges.charges} lemma accSU3_ext {S T : MSSMCharges.Charges}
(hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i) : (hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i) :
accSU3 S = accSU3 T := by accSU3 S = accSU3 T := by
simp only [accSU3, MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue, LinearMap.coe_mk, simp only [accSU3, MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue, LinearMap.coe_mk,
@ -236,7 +236,7 @@ lemma accSU3_ext {S T : MSSMCharges.charges}
/-- The acc for `Y²`. -/ /-- The acc for `Y²`. -/
@[simp] @[simp]
def accYY : MSSMCharges.charges →ₗ[] where def accYY : MSSMCharges.Charges →ₗ[] where
toFun S := ∑ i, ((Q S) i + 8 * (U S) i + 2 * (D S) i + 3 * (L S) i toFun S := ∑ i, ((Q S) i + 8 * (U S) i + 2 * (D S) i + 3 * (L S) i
+ 6 * (E S) i) + 3 * (Hd S + Hu S) + 6 * (E S) i) + 3 * (Hd S + Hu S)
map_add' S T := by map_add' S T := by
@ -256,7 +256,7 @@ def accYY : MSSMCharges.charges →ₗ[] where
ring ring
/-- Extensionality lemma for `accGrav`. -/ /-- Extensionality lemma for `accGrav`. -/
lemma accYY_ext {S T : MSSMCharges.charges} lemma accYY_ext {S T : MSSMCharges.Charges}
(hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i) (hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i)
(hd : Hd S = Hd T) (hu : Hu S = Hu T) : (hd : Hd S = Hd T) (hu : Hu S = Hu T) :
accYY S = accYY T := by accYY S = accYY T := by
@ -270,7 +270,7 @@ lemma accYY_ext {S T : MSSMCharges.charges}
/-- The symmetric bilinear function used to define the quadratic ACC. -/ /-- The symmetric bilinear function used to define the quadratic ACC. -/
@[simps!] @[simps!]
def quadBiLin : BiLinearSymm MSSMCharges.charges := BiLinearSymm.mk₂ ( def quadBiLin : BiLinearSymm MSSMCharges.Charges := BiLinearSymm.mk₂ (
fun (S, T) => ∑ i, (Q S i * Q T i + (- 2) * (U S i * U T i) + fun (S, T) => ∑ i, (Q S i * Q T i + (- 2) * (U S i * U T i) +
D S i * D T i + (- 1) * (L S i * L T i) + E S i * E T i) + D S i * D T i + (- 1) * (L S i * L T i) + E S i * E T i) +
(- Hd S * Hd T + Hu S * Hu T)) (- Hd S * Hd T + Hu S * Hu T))
@ -317,10 +317,10 @@ def quadBiLin : BiLinearSymm MSSMCharges.charges := BiLinearSymm.mk₂ (
/-- The quadratic ACC. -/ /-- The quadratic ACC. -/
@[simp] @[simp]
def accQuad : HomogeneousQuadratic MSSMCharges.charges := quadBiLin.toHomogeneousQuad def accQuad : HomogeneousQuadratic MSSMCharges.Charges := quadBiLin.toHomogeneousQuad
/-- Extensionality lemma for `accQuad`. -/ /-- Extensionality lemma for `accQuad`. -/
lemma accQuad_ext {S T : (MSSMCharges).charges} lemma accQuad_ext {S T : (MSSMCharges).Charges}
(h : ∀ j, ∑ i, ((fun a => a^2) ∘ toSMSpecies j S) i = (h : ∀ j, ∑ i, ((fun a => a^2) ∘ toSMSpecies j S) i =
∑ i, ((fun a => a^2) ∘ toSMSpecies j T) i) ∑ i, ((fun a => a^2) ∘ toSMSpecies j T) i)
(hd : Hd S = Hd T) (hu : Hu S = Hu T) : (hd : Hd S = Hd T) (hu : Hu S = Hu T) :
@ -339,7 +339,7 @@ lemma accQuad_ext {S T : (MSSMCharges).charges}
/-- The function underlying the symmetric trilinear form used to define the cubic ACC. -/ /-- The function underlying the symmetric trilinear form used to define the cubic ACC. -/
@[simp] @[simp]
def cubeTriLinToFun def cubeTriLinToFun
(S : MSSMCharges.charges × MSSMCharges.charges × MSSMCharges.charges) : := (S : MSSMCharges.Charges × MSSMCharges.Charges × MSSMCharges.Charges) : :=
∑ i, (6 * (Q S.1 i * Q S.2.1 i * Q S.2.2 i) ∑ i, (6 * (Q S.1 i * Q S.2.1 i * Q S.2.2 i)
+ 3 * (U S.1 i * U S.2.1 i * U S.2.2 i) + 3 * (U S.1 i * U S.2.1 i * U S.2.2 i)
+ 3 * (D S.1 i * D S.2.1 i * D S.2.2 i) + 3 * (D S.1 i * D S.2.1 i * D S.2.2 i)
@ -349,7 +349,7 @@ def cubeTriLinToFun
+ (2 * Hd S.1 * Hd S.2.1 * Hd S.2.2 + (2 * Hd S.1 * Hd S.2.1 * Hd S.2.2
+ 2 * Hu S.1 * Hu S.2.1 * Hu S.2.2) + 2 * Hu S.1 * Hu S.2.1 * Hu S.2.2)
lemma cubeTriLinToFun_map_smul₁ (a : ) (S T R : MSSMCharges.charges) : lemma cubeTriLinToFun_map_smul₁ (a : ) (S T R : MSSMCharges.Charges) :
cubeTriLinToFun (a • S, T, R) = a * cubeTriLinToFun (S, T, R) := by cubeTriLinToFun (a • S, T, R) = a * cubeTriLinToFun (S, T, R) := by
simp only [cubeTriLinToFun] simp only [cubeTriLinToFun]
rw [mul_add] rw [mul_add]
@ -364,7 +364,7 @@ lemma cubeTriLinToFun_map_smul₁ (a : ) (S T R : MSSMCharges.charges) :
ring ring
lemma cubeTriLinToFun_map_add₁ (S T R L : MSSMCharges.charges) : lemma cubeTriLinToFun_map_add₁ (S T R L : MSSMCharges.Charges) :
cubeTriLinToFun (S + T, R, L) = cubeTriLinToFun (S, R, L) + cubeTriLinToFun (T, R, L) := by cubeTriLinToFun (S + T, R, L) = cubeTriLinToFun (S, R, L) + cubeTriLinToFun (T, R, L) := by
simp only [cubeTriLinToFun] simp only [cubeTriLinToFun]
rw [add_assoc, ← add_assoc (2 * Hd S * Hd R * Hd L + 2 * Hu S * Hu R * Hu L) _ _] rw [add_assoc, ← add_assoc (2 * Hd S * Hd R * Hd L + 2 * Hu S * Hu R * Hu L) _ _]
@ -383,7 +383,7 @@ lemma cubeTriLinToFun_map_add₁ (S T R L : MSSMCharges.charges) :
ring ring
lemma cubeTriLinToFun_swap1 (S T R : MSSMCharges.charges) : lemma cubeTriLinToFun_swap1 (S T R : MSSMCharges.Charges) :
cubeTriLinToFun (S, T, R) = cubeTriLinToFun (T, S, R) := by cubeTriLinToFun (S, T, R) = cubeTriLinToFun (T, S, R) := by
simp only [cubeTriLinToFun, MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue, Hd_apply, simp only [cubeTriLinToFun, MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue, Hd_apply,
Fin.reduceFinMk, Hu_apply] Fin.reduceFinMk, Hu_apply]
@ -393,7 +393,7 @@ lemma cubeTriLinToFun_swap1 (S T R : MSSMCharges.charges) :
ring ring
ring ring
lemma cubeTriLinToFun_swap2 (S T R : MSSMCharges.charges) : lemma cubeTriLinToFun_swap2 (S T R : MSSMCharges.Charges) :
cubeTriLinToFun (S, T, R) = cubeTriLinToFun (S, R, T) := by cubeTriLinToFun (S, T, R) = cubeTriLinToFun (S, R, T) := by
simp only [cubeTriLinToFun, MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue, Hd_apply, simp only [cubeTriLinToFun, MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue, Hd_apply,
Fin.reduceFinMk, Hu_apply] Fin.reduceFinMk, Hu_apply]
@ -405,7 +405,7 @@ lemma cubeTriLinToFun_swap2 (S T R : MSSMCharges.charges) :
/-- The symmetric trilinear form used to define the cubic ACC. -/ /-- The symmetric trilinear form used to define the cubic ACC. -/
@[simps!] @[simps!]
def cubeTriLin : TriLinearSymm MSSMCharges.charges := TriLinearSymm.mk₃ def cubeTriLin : TriLinearSymm MSSMCharges.Charges := TriLinearSymm.mk₃
cubeTriLinToFun cubeTriLinToFun
cubeTriLinToFun_map_smul₁ cubeTriLinToFun_map_smul₁
cubeTriLinToFun_map_add₁ cubeTriLinToFun_map_add₁
@ -414,10 +414,10 @@ def cubeTriLin : TriLinearSymm MSSMCharges.charges := TriLinearSymm.mk₃
/-- The cubic ACC. -/ /-- The cubic ACC. -/
@[simp] @[simp]
def accCube : HomogeneousCubic MSSMCharges.charges := cubeTriLin.toCubic def accCube : HomogeneousCubic MSSMCharges.Charges := cubeTriLin.toCubic
/-- Extensionality lemma for `accCube`. -/ /-- Extensionality lemma for `accCube`. -/
lemma accCube_ext {S T : MSSMCharges.charges} lemma accCube_ext {S T : MSSMCharges.Charges}
(h : ∀ j, ∑ i, ((fun a => a^3) ∘ toSMSpecies j S) i = (h : ∀ j, ∑ i, ((fun a => a^3) ∘ toSMSpecies j S) i =
∑ i, ((fun a => a^3) ∘ toSMSpecies j T) i) ∑ i, ((fun a => a^3) ∘ toSMSpecies j T) i)
(hd : Hd S = Hd T) (hu : Hu S = Hu T) : (hd : Hd S = Hd T) (hu : Hu S = Hu T) :
@ -464,7 +464,7 @@ lemma quadSol (S : MSSMACC.QuadSols) : accQuad S.val = 0 := by
/-- A solution from a charge satisfying the ACCs. -/ /-- A solution from a charge satisfying the ACCs. -/
@[simp] @[simp]
def AnomalyFreeMk (S : MSSMACC.charges) (hg : accGrav S = 0) def AnomalyFreeMk (S : MSSMACC.Charges) (hg : accGrav S = 0)
(hsu2 : accSU2 S = 0) (hsu3 : accSU3 S = 0) (hyy : accYY S = 0) (hsu2 : accSU2 S = 0) (hsu3 : accSU3 S = 0) (hyy : accYY S = 0)
(hquad : accQuad S = 0) (hcube : accCube S = 0) : MSSMACC.Sols := (hquad : accQuad S = 0) (hcube : accCube S = 0) : MSSMACC.Sols :=
⟨⟨⟨S, by ⟨⟨⟨S, by
@ -481,7 +481,7 @@ def AnomalyFreeMk (S : MSSMACC.charges) (hg : accGrav S = 0)
| 0 => exact hquad | 0 => exact hquad
⟩ , by exact hcube ⟩ ⟩ , by exact hcube ⟩
lemma AnomalyFreeMk_val (S : MSSMACC.charges) (hg : accGrav S = 0) lemma AnomalyFreeMk_val (S : MSSMACC.Charges) (hg : accGrav S = 0)
(hsu2 : accSU2 S = 0) (hsu3 : accSU3 S = 0) (hyy : accYY S = 0) (hsu2 : accSU2 S = 0) (hsu3 : accSU3 S = 0) (hyy : accYY S = 0)
(hquad : accQuad S = 0) (hcube : accCube S = 0) : (hquad : accQuad S = 0) (hcube : accCube S = 0) :
(AnomalyFreeMk S hg hsu2 hsu3 hyy hquad hcube).val = S := by (AnomalyFreeMk S hg hsu2 hsu3 hyy hquad hcube).val = S := by
@ -521,7 +521,7 @@ lemma AnomalyFreeMk''_val (S : MSSMACC.QuadSols)
/-- The dot product on the vector space of charges. -/ /-- The dot product on the vector space of charges. -/
@[simps!] @[simps!]
def dot : BiLinearSymm MSSMCharges.charges := BiLinearSymm.mk₂ def dot : BiLinearSymm MSSMCharges.Charges := BiLinearSymm.mk₂
(fun S => ∑ i, (Q S.1 i * Q S.2 i + U S.1 i * U S.2 i + (fun S => ∑ i, (Q S.1 i * Q S.2 i + U S.1 i * U S.2 i +
D S.1 i * D S.2 i + L S.1 i * L S.2 i + E S.1 i * E S.2 i D S.1 i * D S.2 i + L S.1 i * L S.2 i + E S.1 i * E S.2 i
+ N S.1 i * N S.2 i) + Hd S.1 * Hd S.2 + Hu S.1 * Hu S.2) + N S.1 i * N S.2 i) + Hd S.1 * Hd S.2 + Hu S.1 * Hu S.2)

View file

@ -20,7 +20,7 @@ open MSSMACCs
open BigOperators open BigOperators
/-- The hypercharge as an element of `MSSMACC.charges`. -/ /-- The hypercharge as an element of `MSSMACC.charges`. -/
def YAsCharge : MSSMACC.charges := toSpecies.invFun def YAsCharge : MSSMACC.Charges := toSpecies.invFun
⟨fun s => fun i => ⟨fun s => fun i =>
match s, i with match s, i with
| 0, 0 => 1 | 0, 0 => 1

View file

@ -38,14 +38,14 @@ namespace AnomalyFreePerp
/-- A condition for the quad line in the plane spanned by R, Y₃ and B₃ to sit in the cubic, /-- A condition for the quad line in the plane spanned by R, Y₃ and B₃ to sit in the cubic,
and for the cube line to sit in the quad. -/ and for the cube line to sit in the quad. -/
def lineEqProp (R : MSSMACC.AnomalyFreePerp) : Prop := α₁ R = 0 ∧ α₂ R = 0 ∧ α₃ R = 0 def LineEqProp (R : MSSMACC.AnomalyFreePerp) : Prop := α₁ R = 0 ∧ α₂ R = 0 ∧ α₃ R = 0
instance (R : MSSMACC.AnomalyFreePerp) : Decidable (lineEqProp R) := by instance (R : MSSMACC.AnomalyFreePerp) : Decidable (LineEqProp R) := by
apply And.decidable apply And.decidable
/-- A condition on `Sols` which we will show in `linEqPropSol_iff_proj_linEqProp` that is equivalent /-- A condition on `Sols` which we will show in `linEqPropSol_iff_proj_linEqProp` that is equivalent
to the condition that the `proj` of the solution satisfies `lineEqProp`. -/ to the condition that the `proj` of the solution satisfies `lineEqProp`. -/
def lineEqPropSol (R : MSSMACC.Sols) : Prop := def LineEqPropSol (R : MSSMACC.Sols) : Prop :=
cubeTriLin R.val R.val Y₃.val * quadBiLin B₃.val R.val - cubeTriLin R.val R.val Y₃.val * quadBiLin B₃.val R.val -
cubeTriLin R.val R.val B₃.val * quadBiLin Y₃.val R.val = 0 cubeTriLin R.val R.val B₃.val * quadBiLin Y₃.val R.val = 0
@ -54,8 +54,8 @@ equivalent to satisfying `lineEqPropSol`. -/
def lineEqCoeff (T : MSSMACC.Sols) : := dot Y₃.val B₃.val * α₃ (proj T.1.1) def lineEqCoeff (T : MSSMACC.Sols) : := dot Y₃.val B₃.val * α₃ (proj T.1.1)
lemma lineEqPropSol_iff_lineEqCoeff_zero (T : MSSMACC.Sols) : lemma lineEqPropSol_iff_lineEqCoeff_zero (T : MSSMACC.Sols) :
lineEqPropSol T ↔ lineEqCoeff T = 0 := by LineEqPropSol T ↔ lineEqCoeff T = 0 := by
rw [lineEqPropSol, lineEqCoeff, α₃] rw [LineEqPropSol, lineEqCoeff, α₃]
simp only [Fin.isValue, Fin.reduceFinMk, mul_eq_zero, OfNat.ofNat_ne_zero, simp only [Fin.isValue, Fin.reduceFinMk, mul_eq_zero, OfNat.ofNat_ne_zero,
false_or] false_or]
rw [cube_proj_proj_B₃, cube_proj_proj_Y₃, quad_B₃_proj, quad_Y₃_proj] rw [cube_proj_proj_B₃, cube_proj_proj_Y₃, quad_B₃_proj, quad_Y₃_proj]
@ -66,8 +66,8 @@ lemma lineEqPropSol_iff_lineEqCoeff_zero (T : MSSMACC.Sols) :
simp simp
lemma linEqPropSol_iff_proj_linEqProp (R : MSSMACC.Sols) : lemma linEqPropSol_iff_proj_linEqProp (R : MSSMACC.Sols) :
lineEqPropSol R ↔ lineEqProp (proj R.1.1) := by LineEqPropSol R ↔ LineEqProp (proj R.1.1) := by
rw [lineEqPropSol_iff_lineEqCoeff_zero, lineEqCoeff, lineEqProp] rw [lineEqPropSol_iff_lineEqCoeff_zero, lineEqCoeff, LineEqProp]
apply Iff.intro apply Iff.intro
intro h intro h
rw [show dot Y₃.val B₃.val = 108 by rfl] at h rw [show dot Y₃.val B₃.val = 108 by rfl] at h
@ -81,15 +81,15 @@ lemma linEqPropSol_iff_proj_linEqProp (R : MSSMACC.Sols) :
/-- A condition which is satisfied if the plane spanned by `R`, `Y₃` and `B₃` lies /-- A condition which is satisfied if the plane spanned by `R`, `Y₃` and `B₃` lies
entirely in the quadratic surface. -/ entirely in the quadratic surface. -/
def inQuadProp (R : MSSMACC.AnomalyFreePerp) : Prop := def InQuadProp (R : MSSMACC.AnomalyFreePerp) : Prop :=
quadBiLin R.val R.val = 0 ∧ quadBiLin Y₃.val R.val = 0 ∧ quadBiLin B₃.val R.val = 0 quadBiLin R.val R.val = 0 ∧ quadBiLin Y₃.val R.val = 0 ∧ quadBiLin B₃.val R.val = 0
instance (R : MSSMACC.AnomalyFreePerp) : Decidable (inQuadProp R) := by instance (R : MSSMACC.AnomalyFreePerp) : Decidable (InQuadProp R) := by
apply And.decidable apply And.decidable
/-- A condition which is satisfied if the plane spanned by the solutions `R`, `Y₃` and `B₃` /-- A condition which is satisfied if the plane spanned by the solutions `R`, `Y₃` and `B₃`
lies entirely in the quadratic surface. -/ lies entirely in the quadratic surface. -/
def inQuadSolProp (R : MSSMACC.Sols) : Prop := def InQuadSolProp (R : MSSMACC.Sols) : Prop :=
quadBiLin Y₃.val R.val = 0 ∧ quadBiLin B₃.val R.val = 0 quadBiLin Y₃.val R.val = 0 ∧ quadBiLin B₃.val R.val = 0
/-- A rational which has two properties. It is zero for a solution `T` if and only if /-- A rational which has two properties. It is zero for a solution `T` if and only if
@ -98,7 +98,7 @@ def quadCoeff (T : MSSMACC.Sols) : :=
2 * dot Y₃.val B₃.val ^ 2 * 2 * dot Y₃.val B₃.val ^ 2 *
(quadBiLin Y₃.val T.val ^ 2 + quadBiLin B₃.val T.val ^ 2) (quadBiLin Y₃.val T.val ^ 2 + quadBiLin B₃.val T.val ^ 2)
lemma inQuadSolProp_iff_quadCoeff_zero (T : MSSMACC.Sols) : inQuadSolProp T ↔ quadCoeff T = 0 := by lemma inQuadSolProp_iff_quadCoeff_zero (T : MSSMACC.Sols) : InQuadSolProp T ↔ quadCoeff T = 0 := by
apply Iff.intro apply Iff.intro
intro h intro h
rw [quadCoeff, h.1, h.2] rw [quadCoeff, h.1, h.2]
@ -117,8 +117,8 @@ lemma inQuadSolProp_iff_quadCoeff_zero (T : MSSMACC.Sols) : inQuadSolProp T ↔
/-- The conditions `inQuadSolProp R` and `inQuadProp (proj R.1.1)` are equivalent. This is to be /-- The conditions `inQuadSolProp R` and `inQuadProp (proj R.1.1)` are equivalent. This is to be
expected since both `R` and `proj R.1.1` define the same plane with `Y₃` and `B₃`. -/ expected since both `R` and `proj R.1.1` define the same plane with `Y₃` and `B₃`. -/
lemma inQuadSolProp_iff_proj_inQuadProp (R : MSSMACC.Sols) : lemma inQuadSolProp_iff_proj_inQuadProp (R : MSSMACC.Sols) :
inQuadSolProp R ↔ inQuadProp (proj R.1.1) := by InQuadSolProp R ↔ InQuadProp (proj R.1.1) := by
rw [inQuadSolProp, inQuadProp] rw [InQuadSolProp, InQuadProp]
rw [quad_proj, quad_Y₃_proj, quad_B₃_proj] rw [quad_proj, quad_Y₃_proj, quad_B₃_proj]
apply Iff.intro apply Iff.intro
intro h intro h
@ -133,17 +133,17 @@ lemma inQuadSolProp_iff_proj_inQuadProp (R : MSSMACC.Sols) :
/-- A condition which is satisfied if the plane spanned by `R`, `Y₃` and `B₃` lies /-- A condition which is satisfied if the plane spanned by `R`, `Y₃` and `B₃` lies
entirely in the cubic surface. -/ entirely in the cubic surface. -/
def inCubeProp (R : MSSMACC.AnomalyFreePerp) : Prop := def InCubeProp (R : MSSMACC.AnomalyFreePerp) : Prop :=
cubeTriLin R.val R.val R.val = 0 ∧ cubeTriLin R.val R.val B₃.val = 0 ∧ cubeTriLin R.val R.val R.val = 0 ∧ cubeTriLin R.val R.val B₃.val = 0 ∧
cubeTriLin R.val R.val Y₃.val = 0 cubeTriLin R.val R.val Y₃.val = 0
instance (R : MSSMACC.AnomalyFreePerp) : Decidable (inCubeProp R) := by instance (R : MSSMACC.AnomalyFreePerp) : Decidable (InCubeProp R) := by
apply And.decidable apply And.decidable
/-- A condition which is satisfied if the plane spanned by the solutions `R`, `Y₃` and `B₃` /-- A condition which is satisfied if the plane spanned by the solutions `R`, `Y₃` and `B₃`
lies entirely in the cubic surface. -/ lies entirely in the cubic surface. -/
def inCubeSolProp (R : MSSMACC.Sols) : Prop := def InCubeSolProp (R : MSSMACC.Sols) : Prop :=
cubeTriLin R.val R.val B₃.val = 0 ∧ cubeTriLin R.val R.val Y₃.val = 0 cubeTriLin R.val R.val B₃.val = 0 ∧ cubeTriLin R.val R.val Y₃.val = 0
/-- A rational which has two properties. It is zero for a solution `T` if and only if /-- A rational which has two properties. It is zero for a solution `T` if and only if
@ -153,7 +153,7 @@ def cubicCoeff (T : MSSMACC.Sols) : :=
cubeTriLin T.val T.val B₃.val ^ 2 ) cubeTriLin T.val T.val B₃.val ^ 2 )
lemma inCubeSolProp_iff_cubicCoeff_zero (T : MSSMACC.Sols) : lemma inCubeSolProp_iff_cubicCoeff_zero (T : MSSMACC.Sols) :
inCubeSolProp T ↔ cubicCoeff T = 0 := by InCubeSolProp T ↔ cubicCoeff T = 0 := by
apply Iff.intro apply Iff.intro
intro h intro h
rw [cubicCoeff, h.1, h.2] rw [cubicCoeff, h.1, h.2]
@ -170,8 +170,8 @@ lemma inCubeSolProp_iff_cubicCoeff_zero (T : MSSMACC.Sols) :
exact h.symm exact h.symm
lemma inCubeSolProp_iff_proj_inCubeProp (R : MSSMACC.Sols) : lemma inCubeSolProp_iff_proj_inCubeProp (R : MSSMACC.Sols) :
inCubeSolProp R ↔ inCubeProp (proj R.1.1) := by InCubeSolProp R ↔ InCubeProp (proj R.1.1) := by
rw [inCubeSolProp, inCubeProp] rw [InCubeSolProp, InCubeProp]
rw [cube_proj, cube_proj_proj_Y₃, cube_proj_proj_B₃] rw [cube_proj, cube_proj_proj_Y₃, cube_proj_proj_B₃]
apply Iff.intro apply Iff.intro
intro h intro h
@ -186,30 +186,30 @@ lemma inCubeSolProp_iff_proj_inCubeProp (R : MSSMACC.Sols) :
/-- Those charge assignments perpendicular to `Y₃` and `B₃` which satisfy the condition /-- Those charge assignments perpendicular to `Y₃` and `B₃` which satisfy the condition
`lineEqProp`. -/ `lineEqProp`. -/
def inLineEq : Type := {R : MSSMACC.AnomalyFreePerp // lineEqProp R} def InLineEq : Type := {R : MSSMACC.AnomalyFreePerp // LineEqProp R}
/-- Those charge assignments perpendicular to `Y₃` and `B₃` which satisfy the condition /-- Those charge assignments perpendicular to `Y₃` and `B₃` which satisfy the condition
`lineEqProp` and `inQuadProp`. -/ `lineEqProp` and `inQuadProp`. -/
def inQuad : Type := {R : inLineEq // inQuadProp R.val} def InQuad : Type := {R : InLineEq // InQuadProp R.val}
/-- Those charge assignments perpendicular to `Y₃` and `B₃` which satisfy the condition /-- Those charge assignments perpendicular to `Y₃` and `B₃` which satisfy the condition
`lineEqProp`, `inQuadProp` and `inCubeProp`. -/ `lineEqProp`, `inQuadProp` and `inCubeProp`. -/
def inQuadCube : Type := {R : inQuad // inCubeProp R.val.val} def InQuadCube : Type := {R : InQuad // InCubeProp R.val.val}
/-- Those solutions which do not satisfy the condition `lineEqPropSol`. -/ /-- Those solutions which do not satisfy the condition `lineEqPropSol`. -/
def notInLineEqSol : Type := {R : MSSMACC.Sols // ¬ lineEqPropSol R} def NotInLineEqSol : Type := {R : MSSMACC.Sols // ¬ LineEqPropSol R}
/-- Those solutions which satisfy the condition `lineEqPropSol` by not `inQuadSolProp`. -/ /-- Those solutions which satisfy the condition `lineEqPropSol` by not `inQuadSolProp`. -/
def inLineEqSol : Type := {R : MSSMACC.Sols // lineEqPropSol R ∧ ¬ inQuadSolProp R} def InLineEqSol : Type := {R : MSSMACC.Sols // LineEqPropSol R ∧ ¬ InQuadSolProp R}
/-- Those solutions which satisfy the condition `lineEqPropSol` and `inQuadSolProp` but /-- Those solutions which satisfy the condition `lineEqPropSol` and `inQuadSolProp` but
not `inCubeSolProp`. -/ not `inCubeSolProp`. -/
def inQuadSol : Type := {R : MSSMACC.Sols // lineEqPropSol R ∧ inQuadSolProp R ∧ ¬ inCubeSolProp R} def InQuadSol : Type := {R : MSSMACC.Sols // LineEqPropSol R ∧ InQuadSolProp R ∧ ¬ InCubeSolProp R}
/-- Those solutions which satisfy the condition all the conditions `lineEqPropSol`, `inQuadSolProp` /-- Those solutions which satisfy the condition all the conditions `lineEqPropSol`, `inQuadSolProp`
and `inCubeSolProp`. -/ and `inCubeSolProp`. -/
def inQuadCubeSol : Type := def InQuadCubeSol : Type :=
{R : MSSMACC.Sols // lineEqPropSol R ∧ inQuadSolProp R ∧ inCubeSolProp R} {R : MSSMACC.Sols // LineEqPropSol R ∧ InQuadSolProp R ∧ InCubeSolProp R}
/-- Given a `R` perpendicular to `Y₃` and `B₃` a quadratic solution. -/ /-- Given a `R` perpendicular to `Y₃` and `B₃` a quadratic solution. -/
def toSolNSQuad (R : MSSMACC.AnomalyFreePerp) : MSSMACC.QuadSols := def toSolNSQuad (R : MSSMACC.AnomalyFreePerp) : MSSMACC.QuadSols :=
@ -244,7 +244,7 @@ def toSolNS : MSSMACC.AnomalyFreePerp × × × → MSSMACC.Sols := f
def toSolNSProj (T : MSSMACC.Sols) : MSSMACC.AnomalyFreePerp × × × := def toSolNSProj (T : MSSMACC.Sols) : MSSMACC.AnomalyFreePerp × × × :=
(proj T.1.1, (lineEqCoeff T)⁻¹, 0, 0) (proj T.1.1, (lineEqCoeff T)⁻¹, 0, 0)
lemma toSolNS_proj (T : notInLineEqSol) : toSolNS (toSolNSProj T.val) = T.val := by lemma toSolNS_proj (T : NotInLineEqSol) : toSolNS (toSolNSProj T.val) = T.val := by
apply ACCSystem.Sols.ext apply ACCSystem.Sols.ext
rw [toSolNS, toSolNSProj] rw [toSolNS, toSolNSProj]
change (lineEqCoeff T.val)⁻¹ • (toSolNSQuad _).1.1 = _ change (lineEqCoeff T.val)⁻¹ • (toSolNSQuad _).1.1 = _
@ -263,7 +263,7 @@ lemma toSolNS_proj (T : notInLineEqSol) : toSolNS (toSolNSProj T.val) = T.val :=
simp simp
/-- Given a element of `inLineEq × × × `, a solution to the ACCs. -/ /-- Given a element of `inLineEq × × × `, a solution to the ACCs. -/
def inLineEqToSol : inLineEq × × × → MSSMACC.Sols := fun (R, c₁, c₂, c₃) => def inLineEqToSol : InLineEq × × × → MSSMACC.Sols := fun (R, c₁, c₂, c₃) =>
AnomalyFreeMk'' (lineQuad R.val c₁ c₂ c₃) AnomalyFreeMk'' (lineQuad R.val c₁ c₂ c₃)
(by (by
rw [lineQuad_cube] rw [lineQuad_cube]
@ -271,7 +271,7 @@ def inLineEqToSol : inLineEq × × × → MSSMACC.Sols := fun (R, c
simp) simp)
/-- On elements of `inLineEqSol` a right-inverse to `inLineEqSol`. -/ /-- On elements of `inLineEqSol` a right-inverse to `inLineEqSol`. -/
def inLineEqProj (T : inLineEqSol) : inLineEq × × × := def inLineEqProj (T : InLineEqSol) : InLineEq × × × :=
(⟨proj T.val.1.1, (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1 ⟩, (⟨proj T.val.1.1, (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1 ⟩,
(quadCoeff T.val)⁻¹ * quadBiLin B₃.val T.val.val, (quadCoeff T.val)⁻¹ * quadBiLin B₃.val T.val.val,
(quadCoeff T.val)⁻¹ * (- quadBiLin Y₃.val T.val.val), (quadCoeff T.val)⁻¹ * (- quadBiLin Y₃.val T.val.val),
@ -279,14 +279,14 @@ def inLineEqProj (T : inLineEqSol) : inLineEq × × × :=
quadBiLin B₃.val T.val.val * (dot B₃.val T.val.val - dot Y₃.val T.val.val) quadBiLin B₃.val T.val.val * (dot B₃.val T.val.val - dot Y₃.val T.val.val)
- quadBiLin Y₃.val T.val.val * (dot Y₃.val T.val.val - 2 * dot B₃.val T.val.val))) - quadBiLin Y₃.val T.val.val * (dot Y₃.val T.val.val - 2 * dot B₃.val T.val.val)))
lemma inLineEqTo_smul (R : inLineEq) (c₁ c₂ c₃ d : ) : lemma inLineEqTo_smul (R : InLineEq) (c₁ c₂ c₃ d : ) :
inLineEqToSol (R, (d * c₁), (d * c₂), (d * c₃)) = d • inLineEqToSol (R, c₁, c₂, c₃) := by inLineEqToSol (R, (d * c₁), (d * c₂), (d * c₃)) = d • inLineEqToSol (R, c₁, c₂, c₃) := by
apply ACCSystem.Sols.ext apply ACCSystem.Sols.ext
change (lineQuad _ _ _ _).val = _ change (lineQuad _ _ _ _).val = _
rw [lineQuad_smul] rw [lineQuad_smul]
rfl rfl
lemma inLineEqToSol_proj (T : inLineEqSol) : inLineEqToSol (inLineEqProj T) = T.val := by lemma inLineEqToSol_proj (T : InLineEqSol) : inLineEqToSol (inLineEqProj T) = T.val := by
rw [inLineEqProj, inLineEqTo_smul] rw [inLineEqProj, inLineEqTo_smul]
apply ACCSystem.Sols.ext apply ACCSystem.Sols.ext
change _ • (lineQuad _ _ _ _).val = _ change _ • (lineQuad _ _ _ _).val = _
@ -306,7 +306,7 @@ lemma inLineEqToSol_proj (T : inLineEqSol) : inLineEqToSol (inLineEqProj T) = T.
simp simp
/-- Given a element of `inQuad × × × `, a solution to the ACCs. -/ /-- Given a element of `inQuad × × × `, a solution to the ACCs. -/
def inQuadToSol : inQuad × × × → MSSMACC.Sols := fun (R, a₁, a₂, a₃) => def inQuadToSol : InQuad × × × → MSSMACC.Sols := fun (R, a₁, a₂, a₃) =>
AnomalyFreeMk' (lineCube R.val.val a₁ a₂ a₃) AnomalyFreeMk' (lineCube R.val.val a₁ a₂ a₃)
(by (by
erw [planeY₃B₃_quad] erw [planeY₃B₃_quad]
@ -314,7 +314,7 @@ def inQuadToSol : inQuad × × × → MSSMACC.Sols := fun (R, a₁,
simp) simp)
(lineCube_cube R.val.val a₁ a₂ a₃) (lineCube_cube R.val.val a₁ a₂ a₃)
lemma inQuadToSol_smul (R : inQuad) (c₁ c₂ c₃ d : ) : lemma inQuadToSol_smul (R : InQuad) (c₁ c₂ c₃ d : ) :
inQuadToSol (R, (d * c₁), (d * c₂), (d * c₃)) = d • inQuadToSol (R, c₁, c₂, c₃) := by inQuadToSol (R, (d * c₁), (d * c₂), (d * c₃)) = d • inQuadToSol (R, c₁, c₂, c₃) := by
apply ACCSystem.Sols.ext apply ACCSystem.Sols.ext
change (lineCube _ _ _ _).val = _ change (lineCube _ _ _ _).val = _
@ -322,7 +322,7 @@ lemma inQuadToSol_smul (R : inQuad) (c₁ c₂ c₃ d : ) :
rfl rfl
/-- On elements of `inQuadSol` a right-inverse to `inQuadToSol`. -/ /-- On elements of `inQuadSol` a right-inverse to `inQuadToSol`. -/
def inQuadProj (T : inQuadSol) : inQuad × × × := def inQuadProj (T : InQuadSol) : InQuad × × × :=
(⟨⟨proj T.val.1.1, (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1 ⟩, (⟨⟨proj T.val.1.1, (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1 ⟩,
(inQuadSolProp_iff_proj_inQuadProp T.val).mp T.prop.2.1⟩, (inQuadSolProp_iff_proj_inQuadProp T.val).mp T.prop.2.1⟩,
(cubicCoeff T.val)⁻¹ * (cubeTriLin T.val.val T.val.val B₃.val), (cubicCoeff T.val)⁻¹ * (cubeTriLin T.val.val T.val.val B₃.val),
@ -332,7 +332,7 @@ def inQuadProj (T : inQuadSol) : inQuad × × × :=
- cubeTriLin T.val.val T.val.val Y₃.val - cubeTriLin T.val.val T.val.val Y₃.val
* (dot Y₃.val T.val.val - 2 * dot B₃.val T.val.val))) * (dot Y₃.val T.val.val - 2 * dot B₃.val T.val.val)))
lemma inQuadToSol_proj (T : inQuadSol) : inQuadToSol (inQuadProj T) = T.val := by lemma inQuadToSol_proj (T : InQuadSol) : inQuadToSol (inQuadProj T) = T.val := by
rw [inQuadProj, inQuadToSol_smul] rw [inQuadProj, inQuadToSol_smul]
apply ACCSystem.Sols.ext apply ACCSystem.Sols.ext
change _ • (planeY₃B₃ _ _ _ _).val = _ change _ • (planeY₃B₃ _ _ _ _).val = _
@ -352,7 +352,7 @@ lemma inQuadToSol_proj (T : inQuadSol) : inQuadToSol (inQuadProj T) = T.val := b
simp simp
/-- Given a element of `inQuadCube × × × `, a solution to the ACCs. -/ /-- Given a element of `inQuadCube × × × `, a solution to the ACCs. -/
def inQuadCubeToSol : inQuadCube × × × → MSSMACC.Sols := fun (R, b₁, b₂, b₃) => def inQuadCubeToSol : InQuadCube × × × → MSSMACC.Sols := fun (R, b₁, b₂, b₃) =>
AnomalyFreeMk' (planeY₃B₃ R.val.val.val b₁ b₂ b₃) AnomalyFreeMk' (planeY₃B₃ R.val.val.val b₁ b₂ b₃)
(by (by
rw [planeY₃B₃_quad] rw [planeY₃B₃_quad]
@ -363,7 +363,7 @@ def inQuadCubeToSol : inQuadCube × × × → MSSMACC.Sols := fun (R
rw [R.prop.1, R.prop.2.1, R.prop.2.2] rw [R.prop.1, R.prop.2.1, R.prop.2.2]
simp) simp)
lemma inQuadCubeToSol_smul (R : inQuadCube) (c₁ c₂ c₃ d : ) : lemma inQuadCubeToSol_smul (R : InQuadCube) (c₁ c₂ c₃ d : ) :
inQuadCubeToSol (R, (d * c₁), (d * c₂), (d * c₃)) = d • inQuadCubeToSol (R, c₁, c₂, c₃):= by inQuadCubeToSol (R, (d * c₁), (d * c₂), (d * c₃)) = d • inQuadCubeToSol (R, c₁, c₂, c₃):= by
apply ACCSystem.Sols.ext apply ACCSystem.Sols.ext
change (planeY₃B₃ _ _ _ _).val = _ change (planeY₃B₃ _ _ _ _).val = _
@ -371,7 +371,7 @@ lemma inQuadCubeToSol_smul (R : inQuadCube) (c₁ c₂ c₃ d : ) :
rfl rfl
/-- On elements of `inQuadCubeSol` a right-inverse to `inQuadCubeToSol`. -/ /-- On elements of `inQuadCubeSol` a right-inverse to `inQuadCubeToSol`. -/
def inQuadCubeProj (T : inQuadCubeSol) : inQuadCube × × × := def inQuadCubeProj (T : InQuadCubeSol) : InQuadCube × × × :=
(⟨⟨⟨proj T.val.1.1, (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1 ⟩, (⟨⟨⟨proj T.val.1.1, (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1 ⟩,
(inQuadSolProp_iff_proj_inQuadProp T.val).mp T.prop.2.1⟩, (inQuadSolProp_iff_proj_inQuadProp T.val).mp T.prop.2.1⟩,
(inCubeSolProp_iff_proj_inCubeProp T.val).mp T.prop.2.2⟩, (inCubeSolProp_iff_proj_inCubeProp T.val).mp T.prop.2.2⟩,
@ -379,7 +379,7 @@ def inQuadCubeProj (T : inQuadCubeSol) : inQuadCube × × × :=
(dot Y₃.val B₃.val)⁻¹ * (2 * dot B₃.val T.val.val - dot Y₃.val T.val.val), (dot Y₃.val B₃.val)⁻¹ * (2 * dot B₃.val T.val.val - dot Y₃.val T.val.val),
(dot Y₃.val B₃.val)⁻¹ * 1) (dot Y₃.val B₃.val)⁻¹ * 1)
lemma inQuadCubeToSol_proj (T : inQuadCubeSol) : lemma inQuadCubeToSol_proj (T : InQuadCubeSol) :
inQuadCubeToSol (inQuadCubeProj T) = T.val := by inQuadCubeToSol (inQuadCubeProj T) = T.val := by
rw [inQuadCubeProj, inQuadCubeToSol_smul] rw [inQuadCubeProj, inQuadCubeToSol_smul]
apply ACCSystem.Sols.ext apply ACCSystem.Sols.ext
@ -396,64 +396,64 @@ lemma inQuadCubeToSol_proj (T : inQuadCubeSol) :
/-- Given an element of `MSSMACC.AnomalyFreePerp × × × ` a solution. We will /-- Given an element of `MSSMACC.AnomalyFreePerp × × × ` a solution. We will
show that this map is a surjection. -/ show that this map is a surjection. -/
def toSol : MSSMACC.AnomalyFreePerp × × × → MSSMACC.Sols := fun (R, a, b, c) => def toSol : MSSMACC.AnomalyFreePerp × × × → MSSMACC.Sols := fun (R, a, b, c) =>
if h₃ : lineEqProp R ∧ inQuadProp R ∧ inCubeProp R then if h₃ : LineEqProp R ∧ InQuadProp R ∧ InCubeProp R then
inQuadCubeToSol (⟨⟨⟨R, h₃.1⟩, h₃.2.1⟩, h₃.2.2⟩, a, b, c) inQuadCubeToSol (⟨⟨⟨R, h₃.1⟩, h₃.2.1⟩, h₃.2.2⟩, a, b, c)
else else
if h₂ : lineEqProp R ∧ inQuadProp R then if h₂ : LineEqProp R ∧ InQuadProp R then
inQuadToSol (⟨⟨R, h₂.1⟩, h₂.2⟩, a, b, c) inQuadToSol (⟨⟨R, h₂.1⟩, h₂.2⟩, a, b, c)
else else
if h₁ : lineEqProp R then if h₁ : LineEqProp R then
inLineEqToSol (⟨R, h₁⟩, a, b, c) inLineEqToSol (⟨R, h₁⟩, a, b, c)
else else
toSolNS ⟨R, a, b, c⟩ toSolNS ⟨R, a, b, c⟩
lemma toSol_toSolNSProj (T : notInLineEqSol) : lemma toSol_toSolNSProj (T : NotInLineEqSol) :
∃ X, toSol X = T.val := by ∃ X, toSol X = T.val := by
use toSolNSProj T.val use toSolNSProj T.val
have h1 : ¬ lineEqProp (toSolNSProj T.val).1 := have h1 : ¬ LineEqProp (toSolNSProj T.val).1 :=
(linEqPropSol_iff_proj_linEqProp T.val).mpr.mt T.prop (linEqPropSol_iff_proj_linEqProp T.val).mpr.mt T.prop
rw [toSol] rw [toSol]
simp_all simp_all
exact toSolNS_proj T exact toSolNS_proj T
lemma toSol_inLineEq (T : inLineEqSol) : ∃ X, toSol X = T.val := by lemma toSol_inLineEq (T : InLineEqSol) : ∃ X, toSol X = T.val := by
let X := inLineEqProj T let X := inLineEqProj T
use ⟨X.1.val, X.2.1, X.2.2⟩ use ⟨X.1.val, X.2.1, X.2.2⟩
have : ¬ inQuadProp X.1.val := (inQuadSolProp_iff_proj_inQuadProp T.val).mpr.mt T.prop.2 have : ¬ InQuadProp X.1.val := (inQuadSolProp_iff_proj_inQuadProp T.val).mpr.mt T.prop.2
have : lineEqProp X.1.val := (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1 have : LineEqProp X.1.val := (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1
rw [toSol] rw [toSol]
simp_all simp_all
exact inLineEqToSol_proj T exact inLineEqToSol_proj T
lemma toSol_inQuad (T : inQuadSol) : ∃ X, toSol X = T.val := by lemma toSol_inQuad (T : InQuadSol) : ∃ X, toSol X = T.val := by
let X := inQuadProj T let X := inQuadProj T
use ⟨X.1.val.val, X.2.1, X.2.2⟩ use ⟨X.1.val.val, X.2.1, X.2.2⟩
have : ¬ inCubeProp X.1.val.val := (inCubeSolProp_iff_proj_inCubeProp T.val).mpr.mt T.prop.2.2 have : ¬ InCubeProp X.1.val.val := (inCubeSolProp_iff_proj_inCubeProp T.val).mpr.mt T.prop.2.2
have : inQuadProp X.1.val.val := (inQuadSolProp_iff_proj_inQuadProp T.val).mp T.prop.2.1 have : InQuadProp X.1.val.val := (inQuadSolProp_iff_proj_inQuadProp T.val).mp T.prop.2.1
have : lineEqProp X.1.val.val := (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1 have : LineEqProp X.1.val.val := (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1
rw [toSol] rw [toSol]
simp_all simp_all
exact inQuadToSol_proj T exact inQuadToSol_proj T
lemma toSol_inQuadCube (T : inQuadCubeSol) : ∃ X, toSol X = T.val := by lemma toSol_inQuadCube (T : InQuadCubeSol) : ∃ X, toSol X = T.val := by
let X := inQuadCubeProj T let X := inQuadCubeProj T
use ⟨X.1.val.val.val, X.2.1, X.2.2⟩ use ⟨X.1.val.val.val, X.2.1, X.2.2⟩
have : inCubeProp X.1.val.val.val := (inCubeSolProp_iff_proj_inCubeProp T.val).mp T.prop.2.2 have : InCubeProp X.1.val.val.val := (inCubeSolProp_iff_proj_inCubeProp T.val).mp T.prop.2.2
have : inQuadProp X.1.val.val.val := (inQuadSolProp_iff_proj_inQuadProp T.val).mp T.prop.2.1 have : InQuadProp X.1.val.val.val := (inQuadSolProp_iff_proj_inQuadProp T.val).mp T.prop.2.1
have : lineEqProp X.1.val.val.val := (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1 have : LineEqProp X.1.val.val.val := (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1
rw [toSol] rw [toSol]
simp_all simp_all
exact inQuadCubeToSol_proj T exact inQuadCubeToSol_proj T
theorem toSol_surjective : Function.Surjective toSol := by theorem toSol_surjective : Function.Surjective toSol := by
intro T intro T
by_cases h₁ : ¬ lineEqPropSol T by_cases h₁ : ¬ LineEqPropSol T
exact toSol_toSolNSProj ⟨T, h₁⟩ exact toSol_toSolNSProj ⟨T, h₁⟩
simp at h₁ simp at h₁
by_cases h₂ : ¬ inQuadSolProp T by_cases h₂ : ¬ InQuadSolProp T
exact toSol_inLineEq ⟨T, And.intro h₁ h₂⟩ exact toSol_inLineEq ⟨T, And.intro h₁ h₂⟩
simp at h₂ simp at h₂
by_cases h₃ : ¬ inCubeSolProp T by_cases h₃ : ¬ InCubeSolProp T
exact toSol_inQuad ⟨T, And.intro h₁ (And.intro h₂ h₃)⟩ exact toSol_inQuad ⟨T, And.intro h₁ (And.intro h₂ h₃)⟩
simp at h₃ simp at h₃
exact toSol_inQuadCube ⟨T, And.intro h₁ (And.intro h₂ h₃)⟩ exact toSol_inQuadCube ⟨T, And.intro h₁ (And.intro h₂ h₃)⟩

View file

@ -27,14 +27,14 @@ open BigOperators
/-- The group of family permutations is `S₃⁶`-/ /-- The group of family permutations is `S₃⁶`-/
@[simp] @[simp]
def permGroup := Fin 6 → Equiv.Perm (Fin 3) def PermGroup := Fin 6 → Equiv.Perm (Fin 3)
@[simp] @[simp]
instance : Group permGroup := Pi.group instance : Group PermGroup := Pi.group
/-- The image of an element of `permGroup` under the representation on charges. -/ /-- The image of an element of `permGroup` under the representation on charges. -/
@[simps!] @[simps!]
def chargeMap (f : permGroup) : MSSMCharges.charges →ₗ[] MSSMCharges.charges where def chargeMap (f : PermGroup) : MSSMCharges.Charges →ₗ[] MSSMCharges.Charges where
toFun S := toSpecies.symm (fun i => toSMSpecies i S ∘ f i, Prod.snd (toSpecies S)) toFun S := toSpecies.symm (fun i => toSMSpecies i S ∘ f i, Prod.snd (toSpecies S))
map_add' S T := by map_add' S T := by
simp only simp only
@ -52,16 +52,16 @@ def chargeMap (f : permGroup) : MSSMCharges.charges →ₗ[] MSSMCharges.char
rw [(toSMSpecies i).map_smul, toSMSpecies_toSpecies_inv, toSMSpecies_toSpecies_inv] rw [(toSMSpecies i).map_smul, toSMSpecies_toSpecies_inv, toSMSpecies_toSpecies_inv]
rfl rfl
lemma chargeMap_toSpecies (f : permGroup) (S : MSSMCharges.charges) (j : Fin 6) : lemma chargeMap_toSpecies (f : PermGroup) (S : MSSMCharges.Charges) (j : Fin 6) :
toSMSpecies j (chargeMap f S) = toSMSpecies j S ∘ f j := by toSMSpecies j (chargeMap f S) = toSMSpecies j S ∘ f j := by
erw [toSMSpecies_toSpecies_inv] erw [toSMSpecies_toSpecies_inv]
/-- The representation of `permGroup` acting on the vector space of charges. -/ /-- The representation of `permGroup` acting on the vector space of charges. -/
@[simp] @[simp]
def repCharges : Representation (permGroup) (MSSMCharges).charges where def repCharges : Representation PermGroup (MSSMCharges).Charges where
toFun f := chargeMap f⁻¹ toFun f := chargeMap f⁻¹
map_mul' f g :=by map_mul' f g :=by
simp only [permGroup, mul_inv_rev] simp only [PermGroup, mul_inv_rev]
apply LinearMap.ext apply LinearMap.ext
intro S intro S
rw [charges_eq_toSpecies_eq] rw [charges_eq_toSpecies_eq]
@ -81,56 +81,56 @@ def repCharges : Representation (permGroup) (MSSMCharges).charges where
erw [toSMSpecies_toSpecies_inv] erw [toSMSpecies_toSpecies_inv]
rfl rfl
lemma repCharges_toSMSpecies (f : permGroup) (S : MSSMCharges.charges) (j : Fin 6) : lemma repCharges_toSMSpecies (f : PermGroup) (S : MSSMCharges.Charges) (j : Fin 6) :
toSMSpecies j (repCharges f S) = toSMSpecies j S ∘ f⁻¹ j := by toSMSpecies j (repCharges f S) = toSMSpecies j S ∘ f⁻¹ j := by
erw [toSMSpecies_toSpecies_inv] erw [toSMSpecies_toSpecies_inv]
lemma toSpecies_sum_invariant (m : ) (f : permGroup) (S : MSSMCharges.charges) (j : Fin 6) : lemma toSpecies_sum_invariant (m : ) (f : PermGroup) (S : MSSMCharges.Charges) (j : Fin 6) :
∑ i, ((fun a => a ^ m) ∘ toSMSpecies j (repCharges f S)) i = ∑ i, ((fun a => a ^ m) ∘ toSMSpecies j (repCharges f S)) i =
∑ i, ((fun a => a ^ m) ∘ toSMSpecies j S) i := by ∑ i, ((fun a => a ^ m) ∘ toSMSpecies j S) i := by
rw [repCharges_toSMSpecies] rw [repCharges_toSMSpecies]
exact Equiv.sum_comp (f⁻¹ j) ((fun a => a ^ m) ∘ (toSMSpecies j) S) exact Equiv.sum_comp (f⁻¹ j) ((fun a => a ^ m) ∘ (toSMSpecies j) S)
lemma Hd_invariant (f : permGroup) (S : MSSMCharges.charges) : lemma Hd_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
Hd (repCharges f S) = Hd S := rfl Hd (repCharges f S) = Hd S := rfl
lemma Hu_invariant (f : permGroup) (S : MSSMCharges.charges) : lemma Hu_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
Hu (repCharges f S) = Hu S := rfl Hu (repCharges f S) = Hu S := rfl
lemma accGrav_invariant (f : permGroup) (S : MSSMCharges.charges) : lemma accGrav_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
accGrav (repCharges f S) = accGrav S := accGrav (repCharges f S) = accGrav S :=
accGrav_ext accGrav_ext
(by simpa using toSpecies_sum_invariant 1 f S) (by simpa using toSpecies_sum_invariant 1 f S)
(Hd_invariant f S) (Hd_invariant f S)
(Hu_invariant f S) (Hu_invariant f S)
lemma accSU2_invariant (f : permGroup) (S : MSSMCharges.charges) : lemma accSU2_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
accSU2 (repCharges f S) = accSU2 S := accSU2 (repCharges f S) = accSU2 S :=
accSU2_ext accSU2_ext
(by simpa using toSpecies_sum_invariant 1 f S) (by simpa using toSpecies_sum_invariant 1 f S)
(Hd_invariant f S) (Hd_invariant f S)
(Hu_invariant f S) (Hu_invariant f S)
lemma accSU3_invariant (f : permGroup) (S : MSSMCharges.charges) : lemma accSU3_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
accSU3 (repCharges f S) = accSU3 S := accSU3 (repCharges f S) = accSU3 S :=
accSU3_ext accSU3_ext
(by simpa using toSpecies_sum_invariant 1 f S) (by simpa using toSpecies_sum_invariant 1 f S)
lemma accYY_invariant (f : permGroup) (S : MSSMCharges.charges) : lemma accYY_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
accYY (repCharges f S) = accYY S := accYY (repCharges f S) = accYY S :=
accYY_ext accYY_ext
(by simpa using toSpecies_sum_invariant 1 f S) (by simpa using toSpecies_sum_invariant 1 f S)
(Hd_invariant f S) (Hd_invariant f S)
(Hu_invariant f S) (Hu_invariant f S)
lemma accQuad_invariant (f : permGroup) (S : MSSMCharges.charges) : lemma accQuad_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
accQuad (repCharges f S) = accQuad S := accQuad (repCharges f S) = accQuad S :=
accQuad_ext accQuad_ext
(toSpecies_sum_invariant 2 f S) (toSpecies_sum_invariant 2 f S)
(Hd_invariant f S) (Hd_invariant f S)
(Hu_invariant f S) (Hu_invariant f S)
lemma accCube_invariant (f : permGroup) (S : MSSMCharges.charges) : lemma accCube_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
accCube (repCharges f S) = accCube S := accCube (repCharges f S) = accCube S :=
accCube_ext accCube_ext
(fun j => toSpecies_sum_invariant 3 f S j) (fun j => toSpecies_sum_invariant 3 f S j)

View file

@ -26,7 +26,7 @@ open BigOperators
/-- $Y_3$ is the charge which is hypercharge in all families, but with the third /-- $Y_3$ is the charge which is hypercharge in all families, but with the third
family of the opposite sign. -/ family of the opposite sign. -/
def Y₃AsCharge : MSSMACC.charges := toSpecies.symm def Y₃AsCharge : MSSMACC.Charges := toSpecies.symm
⟨fun s => fun i => ⟨fun s => fun i =>
match s, i with match s, i with
| 0, 0 => 1 | 0, 0 => 1

View file

@ -28,7 +28,7 @@ def PureU1Charges (n : ) : ACCSystemCharges := ACCSystemChargesMk n
open BigOperators in open BigOperators in
/-- The gravitational anomaly. -/ /-- The gravitational anomaly. -/
def accGrav (n : ) : ((PureU1Charges n).charges →ₗ[] ) where def accGrav (n : ) : ((PureU1Charges n).Charges →ₗ[] ) where
toFun S := ∑ i : Fin n, S i toFun S := ∑ i : Fin n, S i
map_add' S T := Finset.sum_add_distrib map_add' S T := Finset.sum_add_distrib
map_smul' a S := by map_smul' a S := by
@ -38,7 +38,7 @@ def accGrav (n : ) : ((PureU1Charges n).charges →ₗ[] ) where
/-- The symmetric trilinear form used to define the cubic anomaly. -/ /-- The symmetric trilinear form used to define the cubic anomaly. -/
@[simps!] @[simps!]
def accCubeTriLinSymm {n : } : TriLinearSymm (PureU1Charges n).charges := TriLinearSymm.mk₃ def accCubeTriLinSymm {n : } : TriLinearSymm (PureU1Charges n).Charges := TriLinearSymm.mk₃
(fun S => ∑ i, S.1 i * S.2.1 i * S.2.2 i) (fun S => ∑ i, S.1 i * S.2.1 i * S.2.2 i)
(by (by
intro a S L T intro a S L T
@ -74,11 +74,11 @@ def accCubeTriLinSymm {n : } : TriLinearSymm (PureU1Charges n).charges := Tri
/-- The cubic anomaly equation. -/ /-- The cubic anomaly equation. -/
@[simp] @[simp]
def accCube (n : ) : HomogeneousCubic ((PureU1Charges n).charges) := def accCube (n : ) : HomogeneousCubic ((PureU1Charges n).Charges) :=
(accCubeTriLinSymm).toCubic (accCubeTriLinSymm).toCubic
lemma accCube_explicit (n : ) (S : (PureU1Charges n).charges) : lemma accCube_explicit (n : ) (S : (PureU1Charges n).Charges) :
accCube n S = ∑ i : Fin n, S i ^ 3:= by accCube n S = ∑ i : Fin n, S i ^ 3:= by
rw [accCube, TriLinearSymm.toCubic] rw [accCube, TriLinearSymm.toCubic]
change accCubeTriLinSymm S S S = _ change accCubeTriLinSymm S S S = _
@ -104,7 +104,7 @@ def PureU1 (n : ) : ACCSystem where
/-- An equivalence of vector spaces of charges when the number of fermions is equal. -/ /-- An equivalence of vector spaces of charges when the number of fermions is equal. -/
def pureU1EqCharges {n m : } (h : n = m): def pureU1EqCharges {n m : } (h : n = m):
(PureU1 n).charges ≃ₗ[] (PureU1 m).charges where (PureU1 n).Charges ≃ₗ[] (PureU1 m).Charges where
toFun f := f ∘ Fin.cast h.symm toFun f := f ∘ Fin.cast h.symm
invFun f := f ∘ Fin.cast h invFun f := f ∘ Fin.cast h
map_add' _ _ := rfl map_add' _ _ := rfl
@ -154,7 +154,7 @@ lemma pureU1_anomalyFree_ext {n : } {S T : (PureU1 n.succ).LinSols}
namespace PureU1 namespace PureU1
lemma sum_of_charges {n : } (f : Fin k → (PureU1 n).charges) (j : Fin n) : lemma sum_of_charges {n : } (f : Fin k → (PureU1 n).Charges) (j : Fin n) :
(∑ i : Fin k, (f i)) j = ∑ i : Fin k, (f i) j := by (∑ i : Fin k, (f i)) j = ∑ i : Fin k, (f i) j := by
induction k induction k
simp simp

View file

@ -23,7 +23,7 @@ namespace BasisLinear
/-- The basis elements as charges, defined to have a `1` in the `j`th position and a `-1` in the /-- The basis elements as charges, defined to have a `1` in the `j`th position and a `-1` in the
last position. -/ last position. -/
@[simp] @[simp]
def asCharges (j : Fin n) : (PureU1 n.succ).charges := def asCharges (j : Fin n) : (PureU1 n.succ).Charges :=
(fun i => (fun i =>
if i = j.castSucc then if i = j.castSucc then
1 1

View file

@ -22,15 +22,15 @@ namespace PureU1
variable {n : } variable {n : }
/-- The condition for two rationals to have the same square (equivalent to same abs). -/ /-- The condition for two rationals to have the same square (equivalent to same abs). -/
def constAbsProp : × → Prop := fun s => s.1^2 = s.2^2 def ConstAbsProp : × → Prop := fun s => s.1^2 = s.2^2
/-- The condition on a charge assignment `S` to have constant absolute value among charges. -/ /-- The condition on a charge assignment `S` to have constant absolute value among charges. -/
@[simp] @[simp]
def constAbs (S : (PureU1 n).charges) : Prop := ∀ i j, (S i) ^ 2 = (S j) ^ 2 def ConstAbs (S : (PureU1 n).Charges) : Prop := ∀ i j, (S i) ^ 2 = (S j) ^ 2
lemma constAbs_perm (S : (PureU1 n).charges) (M :(FamilyPermutations n).group) : lemma constAbs_perm (S : (PureU1 n).Charges) (M :(FamilyPermutations n).group) :
constAbs ((FamilyPermutations n).rep M S) ↔ constAbs S := by ConstAbs ((FamilyPermutations n).rep M S) ↔ ConstAbs S := by
simp only [constAbs, PureU1_numberCharges, FamilyPermutations, permGroup, permCharges, simp only [ConstAbs, PureU1_numberCharges, FamilyPermutations, PermGroup, permCharges,
MonoidHom.coe_mk, OneHom.coe_mk, chargeMap_apply] MonoidHom.coe_mk, OneHom.coe_mk, chargeMap_apply]
apply Iff.intro apply Iff.intro
intro h i j intro h i j
@ -40,18 +40,18 @@ lemma constAbs_perm (S : (PureU1 n).charges) (M :(FamilyPermutations n).group) :
intro h i j intro h i j
exact h (M.invFun i) (M.invFun j) exact h (M.invFun i) (M.invFun j)
lemma constAbs_sort {S : (PureU1 n).charges} (CA : constAbs S) : constAbs (sort S) := by lemma constAbs_sort {S : (PureU1 n).Charges} (CA : ConstAbs S) : ConstAbs (sort S) := by
rw [sort] rw [sort]
exact (constAbs_perm S _).mpr CA exact (constAbs_perm S _).mpr CA
/-- The condition for a set of charges to be `sorted`, and have `constAbs`-/ /-- The condition for a set of charges to be `sorted`, and have `constAbs`-/
def constAbsSorted (S : (PureU1 n).charges) : Prop := constAbs S ∧ sorted S def ConstAbsSorted (S : (PureU1 n).Charges) : Prop := ConstAbs S ∧ Sorted S
namespace constAbsSorted namespace ConstAbsSorted
section charges section charges
variable {S : (PureU1 n.succ).charges} {A : (PureU1 n.succ).LinSols} variable {S : (PureU1 n.succ).Charges} {A : (PureU1 n.succ).LinSols}
variable (hS : constAbsSorted S) (hA : constAbsSorted A.val) variable (hS : ConstAbsSorted S) (hA : ConstAbsSorted A.val)
lemma lt_eq {k i : Fin n.succ} (hk : S k ≤ 0) (hik : i ≤ k) : S i = S k := by lemma lt_eq {k i : Fin n.succ} (hk : S k ≤ 0) (hik : i ≤ k) : S i = S k := by
have hSS := hS.2 i k hik have hSS := hS.2 i k hik
@ -98,12 +98,12 @@ lemma is_zero (h0 : S (0 : Fin n.succ) = 0) : S = 0 := by
is defined as a element of `k ∈ Fin n` such that `S k.castSucc` and `S k.succ` are different signs. is defined as a element of `k ∈ Fin n` such that `S k.castSucc` and `S k.succ` are different signs.
-/ -/
@[simp] @[simp]
def boundary (S : (PureU1 n.succ).charges) (k : Fin n) : Prop := S k.castSucc < 0 ∧ 0 < S k.succ def Boundary (S : (PureU1 n.succ).Charges) (k : Fin n) : Prop := S k.castSucc < 0 ∧ 0 < S k.succ
lemma boundary_castSucc {k : Fin n} (hk : boundary S k) : S k.castSucc = S (0 : Fin n.succ) := lemma boundary_castSucc {k : Fin n} (hk : Boundary S k) : S k.castSucc = S (0 : Fin n.succ) :=
(lt_eq hS (le_of_lt hk.left) (by simp : 0 ≤ k.castSucc)).symm (lt_eq hS (le_of_lt hk.left) (by simp : 0 ≤ k.castSucc)).symm
lemma boundary_succ {k : Fin n} (hk : boundary S k) : S k.succ = - S (0 : Fin n.succ) := by lemma boundary_succ {k : Fin n} (hk : Boundary S k) : S k.succ = - S (0 : Fin n.succ) := by
have hn := boundary_castSucc hS hk have hn := boundary_castSucc hS hk
rw [opposite_signs_eq_neg hS (le_of_lt hk.left) (le_of_lt hk.right)] at hn rw [opposite_signs_eq_neg hS (le_of_lt hk.left) (le_of_lt hk.right)] at hn
linear_combination -(1 * hn) linear_combination -(1 * hn)
@ -121,7 +121,7 @@ lemma boundary_accGrav' (k : Fin n) : accGrav n.succ S =
simp simp
lemma boundary_accGrav'' (k : Fin n) (hk : boundary S k) : lemma boundary_accGrav'' (k : Fin n) (hk : Boundary S k) :
accGrav n.succ S = (2 * ↑↑k + 1 - ↑n) * S (0 : Fin n.succ) := by accGrav n.succ S = (2 * ↑↑k + 1 - ↑n) * S (0 : Fin n.succ) := by
rw [boundary_accGrav' k] rw [boundary_accGrav' k]
rw [Fin.sum_univ_add] rw [Fin.sum_univ_add]
@ -139,10 +139,10 @@ lemma boundary_accGrav'' (k : Fin n) (hk : boundary S k) :
/-- We say a `S ∈ charges` has a boundary if there exists a `k ∈ Fin n` which is a boundary. -/ /-- We say a `S ∈ charges` has a boundary if there exists a `k ∈ Fin n` which is a boundary. -/
@[simp] @[simp]
def hasBoundary (S : (PureU1 n.succ).charges) : Prop := def HasBoundary (S : (PureU1 n.succ).Charges) : Prop :=
∃ (k : Fin n), boundary S k ∃ (k : Fin n), Boundary S k
lemma not_hasBoundary_zero_le (hnot : ¬ (hasBoundary S)) (h0 : S (0 : Fin n.succ) < 0) : lemma not_hasBoundary_zero_le (hnot : ¬ (HasBoundary S)) (h0 : S (0 : Fin n.succ) < 0) :
∀ i, S (0 : Fin n.succ) = S i := by ∀ i, S (0 : Fin n.succ) = S i := by
intro ⟨i, hi⟩ intro ⟨i, hi⟩
simp at hnot simp at hnot
@ -154,19 +154,19 @@ lemma not_hasBoundary_zero_le (hnot : ¬ (hasBoundary S)) (h0 : S (0 : Fin n.suc
erw [← hii] at hnott erw [← hii] at hnott
exact (val_le_zero hS (hnott h0)).symm exact (val_le_zero hS (hnott h0)).symm
lemma not_hasBoundry_zero (hnot : ¬ (hasBoundary S)) (i : Fin n.succ) : lemma not_hasBoundry_zero (hnot : ¬ (HasBoundary S)) (i : Fin n.succ) :
S (0 : Fin n.succ) = S i := by S (0 : Fin n.succ) = S i := by
by_cases hi : S (0 : Fin n.succ) < 0 by_cases hi : S (0 : Fin n.succ) < 0
exact not_hasBoundary_zero_le hS hnot hi i exact not_hasBoundary_zero_le hS hnot hi i
simp at hi simp at hi
exact zero_gt hS hi i exact zero_gt hS hi i
lemma not_hasBoundary_grav (hnot : ¬ (hasBoundary S)) : lemma not_hasBoundary_grav (hnot : ¬ (HasBoundary S)) :
accGrav n.succ S = n.succ * S (0 : Fin n.succ) := by accGrav n.succ S = n.succ * S (0 : Fin n.succ) := by
simp [accGrav, ← not_hasBoundry_zero hS hnot] simp [accGrav, ← not_hasBoundry_zero hS hnot]
lemma AFL_hasBoundary (h : A.val (0 : Fin n.succ) ≠ 0) : hasBoundary A.val := by lemma AFL_hasBoundary (h : A.val (0 : Fin n.succ) ≠ 0) : HasBoundary A.val := by
by_contra hn by_contra hn
have h0 := not_hasBoundary_grav hA hn have h0 := not_hasBoundary_grav hA hn
simp [accGrav] at h0 simp [accGrav] at h0
@ -176,8 +176,8 @@ lemma AFL_hasBoundary (h : A.val (0 : Fin n.succ) ≠ 0) : hasBoundary A.val :=
linarith linarith
simp_all simp_all
lemma AFL_odd_noBoundary {A : (PureU1 (2 * n + 1)).LinSols} (h : constAbsSorted A.val) lemma AFL_odd_noBoundary {A : (PureU1 (2 * n + 1)).LinSols} (h : ConstAbsSorted A.val)
(hA : A.val (0 : Fin (2*n +1)) ≠ 0) : ¬ hasBoundary A.val := by (hA : A.val (0 : Fin (2*n +1)) ≠ 0) : ¬ HasBoundary A.val := by
by_contra hn by_contra hn
obtain ⟨k, hk⟩ := hn obtain ⟨k, hk⟩ := hn
have h0 := boundary_accGrav'' h k hk have h0 := boundary_accGrav'' h k hk
@ -190,18 +190,18 @@ lemma AFL_odd_noBoundary {A : (PureU1 (2 * n + 1)).LinSols} (h : constAbsSorted
linear_combination - h0 linear_combination - h0
omega omega
lemma AFL_odd_zero {A : (PureU1 (2 * n + 1)).LinSols} (h : constAbsSorted A.val) : lemma AFL_odd_zero {A : (PureU1 (2 * n + 1)).LinSols} (h : ConstAbsSorted A.val) :
A.val (0 : Fin (2 * n + 1)) = 0 := by A.val (0 : Fin (2 * n + 1)) = 0 := by
by_contra hn by_contra hn
exact (AFL_odd_noBoundary h hn ) (AFL_hasBoundary h hn) exact (AFL_odd_noBoundary h hn ) (AFL_hasBoundary h hn)
theorem AFL_odd (A : (PureU1 (2 * n + 1)).LinSols) (h : constAbsSorted A.val) : theorem AFL_odd (A : (PureU1 (2 * n + 1)).LinSols) (h : ConstAbsSorted A.val) :
A = 0 := by A = 0 := by
apply ACCSystemLinear.LinSols.ext apply ACCSystemLinear.LinSols.ext
exact is_zero h (AFL_odd_zero h) exact is_zero h (AFL_odd_zero h)
lemma AFL_even_Boundary {A : (PureU1 (2 * n.succ)).LinSols} (h : constAbsSorted A.val) lemma AFL_even_Boundary {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted A.val)
(hA : A.val (0 : Fin (2 * n.succ)) ≠ 0) {k : Fin (2 * n + 1)} (hk : boundary A.val k) : (hA : A.val (0 : Fin (2 * n.succ)) ≠ 0) {k : Fin (2 * n + 1)} (hk : Boundary A.val k) :
k.val = n := by k.val = n := by
have h0 := boundary_accGrav'' h k hk have h0 := boundary_accGrav'' h k hk
change ∑ i : Fin (succ (Nat.mul 2 n + 1)), A.val i = _ at h0 change ∑ i : Fin (succ (Nat.mul 2 n + 1)), A.val i = _ at h0
@ -210,7 +210,7 @@ lemma AFL_even_Boundary {A : (PureU1 (2 * n.succ)).LinSols} (h : constAbsSorted
rw [← @Nat.cast_inj ] rw [← @Nat.cast_inj ]
linear_combination h0 / 2 linear_combination h0 / 2
lemma AFL_even_below' {A : (PureU1 (2 * n.succ)).LinSols} (h : constAbsSorted A.val) lemma AFL_even_below' {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted A.val)
(hA : A.val (0 : Fin (2 * n.succ)) ≠ 0) (i : Fin n.succ) : (hA : A.val (0 : Fin (2 * n.succ)) ≠ 0) (i : Fin n.succ) :
A.val (Fin.cast (split_equal n.succ) (Fin.castAdd n.succ i)) = A.val (0 : Fin (2*n.succ)) := by A.val (Fin.cast (split_equal n.succ) (Fin.castAdd n.succ i)) = A.val (0 : Fin (2*n.succ)) := by
obtain ⟨k, hk⟩ := AFL_hasBoundary h hA obtain ⟨k, hk⟩ := AFL_hasBoundary h hA
@ -221,7 +221,7 @@ lemma AFL_even_below' {A : (PureU1 (2 * n.succ)).LinSols} (h : constAbsSorted A.
rw [AFL_even_Boundary h hA hk] rw [AFL_even_Boundary h hA hk]
omega omega
lemma AFL_even_below (A : (PureU1 (2 * n.succ)).LinSols) (h : constAbsSorted A.val) lemma AFL_even_below (A : (PureU1 (2 * n.succ)).LinSols) (h : ConstAbsSorted A.val)
(i : Fin n.succ) : (i : Fin n.succ) :
A.val (Fin.cast (split_equal n.succ) (Fin.castAdd n.succ i)) A.val (Fin.cast (split_equal n.succ) (Fin.castAdd n.succ i))
= A.val (0 : Fin (2*n.succ)) := by = A.val (0 : Fin (2*n.succ)) := by
@ -231,7 +231,7 @@ lemma AFL_even_below (A : (PureU1 (2 * n.succ)).LinSols) (h : constAbsSorted A.v
rfl rfl
exact AFL_even_below' h hA i exact AFL_even_below' h hA i
lemma AFL_even_above' {A : (PureU1 (2 * n.succ)).LinSols} (h : constAbsSorted A.val) lemma AFL_even_above' {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted A.val)
(hA : A.val (0 : Fin (2*n.succ)) ≠ 0) (i : Fin n.succ) : (hA : A.val (0 : Fin (2*n.succ)) ≠ 0) (i : Fin n.succ) :
A.val (Fin.cast (split_equal n.succ) (Fin.natAdd n.succ i)) = A.val (Fin.cast (split_equal n.succ) (Fin.natAdd n.succ i)) =
- A.val (0 : Fin (2*n.succ)) := by - A.val (0 : Fin (2*n.succ)) := by
@ -243,7 +243,7 @@ lemma AFL_even_above' {A : (PureU1 (2 * n.succ)).LinSols} (h : constAbsSorted A.
rw [AFL_even_Boundary h hA hk] rw [AFL_even_Boundary h hA hk]
omega omega
lemma AFL_even_above (A : (PureU1 (2 * n.succ)).LinSols) (h : constAbsSorted A.val) lemma AFL_even_above (A : (PureU1 (2 * n.succ)).LinSols) (h : ConstAbsSorted A.val)
(i : Fin n.succ) : (i : Fin n.succ) :
A.val (Fin.cast (split_equal n.succ) (Fin.natAdd n.succ i)) = A.val (Fin.cast (split_equal n.succ) (Fin.natAdd n.succ i)) =
- A.val (0 : Fin (2*n.succ)) := by - A.val (0 : Fin (2*n.succ)) := by
@ -256,23 +256,23 @@ lemma AFL_even_above (A : (PureU1 (2 * n.succ)).LinSols) (h : constAbsSorted A.v
end charges end charges
end constAbsSorted end ConstAbsSorted
namespace ConstAbs namespace ConstAbs
theorem boundary_value_odd (S : (PureU1 (2 * n + 1)).LinSols) (hs : constAbs S.val) : theorem boundary_value_odd (S : (PureU1 (2 * n + 1)).LinSols) (hs : ConstAbs S.val) :
S = 0 := S = 0 :=
have hS := And.intro (constAbs_sort hs) (sort_sorted S.val) have hS := And.intro (constAbs_sort hs) (sort_sorted S.val)
sortAFL_zero S (constAbsSorted.AFL_odd (sortAFL S) hS) sortAFL_zero S (ConstAbsSorted.AFL_odd (sortAFL S) hS)
theorem boundary_value_even (S : (PureU1 (2 * n.succ)).LinSols) (hs : constAbs S.val) : theorem boundary_value_even (S : (PureU1 (2 * n.succ)).LinSols) (hs : ConstAbs S.val) :
vectorLikeEven S.val := by VectorLikeEven S.val := by
have hS := And.intro (constAbs_sort hs) (sort_sorted S.val) have hS := And.intro (constAbs_sort hs) (sort_sorted S.val)
intro i intro i
have h1 := constAbsSorted.AFL_even_below (sortAFL S) hS have h1 := ConstAbsSorted.AFL_even_below (sortAFL S) hS
have h2 := constAbsSorted.AFL_even_above (sortAFL S) hS have h2 := ConstAbsSorted.AFL_even_above (sortAFL S) hS
rw [sortAFL_val] at h1 h2 rw [sortAFL_val] at h1 h2
rw [h1, h2] rw [h1, h2]
simp simp

View file

@ -135,7 +135,7 @@ lemma δ!₂_δ₂ (j : Fin n) : δ!₂ j = δ₂ j.castSucc := by
end theδs end theδs
/-- The first part of the basis as charges. -/ /-- The first part of the basis as charges. -/
def basisAsCharges (j : Fin n.succ) : (PureU1 (2 * n.succ)).charges := def basisAsCharges (j : Fin n.succ) : (PureU1 (2 * n.succ)).Charges :=
fun i => fun i =>
if i = δ₁ j then if i = δ₁ j then
1 1
@ -146,7 +146,7 @@ def basisAsCharges (j : Fin n.succ) : (PureU1 (2 * n.succ)).charges :=
0 0
/-- The second part of the basis as charges. -/ /-- The second part of the basis as charges. -/
def basis!AsCharges (j : Fin n) : (PureU1 (2 * n.succ)).charges := def basis!AsCharges (j : Fin n) : (PureU1 (2 * n.succ)).Charges :=
fun i => fun i =>
if i = δ!₁ j then if i = δ!₁ j then
1 1
@ -357,13 +357,13 @@ lemma swap!_as_add {S S' : (PureU1 (2 * n.succ)).LinSols} (j : Fin n)
simp simp
/-- A point in the span of the first part of the basis as a charge. -/ /-- A point in the span of the first part of the basis as a charge. -/
def P (f : Fin n.succ → ) : (PureU1 (2 * n.succ)).charges := ∑ i, f i • basisAsCharges i def P (f : Fin n.succ → ) : (PureU1 (2 * n.succ)).Charges := ∑ i, f i • basisAsCharges i
/-- A point in the span of the second part of the basis as a charge. -/ /-- A point in the span of the second part of the basis as a charge. -/
def P! (f : Fin n → ) : (PureU1 (2 * n.succ)).charges := ∑ i, f i • basis!AsCharges i def P! (f : Fin n → ) : (PureU1 (2 * n.succ)).Charges := ∑ i, f i • basis!AsCharges i
/-- A point in the span of the basis as a charge. -/ /-- A point in the span of the basis as a charge. -/
def Pa (f : Fin n.succ → ) (g : Fin n → ) : (PureU1 (2 * n.succ)).charges := P f + P! g def Pa (f : Fin n.succ → ) (g : Fin n → ) : (PureU1 (2 * n.succ)).Charges := P f + P! g
lemma P_δ₁ (f : Fin n.succ → ) (j : Fin n.succ) : P f (δ₁ j) = f j := by lemma P_δ₁ (f : Fin n.succ → ) (j : Fin n.succ) : P f (δ₁ j) = f j := by
rw [P, sum_of_charges] rw [P, sum_of_charges]
@ -727,7 +727,7 @@ lemma span_basis_swap! {S : (PureU1 (2 * n.succ)).LinSols} (j : Fin n)
exact hS exact hS
lemma vectorLikeEven_in_span (S : (PureU1 (2 * n.succ)).LinSols) lemma vectorLikeEven_in_span (S : (PureU1 (2 * n.succ)).LinSols)
(hS : vectorLikeEven S.val) : (hS : VectorLikeEven S.val) :
∃ (M : (FamilyPermutations (2 * n.succ)).group), ∃ (M : (FamilyPermutations (2 * n.succ)).group),
(FamilyPermutations (2 * n.succ)).linSolRep M S (FamilyPermutations (2 * n.succ)).linSolRep M S
∈ Submodule.span (Set.range basis) := by ∈ Submodule.span (Set.range basis) := by

View file

@ -36,11 +36,11 @@ open VectorLikeEvenPlane
/-- A property on `LinSols`, satisfied 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. -/ in the basis through that point is in the cubic. -/
def lineInCubic (S : (PureU1 (2 * n.succ)).LinSols) : Prop := def LineInCubic (S : (PureU1 (2 * n.succ)).LinSols) : Prop :=
∀ (g : Fin n.succ → ) (f : Fin n → ) (_ : S.val = Pa g f) (a b : ) , ∀ (g : Fin n.succ → ) (f : Fin n → ) (_ : S.val = Pa g f) (a b : ) ,
accCube (2 * n.succ) (a • P g + b • P! f) = 0 accCube (2 * n.succ) (a • P g + b • P! f) = 0
lemma lineInCubic_expand {S : (PureU1 (2 * n.succ)).LinSols} (h : lineInCubic S) : lemma lineInCubic_expand {S : (PureU1 (2 * n.succ)).LinSols} (h : LineInCubic S) :
∀ (g : Fin n.succ → ) (f : Fin n → ) (_ : S.val = Pa g f) (a b : ) , ∀ (g : Fin n.succ → ) (f : Fin n → ) (_ : S.val = Pa g f) (a b : ) ,
3 * a * b * (a * accCubeTriLinSymm (P g) (P g) (P! f) 3 * a * b * (a * accCubeTriLinSymm (P g) (P g) (P! f)
+ b * accCubeTriLinSymm (P! f) (P! f) (P g)) = 0 := by + b * accCubeTriLinSymm (P! f) (P! f) (P g)) = 0 := by
@ -60,7 +60,7 @@ lemma lineInCubic_expand {S : (PureU1 (2 * n.succ)).LinSols} (h : lineInCubic S)
for any functions `g : Fin n.succ → ` and `f : Fin n → `, if `S.val = P g + P! f`, for any functions `g : Fin n.succ → ` and `f : Fin n → `, if `S.val = P g + P! f`,
then `accCubeTriLinSymm.toFun (P g, P g, P! f) = 0`. then `accCubeTriLinSymm.toFun (P g, P g, P! f) = 0`.
-/ -/
lemma line_in_cubic_P_P_P! {S : (PureU1 (2 * n.succ)).LinSols} (h : lineInCubic S) : lemma line_in_cubic_P_P_P! {S : (PureU1 (2 * n.succ)).LinSols} (h : LineInCubic S) :
∀ (g : Fin n.succ → ) (f : Fin n → ) (_ : S.val = P g + P! f), ∀ (g : Fin n.succ → ) (f : Fin n → ) (_ : S.val = P g + P! f),
accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by
intro g f hS intro g f hS
@ -68,28 +68,28 @@ lemma line_in_cubic_P_P_P! {S : (PureU1 (2 * n.succ)).LinSols} (h : lineInCubic
(lineInCubic_expand h g f hS 1 2) / 6 (lineInCubic_expand h g f hS 1 2) / 6
/-- We say a `LinSol` satisfies `lineInCubicPerm` if all its permutations satisfy `lineInCubic`. -/ /-- We say a `LinSol` satisfies `lineInCubicPerm` if all its permutations satisfy `lineInCubic`. -/
def lineInCubicPerm (S : (PureU1 (2 * n.succ)).LinSols) : Prop := def LineInCubicPerm (S : (PureU1 (2 * n.succ)).LinSols) : Prop :=
∀ (M : (FamilyPermutations (2 * n.succ)).group ), ∀ (M : (FamilyPermutations (2 * n.succ)).group ),
lineInCubic ((FamilyPermutations (2 * n.succ)).linSolRep M S) LineInCubic ((FamilyPermutations (2 * n.succ)).linSolRep M S)
/-- If `lineInCubicPerm S` then `lineInCubic S`. -/ /-- If `lineInCubicPerm S` then `lineInCubic S`. -/
lemma lineInCubicPerm_self {S : (PureU1 (2 * n.succ)).LinSols} lemma lineInCubicPerm_self {S : (PureU1 (2 * n.succ)).LinSols}
(hS : lineInCubicPerm S) : lineInCubic S := hS 1 (hS : LineInCubicPerm S) : LineInCubic S := hS 1
/-- If `lineInCubicPerm S` then `lineInCubicPerm (M S)` for all permutations `M`. -/ /-- If `lineInCubicPerm S` then `lineInCubicPerm (M S)` for all permutations `M`. -/
lemma lineInCubicPerm_permute {S : (PureU1 (2 * n.succ)).LinSols} lemma lineInCubicPerm_permute {S : (PureU1 (2 * n.succ)).LinSols}
(hS : lineInCubicPerm S) (M' : (FamilyPermutations (2 * n.succ)).group) : (hS : LineInCubicPerm S) (M' : (FamilyPermutations (2 * n.succ)).group) :
lineInCubicPerm ((FamilyPermutations (2 * n.succ)).linSolRep M' S) := by LineInCubicPerm ((FamilyPermutations (2 * n.succ)).linSolRep M' S) := by
rw [lineInCubicPerm] rw [LineInCubicPerm]
intro M intro M
change lineInCubic change LineInCubic
(((FamilyPermutations (2 * n.succ)).linSolRep M * (((FamilyPermutations (2 * n.succ)).linSolRep M *
(FamilyPermutations (2 * n.succ)).linSolRep M') S) (FamilyPermutations (2 * n.succ)).linSolRep M') S)
erw [← (FamilyPermutations (2 * n.succ)).linSolRep.map_mul M M'] erw [← (FamilyPermutations (2 * n.succ)).linSolRep.map_mul M M']
exact hS (M * M') exact hS (M * M')
lemma lineInCubicPerm_swap {S : (PureU1 (2 * n.succ)).LinSols} lemma lineInCubicPerm_swap {S : (PureU1 (2 * n.succ)).LinSols}
(LIC : lineInCubicPerm S) : (LIC : LineInCubicPerm S) :
∀ (j : Fin n) (g : Fin n.succ → ) (f : Fin n → ) (_ : S.val = Pa g f) , ∀ (j : Fin n) (g : Fin n.succ → ) (f : Fin n → ) (_ : S.val = Pa g f) ,
(S.val (δ!₂ j) - S.val (δ!₁ j)) (S.val (δ!₂ j) - S.val (δ!₁ j))
* accCubeTriLinSymm (P g) (P g) (basis!AsCharges j) = 0 := by * accCubeTriLinSymm (P g) (P g) (basis!AsCharges j) = 0 := by
@ -127,8 +127,8 @@ lemma P_P_P!_accCube' {S : (PureU1 (2 * n.succ.succ )).LinSols}
ring ring
lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ)).LinSols} lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ)).LinSols}
(LIC : lineInCubicPerm S) : (LIC : LineInCubicPerm S) :
lineInPlaneProp LineInPlaneProp
((S.val (δ!₂ (Fin.last n))), ((S.val (δ!₁ (Fin.last n))), (S.val δ!₄))) := by ((S.val (δ!₂ (Fin.last n))), ((S.val (δ!₁ (Fin.last n))), (S.val δ!₄))) := by
obtain ⟨g, f, hfg⟩ := span_basis S obtain ⟨g, f, hfg⟩ := span_basis S
have h1 := lineInCubicPerm_swap LIC (Fin.last n) g f hfg have h1 := lineInCubicPerm_swap LIC (Fin.last n) g f hfg
@ -146,8 +146,8 @@ lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ)).LinSols}
exact h1 exact h1
lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ)).LinSols} lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ)).LinSols}
(LIC : lineInCubicPerm S) : lineInPlaneCond S := by (LIC : LineInCubicPerm S) : LineInPlaneCond S := by
refine @Prop_three (2 * n.succ.succ) lineInPlaneProp S (δ!₂ (Fin.last n)) (δ!₁ (Fin.last n)) refine @Prop_three (2 * n.succ.succ) LineInPlaneProp S (δ!₂ (Fin.last n)) (δ!₁ (Fin.last n))
δ!₄ ?_ ?_ ?_ ?_ δ!₄ ?_ ?_ ?_ ?_
simp [Fin.ext_iff, δ!₂, δ!₁] simp [Fin.ext_iff, δ!₂, δ!₁]
simp [Fin.ext_iff, δ!₂, δ!₄] simp [Fin.ext_iff, δ!₂, δ!₄]
@ -157,15 +157,15 @@ lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ)).LinSols}
exact lineInCubicPerm_last_cond (lineInCubicPerm_permute LIC M) exact lineInCubicPerm_last_cond (lineInCubicPerm_permute LIC M)
lemma lineInCubicPerm_constAbs {S : (PureU1 (2 * n.succ.succ)).Sols} lemma lineInCubicPerm_constAbs {S : (PureU1 (2 * n.succ.succ)).Sols}
(LIC : lineInCubicPerm S.1.1) : constAbs S.val := (LIC : LineInCubicPerm S.1.1) : ConstAbs S.val :=
linesInPlane_constAbs_AF S (lineInCubicPerm_last_perm LIC) linesInPlane_constAbs_AF S (lineInCubicPerm_last_perm LIC)
theorem lineInCubicPerm_vectorLike {S : (PureU1 (2 * n.succ.succ)).Sols} theorem lineInCubicPerm_vectorLike {S : (PureU1 (2 * n.succ.succ)).Sols}
(LIC : lineInCubicPerm S.1.1) : vectorLikeEven S.val := (LIC : LineInCubicPerm S.1.1) : VectorLikeEven S.val :=
ConstAbs.boundary_value_even S.1.1 (lineInCubicPerm_constAbs LIC) ConstAbs.boundary_value_even S.1.1 (lineInCubicPerm_constAbs LIC)
theorem lineInCubicPerm_in_plane (S : (PureU1 (2 * n.succ.succ)).Sols) theorem lineInCubicPerm_in_plane (S : (PureU1 (2 * n.succ.succ)).Sols)
(LIC : lineInCubicPerm S.1.1) : ∃ (M : (FamilyPermutations (2 * n.succ.succ)).group), (LIC : LineInCubicPerm S.1.1) : ∃ (M : (FamilyPermutations (2 * n.succ.succ)).group),
(FamilyPermutations (2 * n.succ.succ)).linSolRep M S.1.1 (FamilyPermutations (2 * n.succ.succ)).linSolRep M S.1.1
∈ Submodule.span (Set.range basis) := ∈ Submodule.span (Set.range basis) :=
vectorLikeEven_in_span S.1.1 (lineInCubicPerm_vectorLike LIC) vectorLikeEven_in_span S.1.1 (lineInCubicPerm_vectorLike LIC)

View file

@ -79,13 +79,13 @@ lemma anomalyFree_param {S : (PureU1 (2 * n.succ)).Sols}
/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`. /-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`.
In this case our parameterization above will be able to recover this point. -/ In this case our parameterization above will be able to recover this point. -/
def genericCase (S : (PureU1 (2 * n.succ)).Sols) : Prop := def GenericCase (S : (PureU1 (2 * n.succ)).Sols) : Prop :=
∀ (g : Fin n.succ → ) (f : Fin n → ) (_ : S.val = P g + P! f) , ∀ (g : Fin n.succ → ) (f : Fin n → ) (_ : S.val = P g + P! f) ,
accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0 accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0
lemma genericCase_exists (S : (PureU1 (2 * n.succ)).Sols) lemma genericCase_exists (S : (PureU1 (2 * n.succ)).Sols)
(hs : ∃ (g : Fin n.succ → ) (f : Fin n → ), S.val = P g + P! f ∧ (hs : ∃ (g : Fin n.succ → ) (f : Fin n → ), S.val = P g + P! f ∧
accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0) : genericCase S := by accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0) : GenericCase S := by
intro g f hS hC intro g f hS hC
obtain ⟨g', f', hS', hC'⟩ := hs obtain ⟨g', f', hS', hC'⟩ := hs
rw [hS] at hS' rw [hS] at hS'
@ -94,13 +94,13 @@ lemma genericCase_exists (S : (PureU1 (2 * n.succ)).Sols)
exact hC' hC exact hC' hC
/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) = 0`.-/ /-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) = 0`.-/
def specialCase (S : (PureU1 (2 * n.succ)).Sols) : Prop := def SpecialCase (S : (PureU1 (2 * n.succ)).Sols) : Prop :=
∀ (g : Fin n.succ → ) (f : Fin n → ) (_ : S.val = P g + P! f) , ∀ (g : Fin n.succ → ) (f : Fin n → ) (_ : S.val = P g + P! f) ,
accCubeTriLinSymm (P g) (P g) (P! f) = 0 accCubeTriLinSymm (P g) (P g) (P! f) = 0
lemma specialCase_exists (S : (PureU1 (2 * n.succ)).Sols) lemma specialCase_exists (S : (PureU1 (2 * n.succ)).Sols)
(hs : ∃ (g : Fin n.succ → ) (f : Fin n → ), S.val = P g + P! f ∧ (hs : ∃ (g : Fin n.succ → ) (f : Fin n → ), S.val = P g + P! f ∧
accCubeTriLinSymm (P g) (P g) (P! f) = 0) : specialCase S := by accCubeTriLinSymm (P g) (P g) (P! f) = 0) : SpecialCase S := by
intro g f hS intro g f hS
obtain ⟨g', f', hS', hC'⟩ := hs obtain ⟨g', f', hS', hC'⟩ := hs
rw [hS] at hS' rw [hS] at hS'
@ -109,7 +109,7 @@ lemma specialCase_exists (S : (PureU1 (2 * n.succ)).Sols)
exact hC' exact hC'
lemma generic_or_special (S : (PureU1 (2 * n.succ)).Sols) : lemma generic_or_special (S : (PureU1 (2 * n.succ)).Sols) :
genericCase S specialCase S := by GenericCase S SpecialCase S := by
obtain ⟨g, f, h⟩ := span_basis S.1.1 obtain ⟨g, f, h⟩ := span_basis S.1.1
have h1 : accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0 have h1 : accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0
accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by
@ -118,7 +118,7 @@ lemma generic_or_special (S : (PureU1 (2 * n.succ)).Sols) :
exact Or.inl (genericCase_exists S ⟨g, f, h, h1⟩) exact Or.inl (genericCase_exists S ⟨g, f, h, h1⟩)
exact Or.inr (specialCase_exists S ⟨g, f, h, h1⟩) exact Or.inr (specialCase_exists S ⟨g, f, h, h1⟩)
theorem generic_case {S : (PureU1 (2 * n.succ)).Sols} (h : genericCase S) : theorem generic_case {S : (PureU1 (2 * n.succ)).Sols} (h : GenericCase S) :
∃ g f a, S = parameterization g f a := by ∃ g f a, S = parameterization g f a := by
obtain ⟨g, f, hS⟩ := span_basis S.1.1 obtain ⟨g, f, hS⟩ := span_basis S.1.1
use g, f, (accCubeTriLinSymm (P! f) (P! f) (P g))⁻¹ use g, f, (accCubeTriLinSymm (P! f) (P! f) (P g))⁻¹
@ -136,7 +136,7 @@ theorem generic_case {S : (PureU1 (2 * n.succ)).Sols} (h : genericCase S) :
lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ)).Sols} lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ)).Sols}
(h : specialCase S) : lineInCubic S.1.1 := by (h : SpecialCase S) : LineInCubic S.1.1 := by
intro g f hS a b intro g f hS a b
erw [TriLinearSymm.toCubic_add] erw [TriLinearSymm.toCubic_add]
rw [HomogeneousCubic.map_smul, HomogeneousCubic.map_smul] rw [HomogeneousCubic.map_smul, HomogeneousCubic.map_smul]
@ -155,15 +155,15 @@ lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ)).Sols}
lemma special_case_lineInCubic_perm {S : (PureU1 (2 * n.succ)).Sols} lemma special_case_lineInCubic_perm {S : (PureU1 (2 * n.succ)).Sols}
(h : ∀ (M : (FamilyPermutations (2 * n.succ)).group), (h : ∀ (M : (FamilyPermutations (2 * n.succ)).group),
specialCase ((FamilyPermutations (2 * n.succ)).solAction.toFun S M)) : SpecialCase ((FamilyPermutations (2 * n.succ)).solAction.toFun S M)) :
lineInCubicPerm S.1.1 := by LineInCubicPerm S.1.1 := by
intro M intro M
exact special_case_lineInCubic (h M) exact special_case_lineInCubic (h M)
theorem special_case {S : (PureU1 (2 * n.succ.succ)).Sols} theorem special_case {S : (PureU1 (2 * n.succ.succ)).Sols}
(h : ∀ (M : (FamilyPermutations (2 * n.succ.succ)).group), (h : ∀ (M : (FamilyPermutations (2 * n.succ.succ)).group),
specialCase ((FamilyPermutations (2 * n.succ.succ)).solAction.toFun S M)) : SpecialCase ((FamilyPermutations (2 * n.succ.succ)).solAction.toFun S M)) :
∃ (M : (FamilyPermutations (2 * n.succ.succ)).group), ∃ (M : (FamilyPermutations (2 * n.succ.succ)).group),
((FamilyPermutations (2 * n.succ.succ)).solAction.toFun S M).1.1 ((FamilyPermutations (2 * n.succ.succ)).solAction.toFun S M).1.1
∈ Submodule.span (Set.range basis) := ∈ Submodule.span (Set.range basis) :=

View file

@ -34,17 +34,17 @@ open BigOperators
variable {n : } variable {n : }
/-- The proposition on three rationals to satisfy the `linInPlane` condition. -/ /-- The proposition on three rationals to satisfy the `linInPlane` condition. -/
def lineInPlaneProp : × × → Prop := fun s => def LineInPlaneProp : × × → Prop := fun s =>
s.1 = s.2.1 s.1 = - s.2.1 2 * s.2.2 + s.1 + s.2.1 = 0 s.1 = s.2.1 s.1 = - s.2.1 2 * s.2.2 + s.1 + s.2.1 = 0
/-- The proposition on a `LinSol` to satisfy the `linInPlane` condition. -/ /-- The proposition on a `LinSol` to satisfy the `linInPlane` condition. -/
def lineInPlaneCond (S : (PureU1 (n)).LinSols) : Prop := def LineInPlaneCond (S : (PureU1 (n)).LinSols) : Prop :=
∀ (i1 i2 i3 : Fin (n)) (_ : i1 ≠ i2) (_ : i2 ≠ i3) (_ : i1 ≠ i3), ∀ (i1 i2 i3 : Fin (n)) (_ : i1 ≠ i2) (_ : i2 ≠ i3) (_ : i1 ≠ i3),
lineInPlaneProp (S.val i1, (S.val i2, S.val i3)) LineInPlaneProp (S.val i1, (S.val i2, S.val i3))
lemma lineInPlaneCond_perm {S : (PureU1 (n)).LinSols} (hS : lineInPlaneCond S) lemma lineInPlaneCond_perm {S : (PureU1 (n)).LinSols} (hS : LineInPlaneCond S)
(M : (FamilyPermutations n).group) : (M : (FamilyPermutations n).group) :
lineInPlaneCond ((FamilyPermutations n).linSolRep M S) := by LineInPlaneCond ((FamilyPermutations n).linSolRep M S) := by
intro i1 i2 i3 h1 h2 h3 intro i1 i2 i3 h1 h2 h3
rw [FamilyPermutations_anomalyFreeLinear_apply, FamilyPermutations_anomalyFreeLinear_apply, rw [FamilyPermutations_anomalyFreeLinear_apply, FamilyPermutations_anomalyFreeLinear_apply,
FamilyPermutations_anomalyFreeLinear_apply] FamilyPermutations_anomalyFreeLinear_apply]
@ -53,13 +53,13 @@ lemma lineInPlaneCond_perm {S : (PureU1 (n)).LinSols} (hS : lineInPlaneCond S)
not_false_eq_true] not_false_eq_true]
lemma lineInPlaneCond_eq_last' {S : (PureU1 (n.succ.succ)).LinSols} (hS : lineInPlaneCond S) lemma lineInPlaneCond_eq_last' {S : (PureU1 (n.succ.succ)).LinSols} (hS : LineInPlaneCond S)
(h : ¬ (S.val ((Fin.last n).castSucc))^2 = (S.val ((Fin.last n).succ))^2 ) : (h : ¬ (S.val ((Fin.last n).castSucc))^2 = (S.val ((Fin.last n).succ))^2 ) :
(2 - n) * S.val (Fin.last (n + 1)) = (2 - n) * S.val (Fin.last (n + 1)) =
- (2 - n)* S.val (Fin.castSucc (Fin.last n)) := by - (2 - n)* S.val (Fin.castSucc (Fin.last n)) := by
erw [sq_eq_sq_iff_eq_or_eq_neg] at h erw [sq_eq_sq_iff_eq_or_eq_neg] at h
rw [lineInPlaneCond] at hS rw [LineInPlaneCond] at hS
simp only [lineInPlaneProp] at hS simp only [LineInPlaneProp] at hS
simp [not_or] at h simp [not_or] at h
have h1 (i : Fin n) : S.val i.castSucc.castSucc = have h1 (i : Fin n) : S.val i.castSucc.castSucc =
- (S.val ((Fin.last n).castSucc) + (S.val ((Fin.last n).succ))) / 2 := by - (S.val ((Fin.last n).castSucc) + (S.val ((Fin.last n).succ))) / 2 := by
@ -77,9 +77,9 @@ lemma lineInPlaneCond_eq_last' {S : (PureU1 (n.succ.succ)).LinSols} (hS : lineIn
linear_combination h2 linear_combination h2
lemma lineInPlaneCond_eq_last {S : (PureU1 (n.succ.succ.succ.succ.succ)).LinSols} lemma lineInPlaneCond_eq_last {S : (PureU1 (n.succ.succ.succ.succ.succ)).LinSols}
(hS : lineInPlaneCond S) : constAbsProp ((S.val ((Fin.last n.succ.succ.succ).castSucc)), (hS : LineInPlaneCond S) : ConstAbsProp ((S.val ((Fin.last n.succ.succ.succ).castSucc)),
(S.val ((Fin.last n.succ.succ.succ).succ))) := by (S.val ((Fin.last n.succ.succ.succ).succ))) := by
rw [constAbsProp] rw [ConstAbsProp]
by_contra hn by_contra hn
have h := lineInPlaneCond_eq_last' hS hn have h := lineInPlaneCond_eq_last' hS hn
rw [sq_eq_sq_iff_eq_or_eq_neg] at hn rw [sq_eq_sq_iff_eq_or_eq_neg] at hn
@ -100,25 +100,25 @@ lemma lineInPlaneCond_eq_last {S : (PureU1 (n.succ.succ.succ.succ.succ)).LinSols
simp_all simp_all
lemma linesInPlane_eq_sq {S : (PureU1 (n.succ.succ.succ.succ.succ)).LinSols} lemma linesInPlane_eq_sq {S : (PureU1 (n.succ.succ.succ.succ.succ)).LinSols}
(hS : lineInPlaneCond S) : ∀ (i j : Fin n.succ.succ.succ.succ.succ) (_ : i ≠ j), (hS : LineInPlaneCond S) : ∀ (i j : Fin n.succ.succ.succ.succ.succ) (_ : i ≠ j),
constAbsProp (S.val i, S.val j) := by ConstAbsProp (S.val i, S.val j) := by
have hneq : ((Fin.last n.succ.succ.succ).castSucc) ≠ ((Fin.last n.succ.succ.succ).succ) := by have hneq : ((Fin.last n.succ.succ.succ).castSucc) ≠ ((Fin.last n.succ.succ.succ).succ) := by
simp [Fin.ext_iff] simp [Fin.ext_iff]
refine Prop_two constAbsProp hneq ?_ refine Prop_two ConstAbsProp hneq ?_
intro M intro M
exact lineInPlaneCond_eq_last (lineInPlaneCond_perm hS M) exact lineInPlaneCond_eq_last (lineInPlaneCond_perm hS M)
theorem linesInPlane_constAbs {S : (PureU1 (n.succ.succ.succ.succ.succ)).LinSols} theorem linesInPlane_constAbs {S : (PureU1 (n.succ.succ.succ.succ.succ)).LinSols}
(hS : lineInPlaneCond S) : constAbs S.val := by (hS : LineInPlaneCond S) : ConstAbs S.val := by
intro i j intro i j
by_cases hij : i ≠ j by_cases hij : i ≠ j
exact linesInPlane_eq_sq hS i j hij exact linesInPlane_eq_sq hS i j hij
simp at hij simp at hij
rw [hij] rw [hij]
lemma linesInPlane_four (S : (PureU1 4).Sols) (hS : lineInPlaneCond S.1.1) : lemma linesInPlane_four (S : (PureU1 4).Sols) (hS : LineInPlaneCond S.1.1) :
constAbsProp (S.val (0 : Fin 4), S.val (1 : Fin 4)) := by ConstAbsProp (S.val (0 : Fin 4), S.val (1 : Fin 4)) := by
simp [constAbsProp] simp [ConstAbsProp]
by_contra hn by_contra hn
have hLin := pureU1_linear S.1.1 have hLin := pureU1_linear S.1.1
have hcube := pureU1_cube S have hcube := pureU1_cube S
@ -129,7 +129,7 @@ lemma linesInPlane_four (S : (PureU1 4).Sols) (hS : lineInPlaneCond S.1.1) :
have l012 := hS 0 1 2 (by simp) (by simp) (by simp) have l012 := hS 0 1 2 (by simp) (by simp) (by simp)
have l013 := hS 0 1 3 (by simp) (by simp) (by simp) have l013 := hS 0 1 3 (by simp) (by simp) (by simp)
have l023 := hS 0 2 3 (by simp) (by simp) (by simp) have l023 := hS 0 2 3 (by simp) (by simp) (by simp)
simp_all [lineInPlaneProp] simp_all [LineInPlaneProp]
have h1 : S.val (2 : Fin 4) = S.val (3 : Fin 4) := by have h1 : S.val (2 : Fin 4) = S.val (3 : Fin 4) := by
linear_combination l012 / 2 + -1 * l013 / 2 linear_combination l012 / 2 + -1 * l013 / 2
by_cases h2 : S.val (0 : Fin 4) = S.val (2 : Fin 4) by_cases h2 : S.val (0 : Fin 4) = S.val (2 : Fin 4)
@ -159,18 +159,18 @@ lemma linesInPlane_four (S : (PureU1 4).Sols) (hS : lineInPlaneCond S.1.1) :
lemma linesInPlane_eq_sq_four {S : (PureU1 4).Sols} lemma linesInPlane_eq_sq_four {S : (PureU1 4).Sols}
(hS : lineInPlaneCond S.1.1) : ∀ (i j : Fin 4) (_ : i ≠ j), (hS : LineInPlaneCond S.1.1) : ∀ (i j : Fin 4) (_ : i ≠ j),
constAbsProp (S.val i, S.val j) := by ConstAbsProp (S.val i, S.val j) := by
refine Prop_two constAbsProp (by simp : (0 : Fin 4) ≠ 1) ?_ refine Prop_two ConstAbsProp (by simp : (0 : Fin 4) ≠ 1) ?_
intro M intro M
let S' := (FamilyPermutations 4).solAction.toFun S M let S' := (FamilyPermutations 4).solAction.toFun S M
have hS' : lineInPlaneCond S'.1.1 := have hS' : LineInPlaneCond S'.1.1 :=
(lineInPlaneCond_perm hS M) (lineInPlaneCond_perm hS M)
exact linesInPlane_four S' hS' exact linesInPlane_four S' hS'
lemma linesInPlane_constAbs_four (S : (PureU1 4).Sols) lemma linesInPlane_constAbs_four (S : (PureU1 4).Sols)
(hS : lineInPlaneCond S.1.1) : constAbs S.val := by (hS : LineInPlaneCond S.1.1) : ConstAbs S.val := by
intro i j intro i j
by_cases hij : i ≠ j by_cases hij : i ≠ j
exact linesInPlane_eq_sq_four hS i j hij exact linesInPlane_eq_sq_four hS i j hij
@ -178,7 +178,7 @@ lemma linesInPlane_constAbs_four (S : (PureU1 4).Sols)
rw [hij] rw [hij]
theorem linesInPlane_constAbs_AF (S : (PureU1 (n.succ.succ.succ.succ)).Sols) theorem linesInPlane_constAbs_AF (S : (PureU1 (n.succ.succ.succ.succ)).Sols)
(hS : lineInPlaneCond S.1.1) : constAbs S.val := by (hS : LineInPlaneCond S.1.1) : ConstAbs S.val := by
induction n induction n
exact linesInPlane_constAbs_four S hS exact linesInPlane_constAbs_four S hS
exact linesInPlane_constAbs hS exact linesInPlane_constAbs hS

View file

@ -148,7 +148,7 @@ end theDeltas
section theBasisVectors section theBasisVectors
/-- The first part of the basis as charge assignments. -/ /-- The first part of the basis as charge assignments. -/
def basisAsCharges (j : Fin n) : (PureU1 (2 * n + 1)).charges := def basisAsCharges (j : Fin n) : (PureU1 (2 * n + 1)).Charges :=
fun i => fun i =>
if i = δ₁ j then if i = δ₁ j then
1 1
@ -159,7 +159,7 @@ def basisAsCharges (j : Fin n) : (PureU1 (2 * n + 1)).charges :=
0 0
/-- The second part of the basis as charge assignments. -/ /-- The second part of the basis as charge assignments. -/
def basis!AsCharges (j : Fin n) : (PureU1 (2 * n + 1)).charges := def basis!AsCharges (j : Fin n) : (PureU1 (2 * n + 1)).Charges :=
fun i => fun i =>
if i = δ!₁ j then if i = δ!₁ j then
1 1
@ -351,13 +351,13 @@ lemma swap!_as_add {S S' : (PureU1 (2 * n + 1)).LinSols} (j : Fin n)
simp simp
/-- A point in the span of the first part of the basis as a charge. -/ /-- A point in the span of the first part of the basis as a charge. -/
def P (f : Fin n → ) : (PureU1 (2 * n + 1)).charges := ∑ i, f i • basisAsCharges i def P (f : Fin n → ) : (PureU1 (2 * n + 1)).Charges := ∑ i, f i • basisAsCharges i
/-- A point in the span of the second part of the basis as a charge. -/ /-- A point in the span of the second part of the basis as a charge. -/
def P! (f : Fin n → ) : (PureU1 (2 * n + 1)).charges := ∑ i, f i • basis!AsCharges i def P! (f : Fin n → ) : (PureU1 (2 * n + 1)).Charges := ∑ i, f i • basis!AsCharges i
/-- A point in the span of the basis as a charge. -/ /-- A point in the span of the basis as a charge. -/
def Pa (f : Fin n → ) (g : Fin n → ) : (PureU1 (2 * n + 1)).charges := P f + P! g def Pa (f : Fin n → ) (g : Fin n → ) : (PureU1 (2 * n + 1)).Charges := P f + P! g
lemma P_δ₁ (f : Fin n → ) (j : Fin n) : P f (δ₁ j) = f j := by lemma P_δ₁ (f : Fin n → ) (j : Fin n) : P f (δ₁ j) = f j := by
rw [P, sum_of_charges] rw [P, sum_of_charges]

View file

@ -36,11 +36,11 @@ open VectorLikeOddPlane
/-- A property on `LinSols`, satisfied 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. -/ in the basis through that point is in the cubic. -/
def lineInCubic (S : (PureU1 (2 * n + 1)).LinSols) : Prop := def LineInCubic (S : (PureU1 (2 * n + 1)).LinSols) : Prop :=
∀ (g f : Fin n → ) (_ : S.val = Pa g f) (a b : ) , ∀ (g f : Fin n → ) (_ : S.val = Pa g f) (a b : ) ,
accCube (2 * n + 1) (a • P g + b • P! f) = 0 accCube (2 * n + 1) (a • P g + b • P! f) = 0
lemma lineInCubic_expand {S : (PureU1 (2 * n + 1)).LinSols} (h : lineInCubic S) : lemma lineInCubic_expand {S : (PureU1 (2 * n + 1)).LinSols} (h : LineInCubic S) :
∀ (g : Fin n → ) (f : Fin n → ) (_ : S.val = P g + P! f) (a b : ) , ∀ (g : Fin n → ) (f : Fin n → ) (_ : S.val = P g + P! f) (a b : ) ,
3 * a * b * (a * accCubeTriLinSymm (P g) (P g) (P! f) 3 * a * b * (a * accCubeTriLinSymm (P g) (P g) (P! f)
+ b * accCubeTriLinSymm (P! f) (P! f) (P g)) = 0 := by + b * accCubeTriLinSymm (P! f) (P! f) (P g)) = 0 := by
@ -55,7 +55,7 @@ lemma lineInCubic_expand {S : (PureU1 (2 * n + 1)).LinSols} (h : lineInCubic S)
ring ring
lemma line_in_cubic_P_P_P! {S : (PureU1 (2 * n + 1)).LinSols} (h : lineInCubic S) : lemma line_in_cubic_P_P_P! {S : (PureU1 (2 * n + 1)).LinSols} (h : LineInCubic S) :
∀ (g : Fin n → ) (f : Fin n → ) (_ : S.val = P g + P! f), ∀ (g : Fin n → ) (f : Fin n → ) (_ : S.val = P g + P! f),
accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by
intro g f hS intro g f hS
@ -65,19 +65,19 @@ lemma line_in_cubic_P_P_P! {S : (PureU1 (2 * n + 1)).LinSols} (h : lineInCubic S
/-- We say a `LinSol` satisfies `lineInCubicPerm` if all its permutations satisfy `lineInCubic`. -/ /-- We say a `LinSol` satisfies `lineInCubicPerm` if all its permutations satisfy `lineInCubic`. -/
def lineInCubicPerm (S : (PureU1 (2 * n + 1)).LinSols) : Prop := def LineInCubicPerm (S : (PureU1 (2 * n + 1)).LinSols) : Prop :=
∀ (M : (FamilyPermutations (2 * n + 1)).group ), ∀ (M : (FamilyPermutations (2 * n + 1)).group ),
lineInCubic ((FamilyPermutations (2 * n + 1)).linSolRep M S) LineInCubic ((FamilyPermutations (2 * n + 1)).linSolRep M S)
/-- If `lineInCubicPerm S` then `lineInCubic S`. -/ /-- If `lineInCubicPerm S` then `lineInCubic S`. -/
lemma lineInCubicPerm_self {S : (PureU1 (2 * n + 1)).LinSols} (hS : lineInCubicPerm S) : lemma lineInCubicPerm_self {S : (PureU1 (2 * n + 1)).LinSols} (hS : LineInCubicPerm S) :
lineInCubic S := hS 1 LineInCubic S := hS 1
/-- If `lineInCubicPerm S` then `lineInCubicPerm (M S)` for all permutations `M`. -/ /-- If `lineInCubicPerm S` then `lineInCubicPerm (M S)` for all permutations `M`. -/
lemma lineInCubicPerm_permute {S : (PureU1 (2 * n + 1)).LinSols} lemma lineInCubicPerm_permute {S : (PureU1 (2 * n + 1)).LinSols}
(hS : lineInCubicPerm S) (M' : (FamilyPermutations (2 * n + 1)).group) : (hS : LineInCubicPerm S) (M' : (FamilyPermutations (2 * n + 1)).group) :
lineInCubicPerm ((FamilyPermutations (2 * n + 1)).linSolRep M' S) := by LineInCubicPerm ((FamilyPermutations (2 * n + 1)).linSolRep M' S) := by
rw [lineInCubicPerm] rw [LineInCubicPerm]
intro M intro M
have ht : ((FamilyPermutations (2 * n + 1)).linSolRep M) have ht : ((FamilyPermutations (2 * n + 1)).linSolRep M)
((FamilyPermutations (2 * n + 1)).linSolRep M' S) ((FamilyPermutations (2 * n + 1)).linSolRep M' S)
@ -88,7 +88,7 @@ lemma lineInCubicPerm_permute {S : (PureU1 (2 * n + 1)).LinSols}
lemma lineInCubicPerm_swap {S : (PureU1 (2 * n.succ + 1)).LinSols} lemma lineInCubicPerm_swap {S : (PureU1 (2 * n.succ + 1)).LinSols}
(LIC : lineInCubicPerm S) : (LIC : LineInCubicPerm S) :
∀ (j : Fin n.succ) (g f : Fin n.succ → ) (_ : S.val = Pa g f) , ∀ (j : Fin n.succ) (g f : Fin n.succ → ) (_ : S.val = Pa g f) ,
(S.val (δ!₂ j) - S.val (δ!₁ j)) (S.val (δ!₂ j) - S.val (δ!₁ j))
* accCubeTriLinSymm (P g) (P g) (basis!AsCharges j) = 0 := by * accCubeTriLinSymm (P g) (P g) (basis!AsCharges j) = 0 := by
@ -134,8 +134,8 @@ lemma P_P_P!_accCube' {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
ring ring
lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ+1)).LinSols} lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ+1)).LinSols}
(LIC : lineInCubicPerm S) : (LIC : LineInCubicPerm S) :
lineInPlaneProp ((S.val (δ!₂ 0)), ((S.val (δ!₁ 0)), (S.val δ!₃))) := by LineInPlaneProp ((S.val (δ!₂ 0)), ((S.val (δ!₁ 0)), (S.val δ!₃))) := by
obtain ⟨g, f, hfg⟩ := span_basis S obtain ⟨g, f, hfg⟩ := span_basis S
have h1 := lineInCubicPerm_swap LIC 0 g f hfg have h1 := lineInCubicPerm_swap LIC 0 g f hfg
rw [P_P_P!_accCube' g f hfg] at h1 rw [P_P_P!_accCube' g f hfg] at h1
@ -152,8 +152,8 @@ lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ+1)).LinSols}
linear_combination h1 linear_combination h1
lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ + 1)).LinSols} lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
(LIC : lineInCubicPerm S) : lineInPlaneCond S := by (LIC : LineInCubicPerm S) : LineInPlaneCond S := by
refine @Prop_three (2 * n.succ.succ + 1) lineInPlaneProp S (δ!₂ 0) (δ!₁ 0) refine @Prop_three (2 * n.succ.succ + 1) LineInPlaneProp S (δ!₂ 0) (δ!₁ 0)
δ!₃ ?_ ?_ ?_ ?_ δ!₃ ?_ ?_ ?_ ?_
simp [Fin.ext_iff, δ!₂, δ!₁] simp [Fin.ext_iff, δ!₂, δ!₁]
simp [Fin.ext_iff, δ!₂, δ!₃] simp [Fin.ext_iff, δ!₂, δ!₃]
@ -162,11 +162,11 @@ lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
exact lineInCubicPerm_last_cond (lineInCubicPerm_permute LIC M) exact lineInCubicPerm_last_cond (lineInCubicPerm_permute LIC M)
lemma lineInCubicPerm_constAbs {S : (PureU1 (2 * n.succ.succ + 1)).LinSols} lemma lineInCubicPerm_constAbs {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
(LIC : lineInCubicPerm S) : constAbs S.val := (LIC : LineInCubicPerm S) : ConstAbs S.val :=
linesInPlane_constAbs (lineInCubicPerm_last_perm LIC) linesInPlane_constAbs (lineInCubicPerm_last_perm LIC)
theorem lineInCubicPerm_zero {S : (PureU1 (2 * n.succ.succ + 1)).LinSols} theorem lineInCubicPerm_zero {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
(LIC : lineInCubicPerm S) : S = 0 := (LIC : LineInCubicPerm S) : S = 0 :=
ConstAbs.boundary_value_odd S (lineInCubicPerm_constAbs LIC) ConstAbs.boundary_value_odd S (lineInCubicPerm_constAbs LIC)
end Odd end Odd

View file

@ -77,13 +77,13 @@ lemma anomalyFree_param {S : (PureU1 (2 * n + 1)).Sols}
/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`. /-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`.
In this case our parameterization above will be able to recover this point. -/ In this case our parameterization above will be able to recover this point. -/
def genericCase (S : (PureU1 (2 * n.succ + 1)).Sols) : Prop := def GenericCase (S : (PureU1 (2 * n.succ + 1)).Sols) : Prop :=
∀ (g f : Fin n.succ → ) (_ : S.val = P g + P! f) , ∀ (g f : Fin n.succ → ) (_ : S.val = P g + P! f) ,
accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0 accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0
lemma genericCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols) lemma genericCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols)
(hs : ∃ (g f : Fin n.succ → ), S.val = P g + P! f ∧ (hs : ∃ (g f : Fin n.succ → ), S.val = P g + P! f ∧
accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0) : genericCase S := by accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0) : GenericCase S := by
intro g f hS hC intro g f hS hC
obtain ⟨g', f', hS', hC'⟩ := hs obtain ⟨g', f', hS', hC'⟩ := hs
rw [hS] at hS' rw [hS] at hS'
@ -93,13 +93,13 @@ lemma genericCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols)
/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`. /-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`.
In this case we will show that S is zero if it is true for all permutations. -/ In this case we will show that S is zero if it is true for all permutations. -/
def specialCase (S : (PureU1 (2 * n.succ + 1)).Sols) : Prop := def SpecialCase (S : (PureU1 (2 * n.succ + 1)).Sols) : Prop :=
∀ (g f : Fin n.succ → ) (_ : S.val = P g + P! f) , ∀ (g f : Fin n.succ → ) (_ : S.val = P g + P! f) ,
accCubeTriLinSymm (P g) (P g) (P! f) = 0 accCubeTriLinSymm (P g) (P g) (P! f) = 0
lemma specialCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols) lemma specialCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols)
(hs : ∃ (g f : Fin n.succ → ), S.val = P g + P! f ∧ (hs : ∃ (g f : Fin n.succ → ), S.val = P g + P! f ∧
accCubeTriLinSymm (P g) (P g) (P! f) = 0) : specialCase S := by accCubeTriLinSymm (P g) (P g) (P! f) = 0) : SpecialCase S := by
intro g f hS intro g f hS
obtain ⟨g', f', hS', hC'⟩ := hs obtain ⟨g', f', hS', hC'⟩ := hs
rw [hS] at hS' rw [hS] at hS'
@ -108,7 +108,7 @@ lemma specialCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols)
exact hC' exact hC'
lemma generic_or_special (S : (PureU1 (2 * n.succ + 1)).Sols) : lemma generic_or_special (S : (PureU1 (2 * n.succ + 1)).Sols) :
genericCase S specialCase S := by GenericCase S SpecialCase S := by
obtain ⟨g, f, h⟩ := span_basis S.1.1 obtain ⟨g, f, h⟩ := span_basis S.1.1
have h1 : accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0 have h1 : accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0
accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by
@ -117,7 +117,7 @@ lemma generic_or_special (S : (PureU1 (2 * n.succ + 1)).Sols) :
exact Or.inl (genericCase_exists S ⟨g, f, h, h1⟩) exact Or.inl (genericCase_exists S ⟨g, f, h, h1⟩)
exact Or.inr (specialCase_exists S ⟨g, f, h, h1⟩) exact Or.inr (specialCase_exists S ⟨g, f, h, h1⟩)
theorem generic_case {S : (PureU1 (2 * n.succ + 1)).Sols} (h : genericCase S) : theorem generic_case {S : (PureU1 (2 * n.succ + 1)).Sols} (h : GenericCase S) :
∃ g f a, S = parameterization g f a := by ∃ g f a, S = parameterization g f a := by
obtain ⟨g, f, hS⟩ := span_basis S.1.1 obtain ⟨g, f, hS⟩ := span_basis S.1.1
use g, f, (accCubeTriLinSymm (P! f) (P! f) (P g))⁻¹ use g, f, (accCubeTriLinSymm (P! f) (P! f) (P g))⁻¹
@ -135,8 +135,8 @@ theorem generic_case {S : (PureU1 (2 * n.succ + 1)).Sols} (h : genericCase S) :
lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ + 1)).Sols} lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ + 1)).Sols}
(h : specialCase S) : (h : SpecialCase S) :
lineInCubic S.1.1 := by LineInCubic S.1.1 := by
intro g f hS a b intro g f hS a b
erw [TriLinearSymm.toCubic_add] erw [TriLinearSymm.toCubic_add]
rw [HomogeneousCubic.map_smul, HomogeneousCubic.map_smul] rw [HomogeneousCubic.map_smul, HomogeneousCubic.map_smul]
@ -155,15 +155,15 @@ lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ + 1)).Sols}
lemma special_case_lineInCubic_perm {S : (PureU1 (2 * n.succ + 1)).Sols} lemma special_case_lineInCubic_perm {S : (PureU1 (2 * n.succ + 1)).Sols}
(h : ∀ (M : (FamilyPermutations (2 * n.succ + 1)).group), (h : ∀ (M : (FamilyPermutations (2 * n.succ + 1)).group),
specialCase ((FamilyPermutations (2 * n.succ + 1)).solAction.toFun S M)) : SpecialCase ((FamilyPermutations (2 * n.succ + 1)).solAction.toFun S M)) :
lineInCubicPerm S.1.1 := by LineInCubicPerm S.1.1 := by
intro M intro M
have hM := special_case_lineInCubic (h M) have hM := special_case_lineInCubic (h M)
exact hM exact hM
theorem special_case {S : (PureU1 (2 * n.succ.succ + 1)).Sols} theorem special_case {S : (PureU1 (2 * n.succ.succ + 1)).Sols}
(h : ∀ (M : (FamilyPermutations (2 * n.succ.succ + 1)).group), (h : ∀ (M : (FamilyPermutations (2 * n.succ.succ + 1)).group),
specialCase ((FamilyPermutations (2 * n.succ.succ + 1)).solAction.toFun S M)) : SpecialCase ((FamilyPermutations (2 * n.succ.succ + 1)).solAction.toFun S M)) :
S.1.1 = 0 := by S.1.1 = 0 := by
have ht := special_case_lineInCubic_perm h have ht := special_case_lineInCubic_perm h
exact lineInCubicPerm_zero ht exact lineInCubicPerm_zero ht

View file

@ -24,18 +24,18 @@ namespace PureU1
/-- The permutation group of the n-fermions. -/ /-- The permutation group of the n-fermions. -/
@[simp] @[simp]
def permGroup (n : ) := Equiv.Perm (Fin n) def PermGroup (n : ) := Equiv.Perm (Fin n)
instance {n : } : Group (permGroup n) := by instance {n : } : Group (PermGroup n) := by
simp [permGroup] simp [PermGroup]
infer_instance infer_instance
section Charges section Charges
/-- The image of an element of `permGroup` under the representation on charges. -/ /-- The image of an element of `permGroup` under the representation on charges. -/
@[simps!] @[simps!]
def chargeMap {n : } (f : permGroup n) : def chargeMap {n : } (f : PermGroup n) :
(PureU1 n).charges →ₗ[] (PureU1 n).charges where (PureU1 n).Charges →ₗ[] (PureU1 n).Charges where
toFun S := S ∘ f.toFun toFun S := S ∘ f.toFun
map_add' S T := by map_add' S T := by
funext i funext i
@ -49,10 +49,10 @@ open PureU1Charges in
/-- The representation of `permGroup` acting on the vector space of charges. -/ /-- The representation of `permGroup` acting on the vector space of charges. -/
@[simp] @[simp]
def permCharges {n : } : Representation (permGroup n) (PureU1 n).charges where def permCharges {n : } : Representation (PermGroup n) (PureU1 n).Charges where
toFun f := chargeMap f⁻¹ toFun f := chargeMap f⁻¹
map_mul' f g :=by map_mul' f g :=by
simp only [permGroup, mul_inv_rev] simp only [PermGroup, mul_inv_rev]
apply LinearMap.ext apply LinearMap.ext
intro S intro S
funext i funext i
@ -63,14 +63,14 @@ def permCharges {n : } : Representation (permGroup n) (PureU1 n).charges
funext i funext i
rfl rfl
lemma accGrav_invariant {n : } (f : (permGroup n)) (S : (PureU1 n).charges) : lemma accGrav_invariant {n : } (f : (PermGroup n)) (S : (PureU1 n).Charges) :
PureU1.accGrav n (permCharges f S) = accGrav n S := by PureU1.accGrav n (permCharges f S) = accGrav n S := by
simp [accGrav, permCharges] simp [accGrav, permCharges]
apply (Equiv.Perm.sum_comp _ _ _ ?_) apply (Equiv.Perm.sum_comp _ _ _ ?_)
simp simp
open BigOperators open BigOperators
lemma accCube_invariant {n : } (f : (permGroup n)) (S : (PureU1 n).charges) : lemma accCube_invariant {n : } (f : (PermGroup n)) (S : (PureU1 n).Charges) :
accCube n (permCharges f S) = accCube n S := by accCube n (permCharges f S) = accCube n S := by
rw [accCube_explicit, accCube_explicit] rw [accCube_explicit, accCube_explicit]
change ∑ i : Fin n, ((((fun a => a^3) ∘ S) (f.symm i))) = _ change ∑ i : Fin n, ((((fun a => a^3) ∘ S) (f.symm i))) = _
@ -82,7 +82,7 @@ end Charges
/-- The permutations acting on the ACC system. -/ /-- The permutations acting on the ACC system. -/
@[simp] @[simp]
def FamilyPermutations (n : ) : ACCSystemGroupAction (PureU1 n) where def FamilyPermutations (n : ) : ACCSystemGroupAction (PureU1 n) where
group := permGroup n group := PermGroup n
groupInst := inferInstance groupInst := inferInstance
rep := permCharges rep := permCharges
linearInvariant := by linearInvariant := by
@ -97,7 +97,7 @@ def FamilyPermutations (n : ) : ACCSystemGroupAction (PureU1 n) where
cubicInvariant := accCube_invariant cubicInvariant := accCube_invariant
lemma FamilyPermutations_charges_apply (S : (PureU1 n).charges) lemma FamilyPermutations_charges_apply (S : (PureU1 n).Charges)
(i : Fin n) (f : (FamilyPermutations n).group) : (i : Fin n) (f : (FamilyPermutations n).group) :
((FamilyPermutations n).rep f S) i = S (f.invFun i) := by ((FamilyPermutations n).rep f S) i = S (f.invFun i) := by
rfl rfl

View file

@ -23,29 +23,29 @@ variable {n : }
/-- We say a charge is shorted if for all `i ≤ j`, then `S i ≤ S j`. -/ /-- We say a charge is shorted if for all `i ≤ j`, then `S i ≤ S j`. -/
@[simp] @[simp]
def sorted {n : } (S : (PureU1 n).charges) : Prop := def Sorted {n : } (S : (PureU1 n).Charges) : Prop :=
∀ i j (_ : i ≤ j), S i ≤ S j ∀ i j (_ : i ≤ j), S i ≤ S j
/-- Given a charge assignment `S`, the corresponding sorted charge assignment. -/ /-- Given a charge assignment `S`, the corresponding sorted charge assignment. -/
@[simp] @[simp]
def sort {n : } (S : (PureU1 n).charges) : (PureU1 n).charges := def sort {n : } (S : (PureU1 n).Charges) : (PureU1 n).Charges :=
((FamilyPermutations n).rep (Tuple.sort S).symm S) ((FamilyPermutations n).rep (Tuple.sort S).symm S)
lemma sort_sorted {n : } (S : (PureU1 n).charges) : sorted (sort S) := by lemma sort_sorted {n : } (S : (PureU1 n).Charges) : Sorted (sort S) := by
simp only [sorted, PureU1_numberCharges, sort, FamilyPermutations, permGroup, permCharges, simp only [Sorted, PureU1_numberCharges, sort, FamilyPermutations, PermGroup, permCharges,
MonoidHom.coe_mk, OneHom.coe_mk, chargeMap_apply] MonoidHom.coe_mk, OneHom.coe_mk, chargeMap_apply]
intro i j hij intro i j hij
exact Tuple.monotone_sort S hij exact Tuple.monotone_sort S hij
lemma sort_perm {n : } (S : (PureU1 n).charges) (M :(FamilyPermutations n).group) : lemma sort_perm {n : } (S : (PureU1 n).Charges) (M :(FamilyPermutations n).group) :
sort ((FamilyPermutations n).rep M S) = sort S := sort ((FamilyPermutations n).rep M S) = sort S :=
@Tuple.comp_perm_comp_sort_eq_comp_sort n _ S M⁻¹ @Tuple.comp_perm_comp_sort_eq_comp_sort n _ S M⁻¹
lemma sort_apply {n : } (S : (PureU1 n).charges) (j : Fin n) : lemma sort_apply {n : } (S : (PureU1 n).Charges) (j : Fin n) :
sort S j = S ((Tuple.sort S) j) := by sort S j = S ((Tuple.sort S) j) := by
rfl rfl
lemma sort_zero {n : } (S : (PureU1 n).charges) (hS : sort S = 0) : S = 0 := by lemma sort_zero {n : } (S : (PureU1 n).Charges) (hS : sort S = 0) : S = 0 := by
funext i funext i
have hj : ∀ j, sort S j = 0 := by have hj : ∀ j, sort S j = 0 := by
rw [hS] rw [hS]
@ -57,7 +57,7 @@ lemma sort_zero {n : } (S : (PureU1 n).charges) (hS : sort S = 0) : S = 0 :=
rw [hi] rw [hi]
rfl rfl
lemma sort_projection {n : } (S : (PureU1 n).charges) : sort (sort S) = sort S := lemma sort_projection {n : } (S : (PureU1 n).Charges) : sort (sort S) = sort S :=
sort_perm S (Tuple.sort S).symm sort_perm S (Tuple.sort S).symm
/-- The sort function acting on `LinSols`. -/ /-- The sort function acting on `LinSols`. -/

View file

@ -36,7 +36,7 @@ lemma split_odd (n : ) : n + 1 + n = 2 * n + 1 := by omega
/-- A charge configuration for n even is vector like if when sorted the `i`th element /-- A charge configuration for n even is vector like if when sorted the `i`th element
is equal to the negative of the `n + i`th element. -/ is equal to the negative of the `n + i`th element. -/
@[simp] @[simp]
def vectorLikeEven (S : (PureU1 (2 * n)).charges) : Prop := def VectorLikeEven (S : (PureU1 (2 * n)).Charges) : Prop :=
∀ (i : Fin n), (sort S) (Fin.cast (split_equal n) (Fin.castAdd n i)) ∀ (i : Fin n), (sort S) (Fin.cast (split_equal n) (Fin.castAdd n i))
= - (sort S) (Fin.cast (split_equal n) (Fin.natAdd n i)) = - (sort S) (Fin.cast (split_equal n) (Fin.natAdd n i))

View file

@ -36,17 +36,17 @@ variable {n : }
/-- An equivalence between the set `(SMCharges n).charges` and the set /-- An equivalence between the set `(SMCharges n).charges` and the set
`(Fin 5 → Fin n → )`. -/ `(Fin 5 → Fin n → )`. -/
@[simps!] @[simps!]
def toSpeciesEquiv : (SMCharges n).charges ≃ (Fin 5 → Fin n → ) := def toSpeciesEquiv : (SMCharges n).Charges ≃ (Fin 5 → Fin n → ) :=
((Equiv.curry _ _ _).symm.trans ((@finProdFinEquiv 5 n).arrowCongr (Equiv.refl ))).symm ((Equiv.curry _ _ _).symm.trans ((@finProdFinEquiv 5 n).arrowCongr (Equiv.refl ))).symm
/-- For a given `i ∈ Fin 5`, the projection of a charge onto that species. -/ /-- For a given `i ∈ Fin 5`, the projection of a charge onto that species. -/
@[simps!] @[simps!]
def toSpecies (i : Fin 5) : (SMCharges n).charges →ₗ[] (SMSpecies n).charges where def toSpecies (i : Fin 5) : (SMCharges n).Charges →ₗ[] (SMSpecies n).Charges where
toFun S := toSpeciesEquiv S i toFun S := toSpeciesEquiv S i
map_add' _ _ := by rfl map_add' _ _ := by rfl
map_smul' _ _ := by rfl map_smul' _ _ := by rfl
lemma charges_eq_toSpecies_eq (S T : (SMCharges n).charges) : lemma charges_eq_toSpecies_eq (S T : (SMCharges n).Charges) :
S = T ↔ ∀ i, toSpecies i S = toSpecies i T := by S = T ↔ ∀ i, toSpecies i S = toSpecies i T := by
apply Iff.intro apply Iff.intro
exact fun a i => congrArg (⇑(toSpecies i)) a exact fun a i => congrArg (⇑(toSpecies i)) a
@ -84,7 +84,7 @@ variable {n : }
/-- The gravitational anomaly equation. -/ /-- The gravitational anomaly equation. -/
@[simp] @[simp]
def accGrav : (SMCharges n).charges →ₗ[] where def accGrav : (SMCharges n).Charges →ₗ[] where
toFun S := ∑ i, (6 * Q S i + 3 * U S i + 3 * D S i + 2 * L S i + E S i) toFun S := ∑ i, (6 * Q S i + 3 * U S i + 3 * D S i + 2 * L S i + E S i)
map_add' S T := by map_add' S T := by
simp only simp only
@ -102,7 +102,7 @@ def accGrav : (SMCharges n).charges →ₗ[] where
ring ring
/-- Extensionality lemma for `accGrav`. -/ /-- Extensionality lemma for `accGrav`. -/
lemma accGrav_ext {S T : (SMCharges n).charges} lemma accGrav_ext {S T : (SMCharges n).Charges}
(hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : (hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
accGrav S = accGrav T := by accGrav S = accGrav T := by
simp only [accGrav, SMSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk, simp only [accGrav, SMSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk,
@ -114,7 +114,7 @@ lemma accGrav_ext {S T : (SMCharges n).charges}
/-- The `SU(2)` anomaly equation. -/ /-- The `SU(2)` anomaly equation. -/
@[simp] @[simp]
def accSU2 : (SMCharges n).charges →ₗ[] where def accSU2 : (SMCharges n).Charges →ₗ[] where
toFun S := ∑ i, (3 * Q S i + L S i) toFun S := ∑ i, (3 * Q S i + L S i)
map_add' S T := by map_add' S T := by
simp only simp only
@ -132,7 +132,7 @@ def accSU2 : (SMCharges n).charges →ₗ[] where
ring ring
/-- Extensionality lemma for `accSU2`. -/ /-- Extensionality lemma for `accSU2`. -/
lemma accSU2_ext {S T : (SMCharges n).charges} lemma accSU2_ext {S T : (SMCharges n).Charges}
(hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : (hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
accSU2 S = accSU2 T := by accSU2 S = accSU2 T := by
simp only [accSU2, SMSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk, simp only [accSU2, SMSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk,
@ -143,7 +143,7 @@ lemma accSU2_ext {S T : (SMCharges n).charges}
/-- The `SU(3)` anomaly equations. -/ /-- The `SU(3)` anomaly equations. -/
@[simp] @[simp]
def accSU3 : (SMCharges n).charges →ₗ[] where def accSU3 : (SMCharges n).Charges →ₗ[] where
toFun S := ∑ i, (2 * Q S i + U S i + D S i) toFun S := ∑ i, (2 * Q S i + U S i + D S i)
map_add' S T := by map_add' S T := by
simp only simp only
@ -161,7 +161,7 @@ def accSU3 : (SMCharges n).charges →ₗ[] where
ring ring
/-- Extensionality lemma for `accSU3`. -/ /-- Extensionality lemma for `accSU3`. -/
lemma accSU3_ext {S T : (SMCharges n).charges} lemma accSU3_ext {S T : (SMCharges n).Charges}
(hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : (hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
accSU3 S = accSU3 T := by accSU3 S = accSU3 T := by
simp only [accSU3, SMSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk, simp only [accSU3, SMSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk,
@ -173,7 +173,7 @@ lemma accSU3_ext {S T : (SMCharges n).charges}
/-- The `Y²` anomaly equation. -/ /-- The `Y²` anomaly equation. -/
@[simp] @[simp]
def accYY : (SMCharges n).charges →ₗ[] where def accYY : (SMCharges n).Charges →ₗ[] where
toFun S := ∑ i, (Q S i + 8 * U S i + 2 * D S i + 3 * L S i toFun S := ∑ i, (Q S i + 8 * U S i + 2 * D S i + 3 * L S i
+ 6 * E S i) + 6 * E S i)
map_add' S T := by map_add' S T := by
@ -192,7 +192,7 @@ def accYY : (SMCharges n).charges →ₗ[] where
ring ring
/-- Extensionality lemma for `accYY`. -/ /-- Extensionality lemma for `accYY`. -/
lemma accYY_ext {S T : (SMCharges n).charges} lemma accYY_ext {S T : (SMCharges n).Charges}
(hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : (hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
accYY S = accYY T := by accYY S = accYY T := by
simp only [accYY, SMSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk, simp only [accYY, SMSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk,
@ -204,7 +204,7 @@ lemma accYY_ext {S T : (SMCharges n).charges}
/-- The quadratic bilinear map. -/ /-- The quadratic bilinear map. -/
@[simps!] @[simps!]
def quadBiLin : BiLinearSymm (SMCharges n).charges := BiLinearSymm.mk₂ def quadBiLin : BiLinearSymm (SMCharges n).Charges := BiLinearSymm.mk₂
(fun S => ∑ i, (Q S.1 i * Q S.2 i + (fun S => ∑ i, (Q S.1 i * Q S.2 i +
- 2 * (U S.1 i * U S.2 i) + - 2 * (U S.1 i * U S.2 i) +
D S.1 i * D S.2 i + D S.1 i * D S.2 i +
@ -238,11 +238,11 @@ def quadBiLin : BiLinearSymm (SMCharges n).charges := BiLinearSymm.mk₂
/-- The quadratic anomaly cancellation condition. -/ /-- The quadratic anomaly cancellation condition. -/
@[simp] @[simp]
def accQuad : HomogeneousQuadratic (SMCharges n).charges := def accQuad : HomogeneousQuadratic (SMCharges n).Charges :=
(@quadBiLin n).toHomogeneousQuad (@quadBiLin n).toHomogeneousQuad
/-- Extensionality lemma for `accQuad`. -/ /-- Extensionality lemma for `accQuad`. -/
lemma accQuad_ext {S T : (SMCharges n).charges} lemma accQuad_ext {S T : (SMCharges n).Charges}
(h : ∀ j, ∑ i, ((fun a => a^2) ∘ toSpecies j S) i = (h : ∀ j, ∑ i, ((fun a => a^2) ∘ toSpecies j S) i =
∑ i, ((fun a => a^2) ∘ toSpecies j T) i) : ∑ i, ((fun a => a^2) ∘ toSpecies j T) i) :
accQuad S = accQuad T := by accQuad S = accQuad T := by
@ -258,7 +258,7 @@ lemma accQuad_ext {S T : (SMCharges n).charges}
/-- The trilinear function defining the cubic. -/ /-- The trilinear function defining the cubic. -/
@[simps!] @[simps!]
def cubeTriLin : TriLinearSymm (SMCharges n).charges := TriLinearSymm.mk₃ def cubeTriLin : TriLinearSymm (SMCharges n).Charges := TriLinearSymm.mk₃
(fun S => ∑ i, (6 * ((Q S.1 i) * (Q S.2.1 i) * (Q S.2.2 i)) (fun S => ∑ i, (6 * ((Q S.1 i) * (Q S.2.1 i) * (Q S.2.2 i))
+ 3 * ((U S.1 i) * (U S.2.1 i) * (U S.2.2 i)) + 3 * ((U S.1 i) * (U S.2.1 i) * (U S.2.2 i))
+ 3 * ((D S.1 i) * (D S.2.1 i) * (D S.2.2 i)) + 3 * ((D S.1 i) * (D S.2.1 i) * (D S.2.2 i))
@ -301,10 +301,10 @@ def cubeTriLin : TriLinearSymm (SMCharges n).charges := TriLinearSymm.mk₃
/-- The cubic acc. -/ /-- The cubic acc. -/
@[simp] @[simp]
def accCube : HomogeneousCubic (SMCharges n).charges := cubeTriLin.toCubic def accCube : HomogeneousCubic (SMCharges n).Charges := cubeTriLin.toCubic
/-- Extensionality lemma for `accCube`. -/ /-- Extensionality lemma for `accCube`. -/
lemma accCube_ext {S T : (SMCharges n).charges} lemma accCube_ext {S T : (SMCharges n).Charges}
(h : ∀ j, ∑ i, ((fun a => a^3) ∘ toSpecies j S) i = (h : ∀ j, ∑ i, ((fun a => a^3) ∘ toSpecies j S) i =
∑ i, ((fun a => a^3) ∘ toSpecies j T) i) : ∑ i, ((fun a => a^3) ∘ toSpecies j T) i) :
accCube S = accCube T := by accCube S = accCube T := by

View file

@ -20,8 +20,8 @@ open BigOperators
/-- Given a map of for a generic species, the corresponding map for charges. -/ /-- Given a map of for a generic species, the corresponding map for charges. -/
@[simps!] @[simps!]
def chargesMapOfSpeciesMap {n m : } (f : (SMSpecies n).charges →ₗ[] (SMSpecies m).charges) : def chargesMapOfSpeciesMap {n m : } (f : (SMSpecies n).Charges →ₗ[] (SMSpecies m).Charges) :
(SMCharges n).charges →ₗ[] (SMCharges m).charges where (SMCharges n).Charges →ₗ[] (SMCharges m).Charges where
toFun S := toSpeciesEquiv.symm (fun i => (LinearMap.comp f (toSpecies i)) S) toFun S := toSpeciesEquiv.symm (fun i => (LinearMap.comp f (toSpecies i)) S)
map_add' S T := by map_add' S T := by
rw [charges_eq_toSpecies_eq] rw [charges_eq_toSpecies_eq]
@ -40,7 +40,7 @@ def chargesMapOfSpeciesMap {n m : } (f : (SMSpecies n).charges →ₗ[] (S
/-- The projection of the `m`-family charges onto the first `n`-family charges for species. -/ /-- The projection of the `m`-family charges onto the first `n`-family charges for species. -/
@[simps!] @[simps!]
def speciesFamilyProj {m n : } (h : n ≤ m) : def speciesFamilyProj {m n : } (h : n ≤ m) :
(SMSpecies m).charges →ₗ[] (SMSpecies n).charges where (SMSpecies m).Charges →ₗ[] (SMSpecies n).Charges where
toFun S := S ∘ Fin.castLE h toFun S := S ∘ Fin.castLE h
map_add' S T := by map_add' S T := by
funext i funext i
@ -51,14 +51,14 @@ def speciesFamilyProj {m n : } (h : n ≤ m) :
--rw [show Rat.cast a = a from rfl] --rw [show Rat.cast a = a from rfl]
/-- The projection of the `m`-family charges onto the first `n`-family charges. -/ /-- The projection of the `m`-family charges onto the first `n`-family charges. -/
def familyProjection {m n : } (h : n ≤ m) : (SMCharges m).charges →ₗ[] (SMCharges n).charges := def familyProjection {m n : } (h : n ≤ m) : (SMCharges m).Charges →ₗ[] (SMCharges n).Charges :=
chargesMapOfSpeciesMap (speciesFamilyProj h) chargesMapOfSpeciesMap (speciesFamilyProj h)
/-- For species, the embedding of the `m`-family charges onto the `n`-family charges, with all /-- For species, the embedding of the `m`-family charges onto the `n`-family charges, with all
other charges zero. -/ other charges zero. -/
@[simps!] @[simps!]
def speciesEmbed (m n : ) : def speciesEmbed (m n : ) :
(SMSpecies m).charges →ₗ[] (SMSpecies n).charges where (SMSpecies m).Charges →ₗ[] (SMSpecies n).Charges where
toFun S := fun i => toFun S := fun i =>
if hi : i.val < m then if hi : i.val < m then
S ⟨i, hi⟩ S ⟨i, hi⟩
@ -81,14 +81,14 @@ def speciesEmbed (m n : ) :
/-- The embedding of the `m`-family charges onto the `n`-family charges, with all /-- The embedding of the `m`-family charges onto the `n`-family charges, with all
other charges zero. -/ other charges zero. -/
def familyEmbedding (m n : ) : (SMCharges m).charges →ₗ[] (SMCharges n).charges := def familyEmbedding (m n : ) : (SMCharges m).Charges →ₗ[] (SMCharges n).Charges :=
chargesMapOfSpeciesMap (speciesEmbed m n) chargesMapOfSpeciesMap (speciesEmbed m n)
/-- For species, the embedding 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. -/ a universal manor. -/
@[simps!] @[simps!]
def speciesFamilyUniversial (n : ) : def speciesFamilyUniversial (n : ) :
(SMSpecies 1).charges →ₗ[] (SMSpecies n).charges where (SMSpecies 1).Charges →ₗ[] (SMSpecies n).Charges where
toFun S _ := S ⟨0, by simp⟩ toFun S _ := S ⟨0, by simp⟩
map_add' S T := by map_add' S T := by
funext _ funext _
@ -100,7 +100,7 @@ def speciesFamilyUniversial (n : ) :
/-- The embedding 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. -/ a universal manor. -/
def familyUniversal ( n : ) : (SMCharges 1).charges →ₗ[] (SMCharges n).charges := def familyUniversal ( n : ) : (SMCharges 1).Charges →ₗ[] (SMCharges n).Charges :=
chargesMapOfSpeciesMap (speciesFamilyUniversial n) chargesMapOfSpeciesMap (speciesFamilyUniversial n)
end SM end SM

View file

@ -51,7 +51,7 @@ lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := by
/-- An element of `charges` which satisfies the linear ACCs /-- An element of `charges` which satisfies the linear ACCs
gives us a element of `AnomalyFreeLinear`. -/ gives us a element of `AnomalyFreeLinear`. -/
def chargeToLinear (S : (SMNoGrav n).charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) : def chargeToLinear (S : (SMNoGrav n).Charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) :
(SMNoGrav n).LinSols := (SMNoGrav n).LinSols :=
⟨S, by ⟨S, by
intro i intro i
@ -74,13 +74,13 @@ def quadToAF (S : (SMNoGrav n).QuadSols) (hc : accCube S.val = 0) :
/-- An element of `charges` which satisfies the linear and quadratic ACCs /-- An element of `charges` which satisfies the linear and quadratic ACCs
gives us a element of `AnomalyFreeQuad`. -/ gives us a element of `AnomalyFreeQuad`. -/
def chargeToQuad (S : (SMNoGrav n).charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) : def chargeToQuad (S : (SMNoGrav n).Charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) :
(SMNoGrav n).QuadSols := (SMNoGrav n).QuadSols :=
linearToQuad $ chargeToLinear S hSU2 hSU3 linearToQuad $ chargeToLinear S hSU2 hSU3
/-- An element of `charges` which satisfies the linear, quadratic and cubic ACCs /-- An element of `charges` which satisfies the linear, quadratic and cubic ACCs
gives us a element of `AnomalyFree`. -/ gives us a element of `AnomalyFree`. -/
def chargeToAF (S : (SMNoGrav n).charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) def chargeToAF (S : (SMNoGrav n).Charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0)
(hc : accCube S = 0) : (SMNoGrav n).Sols := (hc : accCube S = 0) : (SMNoGrav n).Sols :=
quadToAF (chargeToQuad S hSU2 hSU3) hc quadToAF (chargeToQuad S hSU2 hSU3) hc

View file

@ -48,7 +48,7 @@ lemma ext {S T : linearParameters} (hQ : S.Q' = T.Q') (hY : S.Y = T.Y) (hE : S.E
/-- The map from the linear parameters to elements of `(SMNoGrav 1).charges`. -/ /-- The map from the linear parameters to elements of `(SMNoGrav 1).charges`. -/
@[simp] @[simp]
def asCharges (S : linearParameters) : (SMNoGrav 1).charges := fun i => def asCharges (S : linearParameters) : (SMNoGrav 1).Charges := fun i =>
match i with match i with
| (0 : Fin 5) => S.Q' | (0 : Fin 5) => S.Q'
| (1 : Fin 5) => S.Y - S.Q' | (1 : Fin 5) => S.Y - S.Q'

View file

@ -26,27 +26,27 @@ open BigOperators
/-- The group of `Sₙ` permutations for each species. -/ /-- The group of `Sₙ` permutations for each species. -/
@[simp] @[simp]
def permGroup (n : ) := ∀ (_ : Fin 5), Equiv.Perm (Fin n) def PermGroup (n : ) := ∀ (_ : Fin 5), Equiv.Perm (Fin n)
variable {n : } variable {n : }
@[simp] @[simp]
instance : Group (permGroup n) := Pi.group instance : Group (PermGroup n) := Pi.group
/-- The image of an element of `permGroup n` under the representation on charges. -/ /-- The image of an element of `permGroup n` under the representation on charges. -/
@[simps!] @[simps!]
def chargeMap (f : permGroup n) : (SMCharges n).charges →ₗ[] (SMCharges n).charges where def chargeMap (f : PermGroup n) : (SMCharges n).Charges →ₗ[] (SMCharges n).Charges where
toFun S := toSpeciesEquiv.symm (fun i => toSpecies i S ∘ f i ) toFun S := toSpeciesEquiv.symm (fun i => toSpecies i S ∘ f i )
map_add' _ _ := rfl map_add' _ _ := rfl
map_smul' _ _ := rfl map_smul' _ _ := rfl
/-- The representation of `(permGroup n)` acting on the vector space of charges. -/ /-- The representation of `(permGroup n)` acting on the vector space of charges. -/
@[simp] @[simp]
def repCharges {n : } : Representation (permGroup n) (SMCharges n).charges where def repCharges {n : } : Representation (PermGroup n) (SMCharges n).Charges where
toFun f := chargeMap f⁻¹ toFun f := chargeMap f⁻¹
map_mul' f g :=by map_mul' f g :=by
simp only [permGroup, mul_inv_rev] simp only [PermGroup, mul_inv_rev]
apply LinearMap.ext apply LinearMap.ext
intro S intro S
rw [charges_eq_toSpecies_eq] rw [charges_eq_toSpecies_eq]
@ -62,12 +62,12 @@ def repCharges {n : } : Representation (permGroup n) (SMCharges n).charge
erw [toSMSpecies_toSpecies_inv] erw [toSMSpecies_toSpecies_inv]
rfl rfl
lemma repCharges_toSpecies (f : permGroup n) (S : (SMCharges n).charges) (j : Fin 5) : lemma repCharges_toSpecies (f : PermGroup n) (S : (SMCharges n).Charges) (j : Fin 5) :
toSpecies j (repCharges f S) = toSpecies j S ∘ f⁻¹ j := by toSpecies j (repCharges f S) = toSpecies j S ∘ f⁻¹ j := by
erw [toSMSpecies_toSpecies_inv] erw [toSMSpecies_toSpecies_inv]
lemma toSpecies_sum_invariant (m : ) (f : permGroup n) (S : (SMCharges n).charges) (j : Fin 5) : lemma toSpecies_sum_invariant (m : ) (f : PermGroup n) (S : (SMCharges n).Charges) (j : Fin 5) :
∑ i, ((fun a => a ^ m) ∘ toSpecies j (repCharges f S)) i = ∑ i, ((fun a => a ^ m) ∘ toSpecies j (repCharges f S)) i =
∑ i, ((fun a => a ^ m) ∘ toSpecies j S) i := by ∑ i, ((fun a => a ^ m) ∘ toSpecies j S) i := by
rw [repCharges_toSpecies] rw [repCharges_toSpecies]
@ -76,34 +76,34 @@ lemma toSpecies_sum_invariant (m : ) (f : permGroup n) (S : (SMCharges n).cha
lemma accGrav_invariant (f : permGroup n) (S : (SMCharges n).charges) : lemma accGrav_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
accGrav (repCharges f S) = accGrav S := accGrav (repCharges f S) = accGrav S :=
accGrav_ext accGrav_ext
(by simpa using toSpecies_sum_invariant 1 f S) (by simpa using toSpecies_sum_invariant 1 f S)
lemma accSU2_invariant (f : permGroup n) (S : (SMCharges n).charges) : lemma accSU2_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
accSU2 (repCharges f S) = accSU2 S := accSU2 (repCharges f S) = accSU2 S :=
accSU2_ext accSU2_ext
(by simpa using toSpecies_sum_invariant 1 f S) (by simpa using toSpecies_sum_invariant 1 f S)
lemma accSU3_invariant (f : permGroup n) (S : (SMCharges n).charges) : lemma accSU3_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
accSU3 (repCharges f S) = accSU3 S := accSU3 (repCharges f S) = accSU3 S :=
accSU3_ext accSU3_ext
(by simpa using toSpecies_sum_invariant 1 f S) (by simpa using toSpecies_sum_invariant 1 f S)
lemma accYY_invariant (f : permGroup n) (S : (SMCharges n).charges) : lemma accYY_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
accYY (repCharges f S) = accYY S := accYY (repCharges f S) = accYY S :=
accYY_ext accYY_ext
(by simpa using toSpecies_sum_invariant 1 f S) (by simpa using toSpecies_sum_invariant 1 f S)
lemma accQuad_invariant (f : permGroup n) (S : (SMCharges n).charges) : lemma accQuad_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
accQuad (repCharges f S) = accQuad S := accQuad (repCharges f S) = accQuad S :=
accQuad_ext accQuad_ext
(toSpecies_sum_invariant 2 f S) (toSpecies_sum_invariant 2 f S)
lemma accCube_invariant (f : permGroup n) (S : (SMCharges n).charges) : lemma accCube_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
accCube (repCharges f S) = accCube S := accCube (repCharges f S) = accCube S :=
accCube_ext (fun j => toSpecies_sum_invariant 3 f S j) accCube_ext (fun j => toSpecies_sum_invariant 3 f S j)

View file

@ -34,17 +34,17 @@ variable {n : }
/-- An equivalence between `(SMνCharges n).charges` and `(Fin 6 → Fin n → )` /-- An equivalence between `(SMνCharges n).charges` and `(Fin 6 → Fin n → )`
splitting the charges into species.-/ splitting the charges into species.-/
@[simps!] @[simps!]
def toSpeciesEquiv : (SMνCharges n).charges ≃ (Fin 6 → Fin n → ) := def toSpeciesEquiv : (SMνCharges n).Charges ≃ (Fin 6 → Fin n → ) :=
((Equiv.curry _ _ _).symm.trans ((@finProdFinEquiv 6 n).arrowCongr (Equiv.refl ))).symm ((Equiv.curry _ _ _).symm.trans ((@finProdFinEquiv 6 n).arrowCongr (Equiv.refl ))).symm
/-- Given an `i ∈ Fin 6`, the projection of charges onto a given species. -/ /-- Given an `i ∈ Fin 6`, the projection of charges onto a given species. -/
@[simps!] @[simps!]
def toSpecies (i : Fin 6) : (SMνCharges n).charges →ₗ[] (SMνSpecies n).charges where def toSpecies (i : Fin 6) : (SMνCharges n).Charges →ₗ[] (SMνSpecies n).Charges where
toFun S := toSpeciesEquiv S i toFun S := toSpeciesEquiv S i
map_add' _ _ := by aesop map_add' _ _ := by aesop
map_smul' _ _ := by aesop map_smul' _ _ := by aesop
lemma charges_eq_toSpecies_eq (S T : (SMνCharges n).charges) : lemma charges_eq_toSpecies_eq (S T : (SMνCharges n).Charges) :
S = T ↔ ∀ i, toSpecies i S = toSpecies i T := by S = T ↔ ∀ i, toSpecies i S = toSpecies i T := by
apply Iff.intro apply Iff.intro
intro h intro h
@ -60,7 +60,7 @@ lemma toSMSpecies_toSpecies_inv (i : Fin 6) (f : (Fin 6 → Fin n → ) ) :
change (toSpeciesEquiv ∘ toSpeciesEquiv.symm ) _ i = f i change (toSpeciesEquiv ∘ toSpeciesEquiv.symm ) _ i = f i
simp simp
lemma toSpecies_one (S : (SMνCharges 1).charges) (j : Fin 6) : lemma toSpecies_one (S : (SMνCharges 1).Charges) (j : Fin 6) :
toSpecies j S ⟨0, by simp⟩ = S j := by toSpecies j S ⟨0, by simp⟩ = S j := by
match j with match j with
| 0 => rfl | 0 => rfl
@ -94,7 +94,7 @@ variable {n : }
/-- The gravitational anomaly equation. -/ /-- The gravitational anomaly equation. -/
@[simp] @[simp]
def accGrav : (SMνCharges n).charges →ₗ[] where def accGrav : (SMνCharges n).Charges →ₗ[] where
toFun S := ∑ i, (6 * Q S i + 3 * U S i + 3 * D S i + 2 * L S i + E S i + N S i) toFun S := ∑ i, (6 * Q S i + 3 * U S i + 3 * D S i + 2 * L S i + E S i + N S i)
map_add' S T := by map_add' S T := by
simp only simp only
@ -112,7 +112,7 @@ def accGrav : (SMνCharges n).charges →ₗ[] where
ring ring
lemma accGrav_decomp (S : (SMνCharges n).charges) : lemma accGrav_decomp (S : (SMνCharges n).Charges) :
accGrav S = 6 * ∑ i, Q S i + 3 * ∑ i, U S i + 3 * ∑ i, D S i + 2 * ∑ i, L S i + ∑ i, E S i + accGrav S = 6 * ∑ i, Q S i + 3 * ∑ i, U S i + 3 * ∑ i, D S i + 2 * ∑ i, L S i + ∑ i, E S i +
∑ i, N S i := by ∑ i, N S i := by
simp only [accGrav, SMνSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk, simp only [accGrav, SMνSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk,
@ -121,7 +121,7 @@ lemma accGrav_decomp (S : (SMνCharges n).charges) :
repeat erw [← Finset.mul_sum] repeat erw [← Finset.mul_sum]
/-- Extensionality lemma for `accGrav`. -/ /-- Extensionality lemma for `accGrav`. -/
lemma accGrav_ext {S T : (SMνCharges n).charges} lemma accGrav_ext {S T : (SMνCharges n).Charges}
(hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : (hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
accGrav S = accGrav T := by accGrav S = accGrav T := by
rw [accGrav_decomp, accGrav_decomp] rw [accGrav_decomp, accGrav_decomp]
@ -130,7 +130,7 @@ lemma accGrav_ext {S T : (SMνCharges n).charges}
/-- The `SU(2)` anomaly equation. -/ /-- The `SU(2)` anomaly equation. -/
@[simp] @[simp]
def accSU2 : (SMνCharges n).charges →ₗ[] where def accSU2 : (SMνCharges n).Charges →ₗ[] where
toFun S := ∑ i, (3 * Q S i + L S i) toFun S := ∑ i, (3 * Q S i + L S i)
map_add' S T := by map_add' S T := by
simp only simp only
@ -147,7 +147,7 @@ def accSU2 : (SMνCharges n).charges →ₗ[] where
-- rw [show Rat.cast a = a from rfl] -- rw [show Rat.cast a = a from rfl]
ring ring
lemma accSU2_decomp (S : (SMνCharges n).charges) : lemma accSU2_decomp (S : (SMνCharges n).Charges) :
accSU2 S = 3 * ∑ i, Q S i + ∑ i, L S i := by accSU2 S = 3 * ∑ i, Q S i + ∑ i, L S i := by
simp only [accSU2, SMνSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk, simp only [accSU2, SMνSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk,
AddHom.coe_mk] AddHom.coe_mk]
@ -155,7 +155,7 @@ lemma accSU2_decomp (S : (SMνCharges n).charges) :
repeat erw [← Finset.mul_sum] repeat erw [← Finset.mul_sum]
/-- Extensionality lemma for `accSU2`. -/ /-- Extensionality lemma for `accSU2`. -/
lemma accSU2_ext {S T : (SMνCharges n).charges} lemma accSU2_ext {S T : (SMνCharges n).Charges}
(hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : (hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
accSU2 S = accSU2 T := by accSU2 S = accSU2 T := by
rw [accSU2_decomp, accSU2_decomp] rw [accSU2_decomp, accSU2_decomp]
@ -163,7 +163,7 @@ lemma accSU2_ext {S T : (SMνCharges n).charges}
/-- The `SU(3)` anomaly equations. -/ /-- The `SU(3)` anomaly equations. -/
@[simp] @[simp]
def accSU3 : (SMνCharges n).charges →ₗ[] where def accSU3 : (SMνCharges n).Charges →ₗ[] where
toFun S := ∑ i, (2 * Q S i + U S i + D S i) toFun S := ∑ i, (2 * Q S i + U S i + D S i)
map_add' S T := by map_add' S T := by
simp only simp only
@ -180,7 +180,7 @@ def accSU3 : (SMνCharges n).charges →ₗ[] where
-- rw [show Rat.cast a = a from rfl] -- rw [show Rat.cast a = a from rfl]
ring ring
lemma accSU3_decomp (S : (SMνCharges n).charges) : lemma accSU3_decomp (S : (SMνCharges n).Charges) :
accSU3 S = 2 * ∑ i, Q S i + ∑ i, U S i + ∑ i, D S i := by accSU3 S = 2 * ∑ i, Q S i + ∑ i, U S i + ∑ i, D S i := by
simp only [accSU3, SMνSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk, simp only [accSU3, SMνSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk,
AddHom.coe_mk] AddHom.coe_mk]
@ -188,7 +188,7 @@ lemma accSU3_decomp (S : (SMνCharges n).charges) :
repeat rw [← Finset.mul_sum] repeat rw [← Finset.mul_sum]
/-- Extensionality lemma for `accSU3`. -/ /-- Extensionality lemma for `accSU3`. -/
lemma accSU3_ext {S T : (SMνCharges n).charges} lemma accSU3_ext {S T : (SMνCharges n).Charges}
(hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : (hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
accSU3 S = accSU3 T := by accSU3 S = accSU3 T := by
rw [accSU3_decomp, accSU3_decomp] rw [accSU3_decomp, accSU3_decomp]
@ -196,7 +196,7 @@ lemma accSU3_ext {S T : (SMνCharges n).charges}
/-- The `Y²` anomaly equation. -/ /-- The `Y²` anomaly equation. -/
@[simp] @[simp]
def accYY : (SMνCharges n).charges →ₗ[] where def accYY : (SMνCharges n).Charges →ₗ[] where
toFun S := ∑ i, (Q S i + 8 * U S i + 2 * D S i + 3 * L S i toFun S := ∑ i, (Q S i + 8 * U S i + 2 * D S i + 3 * L S i
+ 6 * E S i) + 6 * E S i)
map_add' S T := by map_add' S T := by
@ -214,7 +214,7 @@ def accYY : (SMνCharges n).charges →ₗ[] where
-- rw [show Rat.cast a = a from rfl] -- rw [show Rat.cast a = a from rfl]
ring ring
lemma accYY_decomp (S : (SMνCharges n).charges) : lemma accYY_decomp (S : (SMνCharges n).Charges) :
accYY S = ∑ i, Q S i + 8 * ∑ i, U S i + 2 * ∑ i, D S i + 3 * ∑ i, L S i + 6 * ∑ i, E S i := by accYY S = ∑ i, Q S i + 8 * ∑ i, U S i + 2 * ∑ i, D S i + 3 * ∑ i, L S i + 6 * ∑ i, E S i := by
simp only [accYY, SMνSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk, simp only [accYY, SMνSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk,
AddHom.coe_mk] AddHom.coe_mk]
@ -222,7 +222,7 @@ lemma accYY_decomp (S : (SMνCharges n).charges) :
repeat rw [← Finset.mul_sum] repeat rw [← Finset.mul_sum]
/-- Extensionality lemma for `accYY`. -/ /-- Extensionality lemma for `accYY`. -/
lemma accYY_ext {S T : (SMνCharges n).charges} lemma accYY_ext {S T : (SMνCharges n).Charges}
(hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : (hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
accYY S = accYY T := by accYY S = accYY T := by
rw [accYY_decomp, accYY_decomp] rw [accYY_decomp, accYY_decomp]
@ -230,7 +230,7 @@ lemma accYY_ext {S T : (SMνCharges n).charges}
/-- The quadratic bilinear map. -/ /-- The quadratic bilinear map. -/
@[simps!] @[simps!]
def quadBiLin : BiLinearSymm (SMνCharges n).charges := BiLinearSymm.mk₂ def quadBiLin : BiLinearSymm (SMνCharges n).Charges := BiLinearSymm.mk₂
(fun S => ∑ i, (Q S.1 i * Q S.2 i + (fun S => ∑ i, (Q S.1 i * Q S.2 i +
- 2 * (U S.1 i * U S.2 i) + - 2 * (U S.1 i * U S.2 i) +
D S.1 i * D S.2 i + D S.1 i * D S.2 i +
@ -262,7 +262,7 @@ def quadBiLin : BiLinearSymm (SMνCharges n).charges := BiLinearSymm.mk₂
intro i intro i
ring) ring)
lemma quadBiLin_decomp (S T : (SMνCharges n).charges) : lemma quadBiLin_decomp (S T : (SMνCharges n).Charges) :
quadBiLin S T = ∑ i, Q S i * Q T i - 2 * ∑ i, U S i * U T i + quadBiLin S T = ∑ i, Q S i * Q T i - 2 * ∑ i, U S i * U T i +
∑ i, D S i * D T i - ∑ i, L S i * L T i + ∑ i, E S i * E T i := by ∑ i, D S i * D T i - ∑ i, L S i * L T i + ∑ i, E S i * E T i := by
erw [← quadBiLin.toFun_eq_coe] erw [← quadBiLin.toFun_eq_coe]
@ -275,17 +275,17 @@ lemma quadBiLin_decomp (S T : (SMνCharges n).charges) :
/-- The quadratic anomaly cancellation condition. -/ /-- The quadratic anomaly cancellation condition. -/
@[simp] @[simp]
def accQuad : HomogeneousQuadratic (SMνCharges n).charges := def accQuad : HomogeneousQuadratic (SMνCharges n).Charges :=
(@quadBiLin n).toHomogeneousQuad (@quadBiLin n).toHomogeneousQuad
lemma accQuad_decomp (S : (SMνCharges n).charges) : lemma accQuad_decomp (S : (SMνCharges n).Charges) :
accQuad S = ∑ i, (Q S i)^2 - 2 * ∑ i, (U S i)^2 + ∑ i, (D S i)^2 - ∑ i, (L S i)^2 accQuad S = ∑ i, (Q S i)^2 - 2 * ∑ i, (U S i)^2 + ∑ i, (D S i)^2 - ∑ i, (L S i)^2
+ ∑ i, (E S i)^2 := by + ∑ i, (E S i)^2 := by
erw [quadBiLin_decomp] erw [quadBiLin_decomp]
ring_nf ring_nf
/-- Extensionality lemma for `accQuad`. -/ /-- Extensionality lemma for `accQuad`. -/
lemma accQuad_ext {S T : (SMνCharges n).charges} lemma accQuad_ext {S T : (SMνCharges n).Charges}
(h : ∀ j, ∑ i, ((fun a => a^2) ∘ toSpecies j S) i = (h : ∀ j, ∑ i, ((fun a => a^2) ∘ toSpecies j S) i =
∑ i, ((fun a => a^2) ∘ toSpecies j T) i) : ∑ i, ((fun a => a^2) ∘ toSpecies j T) i) :
accQuad S = accQuad T := by accQuad S = accQuad T := by
@ -295,7 +295,7 @@ lemma accQuad_ext {S T : (SMνCharges n).charges}
/-- The symmetric trilinear form used to define the cubic acc. -/ /-- The symmetric trilinear form used to define the cubic acc. -/
@[simps!] @[simps!]
def cubeTriLin : TriLinearSymm (SMνCharges n).charges := TriLinearSymm.mk₃ def cubeTriLin : TriLinearSymm (SMνCharges n).Charges := TriLinearSymm.mk₃
(fun S => ∑ i, (6 * ((Q S.1 i) * (Q S.2.1 i) * (Q S.2.2 i)) (fun S => ∑ i, (6 * ((Q S.1 i) * (Q S.2.1 i) * (Q S.2.2 i))
+ 3 * ((U S.1 i) * (U S.2.1 i) * (U S.2.2 i)) + 3 * ((U S.1 i) * (U S.2.1 i) * (U S.2.2 i))
+ 3 * ((D S.1 i) * (D S.2.1 i) * (D S.2.2 i)) + 3 * ((D S.1 i) * (D S.2.1 i) * (D S.2.2 i))
@ -333,7 +333,7 @@ def cubeTriLin : TriLinearSymm (SMνCharges n).charges := TriLinearSymm.mk₃
intro i intro i
ring) ring)
lemma cubeTriLin_decomp (S T R : (SMνCharges n).charges) : lemma cubeTriLin_decomp (S T R : (SMνCharges n).Charges) :
cubeTriLin S T R = 6 * ∑ i, (Q S i * Q T i * Q R i) + 3 * ∑ i, (U S i * U T i * U R i) + cubeTriLin S T R = 6 * ∑ i, (Q S i * Q T i * Q R i) + 3 * ∑ i, (U S i * U T i * U R i) +
3 * ∑ i, (D S i * D T i * D R i) + 2 * ∑ i, (L S i * L T i * L R i) + 3 * ∑ i, (D S i * D T i * D R i) + 2 * ∑ i, (L S i * L T i * L R i) +
∑ i, (E S i * E T i * E R i) + ∑ i, (N S i * N T i * N R i) := by ∑ i, (E S i * E T i * E R i) + ∑ i, (N S i * N T i * N R i) := by
@ -347,16 +347,16 @@ lemma cubeTriLin_decomp (S T R : (SMνCharges n).charges) :
/-- The cubic ACC. -/ /-- The cubic ACC. -/
@[simp] @[simp]
def accCube : HomogeneousCubic (SMνCharges n).charges := cubeTriLin.toCubic def accCube : HomogeneousCubic (SMνCharges n).Charges := cubeTriLin.toCubic
lemma accCube_decomp (S : (SMνCharges n).charges) : lemma accCube_decomp (S : (SMνCharges n).Charges) :
accCube S = 6 * ∑ i, (Q S i)^3 + 3 * ∑ i, (U S i)^3 + 3 * ∑ i, (D S i)^3 + 2 * ∑ i, (L S i)^3 + accCube S = 6 * ∑ i, (Q S i)^3 + 3 * ∑ i, (U S i)^3 + 3 * ∑ i, (D S i)^3 + 2 * ∑ i, (L S i)^3 +
∑ i, (E S i)^3 + ∑ i, (N S i)^3 := by ∑ i, (E S i)^3 + ∑ i, (N S i)^3 := by
erw [cubeTriLin_decomp] erw [cubeTriLin_decomp]
ring_nf ring_nf
/-- Extensionality lemma for `accCube`. -/ /-- Extensionality lemma for `accCube`. -/
lemma accCube_ext {S T : (SMνCharges n).charges} lemma accCube_ext {S T : (SMνCharges n).Charges}
(h : ∀ j, ∑ i, ((fun a => a^3) ∘ toSpecies j S) i = (h : ∀ j, ∑ i, ((fun a => a^3) ∘ toSpecies j S) i =
∑ i, ((fun a => a^3) ∘ toSpecies j T) i) : ∑ i, ((fun a => a^3) ∘ toSpecies j T) i) :
accCube S = accCube T := by accCube S = accCube T := by

View file

@ -19,8 +19,8 @@ open BigOperators
/-- Given a map of for a generic species, the corresponding map for charges. -/ /-- Given a map of for a generic species, the corresponding map for charges. -/
@[simps!] @[simps!]
def chargesMapOfSpeciesMap {n m : } (f : (SMνSpecies n).charges →ₗ[] (SMνSpecies m).charges) : def chargesMapOfSpeciesMap {n m : } (f : (SMνSpecies n).Charges →ₗ[] (SMνSpecies m).Charges) :
(SMνCharges n).charges →ₗ[] (SMνCharges m).charges where (SMνCharges n).Charges →ₗ[] (SMνCharges m).Charges where
toFun S := toSpeciesEquiv.symm (fun i => (LinearMap.comp f (toSpecies i)) S) toFun S := toSpeciesEquiv.symm (fun i => (LinearMap.comp f (toSpecies i)) S)
map_add' S T := by map_add' S T := by
rw [charges_eq_toSpecies_eq] rw [charges_eq_toSpecies_eq]
@ -37,28 +37,28 @@ def chargesMapOfSpeciesMap {n m : } (f : (SMνSpecies n).charges →ₗ[]
rfl rfl
lemma chargesMapOfSpeciesMap_toSpecies {n m : } lemma chargesMapOfSpeciesMap_toSpecies {n m : }
(f : (SMνSpecies n).charges →ₗ[] (SMνSpecies m).charges) (f : (SMνSpecies n).Charges →ₗ[] (SMνSpecies m).Charges)
(S : (SMνCharges n).charges) (j : Fin 6) : (S : (SMνCharges n).Charges) (j : Fin 6) :
toSpecies j (chargesMapOfSpeciesMap f S) = (LinearMap.comp f (toSpecies j)) S := by toSpecies j (chargesMapOfSpeciesMap f S) = (LinearMap.comp f (toSpecies j)) S := by
erw [toSMSpecies_toSpecies_inv] erw [toSMSpecies_toSpecies_inv]
/-- The projection of the `m`-family charges onto the first `n`-family charges for species. -/ /-- The projection of the `m`-family charges onto the first `n`-family charges for species. -/
@[simps!] @[simps!]
def speciesFamilyProj {m n : } (h : n ≤ m) : def speciesFamilyProj {m n : } (h : n ≤ m) :
(SMνSpecies m).charges →ₗ[] (SMνSpecies n).charges where (SMνSpecies m).Charges →ₗ[] (SMνSpecies n).Charges where
toFun S := S ∘ Fin.castLE h toFun S := S ∘ Fin.castLE h
map_add' _ _ := rfl map_add' _ _ := rfl
map_smul' _ _ := rfl map_smul' _ _ := rfl
/-- The projection of the `m`-family charges onto the first `n`-family charges. -/ /-- The projection of the `m`-family charges onto the first `n`-family charges. -/
def familyProjection {m n : } (h : n ≤ m) : (SMνCharges m).charges →ₗ[] (SMνCharges n).charges := def familyProjection {m n : } (h : n ≤ m) : (SMνCharges m).Charges →ₗ[] (SMνCharges n).Charges :=
chargesMapOfSpeciesMap (speciesFamilyProj h) chargesMapOfSpeciesMap (speciesFamilyProj h)
/-- For species, the embedding of the `m`-family charges onto the `n`-family charges, with all /-- For species, the embedding of the `m`-family charges onto the `n`-family charges, with all
other charges zero. -/ other charges zero. -/
@[simps!] @[simps!]
def speciesEmbed (m n : ) : def speciesEmbed (m n : ) :
(SMνSpecies m).charges →ₗ[] (SMνSpecies n).charges where (SMνSpecies m).Charges →ₗ[] (SMνSpecies n).Charges where
toFun S := fun i => toFun S := fun i =>
if hi : i.val < m then if hi : i.val < m then
S ⟨i, hi⟩ S ⟨i, hi⟩
@ -82,29 +82,29 @@ def speciesEmbed (m n : ) :
/-- The embedding of the `m`-family charges onto the `n`-family charges, with all /-- The embedding of the `m`-family charges onto the `n`-family charges, with all
other charges zero. -/ other charges zero. -/
def familyEmbedding (m n : ) : (SMνCharges m).charges →ₗ[] (SMνCharges n).charges := def familyEmbedding (m n : ) : (SMνCharges m).Charges →ₗ[] (SMνCharges n).Charges :=
chargesMapOfSpeciesMap (speciesEmbed m n) chargesMapOfSpeciesMap (speciesEmbed m n)
/-- For species, the embedding 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. -/ a universal manor. -/
@[simps!] @[simps!]
def speciesFamilyUniversial (n : ) : def speciesFamilyUniversial (n : ) :
(SMνSpecies 1).charges →ₗ[] (SMνSpecies n).charges where (SMνSpecies 1).Charges →ₗ[] (SMνSpecies n).Charges where
toFun S _ := S ⟨0, by simp⟩ toFun S _ := S ⟨0, by simp⟩
map_add' S T := rfl map_add' S T := rfl
map_smul' a S := rfl map_smul' a S := rfl
/-- The embedding 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. -/ a universal manor. -/
def familyUniversal (n : ) : (SMνCharges 1).charges →ₗ[] (SMνCharges n).charges := def familyUniversal (n : ) : (SMνCharges 1).Charges →ₗ[] (SMνCharges n).Charges :=
chargesMapOfSpeciesMap (speciesFamilyUniversial n) chargesMapOfSpeciesMap (speciesFamilyUniversial n)
lemma toSpecies_familyUniversal {n : } (j : Fin 6) (S : (SMνCharges 1).charges) lemma toSpecies_familyUniversal {n : } (j : Fin 6) (S : (SMνCharges 1).Charges)
(i : Fin n) : toSpecies j (familyUniversal n S) i = toSpecies j S ⟨0, by simp⟩ := by (i : Fin n) : toSpecies j (familyUniversal n S) i = toSpecies j S ⟨0, by simp⟩ := by
erw [chargesMapOfSpeciesMap_toSpecies] erw [chargesMapOfSpeciesMap_toSpecies]
rfl rfl
lemma sum_familyUniversal {n : } (m : ) (S : (SMνCharges 1).charges) (j : Fin 6) : lemma sum_familyUniversal {n : } (m : ) (S : (SMνCharges 1).Charges) (j : Fin 6) :
∑ i, ((fun a => a ^ m) ∘ toSpecies j (familyUniversal n S)) i = ∑ i, ((fun a => a ^ m) ∘ toSpecies j (familyUniversal n S)) i =
n * (toSpecies j S ⟨0, by simp⟩) ^ m := by n * (toSpecies j S ⟨0, by simp⟩) ^ m := by
simp only [SMνSpecies_numberCharges, Function.comp_apply, toSpecies_apply, Fin.zero_eta, simp only [SMνSpecies_numberCharges, Function.comp_apply, toSpecies_apply, Fin.zero_eta,
@ -119,12 +119,12 @@ lemma sum_familyUniversal {n : } (m : ) (S : (SMνCharges 1).charges) (j :
intro i _ intro i _
erw [toSpecies_familyUniversal] erw [toSpecies_familyUniversal]
lemma sum_familyUniversal_one {n : } (S : (SMνCharges 1).charges) (j : Fin 6) : lemma sum_familyUniversal_one {n : } (S : (SMνCharges 1).Charges) (j : Fin 6) :
∑ i, toSpecies j (familyUniversal n S) i = n * (toSpecies j S ⟨0, by simp⟩) := by ∑ i, toSpecies j (familyUniversal n S) i = n * (toSpecies j S ⟨0, by simp⟩) := by
simpa using @sum_familyUniversal n 1 S j simpa using @sum_familyUniversal n 1 S j
lemma sum_familyUniversal_two {n : } (S : (SMνCharges 1).charges) lemma sum_familyUniversal_two {n : } (S : (SMνCharges 1).Charges)
(T : (SMνCharges n).charges) (j : Fin 6) : (T : (SMνCharges n).Charges) (j : Fin 6) :
∑ i, (toSpecies j (familyUniversal n S) i * toSpecies j T i) = ∑ i, (toSpecies j (familyUniversal n S) i * toSpecies j T i) =
(toSpecies j S ⟨0, by simp⟩) * ∑ i, toSpecies j T i := by (toSpecies j S ⟨0, by simp⟩) * ∑ i, toSpecies j T i := by
simp only [SMνSpecies_numberCharges, toSpecies_apply, Fin.zero_eta, Fin.isValue] simp only [SMνSpecies_numberCharges, toSpecies_apply, Fin.zero_eta, Fin.isValue]
@ -135,8 +135,8 @@ lemma sum_familyUniversal_two {n : } (S : (SMνCharges 1).charges)
erw [toSpecies_familyUniversal] erw [toSpecies_familyUniversal]
rfl rfl
lemma sum_familyUniversal_three {n : } (S : (SMνCharges 1).charges) lemma sum_familyUniversal_three {n : } (S : (SMνCharges 1).Charges)
(T L : (SMνCharges n).charges) (j : Fin 6) : (T L : (SMνCharges n).Charges) (j : Fin 6) :
∑ i, (toSpecies j (familyUniversal n S) i * toSpecies j T i * toSpecies j L i) = ∑ i, (toSpecies j (familyUniversal n S) i * toSpecies j T i * toSpecies j L i) =
(toSpecies j S ⟨0, by simp⟩) * ∑ i, toSpecies j T i * toSpecies j L i := by (toSpecies j S ⟨0, by simp⟩) * ∑ i, toSpecies j T i * toSpecies j L i := by
simp only [SMνSpecies_numberCharges, toSpecies_apply, Fin.zero_eta, Fin.isValue] simp only [SMνSpecies_numberCharges, toSpecies_apply, Fin.zero_eta, Fin.isValue]
@ -148,7 +148,7 @@ lemma sum_familyUniversal_three {n : } (S : (SMνCharges 1).charges)
simp only [SMνSpecies_numberCharges, Fin.zero_eta, Fin.isValue, toSpecies_apply] simp only [SMνSpecies_numberCharges, Fin.zero_eta, Fin.isValue, toSpecies_apply]
ring ring
lemma familyUniversal_accGrav (S : (SMνCharges 1).charges) : lemma familyUniversal_accGrav (S : (SMνCharges 1).Charges) :
accGrav (familyUniversal n S) = n * (accGrav S) := by accGrav (familyUniversal n S) = n * (accGrav S) := by
rw [accGrav_decomp, accGrav_decomp] rw [accGrav_decomp, accGrav_decomp]
repeat rw [sum_familyUniversal_one] repeat rw [sum_familyUniversal_one]
@ -156,7 +156,7 @@ lemma familyUniversal_accGrav (S : (SMνCharges 1).charges) :
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton] Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
ring ring
lemma familyUniversal_accSU2 (S : (SMνCharges 1).charges) : lemma familyUniversal_accSU2 (S : (SMνCharges 1).Charges) :
accSU2 (familyUniversal n S) = n * (accSU2 S) := by accSU2 (familyUniversal n S) = n * (accSU2 S) := by
rw [accSU2_decomp, accSU2_decomp] rw [accSU2_decomp, accSU2_decomp]
repeat rw [sum_familyUniversal_one] repeat rw [sum_familyUniversal_one]
@ -164,7 +164,7 @@ lemma familyUniversal_accSU2 (S : (SMνCharges 1).charges) :
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton] Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
ring ring
lemma familyUniversal_accSU3 (S : (SMνCharges 1).charges) : lemma familyUniversal_accSU3 (S : (SMνCharges 1).Charges) :
accSU3 (familyUniversal n S) = n * (accSU3 S) := by accSU3 (familyUniversal n S) = n * (accSU3 S) := by
rw [accSU3_decomp, accSU3_decomp] rw [accSU3_decomp, accSU3_decomp]
repeat rw [sum_familyUniversal_one] repeat rw [sum_familyUniversal_one]
@ -172,7 +172,7 @@ lemma familyUniversal_accSU3 (S : (SMνCharges 1).charges) :
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton] Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
ring ring
lemma familyUniversal_accYY (S : (SMνCharges 1).charges) : lemma familyUniversal_accYY (S : (SMνCharges 1).Charges) :
accYY (familyUniversal n S) = n * (accYY S) := by accYY (familyUniversal n S) = n * (accYY S) := by
rw [accYY_decomp, accYY_decomp] rw [accYY_decomp, accYY_decomp]
repeat rw [sum_familyUniversal_one] repeat rw [sum_familyUniversal_one]
@ -180,7 +180,7 @@ lemma familyUniversal_accYY (S : (SMνCharges 1).charges) :
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton] Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
ring ring
lemma familyUniversal_quadBiLin (S : (SMνCharges 1).charges) (T : (SMνCharges n).charges) : lemma familyUniversal_quadBiLin (S : (SMνCharges 1).Charges) (T : (SMνCharges n).Charges) :
quadBiLin (familyUniversal n S) T = quadBiLin (familyUniversal n S) T =
S (0 : Fin 6) * ∑ i, Q T i - 2 * S (1 : Fin 6) * ∑ i, U T i + S (2 : Fin 6) *∑ i, D T i - S (0 : Fin 6) * ∑ i, Q T i - 2 * S (1 : Fin 6) * ∑ i, U T i + S (2 : Fin 6) *∑ i, D T i -
S (3 : Fin 6) * ∑ i, L T i + S (4 : Fin 6) * ∑ i, E T i := by S (3 : Fin 6) * ∑ i, L T i + S (4 : Fin 6) * ∑ i, E T i := by
@ -191,7 +191,7 @@ lemma familyUniversal_quadBiLin (S : (SMνCharges 1).charges) (T : (SMνCharges
sub_right_inj] sub_right_inj]
ring ring
lemma familyUniversal_accQuad (S : (SMνCharges 1).charges) : lemma familyUniversal_accQuad (S : (SMνCharges 1).Charges) :
accQuad (familyUniversal n S) = n * (accQuad S) := by accQuad (familyUniversal n S) = n * (accQuad S) := by
rw [accQuad_decomp] rw [accQuad_decomp]
repeat erw [sum_familyUniversal] repeat erw [sum_familyUniversal]
@ -200,7 +200,7 @@ lemma familyUniversal_accQuad (S : (SMνCharges 1).charges) :
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton] Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
ring ring
lemma familyUniversal_cubeTriLin (S : (SMνCharges 1).charges) (T R : (SMνCharges n).charges) : lemma familyUniversal_cubeTriLin (S : (SMνCharges 1).Charges) (T R : (SMνCharges n).Charges) :
cubeTriLin (familyUniversal n S) T R = 6 * S (0 : Fin 6) * ∑ i, (Q T i * Q R i) + cubeTriLin (familyUniversal n S) T R = 6 * S (0 : Fin 6) * ∑ i, (Q T i * Q R i) +
3 * S (1 : Fin 6) * ∑ i, (U T i * U R i) + 3 * S (2 : Fin 6) * ∑ i, (D T i * D R i) 3 * S (1 : Fin 6) * ∑ i, (U T i * U R i) + 3 * S (2 : Fin 6) * ∑ i, (D T i * D R i)
+ 2 * S (3 : Fin 6) * ∑ i, (L T i * L R i) + + 2 * S (3 : Fin 6) * ∑ i, (L T i * L R i) +
@ -211,7 +211,7 @@ lemma familyUniversal_cubeTriLin (S : (SMνCharges 1).charges) (T R : (SMνCharg
simp only [Fin.isValue, SMνSpecies_numberCharges, toSpecies_apply, add_left_inj] simp only [Fin.isValue, SMνSpecies_numberCharges, toSpecies_apply, add_left_inj]
ring ring
lemma familyUniversal_cubeTriLin' (S T : (SMνCharges 1).charges) (R : (SMνCharges n).charges) : lemma familyUniversal_cubeTriLin' (S T : (SMνCharges 1).Charges) (R : (SMνCharges n).Charges) :
cubeTriLin (familyUniversal n S) (familyUniversal n T) R = cubeTriLin (familyUniversal n S) (familyUniversal n T) R =
6 * S (0 : Fin 6) * T (0 : Fin 6) * ∑ i, Q R i + 6 * S (0 : Fin 6) * T (0 : Fin 6) * ∑ i, Q R i +
3 * S (1 : Fin 6) * T (1 : Fin 6) * ∑ i, U R i 3 * S (1 : Fin 6) * T (1 : Fin 6) * ∑ i, U R i
@ -224,7 +224,7 @@ lemma familyUniversal_cubeTriLin' (S T : (SMνCharges 1).charges) (R : (SMνChar
simp only [Fin.isValue, SMνSpecies_numberCharges, toSpecies_apply] simp only [Fin.isValue, SMνSpecies_numberCharges, toSpecies_apply]
ring ring
lemma familyUniversal_accCube (S : (SMνCharges 1).charges) : lemma familyUniversal_accCube (S : (SMνCharges 1).Charges) :
accCube (familyUniversal n S) = n * (accCube S) := by accCube (familyUniversal n S) = n * (accCube S) := by
rw [accCube_decomp] rw [accCube_decomp]
repeat erw [sum_familyUniversal] repeat erw [sum_familyUniversal]

View file

@ -53,7 +53,7 @@ lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := by
/-- An element of `charges` which satisfies the linear ACCs /-- An element of `charges` which satisfies the linear ACCs
gives us a element of `LinSols`. -/ gives us a element of `LinSols`. -/
def chargeToLinear (S : (SMNoGrav n).charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) : def chargeToLinear (S : (SMNoGrav n).Charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) :
(SMNoGrav n).LinSols := (SMNoGrav n).LinSols :=
⟨S, by ⟨S, by
intro i intro i
@ -76,13 +76,13 @@ def quadToAF (S : (SMNoGrav n).QuadSols) (hc : accCube S.val = 0) :
/-- An element of `charges` which satisfies the linear and quadratic ACCs /-- An element of `charges` which satisfies the linear and quadratic ACCs
gives us a element of `QuadSols`. -/ gives us a element of `QuadSols`. -/
def chargeToQuad (S : (SMNoGrav n).charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) : def chargeToQuad (S : (SMNoGrav n).Charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) :
(SMNoGrav n).QuadSols := (SMNoGrav n).QuadSols :=
linearToQuad $ chargeToLinear S hSU2 hSU3 linearToQuad $ chargeToLinear S hSU2 hSU3
/-- An element of `charges` which satisfies the linear, quadratic and cubic ACCs /-- An element of `charges` which satisfies the linear, quadratic and cubic ACCs
gives us a element of `Sols`. -/ gives us a element of `Sols`. -/
def chargeToAF (S : (SMNoGrav n).charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) def chargeToAF (S : (SMNoGrav n).Charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0)
(hc : accCube S = 0) : (SMNoGrav n).Sols := (hc : accCube S = 0) : (SMNoGrav n).Sols :=
quadToAF (chargeToQuad S hSU2 hSU3) hc quadToAF (chargeToQuad S hSU2 hSU3) hc
@ -95,7 +95,7 @@ def linearToAF (S : (SMNoGrav n).LinSols)
/-- The permutations acting on the ACC system corresponding to the SM with RHN, /-- The permutations acting on the ACC system corresponding to the SM with RHN,
and no gravitational anomaly. -/ and no gravitational anomaly. -/
def perm (n : ) : ACCSystemGroupAction (SMNoGrav n) where def perm (n : ) : ACCSystemGroupAction (SMNoGrav n) where
group := permGroup n group := PermGroup n
groupInst := inferInstance groupInst := inferInstance
rep := repCharges rep := repCharges
linearInvariant := by linearInvariant := by

View file

@ -60,7 +60,7 @@ lemma cubeSol (S : (SM n).Sols) : accCube S.val = 0 := by
/-- An element of `charges` which satisfies the linear ACCs /-- An element of `charges` which satisfies the linear ACCs
gives us a element of `LinSols`. -/ gives us a element of `LinSols`. -/
def chargeToLinear (S : (SM n).charges) (hGrav : accGrav S = 0) def chargeToLinear (S : (SM n).Charges) (hGrav : accGrav S = 0)
(hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) : (SM n).LinSols := (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) : (SM n).LinSols :=
⟨S, by ⟨S, by
intro i intro i
@ -84,14 +84,14 @@ def quadToAF (S : (SM n).QuadSols) (hc : accCube S.val = 0) :
/-- An element of `charges` which satisfies the linear and quadratic ACCs /-- An element of `charges` which satisfies the linear and quadratic ACCs
gives us a element of `QuadSols`. -/ gives us a element of `QuadSols`. -/
def chargeToQuad (S : (SM n).charges) (hGrav : accGrav S = 0) def chargeToQuad (S : (SM n).Charges) (hGrav : accGrav S = 0)
(hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) : (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) :
(SM n).QuadSols := (SM n).QuadSols :=
linearToQuad $ chargeToLinear S hGrav hSU2 hSU3 linearToQuad $ chargeToLinear S hGrav hSU2 hSU3
/-- An element of `charges` which satisfies the linear, quadratic and cubic ACCs /-- An element of `charges` which satisfies the linear, quadratic and cubic ACCs
gives us a element of `Sols`. -/ gives us a element of `Sols`. -/
def chargeToAF (S : (SM n).charges) (hGrav : accGrav S = 0) (hSU2 : accSU2 S = 0) def chargeToAF (S : (SM n).Charges) (hGrav : accGrav S = 0) (hSU2 : accSU2 S = 0)
(hSU3 : accSU3 S = 0) (hc : accCube S = 0) : (SM n).Sols := (hSU3 : accSU3 S = 0) (hc : accCube S = 0) : (SM n).Sols :=
quadToAF (chargeToQuad S hGrav hSU2 hSU3) hc quadToAF (chargeToQuad S hGrav hSU2 hSU3) hc
@ -103,7 +103,7 @@ def linearToAF (S : (SM n).LinSols)
/-- The permutations acting on the ACC system corresponding to the SM with RHN. -/ /-- The permutations acting on the ACC system corresponding to the SM with RHN. -/
def perm (n : ) : ACCSystemGroupAction (SM n) where def perm (n : ) : ACCSystemGroupAction (SM n) where
group := permGroup n group := PermGroup n
groupInst := inferInstance groupInst := inferInstance
rep := repCharges rep := repCharges
linearInvariant := by linearInvariant := by

View file

@ -24,102 +24,102 @@ open BigOperators
namespace PlaneSeven namespace PlaneSeven
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₀ : (SM 3).charges := toSpeciesEquiv.invFun ( fun s => fun i => def B₀ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with match s, i with
| 0, 0 => 1 | 0, 0 => 1
| 0, 1 => - 1 | 0, 1 => - 1
| _, _ => 0 | _, _ => 0
) )
lemma B₀_cubic (S T : (SM 3).charges) : cubeTriLin B₀ S T = lemma B₀_cubic (S T : (SM 3).Charges) : cubeTriLin B₀ S T =
6 * (S (0 : Fin 18) * T (0 : Fin 18) - S (1 : Fin 18) * T (1 : Fin 18)) := by 6 * (S (0 : Fin 18) * T (0 : Fin 18) - S (1 : Fin 18) * T (1 : Fin 18)) := by
simp [Fin.sum_univ_three, B₀, Fin.divNat, Fin.modNat, finProdFinEquiv] simp [Fin.sum_univ_three, B₀, Fin.divNat, Fin.modNat, finProdFinEquiv]
ring ring
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₁ : (SM 3).charges := toSpeciesEquiv.invFun ( fun s => fun i => def B₁ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with match s, i with
| 1, 0 => 1 | 1, 0 => 1
| 1, 1 => - 1 | 1, 1 => - 1
| _, _ => 0 | _, _ => 0
) )
lemma B₁_cubic (S T : (SM 3).charges) : cubeTriLin B₁ S T = lemma B₁_cubic (S T : (SM 3).Charges) : cubeTriLin B₁ S T =
3 * (S (3 : Fin 18) * T (3 : Fin 18) - S (4 : Fin 18) * T (4 : Fin 18)) := by 3 * (S (3 : Fin 18) * T (3 : Fin 18) - S (4 : Fin 18) * T (4 : Fin 18)) := by
simp [Fin.sum_univ_three, B₁, Fin.divNat, Fin.modNat, finProdFinEquiv] simp [Fin.sum_univ_three, B₁, Fin.divNat, Fin.modNat, finProdFinEquiv]
ring ring
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₂ : (SM 3).charges := toSpeciesEquiv.invFun ( fun s => fun i => def B₂ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with match s, i with
| 2, 0 => 1 | 2, 0 => 1
| 2, 1 => - 1 | 2, 1 => - 1
| _, _ => 0 | _, _ => 0
) )
lemma B₂_cubic (S T : (SM 3).charges) : cubeTriLin B₂ S T = lemma B₂_cubic (S T : (SM 3).Charges) : cubeTriLin B₂ S T =
3 * (S (6 : Fin 18) * T (6 : Fin 18) - S (7 : Fin 18) * T (7 : Fin 18)) := by 3 * (S (6 : Fin 18) * T (6 : Fin 18) - S (7 : Fin 18) * T (7 : Fin 18)) := by
simp [Fin.sum_univ_three, B₂, Fin.divNat, Fin.modNat, finProdFinEquiv] simp [Fin.sum_univ_three, B₂, Fin.divNat, Fin.modNat, finProdFinEquiv]
ring ring
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₃ : (SM 3).charges := toSpeciesEquiv.invFun ( fun s => fun i => def B₃ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with match s, i with
| 3, 0 => 1 | 3, 0 => 1
| 3, 1 => - 1 | 3, 1 => - 1
| _, _ => 0 | _, _ => 0
) )
lemma B₃_cubic (S T : (SM 3).charges) : cubeTriLin B₃ S T = lemma B₃_cubic (S T : (SM 3).Charges) : cubeTriLin B₃ S T =
2 * (S (9 : Fin 18) * T (9 : Fin 18) - S (10 : Fin 18) * T (10 : Fin 18)) := by 2 * (S (9 : Fin 18) * T (9 : Fin 18) - S (10 : Fin 18) * T (10 : Fin 18)) := by
simp [Fin.sum_univ_three, B₃, Fin.divNat, Fin.modNat, finProdFinEquiv] simp [Fin.sum_univ_three, B₃, Fin.divNat, Fin.modNat, finProdFinEquiv]
ring_nf ring_nf
rfl rfl
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₄ : (SM 3).charges := toSpeciesEquiv.invFun ( fun s => fun i => def B₄ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with match s, i with
| 4, 0 => 1 | 4, 0 => 1
| 4, 1 => - 1 | 4, 1 => - 1
| _, _ => 0 | _, _ => 0
) )
lemma B₄_cubic (S T : (SM 3).charges) : cubeTriLin B₄ S T = lemma B₄_cubic (S T : (SM 3).Charges) : cubeTriLin B₄ S T =
(S (12 : Fin 18) * T (12 : Fin 18) - S (13 : Fin 18) * T (13 : Fin 18)) := by (S (12 : Fin 18) * T (12 : Fin 18) - S (13 : Fin 18) * T (13 : Fin 18)) := by
simp [Fin.sum_univ_three, B₄, Fin.divNat, Fin.modNat, finProdFinEquiv] simp [Fin.sum_univ_three, B₄, Fin.divNat, Fin.modNat, finProdFinEquiv]
ring_nf ring_nf
rfl rfl
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₅ : (SM 3).charges := toSpeciesEquiv.invFun ( fun s => fun i => def B₅ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with match s, i with
| 5, 0 => 1 | 5, 0 => 1
| 5, 1 => - 1 | 5, 1 => - 1
| _, _ => 0 | _, _ => 0
) )
lemma B₅_cubic (S T : (SM 3).charges) : cubeTriLin B₅ S T = lemma B₅_cubic (S T : (SM 3).Charges) : cubeTriLin B₅ S T =
(S (15 : Fin 18) * T (15 : Fin 18) - S (16 : Fin 18) * T (16 : Fin 18)) := by (S (15 : Fin 18) * T (15 : Fin 18) - S (16 : Fin 18) * T (16 : Fin 18)) := by
simp [Fin.sum_univ_three, B₅, Fin.divNat, Fin.modNat, finProdFinEquiv] simp [Fin.sum_univ_three, B₅, Fin.divNat, Fin.modNat, finProdFinEquiv]
ring_nf ring_nf
rfl rfl
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₆ : (SM 3).charges := toSpeciesEquiv.invFun ( fun s => fun i => def B₆ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with match s, i with
| 1, 2 => 1 | 1, 2 => 1
| 2, 2 => -1 | 2, 2 => -1
| _, _ => 0 | _, _ => 0
) )
lemma B₆_cubic (S T : (SM 3).charges) : cubeTriLin B₆ S T = lemma B₆_cubic (S T : (SM 3).Charges) : cubeTriLin B₆ S T =
3* (S (5 : Fin 18) * T (5 : Fin 18) - S (8 : Fin 18) * T (8 : Fin 18)) := by 3* (S (5 : Fin 18) * T (5 : Fin 18) - S (8 : Fin 18) * T (8 : Fin 18)) := by
simp [Fin.sum_univ_three, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv] simp [Fin.sum_univ_three, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
ring_nf ring_nf
/-- The charge assignments forming a basis of the plane. -/ /-- The charge assignments forming a basis of the plane. -/
@[simp] @[simp]
def B : Fin 7 → (SM 3).charges := fun i => def B : Fin 7 → (SM 3).Charges := fun i =>
match i with match i with
| 0 => B₀ | 0 => B₀
| 1 => B₁ | 1 => B₁
@ -129,7 +129,7 @@ def B : Fin 7 → (SM 3).charges := fun i =>
| 5 => B₅ | 5 => B₅
| 6 => B₆ | 6 => B₆
lemma B₀_Bi_cubic {i : Fin 7} (hi : 0 ≠ i) (S : (SM 3).charges) : lemma B₀_Bi_cubic {i : Fin 7} (hi : 0 ≠ i) (S : (SM 3).Charges) :
cubeTriLin (B 0) (B i) S = 0 := by cubeTriLin (B 0) (B i) S = 0 := by
change cubeTriLin B₀ (B i) S = 0 change cubeTriLin B₀ (B i) S = 0
rw [B₀_cubic] rw [B₀_cubic]
@ -138,7 +138,7 @@ lemma B₀_Bi_cubic {i : Fin 7} (hi : 0 ≠ i) (S : (SM 3).charges) :
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv] simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
lemma B₁_Bi_cubic {i : Fin 7} (hi : 1 ≠ i) (S : (SM 3).charges) : lemma B₁_Bi_cubic {i : Fin 7} (hi : 1 ≠ i) (S : (SM 3).Charges) :
cubeTriLin (B 1) (B i) S = 0 := by cubeTriLin (B 1) (B i) S = 0 := by
change cubeTriLin B₁ (B i) S = 0 change cubeTriLin B₁ (B i) S = 0
rw [B₁_cubic] rw [B₁_cubic]
@ -146,7 +146,7 @@ lemma B₁_Bi_cubic {i : Fin 7} (hi : 1 ≠ i) (S : (SM 3).charges) :
simp at hi <;> simp at hi <;>
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv] simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
lemma B₂_Bi_cubic {i : Fin 7} (hi : 2 ≠ i) (S : (SM 3).charges) : lemma B₂_Bi_cubic {i : Fin 7} (hi : 2 ≠ i) (S : (SM 3).Charges) :
cubeTriLin (B 2) (B i) S = 0 := by cubeTriLin (B 2) (B i) S = 0 := by
change cubeTriLin B₂ (B i) S = 0 change cubeTriLin B₂ (B i) S = 0
rw [B₂_cubic] rw [B₂_cubic]
@ -154,7 +154,7 @@ lemma B₂_Bi_cubic {i : Fin 7} (hi : 2 ≠ i) (S : (SM 3).charges) :
simp at hi <;> simp at hi <;>
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv] simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
lemma B₃_Bi_cubic {i : Fin 7} (hi : 3 ≠ i) (S : (SM 3).charges) : lemma B₃_Bi_cubic {i : Fin 7} (hi : 3 ≠ i) (S : (SM 3).Charges) :
cubeTriLin (B 3) (B i) S = 0 := by cubeTriLin (B 3) (B i) S = 0 := by
change cubeTriLin (B₃) (B i) S = 0 change cubeTriLin (B₃) (B i) S = 0
rw [B₃_cubic] rw [B₃_cubic]
@ -162,7 +162,7 @@ lemma B₃_Bi_cubic {i : Fin 7} (hi : 3 ≠ i) (S : (SM 3).charges) :
simp at hi <;> simp at hi <;>
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv] simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
lemma B₄_Bi_cubic {i : Fin 7} (hi : 4 ≠ i) (S : (SM 3).charges) : lemma B₄_Bi_cubic {i : Fin 7} (hi : 4 ≠ i) (S : (SM 3).Charges) :
cubeTriLin (B 4) (B i) S = 0 := by cubeTriLin (B 4) (B i) S = 0 := by
change cubeTriLin (B₄) (B i) S = 0 change cubeTriLin (B₄) (B i) S = 0
rw [B₄_cubic] rw [B₄_cubic]
@ -170,7 +170,7 @@ lemma B₄_Bi_cubic {i : Fin 7} (hi : 4 ≠ i) (S : (SM 3).charges) :
simp at hi <;> simp at hi <;>
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv] simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
lemma B₅_Bi_cubic {i : Fin 7} (hi : 5 ≠ i) (S : (SM 3).charges) : lemma B₅_Bi_cubic {i : Fin 7} (hi : 5 ≠ i) (S : (SM 3).Charges) :
cubeTriLin (B 5) (B i) S = 0 := by cubeTriLin (B 5) (B i) S = 0 := by
change cubeTriLin (B₅) (B i) S = 0 change cubeTriLin (B₅) (B i) S = 0
rw [B₅_cubic] rw [B₅_cubic]
@ -178,7 +178,7 @@ lemma B₅_Bi_cubic {i : Fin 7} (hi : 5 ≠ i) (S : (SM 3).charges) :
simp at hi <;> simp at hi <;>
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv] simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
lemma B₆_Bi_cubic {i : Fin 7} (hi : 6 ≠ i) (S : (SM 3).charges) : lemma B₆_Bi_cubic {i : Fin 7} (hi : 6 ≠ i) (S : (SM 3).Charges) :
cubeTriLin (B 6) (B i) S = 0 := by cubeTriLin (B 6) (B i) S = 0 := by
change cubeTriLin (B₆) (B i) S = 0 change cubeTriLin (B₆) (B i) S = 0
rw [B₆_cubic] rw [B₆_cubic]
@ -186,7 +186,7 @@ lemma B₆_Bi_cubic {i : Fin 7} (hi : 6 ≠ i) (S : (SM 3).charges) :
simp at hi <;> simp at hi <;>
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv] simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
lemma Bi_Bj_ne_cubic {i j : Fin 7} (h : i ≠ j) (S : (SM 3).charges) : lemma Bi_Bj_ne_cubic {i j : Fin 7} (h : i ≠ j) (S : (SM 3).Charges) :
cubeTriLin (B i) (B j) S = 0 := by cubeTriLin (B i) (B j) S = 0 := by
fin_cases i fin_cases i
exact B₀_Bi_cubic h S exact B₀_Bi_cubic h S
@ -279,7 +279,7 @@ theorem B_in_accCube (f : Fin 7 → ) : accCube (∑ i, f i • B i) = 0 := b
rw [Bi_Bj_Bk_cubic] rw [Bi_Bj_Bk_cubic]
simp simp
lemma B_sum_is_sol (f : Fin 7 → ) : (SM 3).isSolution (∑ i, f i • B i) := by lemma B_sum_is_sol (f : Fin 7 → ) : (SM 3).IsSolution (∑ i, f i • B i) := by
let X := chargeToAF (∑ i, f i • B i) (by let X := chargeToAF (∑ i, f i • B i) (by
rw [map_sum] rw [map_sum]
apply Fintype.sum_eq_zero apply Fintype.sum_eq_zero
@ -337,8 +337,8 @@ theorem basis_linear_independent : LinearIndependent B := by
end PlaneSeven end PlaneSeven
theorem seven_dim_plane_exists : ∃ (B : Fin 7 → (SM 3).charges), theorem seven_dim_plane_exists : ∃ (B : Fin 7 → (SM 3).Charges),
LinearIndependent B ∧ ∀ (f : Fin 7 → ), (SM 3).isSolution (∑ i, f i • B i) := by LinearIndependent B ∧ ∀ (f : Fin 7 → ), (SM 3).IsSolution (∑ i, f i • B i) := by
use PlaneSeven.B use PlaneSeven.B
apply And.intro apply And.intro
exact PlaneSeven.basis_linear_independent exact PlaneSeven.basis_linear_independent

View file

@ -26,26 +26,26 @@ open BigOperators
/-- The group of `Sₙ` permutations for each species. -/ /-- The group of `Sₙ` permutations for each species. -/
@[simp] @[simp]
def permGroup (n : ) := Fin 6 → Equiv.Perm (Fin n) def PermGroup (n : ) := Fin 6 → Equiv.Perm (Fin n)
variable {n : } variable {n : }
@[simp] @[simp]
instance : Group (permGroup n) := Pi.group instance : Group (PermGroup n) := Pi.group
/-- The image of an element of `permGroup n` under the representation on charges. -/ /-- The image of an element of `permGroup n` under the representation on charges. -/
@[simps!] @[simps!]
def chargeMap (f : permGroup n) : (SMνCharges n).charges →ₗ[] (SMνCharges n).charges where def chargeMap (f : PermGroup n) : (SMνCharges n).Charges →ₗ[] (SMνCharges n).Charges where
toFun S := toSpeciesEquiv.symm (fun i => toSpecies i S ∘ f i ) toFun S := toSpeciesEquiv.symm (fun i => toSpecies i S ∘ f i )
map_add' _ _ := rfl map_add' _ _ := rfl
map_smul' _ _ := rfl map_smul' _ _ := rfl
/-- The representation of `(permGroup n)` acting on the vector space of charges. -/ /-- The representation of `(permGroup n)` acting on the vector space of charges. -/
@[simp] @[simp]
def repCharges {n : } : Representation (permGroup n) (SMνCharges n).charges where def repCharges {n : } : Representation (PermGroup n) (SMνCharges n).Charges where
toFun f := chargeMap f⁻¹ toFun f := chargeMap f⁻¹
map_mul' f g := by map_mul' f g := by
simp only [permGroup, mul_inv_rev] simp only [PermGroup, mul_inv_rev]
apply LinearMap.ext apply LinearMap.ext
intro S intro S
rw [charges_eq_toSpecies_eq] rw [charges_eq_toSpecies_eq]
@ -61,49 +61,49 @@ def repCharges {n : } : Representation (permGroup n) (SMνCharges n).char
erw [toSMSpecies_toSpecies_inv] erw [toSMSpecies_toSpecies_inv]
rfl rfl
lemma repCharges_toSpecies (f : permGroup n) (S : (SMνCharges n).charges) (j : Fin 6) : lemma repCharges_toSpecies (f : PermGroup n) (S : (SMνCharges n).Charges) (j : Fin 6) :
toSpecies j (repCharges f S) = toSpecies j S ∘ f⁻¹ j := by toSpecies j (repCharges f S) = toSpecies j S ∘ f⁻¹ j := by
erw [toSMSpecies_toSpecies_inv] erw [toSMSpecies_toSpecies_inv]
lemma toSpecies_sum_invariant (m : ) (f : permGroup n) (S : (SMνCharges n).charges) (j : Fin 6) : lemma toSpecies_sum_invariant (m : ) (f : PermGroup n) (S : (SMνCharges n).Charges) (j : Fin 6) :
∑ i, ((fun a => a ^ m) ∘ toSpecies j (repCharges f S)) i = ∑ i, ((fun a => a ^ m) ∘ toSpecies j (repCharges f S)) i =
∑ i, ((fun a => a ^ m) ∘ toSpecies j S) i := by ∑ i, ((fun a => a ^ m) ∘ toSpecies j S) i := by
erw [repCharges_toSpecies] erw [repCharges_toSpecies]
change ∑ i : Fin n, ((fun a => a ^ m) ∘ _) (⇑(f⁻¹ _) i) = ∑ i : Fin n, ((fun a => a ^ m) ∘ _) i change ∑ i : Fin n, ((fun a => a ^ m) ∘ _) (⇑(f⁻¹ _) i) = ∑ i : Fin n, ((fun a => a ^ m) ∘ _) i
refine Equiv.Perm.sum_comp _ _ _ ?_ refine Equiv.Perm.sum_comp _ _ _ ?_
simp only [permGroup, Fin.isValue, Pi.inv_apply, ne_eq, coe_univ, Set.subset_univ] simp only [PermGroup, Fin.isValue, Pi.inv_apply, ne_eq, coe_univ, Set.subset_univ]
lemma accGrav_invariant (f : permGroup n) (S : (SMνCharges n).charges) : lemma accGrav_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
accGrav (repCharges f S) = accGrav S := accGrav (repCharges f S) = accGrav S :=
accGrav_ext accGrav_ext
(by simpa using toSpecies_sum_invariant 1 f S) (by simpa using toSpecies_sum_invariant 1 f S)
lemma accSU2_invariant (f : permGroup n) (S : (SMνCharges n).charges) : lemma accSU2_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
accSU2 (repCharges f S) = accSU2 S := accSU2 (repCharges f S) = accSU2 S :=
accSU2_ext accSU2_ext
(by simpa using toSpecies_sum_invariant 1 f S) (by simpa using toSpecies_sum_invariant 1 f S)
lemma accSU3_invariant (f : permGroup n) (S : (SMνCharges n).charges) : lemma accSU3_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
accSU3 (repCharges f S) = accSU3 S := accSU3 (repCharges f S) = accSU3 S :=
accSU3_ext accSU3_ext
(by simpa using toSpecies_sum_invariant 1 f S) (by simpa using toSpecies_sum_invariant 1 f S)
lemma accYY_invariant (f : permGroup n) (S : (SMνCharges n).charges) : lemma accYY_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
accYY (repCharges f S) = accYY S := accYY (repCharges f S) = accYY S :=
accYY_ext accYY_ext
(by simpa using toSpecies_sum_invariant 1 f S) (by simpa using toSpecies_sum_invariant 1 f S)
lemma accQuad_invariant (f : permGroup n) (S : (SMνCharges n).charges) : lemma accQuad_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
accQuad (repCharges f S) = accQuad S := accQuad (repCharges f S) = accQuad S :=
accQuad_ext accQuad_ext
(toSpecies_sum_invariant 2 f S) (toSpecies_sum_invariant 2 f S)
lemma accCube_invariant (f : permGroup n) (S : (SMνCharges n).charges) : lemma accCube_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
accCube (repCharges f S) = accCube S := accCube (repCharges f S) = accCube S :=
accCube_ext accCube_ext
(by simpa using toSpecies_sum_invariant 3 f S) (by simpa using toSpecies_sum_invariant 3 f S)

View file

@ -57,7 +57,7 @@ namespace BL
variable {n : } variable {n : }
lemma on_quadBiLin (S : (PlusU1 n).charges) : lemma on_quadBiLin (S : (PlusU1 n).Charges) :
quadBiLin (BL n).val S = 1/2 * accYY S + 3/2 * accSU2 S - 2 * accSU3 S := by quadBiLin (BL n).val S = 1/2 * accYY S + 3/2 * accSU2 S - 2 * accSU3 S := by
erw [familyUniversal_quadBiLin] erw [familyUniversal_quadBiLin]
rw [accYY_decomp, accSU2_decomp, accSU3_decomp] rw [accYY_decomp, accSU2_decomp, accSU3_decomp]
@ -90,7 +90,7 @@ lemma addQuad_zero (S : (PlusU1 n).QuadSols) (a : ): addQuad S a 0 = a • S
simp [addQuad, linearToQuad] simp [addQuad, linearToQuad]
rfl rfl
lemma on_cubeTriLin (S : (PlusU1 n).charges) : lemma on_cubeTriLin (S : (PlusU1 n).Charges) :
cubeTriLin (BL n).val (BL n).val S = 9 * accGrav S - 24 * accSU3 S := by cubeTriLin (BL n).val (BL n).val S = 9 * accGrav S - 24 * accSU3 S := by
erw [familyUniversal_cubeTriLin'] erw [familyUniversal_cubeTriLin']
rw [accGrav_decomp, accSU3_decomp] rw [accGrav_decomp, accSU3_decomp]

View file

@ -70,7 +70,7 @@ lemma cubeSol (S : (PlusU1 n).Sols) : accCube S.val = 0 := by
/-- An element of `charges` which satisfies the linear ACCs /-- An element of `charges` which satisfies the linear ACCs
gives us a element of `LinSols`. -/ gives us a element of `LinSols`. -/
def chargeToLinear (S : (PlusU1 n).charges) (hGrav : accGrav S = 0) def chargeToLinear (S : (PlusU1 n).Charges) (hGrav : accGrav S = 0)
(hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) (hYY : accYY S = 0) : (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) (hYY : accYY S = 0) :
(PlusU1 n).LinSols := (PlusU1 n).LinSols :=
⟨S, by ⟨S, by
@ -99,14 +99,14 @@ def quadToAF (S : (PlusU1 n).QuadSols) (hc : accCube S.val = 0) :
/-- An element of `charges` which satisfies the linear and quadratic ACCs /-- An element of `charges` which satisfies the linear and quadratic ACCs
gives us a element of `QuadSols`. -/ gives us a element of `QuadSols`. -/
def chargeToQuad (S : (PlusU1 n).charges) (hGrav : accGrav S = 0) def chargeToQuad (S : (PlusU1 n).Charges) (hGrav : accGrav S = 0)
(hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) (hYY : accYY S = 0) (hQ : accQuad S = 0) : (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) (hYY : accYY S = 0) (hQ : accQuad S = 0) :
(PlusU1 n).QuadSols := (PlusU1 n).QuadSols :=
linearToQuad (chargeToLinear S hGrav hSU2 hSU3 hYY) hQ linearToQuad (chargeToLinear S hGrav hSU2 hSU3 hYY) hQ
/-- An element of `charges` which satisfies the linear, quadratic and cubic ACCs /-- An element of `charges` which satisfies the linear, quadratic and cubic ACCs
gives us a element of `Sols`. -/ gives us a element of `Sols`. -/
def chargeToAF (S : (PlusU1 n).charges) (hGrav : accGrav S = 0) (hSU2 : accSU2 S = 0) def chargeToAF (S : (PlusU1 n).Charges) (hGrav : accGrav S = 0) (hSU2 : accSU2 S = 0)
(hSU3 : accSU3 S = 0) (hYY : accYY S = 0) (hQ : accQuad S = 0) (hc : accCube S = 0) : (hSU3 : accSU3 S = 0) (hYY : accYY S = 0) (hQ : accQuad S = 0) (hc : accCube S = 0) :
(PlusU1 n).Sols := (PlusU1 n).Sols :=
quadToAF (chargeToQuad S hGrav hSU2 hSU3 hYY hQ) hc quadToAF (chargeToQuad S hGrav hSU2 hSU3 hYY hQ) hc
@ -119,7 +119,7 @@ def linearToAF (S : (PlusU1 n).LinSols) (hQ : accQuad S.val = 0)
/-- The permutations acting on the ACC system corresponding to the SM with RHN. -/ /-- The permutations acting on the ACC system corresponding to the SM with RHN. -/
def perm (n : ) : ACCSystemGroupAction (PlusU1 n) where def perm (n : ) : ACCSystemGroupAction (PlusU1 n) where
group := permGroup n group := PermGroup n
groupInst := inferInstance groupInst := inferInstance
rep := repCharges rep := repCharges
linearInvariant := by linearInvariant := by

View file

@ -23,11 +23,11 @@ open BigOperators
/-- A proposition which is true if for a given `n` a plane of charges of dimension `n` exists /-- A proposition which is true if for a given `n` a plane of charges of dimension `n` exists
in which each point is a solution. -/ in which each point is a solution. -/
def existsPlane (n : ) : Prop := ∃ (B : Fin n → (PlusU1 3).charges), def ExistsPlane (n : ) : Prop := ∃ (B : Fin n → (PlusU1 3).Charges),
LinearIndependent B ∧ ∀ (f : Fin n → ), (PlusU1 3).isSolution (∑ i, f i • B i) LinearIndependent B ∧ ∀ (f : Fin n → ), (PlusU1 3).IsSolution (∑ i, f i • B i)
lemma exists_plane_exists_basis {n : } (hE : existsPlane n) : lemma exists_plane_exists_basis {n : } (hE : ExistsPlane n) :
∃ (B : Fin 11 ⊕ Fin n → (PlusU1 3).charges), LinearIndependent B := by ∃ (B : Fin 11 ⊕ Fin n → (PlusU1 3).Charges), LinearIndependent B := by
obtain ⟨E, hE1, hE2⟩ := hE obtain ⟨E, hE1, hE2⟩ := hE
obtain ⟨B, hB1, hB2⟩ := eleven_dim_plane_of_no_sols_exists obtain ⟨B, hB1, hB2⟩ := eleven_dim_plane_of_no_sols_exists
let Y := Sum.elim B E let Y := Sum.elim B E
@ -59,11 +59,11 @@ lemma exists_plane_exists_basis {n : } (hE : existsPlane n) :
| Sum.inr i => exact h4 i | Sum.inr i => exact h4 i
theorem plane_exists_dim_le_7 {n : } (hn : existsPlane n) : n ≤ 7 := by theorem plane_exists_dim_le_7 {n : } (hn : ExistsPlane n) : n ≤ 7 := by
obtain ⟨B, hB⟩ := exists_plane_exists_basis hn obtain ⟨B, hB⟩ := exists_plane_exists_basis hn
have h1 := LinearIndependent.fintype_card_le_finrank hB have h1 := LinearIndependent.fintype_card_le_finrank hB
simp at h1 simp at h1
rw [show FiniteDimensional.finrank (PlusU1 3).charges = 18 from rw [show FiniteDimensional.finrank (PlusU1 3).Charges = 18 from
FiniteDimensional.finrank_fin_fun ] at h1 FiniteDimensional.finrank_fin_fun ] at h1
exact Nat.le_of_add_le_add_left h1 exact Nat.le_of_add_le_add_left h1

View file

@ -55,7 +55,7 @@ namespace Y
variable {n : } variable {n : }
lemma on_quadBiLin (S : (PlusU1 n).charges) : lemma on_quadBiLin (S : (PlusU1 n).Charges) :
quadBiLin (Y n).val S = accYY S := by quadBiLin (Y n).val S = accYY S := by
erw [familyUniversal_quadBiLin] erw [familyUniversal_quadBiLin]
rw [accYY_decomp] rw [accYY_decomp]
@ -89,7 +89,7 @@ lemma addQuad_zero (S : (PlusU1 n).QuadSols) (a : ): addQuad S a 0 = a • S
simp [addQuad, linearToQuad] simp [addQuad, linearToQuad]
rfl rfl
lemma on_cubeTriLin (S : (PlusU1 n).charges) : lemma on_cubeTriLin (S : (PlusU1 n).Charges) :
cubeTriLin (Y n).val (Y n).val S = 6 * accYY S := by cubeTriLin (Y n).val (Y n).val S = 6 * accYY S := by
erw [familyUniversal_cubeTriLin'] erw [familyUniversal_cubeTriLin']
rw [accYY_decomp] rw [accYY_decomp]
@ -103,7 +103,7 @@ lemma on_cubeTriLin_AFL (S : (PlusU1 n).LinSols) :
rw [YYsol S] rw [YYsol S]
simp simp
lemma on_cubeTriLin' (S : (PlusU1 n).charges) : lemma on_cubeTriLin' (S : (PlusU1 n).Charges) :
cubeTriLin (Y n).val S S = 6 * accQuad S := by cubeTriLin (Y n).val S S = 6 * accQuad S := by
erw [familyUniversal_cubeTriLin] erw [familyUniversal_cubeTriLin]
rw [accQuad_decomp] rw [accQuad_decomp]

View file

@ -26,40 +26,40 @@ open BigOperators
namespace ElevenPlane namespace ElevenPlane
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₀ : (PlusU1 3).charges := ![1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0] def B₀ : (PlusU1 3).Charges := ![1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0]
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₁ : (PlusU1 3).charges := ![0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0] def B₁ : (PlusU1 3).Charges := ![0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0]
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₂ : (PlusU1 3).charges := ![0, 0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0] def B₂ : (PlusU1 3).Charges := ![0, 0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0]
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₃ : (PlusU1 3).charges := ![0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0] def B₃ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0]
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₄ : (PlusU1 3).charges := ![0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0] def B₄ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0]
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₅ : (PlusU1 3).charges := ![0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0] def B₅ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0]
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₆ : (PlusU1 3).charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0, 0] def B₆ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0, 0]
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₇ : (PlusU1 3).charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0] def B₇ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0]
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₈ : (PlusU1 3).charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0] def B₈ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0]
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₉ : (PlusU1 3).charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 2, 0] def B₉ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 2, 0]
/-- A charge assignment forming one of the basis elements of the plane. -/ /-- A charge assignment forming one of the basis elements of the plane. -/
def B₁₀ : (PlusU1 3).charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1] def B₁₀ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1]
/-- The charge assignment forming a basis of the plane. -/ /-- The charge assignment forming a basis of the plane. -/
def B : Fin 11 → (PlusU1 3).charges := fun i => def B : Fin 11 → (PlusU1 3).Charges := fun i =>
match i with match i with
| 0 => B₀ | 0 => B₀
| 1 => B₁ | 1 => B₁
@ -105,7 +105,7 @@ lemma on_accQuad (f : Fin 11 → ) :
ring ring
lemma isSolution_quadCoeff_f_sq_zero (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) lemma isSolution_quadCoeff_f_sq_zero (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i))
(k : Fin 11) : quadCoeff k * (f k)^2 = 0 := by (k : Fin 11) : quadCoeff k * (f k)^2 = 0 := by
obtain ⟨S, hS⟩ := hS obtain ⟨S, hS⟩ := hS
have hQ := quadSol S.1 have hQ := quadSol S.1
@ -120,46 +120,46 @@ lemma isSolution_quadCoeff_f_sq_zero (f : Fin 11 → ) (hS : (PlusU1 3).isSol
fin_cases i <;> rfl fin_cases i <;> rfl
exact sq_nonneg (f i) exact sq_nonneg (f i)
lemma isSolution_f0 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 0 = 0 := by lemma isSolution_f0 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 0 = 0 := by
simpa using (isSolution_quadCoeff_f_sq_zero f hS 0) simpa using (isSolution_quadCoeff_f_sq_zero f hS 0)
lemma isSolution_f1 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 1 = 0 := by lemma isSolution_f1 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 1 = 0 := by
simpa using (isSolution_quadCoeff_f_sq_zero f hS 1) simpa using (isSolution_quadCoeff_f_sq_zero f hS 1)
lemma isSolution_f2 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 2 = 0 := by lemma isSolution_f2 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 2 = 0 := by
simpa using (isSolution_quadCoeff_f_sq_zero f hS 2) simpa using (isSolution_quadCoeff_f_sq_zero f hS 2)
lemma isSolution_f3 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 3 = 0 := by lemma isSolution_f3 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 3 = 0 := by
simpa using (isSolution_quadCoeff_f_sq_zero f hS 3) simpa using (isSolution_quadCoeff_f_sq_zero f hS 3)
lemma isSolution_f4 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 4 = 0 := by lemma isSolution_f4 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 4 = 0 := by
simpa using (isSolution_quadCoeff_f_sq_zero f hS 4) simpa using (isSolution_quadCoeff_f_sq_zero f hS 4)
lemma isSolution_f5 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 5 = 0 := by lemma isSolution_f5 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 5 = 0 := by
have h := isSolution_quadCoeff_f_sq_zero f hS 5 have h := isSolution_quadCoeff_f_sq_zero f hS 5
rw [mul_eq_zero] at h rw [mul_eq_zero] at h
change 1 = 0 _ = _ at h change 1 = 0 _ = _ at h
simpa using h simpa using h
lemma isSolution_f6 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 6 = 0 := by lemma isSolution_f6 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 6 = 0 := by
have h := isSolution_quadCoeff_f_sq_zero f hS 6 have h := isSolution_quadCoeff_f_sq_zero f hS 6
rw [mul_eq_zero] at h rw [mul_eq_zero] at h
change 1 = 0 _ = _ at h change 1 = 0 _ = _ at h
simpa using h simpa using h
lemma isSolution_f7 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 7 = 0 := by lemma isSolution_f7 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 7 = 0 := by
have h := isSolution_quadCoeff_f_sq_zero f hS 7 have h := isSolution_quadCoeff_f_sq_zero f hS 7
rw [mul_eq_zero] at h rw [mul_eq_zero] at h
change 1 = 0 _ = _ at h change 1 = 0 _ = _ at h
simpa using h simpa using h
lemma isSolution_f8 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 8 = 0 := by lemma isSolution_f8 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 8 = 0 := by
have h := isSolution_quadCoeff_f_sq_zero f hS 8 have h := isSolution_quadCoeff_f_sq_zero f hS 8
rw [mul_eq_zero] at h rw [mul_eq_zero] at h
change 1 = 0 _ = _ at h change 1 = 0 _ = _ at h
simpa using h simpa using h
lemma isSolution_sum_part (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : lemma isSolution_sum_part (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
∑ i, f i • B i = f 9 • B₉ + f 10 • B₁₀ := by ∑ i, f i • B i = f 9 • B₉ + f 10 • B₁₀ := by
rw [Fin.sum_univ_castSucc, Fin.sum_univ_castSucc, Fin.sum_univ_castSucc, Fin.sum_univ_eight] rw [Fin.sum_univ_castSucc, Fin.sum_univ_castSucc, Fin.sum_univ_castSucc, Fin.sum_univ_eight]
change f 0 • B 0 + f 1 • B 1 + f 2 • B 2 + f 3 • B 3 + f 4 • B 4 + f 5 • B 5 + f 6 • B 6 + change f 0 • B 0 + f 1 • B 1 + f 2 • B 2 + f 3 • B 3 + f 4 • B 4 + f 5 • B 5 + f 6 • B 6 +
@ -170,7 +170,7 @@ lemma isSolution_sum_part (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑
simp simp
rfl rfl
lemma isSolution_grav (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : lemma isSolution_grav (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
f 10 = - 3 * f 9 := by f 10 = - 3 * f 9 := by
have hx := isSolution_sum_part f hS have hx := isSolution_sum_part f hS
obtain ⟨S, hS'⟩ := hS obtain ⟨S, hS'⟩ := hS
@ -183,12 +183,12 @@ lemma isSolution_grav (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i,
simp at hg simp at hg
linear_combination hg linear_combination hg
lemma isSolution_sum_part' (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : lemma isSolution_sum_part' (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
∑ i, f i • B i = f 9 • B₉ + (- 3 * f 9) • B₁₀ := by ∑ i, f i • B i = f 9 • B₉ + (- 3 * f 9) • B₁₀ := by
rw [isSolution_sum_part f hS] rw [isSolution_sum_part f hS]
rw [isSolution_grav f hS] rw [isSolution_grav f hS]
lemma isSolution_f9 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : lemma isSolution_f9 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
f 9 = 0 := by f 9 = 0 := by
have hx := isSolution_sum_part' f hS have hx := isSolution_sum_part' f hS
obtain ⟨S, hS'⟩ := hS obtain ⟨S, hS'⟩ := hS
@ -211,12 +211,12 @@ lemma isSolution_f9 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i
rw [h1] at hc rw [h1] at hc
simpa using hc simpa using hc
lemma isSolution_f10 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : lemma isSolution_f10 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
f 10 = 0 := by f 10 = 0 := by
rw [isSolution_grav f hS, isSolution_f9 f hS] rw [isSolution_grav f hS, isSolution_f9 f hS]
simp simp
lemma isSolution_f_zero (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) lemma isSolution_f_zero (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i))
(k : Fin 11) : f k = 0 := by (k : Fin 11) : f k = 0 := by
fin_cases k fin_cases k
exact isSolution_f0 f hS exact isSolution_f0 f hS
@ -232,7 +232,7 @@ lemma isSolution_f_zero (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i,
exact isSolution_f10 f hS exact isSolution_f10 f hS
lemma isSolution_only_if_zero (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : lemma isSolution_only_if_zero (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
∑ i, f i • B i = 0 := by ∑ i, f i • B i = 0 := by
rw [isSolution_sum_part f hS] rw [isSolution_sum_part f hS]
rw [isSolution_grav f hS] rw [isSolution_grav f hS]
@ -244,7 +244,7 @@ theorem basis_linear_independent : LinearIndependent B := by
apply Fintype.linearIndependent_iff.mpr apply Fintype.linearIndependent_iff.mpr
intro f h intro f h
let X : (PlusU1 3).Sols := chargeToAF 0 (by rfl) (by rfl) (by rfl) (by rfl) (by rfl) (by rfl) let X : (PlusU1 3).Sols := chargeToAF 0 (by rfl) (by rfl) (by rfl) (by rfl) (by rfl) (by rfl)
have hS : (PlusU1 3).isSolution (∑ i, f i • B i) := by have hS : (PlusU1 3).IsSolution (∑ i, f i • B i) := by
use X use X
rw [h] rw [h]
rfl rfl
@ -252,9 +252,9 @@ theorem basis_linear_independent : LinearIndependent B := by
end ElevenPlane end ElevenPlane
theorem eleven_dim_plane_of_no_sols_exists : ∃ (B : Fin 11 → (PlusU1 3).charges), theorem eleven_dim_plane_of_no_sols_exists : ∃ (B : Fin 11 → (PlusU1 3).Charges),
LinearIndependent B ∧ LinearIndependent B ∧
∀ (f : Fin 11 → ), (PlusU1 3).isSolution (∑ i, f i • B i) → ∑ i, f i • B i = 0 := by ∀ (f : Fin 11 → ), (PlusU1 3).IsSolution (∑ i, f i • B i) → ∑ i, f i • B i = 0 := by
use ElevenPlane.B use ElevenPlane.B
apply And.intro apply And.intro
exact ElevenPlane.basis_linear_independent exact ElevenPlane.basis_linear_independent

View file

@ -22,8 +22,8 @@ open ComplexConjugate
noncomputable section noncomputable section
/-- The potential of the two Higgs doublet model. -/ /-- The potential of the two Higgs doublet model. -/
def potential (Φ1 Φ2 : higgsField) def potential (Φ1 Φ2 : HiggsField)
(m11Sq m22Sq lam₁ lam₂ lam₃ lam₄ : ) (m12Sq lam₅ lam₆ lam₇ : ) (x : spaceTime) : := (m11Sq m22Sq lam₁ lam₂ lam₃ lam₄ : ) (m12Sq lam₅ lam₆ lam₇ : ) (x : SpaceTime) : :=
m11Sq * Φ1.normSq x + m22Sq * Φ2.normSq x m11Sq * Φ1.normSq x + m22Sq * Φ2.normSq x
- (m12Sq * (Φ1.innerProd Φ2) x + conj m12Sq * (Φ2.innerProd Φ1) x).re - (m12Sq * (Φ1.innerProd Φ2) x + conj m12Sq * (Φ2.innerProd Φ1) x).re
+ 1/2 * lam₁ * Φ1.normSq x ^ 2 + 1/2 * lam₂ * Φ2.normSq x ^ 2 + 1/2 * lam₁ * Φ1.normSq x ^ 2 + 1/2 * lam₂ * Φ2.normSq x ^ 2

View file

@ -242,7 +242,7 @@ instance externalDecidable [IsFinitePreFeynmanRule P] (v : P.VertexLabel) :
`F : Over (P.HalfEdgeLabel × P.EdgeLabel × P.VertexLabel)` for it to be a Feynman diagram. `F : Over (P.HalfEdgeLabel × P.EdgeLabel × P.VertexLabel)` for it to be a Feynman diagram.
This condition corresponds to the vertices of `F` having the correct half-edges associated This condition corresponds to the vertices of `F` having the correct half-edges associated
with them. -/ with them. -/
def diagramVertexProp {𝓔 𝓥 : Type} (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) def DiagramVertexProp {𝓔 𝓥 : Type} (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥))
(f : 𝓥 ⟶ P.VertexLabel) := (f : 𝓥 ⟶ P.VertexLabel) :=
∀ v, IsIsomorphic (P.vertexLabelMap (f v)) ((P.preimageVertex v).obj F) ∀ v, IsIsomorphic (P.vertexLabelMap (f v)) ((P.preimageVertex v).obj F)
@ -250,12 +250,12 @@ def diagramVertexProp {𝓔 𝓥 : Type} (F : Over (P.HalfEdgeLabel × 𝓔 ×
`F : Over (P.HalfEdgeLabel × P.EdgeLabel × P.VertexLabel)` for it to be a Feynman diagram. `F : Over (P.HalfEdgeLabel × P.EdgeLabel × P.VertexLabel)` for it to be a Feynman diagram.
This condition corresponds to the edges of `F` having the correct half-edges associated This condition corresponds to the edges of `F` having the correct half-edges associated
with them. -/ with them. -/
def diagramEdgeProp {𝓔 𝓥 : Type} (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) def DiagramEdgeProp {𝓔 𝓥 : Type} (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥))
(f : 𝓔 ⟶ P.EdgeLabel) := (f : 𝓔 ⟶ P.EdgeLabel) :=
∀ v, IsIsomorphic (P.edgeLabelMap (f v)) ((P.preimageEdge v).obj F) ∀ v, IsIsomorphic (P.edgeLabelMap (f v)) ((P.preimageEdge v).obj F)
lemma diagramVertexProp_iff {𝓔 𝓥 : Type} (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) lemma diagramVertexProp_iff {𝓔 𝓥 : Type} (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥))
(f : 𝓥 ⟶ P.VertexLabel) : P.diagramVertexProp F f ↔ (f : 𝓥 ⟶ P.VertexLabel) : P.DiagramVertexProp F f ↔
∀ v, ∃ (κ : (P.vertexLabelMap (f v)).left → ((P.preimageVertex v).obj F).left), ∀ v, ∃ (κ : (P.vertexLabelMap (f v)).left → ((P.preimageVertex v).obj F).left),
Function.Bijective κ Function.Bijective κ
∧ ((P.preimageVertex v).obj F).hom ∘ κ = (P.vertexLabelMap (f v)).hom := by ∧ ((P.preimageVertex v).obj F).hom ∘ κ = (P.vertexLabelMap (f v)).hom := by
@ -270,7 +270,7 @@ lemma diagramVertexProp_iff {𝓔 𝓥 : Type} (F : Over (P.HalfEdgeLabel ×
exact ⟨f, fm, hf1, hf2⟩ exact ⟨f, fm, hf1, hf2⟩
lemma diagramEdgeProp_iff {𝓔 𝓥 : Type} (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) lemma diagramEdgeProp_iff {𝓔 𝓥 : Type} (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥))
(f : 𝓔 ⟶ P.EdgeLabel) : P.diagramEdgeProp F f ↔ (f : 𝓔 ⟶ P.EdgeLabel) : P.DiagramEdgeProp F f ↔
∀ v, ∃ (κ : (P.edgeLabelMap (f v)).left → ((P.preimageEdge v).obj F).left), ∀ v, ∃ (κ : (P.edgeLabelMap (f v)).left → ((P.preimageEdge v).obj F).left),
Function.Bijective κ Function.Bijective κ
∧ ((P.preimageEdge v).obj F).hom ∘ κ = (P.edgeLabelMap (f v)).hom := by ∧ ((P.preimageEdge v).obj F).hom ∘ κ = (P.edgeLabelMap (f v)).hom := by
@ -287,7 +287,7 @@ lemma diagramEdgeProp_iff {𝓔 𝓥 : Type} (F : Over (P.HalfEdgeLabel × 𝓔
instance diagramVertexPropDecidable instance diagramVertexPropDecidable
[IsFinitePreFeynmanRule P] {𝓔 𝓥 : Type} [Fintype 𝓥] [DecidableEq 𝓥] [IsFinitePreFeynmanRule P] {𝓔 𝓥 : Type} [Fintype 𝓥] [DecidableEq 𝓥]
(F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) [DecidableEq F.left] [Fintype F.left] (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) [DecidableEq F.left] [Fintype F.left]
(f : 𝓥 ⟶ P.VertexLabel) : Decidable (P.diagramVertexProp F f) := (f : 𝓥 ⟶ P.VertexLabel) : Decidable (P.DiagramVertexProp F f) :=
@decidable_of_decidable_of_iff _ _ @decidable_of_decidable_of_iff _ _
(@Fintype.decidableForallFintype _ _ (fun _ => @Fintype.decidableExistsFintype _ _ (@Fintype.decidableForallFintype _ _ (fun _ => @Fintype.decidableExistsFintype _ _
(fun _ => @And.decidable _ _ _ (@Fintype.decidablePiFintype _ _ (fun _ => @And.decidable _ _ _ (@Fintype.decidablePiFintype _ _
@ -297,7 +297,7 @@ instance diagramVertexPropDecidable
instance diagramEdgePropDecidable instance diagramEdgePropDecidable
[IsFinitePreFeynmanRule P] {𝓔 𝓥 : Type} [Fintype 𝓔] [DecidableEq 𝓔] [IsFinitePreFeynmanRule P] {𝓔 𝓥 : Type} [Fintype 𝓔] [DecidableEq 𝓔]
(F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) [DecidableEq F.left] [Fintype F.left] (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) [DecidableEq F.left] [Fintype F.left]
(f : 𝓔 ⟶ P.EdgeLabel) : Decidable (P.diagramEdgeProp F f) := (f : 𝓔 ⟶ P.EdgeLabel) : Decidable (P.DiagramEdgeProp F f) :=
@decidable_of_decidable_of_iff _ _ @decidable_of_decidable_of_iff _ _
(@Fintype.decidableForallFintype _ _ (fun _ => @Fintype.decidableExistsFintype _ _ (@Fintype.decidableForallFintype _ _ (fun _ => @Fintype.decidableExistsFintype _ _
(fun _ => @And.decidable _ _ _ (@Fintype.decidablePiFintype _ _ (fun _ => @And.decidable _ _ _ (@Fintype.decidablePiFintype _ _
@ -327,9 +327,9 @@ structure FeynmanDiagram (P : PreFeynmanRule) where
belongs to, and the vertex they belong to. -/ belongs to, and the vertex they belong to. -/
𝓱𝓔To𝓔𝓥 : Over (P.HalfEdgeLabel × 𝓔𝓞.left × 𝓥𝓞.left) 𝓱𝓔To𝓔𝓥 : Over (P.HalfEdgeLabel × 𝓔𝓞.left × 𝓥𝓞.left)
/-- Each edge has the correct type of half edges. -/ /-- Each edge has the correct type of half edges. -/
𝓔Cond : P.diagramEdgeProp 𝓱𝓔To𝓔𝓥 𝓔𝓞.hom 𝓔Cond : P.DiagramEdgeProp 𝓱𝓔To𝓔𝓥 𝓔𝓞.hom
/-- Each vertex has the correct sort of half edges. -/ /-- Each vertex has the correct sort of half edges. -/
𝓥Cond : P.diagramVertexProp 𝓱𝓔To𝓔𝓥 𝓥𝓞.hom 𝓥Cond : P.DiagramVertexProp 𝓱𝓔To𝓔𝓥 𝓥𝓞.hom
namespace FeynmanDiagram namespace FeynmanDiagram
@ -362,8 +362,8 @@ def 𝓱𝓔To𝓥 : Over F.𝓥 := P.toVertex.obj F.𝓱𝓔To𝓔𝓥
/-- The condition which must be satisfied by maps to form a Feynman diagram. -/ /-- The condition which must be satisfied by maps to form a Feynman diagram. -/
def Cond {𝓔 𝓥 𝓱𝓔 : Type} (𝓔𝓞 : 𝓔 → P.EdgeLabel) (𝓥𝓞 : 𝓥 → P.VertexLabel) def Cond {𝓔 𝓥 𝓱𝓔 : Type} (𝓔𝓞 : 𝓔 → P.EdgeLabel) (𝓥𝓞 : 𝓥 → P.VertexLabel)
(𝓱𝓔To𝓔𝓥 : 𝓱𝓔 → P.HalfEdgeLabel × 𝓔 × 𝓥) : Prop := (𝓱𝓔To𝓔𝓥 : 𝓱𝓔 → P.HalfEdgeLabel × 𝓔 × 𝓥) : Prop :=
P.diagramEdgeProp (Over.mk 𝓱𝓔To𝓔𝓥) 𝓔𝓞 P.DiagramEdgeProp (Over.mk 𝓱𝓔To𝓔𝓥) 𝓔𝓞
P.diagramVertexProp (Over.mk 𝓱𝓔To𝓔𝓥) 𝓥𝓞 P.DiagramVertexProp (Over.mk 𝓱𝓔To𝓔𝓥) 𝓥𝓞
lemma cond_self : Cond F.𝓔𝓞.hom F.𝓥𝓞.hom F.𝓱𝓔To𝓔𝓥.hom := lemma cond_self : Cond F.𝓔𝓞.hom F.𝓥𝓞.hom F.𝓱𝓔To𝓔𝓥.hom :=
⟨F.𝓔Cond, F.𝓥Cond⟩ ⟨F.𝓔Cond, F.𝓥Cond⟩
@ -716,12 +716,12 @@ A feynman diagram is connected if its simple graph is connected.
/-- A relation on the vertices of Feynman diagrams. The proposition is true if the two /-- A relation on the vertices of Feynman diagrams. The proposition is true if the two
vertices are not equal and are connected by a single edge. -/ vertices are not equal and are connected by a single edge. -/
@[simp] @[simp]
def adjRelation : F.𝓥 → F.𝓥 → Prop := fun x y => def AdjRelation : F.𝓥 → F.𝓥 → Prop := fun x y =>
x ≠ y ∧ x ≠ y ∧
∃ (a b : F.𝓱𝓔), ((F.𝓱𝓔To𝓔𝓥.hom a).2.1 = (F.𝓱𝓔To𝓔𝓥.hom b).2.1 ∃ (a b : F.𝓱𝓔), ((F.𝓱𝓔To𝓔𝓥.hom a).2.1 = (F.𝓱𝓔To𝓔𝓥.hom b).2.1
∧ (F.𝓱𝓔To𝓔𝓥.hom a).2.2 = x ∧ (F.𝓱𝓔To𝓔𝓥.hom b).2.2 = y) ∧ (F.𝓱𝓔To𝓔𝓥.hom a).2.2 = x ∧ (F.𝓱𝓔To𝓔𝓥.hom b).2.2 = y)
instance [IsFiniteDiagram F] : DecidableRel F.adjRelation := fun _ _ => instance [IsFiniteDiagram F] : DecidableRel F.AdjRelation := fun _ _ =>
@And.decidable _ _ _ $ @And.decidable _ _ _ $
@Fintype.decidableExistsFintype _ _ (fun _ => @Fintype.decidableExistsFintype _ _ ( @Fintype.decidableExistsFintype _ _ (fun _ => @Fintype.decidableExistsFintype _ _ (
fun _ => @And.decidable _ _ (instDecidableEq𝓔OfIsFiniteDiagram _ _) $ fun _ => @And.decidable _ _ (instDecidableEq𝓔OfIsFiniteDiagram _ _) $
@ -731,13 +731,13 @@ instance [IsFiniteDiagram F] : DecidableRel F.adjRelation := fun _ _ =>
/-- From a Feynman diagram the simple graph showing those vertices which are connected. -/ /-- From a Feynman diagram the simple graph showing those vertices which are connected. -/
def toSimpleGraph : SimpleGraph F.𝓥 where def toSimpleGraph : SimpleGraph F.𝓥 where
Adj := adjRelation F Adj := AdjRelation F
symm := by symm := by
intro x y h intro x y h
apply And.intro (Ne.symm h.1) apply And.intro (Ne.symm h.1)
obtain ⟨a, b, hab⟩ := h.2 obtain ⟨a, b, hab⟩ := h.2
use b, a use b, a
simp_all only [adjRelation, ne_eq, and_self] simp_all only [AdjRelation, ne_eq, and_self]
loopless := by loopless := by
intro x h intro x h
simp at h simp at h

View file

@ -68,10 +68,10 @@ def phaseShift (a b c : ) : unitaryGroup (Fin 3) :=
lemma phaseShift_coe_matrix (a b c : ) : ↑(phaseShift a b c) = phaseShiftMatrix a b c := rfl lemma phaseShift_coe_matrix (a b c : ) : ↑(phaseShift a b c) = phaseShiftMatrix a b c := rfl
/-- The equivalence relation between CKM matrices. -/ /-- The equivalence relation between CKM matrices. -/
def phaseShiftRelation (U V : unitaryGroup (Fin 3) ) : Prop := def PhaseShiftRelation (U V : unitaryGroup (Fin 3) ) : Prop :=
∃ a b c e f g , U = phaseShift a b c * V * phaseShift e f g ∃ a b c e f g , U = phaseShift a b c * V * phaseShift e f g
lemma phaseShiftRelation_refl (U : unitaryGroup (Fin 3) ) : phaseShiftRelation U U := by lemma phaseShiftRelation_refl (U : unitaryGroup (Fin 3) ) : PhaseShiftRelation U U := by
use 0, 0, 0, 0, 0, 0 use 0, 0, 0, 0, 0, 0
rw [Subtype.ext_iff_val] rw [Subtype.ext_iff_val]
simp only [Submonoid.coe_mul, phaseShift_coe_matrix, ofReal_zero, mul_zero, exp_zero] simp only [Submonoid.coe_mul, phaseShift_coe_matrix, ofReal_zero, mul_zero, exp_zero]
@ -79,7 +79,7 @@ lemma phaseShiftRelation_refl (U : unitaryGroup (Fin 3) ) : phaseShiftRelatio
simp only [one_mul, mul_one] simp only [one_mul, mul_one]
lemma phaseShiftRelation_symm {U V : unitaryGroup (Fin 3) } : lemma phaseShiftRelation_symm {U V : unitaryGroup (Fin 3) } :
phaseShiftRelation U V → phaseShiftRelation V U := by PhaseShiftRelation U V → PhaseShiftRelation V U := by
intro h intro h
obtain ⟨a, b, c, e, f, g, h⟩ := h obtain ⟨a, b, c, e, f, g, h⟩ := h
use (- a), (- b), (- c), (- e), (- f), (- g) use (- a), (- b), (- c), (- e), (- f), (- g)
@ -93,7 +93,7 @@ lemma phaseShiftRelation_symm {U V : unitaryGroup (Fin 3) } :
mul_one] mul_one]
lemma phaseShiftRelation_trans {U V W : unitaryGroup (Fin 3) } : lemma phaseShiftRelation_trans {U V W : unitaryGroup (Fin 3) } :
phaseShiftRelation U V → phaseShiftRelation V W → phaseShiftRelation U W := by PhaseShiftRelation U V → PhaseShiftRelation V W → PhaseShiftRelation U W := by
intro hUV hVW intro hUV hVW
obtain ⟨a, b, c, e, f, g, hUV⟩ := hUV obtain ⟨a, b, c, e, f, g, hUV⟩ := hUV
obtain ⟨d, i, j, k, l, m, hVW⟩ := hVW obtain ⟨d, i, j, k, l, m, hVW⟩ := hVW
@ -108,7 +108,7 @@ lemma phaseShiftRelation_trans {U V W : unitaryGroup (Fin 3) } :
rw [add_comm k e, add_comm l f, add_comm m g] rw [add_comm k e, add_comm l f, add_comm m g]
lemma phaseShiftRelation_equiv : Equivalence phaseShiftRelation where lemma phaseShiftRelation_equiv : Equivalence PhaseShiftRelation where
refl := phaseShiftRelation_refl refl := phaseShiftRelation_refl
symm := phaseShiftRelation_symm symm := phaseShiftRelation_symm
trans := phaseShiftRelation_trans trans := phaseShiftRelation_trans
@ -149,7 +149,7 @@ scoped[CKMMatrix] notation (name := ts_element) "[" V "]ts" => V.1 2 1
/-- The `tb`th element of the CKM matrix. -/ /-- The `tb`th element of the CKM matrix. -/
scoped[CKMMatrix] notation (name := tb_element) "[" V "]tb" => V.1 2 2 scoped[CKMMatrix] notation (name := tb_element) "[" V "]tb" => V.1 2 2
instance CKMMatrixSetoid : Setoid CKMMatrix := ⟨phaseShiftRelation, phaseShiftRelation_equiv⟩ instance CKMMatrixSetoid : Setoid CKMMatrix := ⟨PhaseShiftRelation, phaseShiftRelation_equiv⟩
/-- The matrix obtained from `V` by shifting the phases of the fermions. -/ /-- The matrix obtained from `V` by shifting the phases of the fermions. -/
@[simps!] @[simps!]

View file

@ -160,7 +160,7 @@ variable (a b c d e f : )
/-- A proposition which is satisfied by a CKM matrix if its `ud`, `us`, `cb` and `tb` elements /-- A proposition which is satisfied by a CKM matrix if its `ud`, `us`, `cb` and `tb` elements
are positive and real, and there is no phase difference between the `t`th-row and are positive and real, and there is no phase difference between the `t`th-row and
the cross product of the conjugates of the `u`th and `c`th rows. -/ the cross product of the conjugates of the `u`th and `c`th rows. -/
def fstRowThdColRealCond (U : CKMMatrix) : Prop := [U]ud = VudAbs ⟦U⟧ ∧ [U]us = VusAbs ⟦U⟧ def FstRowThdColRealCond (U : CKMMatrix) : Prop := [U]ud = VudAbs ⟦U⟧ ∧ [U]us = VusAbs ⟦U⟧
∧ [U]cb = VcbAbs ⟦U⟧ ∧ [U]tb = VtbAbs ⟦U⟧ ∧ [U]t = conj [U]u ×₃ conj [U]c ∧ [U]cb = VcbAbs ⟦U⟧ ∧ [U]tb = VtbAbs ⟦U⟧ ∧ [U]t = conj [U]u ×₃ conj [U]c
/-- A proposition which is satisfied by a CKM matrix `ub` is one, `ud`, `us` and `cb` are zero, /-- A proposition which is satisfied by a CKM matrix `ub` is one, `ud`, `us` and `cb` are zero,
@ -226,7 +226,7 @@ lemma ubOnePhaseCond_shift_solution {V : CKMMatrix} (h1 : a + f = - arg [V]ub)
-- rename -- rename
lemma fstRowThdColRealCond_holds_up_to_equiv (V : CKMMatrix) : lemma fstRowThdColRealCond_holds_up_to_equiv (V : CKMMatrix) :
∃ (U : CKMMatrix), V ≈ U ∧ fstRowThdColRealCond U:= by ∃ (U : CKMMatrix), V ≈ U ∧ FstRowThdColRealCond U:= by
obtain ⟨τ, hτ⟩ := V.uRow_cross_cRow_eq_tRow obtain ⟨τ, hτ⟩ := V.uRow_cross_cRow_eq_tRow
let U : CKMMatrix := phaseShiftApply V let U : CKMMatrix := phaseShiftApply V
0 0
@ -263,7 +263,7 @@ lemma fstRowThdColRealCond_holds_up_to_equiv (V : CKMMatrix) :
lemma ubOnePhaseCond_hold_up_to_equiv_of_ub_one {V : CKMMatrix} (hb : ¬ ([V]ud ≠ 0 [V]us ≠ 0)) lemma ubOnePhaseCond_hold_up_to_equiv_of_ub_one {V : CKMMatrix} (hb : ¬ ([V]ud ≠ 0 [V]us ≠ 0))
(hV : fstRowThdColRealCond V) : (hV : FstRowThdColRealCond V) :
∃ (U : CKMMatrix), V ≈ U ∧ ubOnePhaseCond U:= by ∃ (U : CKMMatrix), V ≈ U ∧ ubOnePhaseCond U:= by
let U : CKMMatrix := phaseShiftApply V 0 0 (- Real.pi + arg [V]cd + arg [V]cs + arg [V]ub) let U : CKMMatrix := phaseShiftApply V 0 0 (- Real.pi + arg [V]cd + arg [V]cs + arg [V]ub)
(Real.pi - arg [V]cd ) (- arg [V]cs) (- arg [V]ub ) (Real.pi - arg [V]cd ) (- arg [V]cs) (- arg [V]ub )
@ -321,7 +321,7 @@ lemma ubOnePhaseCond_hold_up_to_equiv_of_ub_one {V : CKMMatrix} (hb : ¬ ([V]ud
lemma cd_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0) lemma cd_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0)
(hV : fstRowThdColRealCond V) : (hV : FstRowThdColRealCond V) :
[V]cd = (- VtbAbs ⟦V⟧ * VusAbs ⟦V⟧ / (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2)) + [V]cd = (- VtbAbs ⟦V⟧ * VusAbs ⟦V⟧ / (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2)) +
(- VubAbs ⟦V⟧ * VudAbs ⟦V⟧ * VcbAbs ⟦V⟧ / (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2 )) (- VubAbs ⟦V⟧ * VudAbs ⟦V⟧ * VcbAbs ⟦V⟧ / (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2 ))
* cexp (- arg [V]ub * I) := by * cexp (- arg [V]ub * I) := by
@ -341,7 +341,7 @@ lemma cd_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠
ring_nf ring_nf
lemma cs_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0) lemma cs_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0)
(hV : fstRowThdColRealCond V) : (hV : FstRowThdColRealCond V) :
[V]cs = (VtbAbs ⟦V⟧ * VudAbs ⟦V⟧ / (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2)) [V]cs = (VtbAbs ⟦V⟧ * VudAbs ⟦V⟧ / (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2))
+ (- VubAbs ⟦V⟧ * VusAbs ⟦V⟧ * VcbAbs ⟦V⟧/ (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2)) + (- VubAbs ⟦V⟧ * VusAbs ⟦V⟧ * VcbAbs ⟦V⟧/ (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2))
* cexp (- arg [V]ub * I) := by * cexp (- arg [V]ub * I) := by

View file

@ -510,7 +510,7 @@ lemma on_param_sin_θ₂₃_eq_zero {V : CKMMatrix} (δ₁₃ : ) (h : Real.s
lemma eq_standParam_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0) lemma eq_standParam_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0)
(hV : fstRowThdColRealCond V) : V = standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) (- arg [V]ub) := by (hV : FstRowThdColRealCond V) : V = standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) (- arg [V]ub) := by
have hb' : VubAbs ⟦V⟧ ≠ 1 := by have hb' : VubAbs ⟦V⟧ ≠ 1 := by
rw [ud_us_neq_zero_iff_ub_neq_one] at hb rw [ud_us_neq_zero_iff_ub_neq_one] at hb
simp [VAbs, hb] simp [VAbs, hb]

View file

@ -14,7 +14,7 @@ and the vector space of 2×2-complex self-adjoint matrices.
In this file we define this linear equivalence in `toSelfAdjointMatrix`. In this file we define this linear equivalence in `toSelfAdjointMatrix`.
-/ -/
namespace spaceTime namespace SpaceTime
open Matrix open Matrix
open MatrixGroups open MatrixGroups
@ -22,11 +22,11 @@ open Complex
/-- A 2×2-complex matrix formed from a space-time point. -/ /-- A 2×2-complex matrix formed from a space-time point. -/
@[simp] @[simp]
def toMatrix (x : spaceTime) : Matrix (Fin 2) (Fin 2) := def toMatrix (x : SpaceTime) : Matrix (Fin 2) (Fin 2) :=
!![x 0 + x 3, x 1 - x 2 * I; x 1 + x 2 * I, x 0 - x 3] !![x 0 + x 3, x 1 - x 2 * I; x 1 + x 2 * I, x 0 - x 3]
/-- The matrix `x.toMatrix` for `x ∈ spaceTime` is self adjoint. -/ /-- The matrix `x.toMatrix` for `x ∈ spaceTime` is self adjoint. -/
lemma toMatrix_isSelfAdjoint (x : spaceTime) : IsSelfAdjoint x.toMatrix := by lemma toMatrix_isSelfAdjoint (x : SpaceTime) : IsSelfAdjoint x.toMatrix := by
rw [isSelfAdjoint_iff, star_eq_conjTranspose, ← Matrix.ext_iff] rw [isSelfAdjoint_iff, star_eq_conjTranspose, ← Matrix.ext_iff]
intro i j intro i j
fin_cases i <;> fin_cases j <;> fin_cases i <;> fin_cases j <;>
@ -35,17 +35,17 @@ lemma toMatrix_isSelfAdjoint (x : spaceTime) : IsSelfAdjoint x.toMatrix := by
/-- A self-adjoint matrix formed from a space-time point. -/ /-- A self-adjoint matrix formed from a space-time point. -/
@[simps!] @[simps!]
def toSelfAdjointMatrix' (x : spaceTime) : selfAdjoint (Matrix (Fin 2) (Fin 2) ) := def toSelfAdjointMatrix' (x : SpaceTime) : selfAdjoint (Matrix (Fin 2) (Fin 2) ) :=
⟨x.toMatrix, toMatrix_isSelfAdjoint x⟩ ⟨x.toMatrix, toMatrix_isSelfAdjoint x⟩
/-- A self-adjoint matrix formed from a space-time point. -/ /-- A self-adjoint matrix formed from a space-time point. -/
@[simp] @[simp]
noncomputable def fromSelfAdjointMatrix' (x : selfAdjoint (Matrix (Fin 2) (Fin 2) )) : spaceTime := noncomputable def fromSelfAdjointMatrix' (x : selfAdjoint (Matrix (Fin 2) (Fin 2) )) : SpaceTime :=
![1/2 * (x.1 0 0 + x.1 1 1).re, (x.1 1 0).re, (x.1 1 0).im , (x.1 0 0 - x.1 1 1).re/2] ![1/2 * (x.1 0 0 + x.1 1 1).re, (x.1 1 0).re, (x.1 1 0).im , (x.1 0 0 - x.1 1 1).re/2]
/-- The linear equivalence between the vector-space `spaceTime` and self-adjoint /-- The linear equivalence between the vector-space `spaceTime` and self-adjoint
2×2-complex matrices. -/ 2×2-complex matrices. -/
noncomputable def toSelfAdjointMatrix : spaceTime ≃ₗ[] selfAdjoint (Matrix (Fin 2) (Fin 2) ) where noncomputable def toSelfAdjointMatrix : SpaceTime ≃ₗ[] selfAdjoint (Matrix (Fin 2) (Fin 2) ) where
toFun := toSelfAdjointMatrix' toFun := toSelfAdjointMatrix'
invFun := fromSelfAdjointMatrix' invFun := fromSelfAdjointMatrix'
left_inv x := by left_inv x := by
@ -53,7 +53,7 @@ noncomputable def toSelfAdjointMatrix : spaceTime ≃ₗ[] selfAdjoint (Matri
cons_val_zero, empty_val', cons_val_fin_one, cons_val_one, head_cons, head_fin_const, cons_val_zero, empty_val', cons_val_fin_one, cons_val_one, head_cons, head_fin_const,
add_add_sub_cancel, add_re, ofReal_re, mul_re, I_re, mul_zero, ofReal_im, I_im, mul_one, add_add_sub_cancel, add_re, ofReal_re, mul_re, I_re, mul_zero, ofReal_im, I_im, mul_one,
sub_self, add_zero, add_im, mul_im, zero_add, add_sub_sub_cancel, half_add_self] sub_self, add_zero, add_im, mul_im, zero_add, add_sub_sub_cancel, half_add_self]
field_simp [spaceTime] field_simp [SpaceTime]
ext1 x ext1 x
fin_cases x <;> rfl fin_cases x <;> rfl
right_inv x := by right_inv x := by
@ -83,10 +83,10 @@ noncomputable def toSelfAdjointMatrix : spaceTime ≃ₗ[] selfAdjoint (Matri
field_simp [fromSelfAdjointMatrix', toMatrix, conj_ofReal, smul_apply] field_simp [fromSelfAdjointMatrix', toMatrix, conj_ofReal, smul_apply]
<;> ring <;> ring
lemma det_eq_ηLin (x : spaceTime) : det (toSelfAdjointMatrix x).1 = ηLin x x := by lemma det_eq_ηLin (x : SpaceTime) : det (toSelfAdjointMatrix x).1 = ηLin x x := by
simp [toSelfAdjointMatrix, ηLin_expand] simp [toSelfAdjointMatrix, ηLin_expand]
ring_nf ring_nf
simp only [Fin.isValue, I_sq, mul_neg, mul_one] simp only [Fin.isValue, I_sq, mul_neg, mul_one]
ring ring
end spaceTime end SpaceTime

View file

@ -17,21 +17,21 @@ This file introduce 4d Minkowski spacetime.
noncomputable section noncomputable section
/-- The space-time -/ /-- The space-time -/
def spaceTime : Type := Fin 4 → def SpaceTime : Type := Fin 4 →
/-- Give spacetime the structure of an additive commutative monoid. -/ /-- Give spacetime the structure of an additive commutative monoid. -/
instance : AddCommMonoid spaceTime := Pi.addCommMonoid instance : AddCommMonoid SpaceTime := Pi.addCommMonoid
/-- Give spacetime the structure of a module over the reals. -/ /-- Give spacetime the structure of a module over the reals. -/
instance : Module spaceTime := Pi.module _ _ _ instance : Module SpaceTime := Pi.module _ _ _
instance euclideanNormedAddCommGroup : NormedAddCommGroup spaceTime := Pi.normedAddCommGroup instance euclideanNormedAddCommGroup : NormedAddCommGroup SpaceTime := Pi.normedAddCommGroup
instance euclideanNormedSpace : NormedSpace spaceTime := Pi.normedSpace instance euclideanNormedSpace : NormedSpace SpaceTime := Pi.normedSpace
namespace spaceTime namespace SpaceTime
open Manifold open Manifold
open Matrix open Matrix
@ -40,15 +40,15 @@ open ComplexConjugate
/-- The space part of spacetime. -/ /-- The space part of spacetime. -/
@[simp] @[simp]
def space (x : spaceTime) : EuclideanSpace (Fin 3) := ![x 1, x 2, x 3] def space (x : SpaceTime) : EuclideanSpace (Fin 3) := ![x 1, x 2, x 3]
/-- The structure of a smooth manifold on spacetime. -/ /-- The structure of a smooth manifold on spacetime. -/
def asSmoothManifold : ModelWithCorners spaceTime spaceTime := 𝓘(, spaceTime) def asSmoothManifold : ModelWithCorners SpaceTime SpaceTime := 𝓘(, SpaceTime)
instance : ChartedSpace spaceTime spaceTime := chartedSpaceSelf spaceTime instance : ChartedSpace SpaceTime SpaceTime := chartedSpaceSelf SpaceTime
/-- The standard basis for spacetime. -/ /-- The standard basis for spacetime. -/
def stdBasis : Basis (Fin 4) spaceTime := Pi.basisFun (Fin 4) def stdBasis : Basis (Fin 4) SpaceTime := Pi.basisFun (Fin 4)
lemma stdBasis_apply (μ ν : Fin 4) : stdBasis μ ν = if μ = ν then 1 else 0 := by lemma stdBasis_apply (μ ν : Fin 4) : stdBasis μ ν = if μ = ν then 1 else 0 := by
erw [stdBasis, Pi.basisFun_apply, LinearMap.stdBasis_apply'] erw [stdBasis, Pi.basisFun_apply, LinearMap.stdBasis_apply']
@ -85,16 +85,16 @@ lemma stdBasis_mulVec (μ ν : Fin 4) (Λ : Matrix (Fin 4) (Fin 4) ) :
rw [stdBasis_apply, if_neg (Ne.symm h)] rw [stdBasis_apply, if_neg (Ne.symm h)]
exact CommMonoidWithZero.mul_zero (Λ ν x) exact CommMonoidWithZero.mul_zero (Λ ν x)
lemma explicit (x : spaceTime) : x = ![x 0, x 1, x 2, x 3] := by lemma explicit (x : SpaceTime) : x = ![x 0, x 1, x 2, x 3] := by
funext i funext i
fin_cases i <;> rfl fin_cases i <;> rfl
@[simp] @[simp]
lemma add_apply (x y : spaceTime) (i : Fin 4) : (x + y) i = x i + y i := rfl lemma add_apply (x y : SpaceTime) (i : Fin 4) : (x + y) i = x i + y i := rfl
@[simp] @[simp]
lemma smul_apply (x : spaceTime) (a : ) (i : Fin 4) : (a • x) i = a * x i := rfl lemma smul_apply (x : SpaceTime) (a : ) (i : Fin 4) : (a • x) i = a * x i := rfl
end spaceTime end SpaceTime
end end

View file

@ -13,18 +13,18 @@ We define
- `FourVelocity` : as a space-time element with norm 1 and future pointing. - `FourVelocity` : as a space-time element with norm 1 and future pointing.
-/ -/
open spaceTime open SpaceTime
/-- A `spaceTime` vector for which `ηLin x x = 1`. -/ /-- A `spaceTime` vector for which `ηLin x x = 1`. -/
@[simp] @[simp]
def PreFourVelocity : Set spaceTime := def PreFourVelocity : Set SpaceTime :=
fun x => ηLin x x = 1 fun x => ηLin x x = 1
instance : TopologicalSpace PreFourVelocity := by instance : TopologicalSpace PreFourVelocity := by
exact instTopologicalSpaceSubtype exact instTopologicalSpaceSubtype
namespace PreFourVelocity namespace PreFourVelocity
lemma mem_PreFourVelocity_iff {x : spaceTime} : x ∈ PreFourVelocity ↔ ηLin x x = 1 := by lemma mem_PreFourVelocity_iff {x : SpaceTime} : x ∈ PreFourVelocity ↔ ηLin x x = 1 := by
rfl rfl
/-- The negative of a `PreFourVelocity` as a `PreFourVelocity`. -/ /-- The negative of a `PreFourVelocity` as a `PreFourVelocity`. -/
@ -228,7 +228,7 @@ lemma η_pos (u v : FourVelocity) : 0 < ηLin u v := by
lemma one_plus_ne_zero (u v : FourVelocity) : 1 + ηLin u v ≠ 0 := by lemma one_plus_ne_zero (u v : FourVelocity) : 1 + ηLin u v ≠ 0 := by
linarith [η_pos u v] linarith [η_pos u v]
lemma η_continuous (u : spaceTime) : Continuous (fun (a : FourVelocity) => ηLin u a) := by lemma η_continuous (u : SpaceTime) : Continuous (fun (a : FourVelocity) => ηLin u a) := by
simp only [ηLin_expand] simp only [ηLin_expand]
refine Continuous.add ?_ ?_ refine Continuous.add ?_ ?_
refine Continuous.add ?_ ?_ refine Continuous.add ?_ ?_

View file

@ -19,7 +19,7 @@ We define
-/ -/
namespace spaceTime namespace SpaceTime
open Matrix open Matrix
open TensorProduct open TensorProduct
@ -96,7 +96,7 @@ lemma space_comps (Λ : lorentzAlgebra) (i j : Fin 3) :
end lorentzAlgebra end lorentzAlgebra
@[simps!] @[simps!]
instance spaceTimeAsLieRingModule : LieRingModule lorentzAlgebra spaceTime where instance spaceTimeAsLieRingModule : LieRingModule lorentzAlgebra SpaceTime where
bracket Λ x := Λ.1.mulVec x bracket Λ x := Λ.1.mulVec x
add_lie Λ1 Λ2 x := by add_lie Λ1 Λ2 x := by
simp [add_mulVec] simp [add_mulVec]
@ -107,7 +107,7 @@ instance spaceTimeAsLieRingModule : LieRingModule lorentzAlgebra spaceTime where
simp [mulVec_add, Bracket.bracket, sub_mulVec] simp [mulVec_add, Bracket.bracket, sub_mulVec]
@[simps!] @[simps!]
instance spaceTimeAsLieModule : LieModule lorentzAlgebra spaceTime where instance spaceTimeAsLieModule : LieModule lorentzAlgebra SpaceTime where
smul_lie r Λ x := by smul_lie r Λ x := by
simp [Bracket.bracket, smul_mulVec_assoc] simp [Bracket.bracket, smul_mulVec_assoc]
lie_smul r Λ x := by lie_smul r Λ x := by
@ -115,4 +115,4 @@ instance spaceTimeAsLieModule : LieModule lorentzAlgebra spaceTime where
rw [mulVec_smul] rw [mulVec_smul]
end spaceTime end SpaceTime

View file

@ -11,7 +11,7 @@ We define the standard basis of the Lorentz group.
-/ -/
namespace spaceTime namespace SpaceTime
namespace lorentzAlgebra namespace lorentzAlgebra
open Matrix open Matrix
@ -145,4 +145,4 @@ theorem finrank_eq_six : FiniteDimensional.finrank lorentzAlgebra = 6 := by
end lorentzAlgebra end lorentzAlgebra
end spaceTime end SpaceTime

View file

@ -26,7 +26,7 @@ identity.
noncomputable section noncomputable section
namespace spaceTime namespace SpaceTime
open Manifold open Manifold
open Matrix open Matrix
@ -44,14 +44,14 @@ These matrices form the Lorentz group, which we will define in the next section
/-- We say a matrix `Λ` preserves `ηLin` if for all `x` and `y`, /-- We say a matrix `Λ` preserves `ηLin` if for all `x` and `y`,
`ηLin (Λ *ᵥ x) (Λ *ᵥ y) = ηLin x y`. -/ `ηLin (Λ *ᵥ x) (Λ *ᵥ y) = ηLin x y`. -/
def PreservesηLin (Λ : Matrix (Fin 4) (Fin 4) ) : Prop := def PreservesηLin (Λ : Matrix (Fin 4) (Fin 4) ) : Prop :=
∀ (x y : spaceTime), ηLin (Λ *ᵥ x) (Λ *ᵥ y) = ηLin x y ∀ (x y : SpaceTime), ηLin (Λ *ᵥ x) (Λ *ᵥ y) = ηLin x y
namespace PreservesηLin namespace PreservesηLin
variable (Λ : Matrix (Fin 4) (Fin 4) ) variable (Λ : Matrix (Fin 4) (Fin 4) )
lemma iff_norm : PreservesηLin Λ ↔ lemma iff_norm : PreservesηLin Λ ↔
∀ (x : spaceTime), ηLin (Λ *ᵥ x) (Λ *ᵥ x) = ηLin x x := by ∀ (x : SpaceTime), ηLin (Λ *ᵥ x) (Λ *ᵥ x) = ηLin x x := by
refine Iff.intro (fun h x => h x x) (fun h x y => ?_) refine Iff.intro (fun h x => h x x) (fun h x y => ?_)
have hp := h (x + y) have hp := h (x + y)
have hn := h (x - y) have hn := h (x - y)
@ -75,7 +75,7 @@ lemma iff_det_selfAdjoint : PreservesηLin Λ ↔
simpa [det_eq_ηLin] using h1 simpa [det_eq_ηLin] using h1
lemma iff_on_right : PreservesηLin Λ ↔ lemma iff_on_right : PreservesηLin Λ ↔
∀ (x y : spaceTime), ηLin x ((η * Λᵀ * η * Λ) *ᵥ y) = ηLin x y := by ∀ (x y : SpaceTime), ηLin x ((η * Λᵀ * η * Λ) *ᵥ y) = ηLin x y := by
apply Iff.intro apply Iff.intro
intro h x y intro h x y
have h1 := h x y have h1 := h x y
@ -138,10 +138,10 @@ We show that the Lorentz group is indeed a group.
-/ -/
/-- The Lorentz group is the subset of matrices which preserve ηLin. -/ /-- The Lorentz group is the subset of matrices which preserve ηLin. -/
def lorentzGroup : Type := {Λ : Matrix (Fin 4) (Fin 4) // PreservesηLin Λ} def LorentzGroup : Type := {Λ : Matrix (Fin 4) (Fin 4) // PreservesηLin Λ}
@[simps mul_coe one_coe inv div] @[simps mul_coe one_coe inv div]
instance lorentzGroupIsGroup : Group lorentzGroup where instance lorentzGroupIsGroup : Group LorentzGroup where
mul A B := ⟨A.1 * B.1, PreservesηLin.mul A.2 B.2⟩ mul A B := ⟨A.1 * B.1, PreservesηLin.mul A.2 B.2⟩
mul_assoc A B C := by mul_assoc A B C := by
apply Subtype.eq apply Subtype.eq
@ -160,20 +160,20 @@ instance lorentzGroupIsGroup : Group lorentzGroup where
exact (PreservesηLin.iff_matrix A.1).mp A.2 exact (PreservesηLin.iff_matrix A.1).mp A.2
/-- Notation for the Lorentz group. -/ /-- Notation for the Lorentz group. -/
scoped[spaceTime] notation (name := lorentzGroup_notation) "𝓛" => lorentzGroup scoped[SpaceTime] notation (name := lorentzGroup_notation) "𝓛" => LorentzGroup
/-- `lorentzGroup` has the subtype topology. -/ /-- `lorentzGroup` has the subtype topology. -/
instance : TopologicalSpace lorentzGroup := instTopologicalSpaceSubtype instance : TopologicalSpace LorentzGroup := instTopologicalSpaceSubtype
namespace lorentzGroup namespace LorentzGroup
lemma coe_inv (A : 𝓛) : (A⁻¹).1 = A.1⁻¹:= by lemma coe_inv (A : LorentzGroup) : (A⁻¹).1 = A.1⁻¹:= by
refine (inv_eq_left_inv ?h).symm refine (inv_eq_left_inv ?h).symm
exact (PreservesηLin.iff_matrix A.1).mp A.2 exact (PreservesηLin.iff_matrix A.1).mp A.2
/-- The transpose of an matrix in the Lorentz group is an element of the Lorentz group. -/ /-- The transpose of an matrix in the Lorentz group is an element of the Lorentz group. -/
def transpose (Λ : lorentzGroup) : lorentzGroup := ⟨Λ.1ᵀ, (PreservesηLin.iff_transpose Λ.1).mp Λ.2⟩ def transpose (Λ : LorentzGroup) : LorentzGroup := ⟨Λ.1ᵀ, (PreservesηLin.iff_transpose Λ.1).mp Λ.2⟩
/-! /-!
@ -186,7 +186,7 @@ embedding.
-/ -/
/-- The homomorphism of the Lorentz group into `GL (Fin 4) `. -/ /-- The homomorphism of the Lorentz group into `GL (Fin 4) `. -/
def toGL : 𝓛 →* GL (Fin 4) where def toGL : LorentzGroup →* GL (Fin 4) where
toFun A := ⟨A.1, (A⁻¹).1, mul_eq_one_comm.mpr $ (PreservesηLin.iff_matrix A.1).mp A.2, toFun A := ⟨A.1, (A⁻¹).1, mul_eq_one_comm.mpr $ (PreservesηLin.iff_matrix A.1).mp A.2,
(PreservesηLin.iff_matrix A.1).mp A.2⟩ (PreservesηLin.iff_matrix A.1).mp A.2⟩
map_one' := by map_one' := by
@ -206,7 +206,7 @@ lemma toGL_injective : Function.Injective toGL := by
/-- The homomorphism from the Lorentz Group into the monoid of matrices times the opposite of /-- The homomorphism from the Lorentz Group into the monoid of matrices times the opposite of
the monoid of matrices. -/ the monoid of matrices. -/
@[simps!] @[simps!]
def toProd : 𝓛 →* (Matrix (Fin 4) (Fin 4) ) × (Matrix (Fin 4) (Fin 4) )ᵐᵒᵖ := def toProd : LorentzGroup →* (Matrix (Fin 4) (Fin 4) ) × (Matrix (Fin 4) (Fin 4) )ᵐᵒᵖ :=
MonoidHom.comp (Units.embedProduct _) toGL MonoidHom.comp (Units.embedProduct _) toGL
lemma toProd_eq_transpose_η : toProd A = (A.1, ⟨η * A.1ᵀ * η⟩) := rfl lemma toProd_eq_transpose_η : toProd A = (A.1, ⟨η * A.1ᵀ * η⟩) := rfl
@ -248,11 +248,11 @@ lemma toGL_embedding : Embedding toGL.toFun where
exact exists_exists_and_eq_and exact exists_exists_and_eq_and
instance : TopologicalGroup 𝓛 := Inducing.topologicalGroup toGL toGL_embedding.toInducing instance : TopologicalGroup LorentzGroup := Inducing.topologicalGroup toGL toGL_embedding.toInducing
end lorentzGroup end LorentzGroup
end spaceTime end SpaceTime

View file

@ -28,14 +28,14 @@ A boost is the special case of a generalised boost when `u = stdBasis 0`.
-/ -/
noncomputable section noncomputable section
namespace spaceTime namespace SpaceTime
namespace lorentzGroup namespace LorentzGroup
open FourVelocity open FourVelocity
/-- An auxillary linear map used in the definition of a generalised boost. -/ /-- An auxillary linear map used in the definition of a generalised boost. -/
def genBoostAux₁ (u v : FourVelocity) : spaceTime →ₗ[] spaceTime where def genBoostAux₁ (u v : FourVelocity) : SpaceTime →ₗ[] SpaceTime where
toFun x := (2 * ηLin x u) • v.1.1 toFun x := (2 * ηLin x u) • v.1.1
map_add' x y := by map_add' x y := by
simp only [map_add, LinearMap.add_apply] simp only [map_add, LinearMap.add_apply]
@ -46,7 +46,7 @@ def genBoostAux₁ (u v : FourVelocity) : spaceTime →ₗ[] spaceTime where
rw [← mul_assoc, mul_comm 2 c, mul_assoc, mul_smul] rw [← mul_assoc, mul_comm 2 c, mul_assoc, mul_smul]
/-- An auxillary linear map used in the definition of a genearlised boost. -/ /-- An auxillary linear map used in the definition of a genearlised boost. -/
def genBoostAux₂ (u v : FourVelocity) : spaceTime →ₗ[] spaceTime where def genBoostAux₂ (u v : FourVelocity) : SpaceTime →ₗ[] SpaceTime where
toFun x := - (ηLin x (u + v) / (1 + ηLin u v)) • (u + v) toFun x := - (ηLin x (u + v) / (1 + ηLin u v)) • (u + v)
map_add' x y := by map_add' x y := by
simp only simp only
@ -62,7 +62,7 @@ def genBoostAux₂ (u v : FourVelocity) : spaceTime →ₗ[] spaceTime where
/-- An generalised boost. This is a Lorentz transformation which takes the four velocity `u` /-- An generalised boost. This is a Lorentz transformation which takes the four velocity `u`
to `v`. -/ to `v`. -/
def genBoost (u v : FourVelocity) : spaceTime →ₗ[] spaceTime := def genBoost (u v : FourVelocity) : SpaceTime →ₗ[] SpaceTime :=
LinearMap.id + genBoostAux₁ u v + genBoostAux₂ u v LinearMap.id + genBoostAux₁ u v + genBoostAux₂ u v
namespace genBoost namespace genBoost
@ -91,7 +91,7 @@ lemma self (u : FourVelocity) : genBoost u u = LinearMap.id := by
def toMatrix (u v : FourVelocity) : Matrix (Fin 4) (Fin 4) := def toMatrix (u v : FourVelocity) : Matrix (Fin 4) (Fin 4) :=
LinearMap.toMatrix stdBasis stdBasis (genBoost u v) LinearMap.toMatrix stdBasis stdBasis (genBoost u v)
lemma toMatrix_mulVec (u v : FourVelocity) (x : spaceTime) : lemma toMatrix_mulVec (u v : FourVelocity) (x : SpaceTime) :
(toMatrix u v).mulVec x = genBoost u v x := (toMatrix u v).mulVec x = genBoost u v x :=
LinearMap.toMatrix_mulVec_repr stdBasis stdBasis (genBoost u v) x LinearMap.toMatrix_mulVec_repr stdBasis stdBasis (genBoost u v) x
@ -143,7 +143,7 @@ lemma toMatrix_PreservesηLin (u v : FourVelocity) : PreservesηLin (toMatrix u
ring ring
/-- A generalised boost as an element of the Lorentz Group. -/ /-- A generalised boost as an element of the Lorentz Group. -/
def toLorentz (u v : FourVelocity) : lorentzGroup := def toLorentz (u v : FourVelocity) : LorentzGroup :=
⟨toMatrix u v, toMatrix_PreservesηLin u v⟩ ⟨toMatrix u v, toMatrix_PreservesηLin u v⟩
lemma toLorentz_continuous (u : FourVelocity) : Continuous (toLorentz u) := by lemma toLorentz_continuous (u : FourVelocity) : Continuous (toLorentz u) := by
@ -171,8 +171,8 @@ lemma isProper (u v : FourVelocity) : IsProper (toLorentz u v) :=
end genBoost end genBoost
end lorentzGroup end LorentzGroup
end spaceTime end SpaceTime
end end

View file

@ -20,19 +20,19 @@ matrices.
noncomputable section noncomputable section
namespace spaceTime namespace SpaceTime
open Manifold open Manifold
open Matrix open Matrix
open Complex open Complex
open ComplexConjugate open ComplexConjugate
namespace lorentzGroup namespace LorentzGroup
open PreFourVelocity open PreFourVelocity
/-- The first column of a lorentz matrix as a `PreFourVelocity`. -/ /-- The first column of a lorentz matrix as a `PreFourVelocity`. -/
@[simp] @[simp]
def fstCol (Λ : lorentzGroup) : PreFourVelocity := ⟨Λ.1 *ᵥ stdBasis 0, by def fstCol (Λ : LorentzGroup) : PreFourVelocity := ⟨Λ.1 *ᵥ stdBasis 0, by
rw [mem_PreFourVelocity_iff, ηLin_expand] rw [mem_PreFourVelocity_iff, ηLin_expand]
simp only [Fin.isValue, stdBasis_mulVec] simp only [Fin.isValue, stdBasis_mulVec]
have h00 := congrFun (congrFun ((PreservesηLin.iff_matrix Λ.1).mp Λ.2) 0) 0 have h00 := congrFun (congrFun ((PreservesηLin.iff_matrix Λ.1).mp Λ.2) 0) 0
@ -46,18 +46,18 @@ def fstCol (Λ : lorentzGroup) : PreFourVelocity := ⟨Λ.1 *ᵥ stdBasis 0, by
exact h00⟩ exact h00⟩
/-- 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 (Λ : lorentzGroup) : Prop := 0 ≤ Λ.1 0 0 def IsOrthochronous (Λ : LorentzGroup) : Prop := 0 ≤ Λ.1 0 0
lemma IsOrthochronous_iff_transpose (Λ : lorentzGroup) : lemma IsOrthochronous_iff_transpose (Λ : LorentzGroup) :
IsOrthochronous Λ ↔ IsOrthochronous (transpose Λ) := by rfl IsOrthochronous Λ ↔ IsOrthochronous (transpose Λ) := by rfl
lemma IsOrthochronous_iff_fstCol_IsFourVelocity (Λ : lorentzGroup) : lemma IsOrthochronous_iff_fstCol_IsFourVelocity (Λ : LorentzGroup) :
IsOrthochronous Λ ↔ IsFourVelocity (fstCol Λ) := by IsOrthochronous Λ ↔ IsFourVelocity (fstCol Λ) := by
simp [IsOrthochronous, IsFourVelocity] simp [IsOrthochronous, IsFourVelocity]
rw [stdBasis_mulVec] rw [stdBasis_mulVec]
/-- The continuous map taking a Lorentz transformation to its `0 0` element. -/ /-- The continuous map taking a Lorentz transformation to its `0 0` element. -/
def mapZeroZeroComp : C(lorentzGroup, ) := ⟨fun Λ => Λ.1 0 0, def mapZeroZeroComp : C(LorentzGroup, ) := ⟨fun Λ => Λ.1 0 0,
Continuous.matrix_elem (continuous_iff_le_induced.mpr fun _ a => a) 0 0⟩ Continuous.matrix_elem (continuous_iff_le_induced.mpr fun _ a => a) 0 0⟩
/-- An auxillary function used in the definition of `orthchroMapReal`. -/ /-- An auxillary function used in the definition of `orthchroMapReal`. -/
@ -78,10 +78,10 @@ lemma stepFunction_continuous : Continuous stepFunction := by
/-- The continuous map from `lorentzGroup` to `` wh /-- The continuous map from `lorentzGroup` to `` wh
taking Orthochronous elements to `1` and non-orthochronous to `-1`. -/ taking Orthochronous elements to `1` and non-orthochronous to `-1`. -/
def orthchroMapReal : C(lorentzGroup, ) := ContinuousMap.comp def orthchroMapReal : C(LorentzGroup, ) := ContinuousMap.comp
⟨stepFunction, stepFunction_continuous⟩ mapZeroZeroComp ⟨stepFunction, stepFunction_continuous⟩ mapZeroZeroComp
lemma orthchroMapReal_on_IsOrthochronous {Λ : lorentzGroup} (h : IsOrthochronous Λ) : lemma orthchroMapReal_on_IsOrthochronous {Λ : LorentzGroup} (h : IsOrthochronous Λ) :
orthchroMapReal Λ = 1 := by orthchroMapReal Λ = 1 := by
rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h
simp only [IsFourVelocity] at h simp only [IsFourVelocity] at h
@ -92,7 +92,7 @@ lemma orthchroMapReal_on_IsOrthochronous {Λ : lorentzGroup} (h : IsOrthochronou
rw [stepFunction, if_neg h1, if_pos h] rw [stepFunction, if_neg h1, if_pos h]
lemma orthchroMapReal_on_not_IsOrthochronous {Λ : lorentzGroup} (h : ¬ IsOrthochronous Λ) : lemma orthchroMapReal_on_not_IsOrthochronous {Λ : LorentzGroup} (h : ¬ IsOrthochronous Λ) :
orthchroMapReal Λ = - 1 := by orthchroMapReal Λ = - 1 := by
rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h
rw [not_IsFourVelocity_iff, zero_nonpos_iff] at h rw [not_IsFourVelocity_iff, zero_nonpos_iff] at h
@ -100,7 +100,7 @@ lemma orthchroMapReal_on_not_IsOrthochronous {Λ : lorentzGroup} (h : ¬ IsOrtho
change stepFunction (Λ.1 0 0) = - 1 change stepFunction (Λ.1 0 0) = - 1
rw [stepFunction, if_pos h] rw [stepFunction, if_pos h]
lemma orthchroMapReal_minus_one_or_one (Λ : lorentzGroup) : lemma orthchroMapReal_minus_one_or_one (Λ : LorentzGroup) :
orthchroMapReal Λ = -1 orthchroMapReal Λ = 1 := by orthchroMapReal Λ = -1 orthchroMapReal Λ = 1 := by
by_cases h : IsOrthochronous Λ by_cases h : IsOrthochronous Λ
apply Or.inr $ orthchroMapReal_on_IsOrthochronous h apply Or.inr $ orthchroMapReal_on_IsOrthochronous h
@ -109,21 +109,21 @@ lemma orthchroMapReal_minus_one_or_one (Λ : lorentzGroup) :
local notation "ℤ₂" => Multiplicative (ZMod 2) local notation "ℤ₂" => Multiplicative (ZMod 2)
/-- A continuous map from `lorentzGroup` to `ℤ₂` whose kernal are the Orthochronous elements. -/ /-- A continuous map from `lorentzGroup` to `ℤ₂` whose kernal are the Orthochronous elements. -/
def orthchroMap : C(lorentzGroup, ℤ₂) := def orthchroMap : C(LorentzGroup, ℤ₂) :=
ContinuousMap.comp coeFor₂ { ContinuousMap.comp coeFor₂ {
toFun := fun Λ => ⟨orthchroMapReal Λ, orthchroMapReal_minus_one_or_one Λ⟩, toFun := fun Λ => ⟨orthchroMapReal Λ, orthchroMapReal_minus_one_or_one Λ⟩,
continuous_toFun := Continuous.subtype_mk (ContinuousMap.continuous orthchroMapReal) _} continuous_toFun := Continuous.subtype_mk (ContinuousMap.continuous orthchroMapReal) _}
lemma orthchroMap_IsOrthochronous {Λ : lorentzGroup} (h : IsOrthochronous Λ) : lemma orthchroMap_IsOrthochronous {Λ : LorentzGroup} (h : IsOrthochronous Λ) :
orthchroMap Λ = 1 := by orthchroMap Λ = 1 := by
simp [orthchroMap, orthchroMapReal_on_IsOrthochronous h] simp [orthchroMap, orthchroMapReal_on_IsOrthochronous h]
lemma orthchroMap_not_IsOrthochronous {Λ : lorentzGroup} (h : ¬ IsOrthochronous Λ) : lemma orthchroMap_not_IsOrthochronous {Λ : LorentzGroup} (h : ¬ IsOrthochronous Λ) :
orthchroMap Λ = Additive.toMul (1 : ZMod 2) := by orthchroMap Λ = Additive.toMul (1 : ZMod 2) := by
simp [orthchroMap, orthchroMapReal_on_not_IsOrthochronous h] simp [orthchroMap, orthchroMapReal_on_not_IsOrthochronous h]
rfl rfl
lemma zero_zero_mul (Λ Λ' : lorentzGroup) : lemma zero_zero_mul (Λ Λ' : LorentzGroup) :
(Λ * Λ').1 0 0 = (fstCol (transpose Λ)).1 0 * (fstCol Λ').1 0 + (Λ * Λ').1 0 0 = (fstCol (transpose Λ)).1 0 * (fstCol Λ').1 0 +
⟪(fstCol (transpose Λ)).1.space, (fstCol Λ').1.space⟫_ := by ⟪(fstCol (transpose Λ)).1.space, (fstCol Λ').1.space⟫_ := by
simp only [Fin.isValue, lorentzGroupIsGroup_mul_coe, mul_apply, Fin.sum_univ_four, fstCol, simp only [Fin.isValue, lorentzGroupIsGroup_mul_coe, mul_apply, Fin.sum_univ_four, fstCol,
@ -132,21 +132,21 @@ lemma zero_zero_mul (Λ Λ' : lorentzGroup) :
cons_val_one, head_cons, cons_val_two, tail_cons] cons_val_one, head_cons, cons_val_two, tail_cons]
ring ring
lemma mul_othchron_of_othchron_othchron {Λ Λ' : lorentzGroup} (h : IsOrthochronous Λ) lemma mul_othchron_of_othchron_othchron {Λ Λ' : LorentzGroup} (h : IsOrthochronous Λ)
(h' : IsOrthochronous Λ') : IsOrthochronous (Λ * Λ') := by (h' : IsOrthochronous Λ') : IsOrthochronous (Λ * Λ') := by
rw [IsOrthochronous_iff_transpose] at h rw [IsOrthochronous_iff_transpose] at h
rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h h' rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h h'
rw [IsOrthochronous, zero_zero_mul] rw [IsOrthochronous, zero_zero_mul]
exact euclid_norm_IsFourVelocity_IsFourVelocity h h' exact euclid_norm_IsFourVelocity_IsFourVelocity h h'
lemma mul_othchron_of_not_othchron_not_othchron {Λ Λ' : lorentzGroup} (h : ¬ IsOrthochronous Λ) lemma mul_othchron_of_not_othchron_not_othchron {Λ Λ' : LorentzGroup} (h : ¬ IsOrthochronous Λ)
(h' : ¬ IsOrthochronous Λ') : IsOrthochronous (Λ * Λ') := by (h' : ¬ IsOrthochronous Λ') : IsOrthochronous (Λ * Λ') := by
rw [IsOrthochronous_iff_transpose] at h rw [IsOrthochronous_iff_transpose] at h
rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h h' rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h h'
rw [IsOrthochronous, zero_zero_mul] rw [IsOrthochronous, zero_zero_mul]
exact euclid_norm_not_IsFourVelocity_not_IsFourVelocity h h' exact euclid_norm_not_IsFourVelocity_not_IsFourVelocity h h'
lemma mul_not_othchron_of_othchron_not_othchron {Λ Λ' : lorentzGroup} (h : IsOrthochronous Λ) lemma mul_not_othchron_of_othchron_not_othchron {Λ Λ' : LorentzGroup} (h : IsOrthochronous Λ)
(h' : ¬ IsOrthochronous Λ') : ¬ IsOrthochronous (Λ * Λ') := by (h' : ¬ IsOrthochronous Λ') : ¬ IsOrthochronous (Λ * Λ') := by
rw [IsOrthochronous_iff_transpose] at h rw [IsOrthochronous_iff_transpose] at h
rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h h' rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h h'
@ -156,7 +156,7 @@ lemma mul_not_othchron_of_othchron_not_othchron {Λ Λ' : lorentzGroup} (h : Is
rw [zero_zero_mul] rw [zero_zero_mul]
exact euclid_norm_IsFourVelocity_not_IsFourVelocity h h' exact euclid_norm_IsFourVelocity_not_IsFourVelocity h h'
lemma mul_not_othchron_of_not_othchron_othchron {Λ Λ' : lorentzGroup} (h : ¬ IsOrthochronous Λ) lemma mul_not_othchron_of_not_othchron_othchron {Λ Λ' : LorentzGroup} (h : ¬ IsOrthochronous Λ)
(h' : IsOrthochronous Λ') : ¬ IsOrthochronous (Λ * Λ') := by (h' : IsOrthochronous Λ') : ¬ IsOrthochronous (Λ * Λ') := by
rw [IsOrthochronous_iff_transpose] at h rw [IsOrthochronous_iff_transpose] at h
rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h h' rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h h'
@ -167,7 +167,7 @@ lemma mul_not_othchron_of_not_othchron_othchron {Λ Λ' : lorentzGroup} (h : ¬
exact euclid_norm_not_IsFourVelocity_IsFourVelocity h h' exact euclid_norm_not_IsFourVelocity_IsFourVelocity h h'
/-- The homomorphism from `lorentzGroup` to `ℤ₂` whose kernel are the Orthochronous elements. -/ /-- The homomorphism from `lorentzGroup` to `ℤ₂` whose kernel are the Orthochronous elements. -/
def orthchroRep : lorentzGroup →* ℤ₂ where def orthchroRep : LorentzGroup →* ℤ₂ where
toFun := orthchroMap toFun := orthchroMap
map_one' := orthchroMap_IsOrthochronous (by simp [IsOrthochronous]) map_one' := orthchroMap_IsOrthochronous (by simp [IsOrthochronous])
map_mul' Λ Λ' := by map_mul' Λ Λ' := by
@ -187,7 +187,7 @@ def orthchroRep : lorentzGroup →* ℤ₂ where
orthchroMap_IsOrthochronous (mul_othchron_of_not_othchron_not_othchron h h')] orthchroMap_IsOrthochronous (mul_othchron_of_not_othchron_not_othchron h h')]
rfl rfl
end lorentzGroup end LorentzGroup
end spaceTime end SpaceTime
end end

View file

@ -14,14 +14,14 @@ We define the give a series of lemmas related to the determinant of the lorentz
noncomputable section noncomputable section
namespace spaceTime namespace SpaceTime
open Manifold open Manifold
open Matrix open Matrix
open Complex open Complex
open ComplexConjugate open ComplexConjugate
namespace lorentzGroup namespace LorentzGroup
/-- The determinant of a member of the lorentz group is `1` or `-1`. -/ /-- The determinant of a member of the lorentz group is `1` or `-1`. -/
lemma det_eq_one_or_neg_one (Λ : 𝓛) : Λ.1.det = 1 Λ.1.det = -1 := by lemma det_eq_one_or_neg_one (Λ : 𝓛) : Λ.1.det = 1 Λ.1.det = -1 := by
@ -49,14 +49,14 @@ def coeFor₂ : C(({-1, 1} : Set ), ℤ₂) where
/-- The continuous map taking a lorentz matrix to its determinant. -/ /-- The continuous map taking a lorentz matrix to its determinant. -/
def detContinuous : C(𝓛, ℤ₂) := def detContinuous : C(𝓛, ℤ₂) :=
ContinuousMap.comp coeFor₂ { ContinuousMap.comp coeFor₂ {
toFun := fun Λ => ⟨Λ.1.det, Or.symm (lorentzGroup.det_eq_one_or_neg_one _)⟩, toFun := fun Λ => ⟨Λ.1.det, Or.symm (LorentzGroup.det_eq_one_or_neg_one _)⟩,
continuous_toFun := by continuous_toFun := by
refine Continuous.subtype_mk ?_ _ refine Continuous.subtype_mk ?_ _
apply Continuous.matrix_det $ apply Continuous.matrix_det $
Continuous.comp' (continuous_iff_le_induced.mpr fun U a => a) continuous_id' Continuous.comp' (continuous_iff_le_induced.mpr fun U a => a) continuous_id'
} }
lemma detContinuous_eq_iff_det_eq (Λ Λ' : lorentzGroup) : lemma detContinuous_eq_iff_det_eq (Λ Λ' : LorentzGroup) :
detContinuous Λ = detContinuous Λ' ↔ Λ.1.det = Λ'.1.det := by detContinuous Λ = detContinuous Λ' ↔ Λ.1.det = Λ'.1.det := by
apply Iff.intro apply Iff.intro
intro h intro h
@ -90,7 +90,7 @@ def detRep : 𝓛 →* ℤ₂ where
lemma detRep_continuous : Continuous detRep := detContinuous.2 lemma detRep_continuous : Continuous detRep := detContinuous.2
lemma det_on_connected_component {Λ Λ' : lorentzGroup} (h : Λ' ∈ connectedComponent Λ) : lemma det_on_connected_component {Λ Λ' : LorentzGroup} (h : Λ' ∈ connectedComponent Λ) :
Λ.1.det = Λ'.1.det := by Λ.1.det = Λ'.1.det := by
obtain ⟨s, hs, hΛ'⟩ := h obtain ⟨s, hs, hΛ'⟩ := h
let f : ContinuousMap s ℤ₂ := ContinuousMap.restrict s detContinuous let f : ContinuousMap s ℤ₂ := ContinuousMap.restrict s detContinuous
@ -99,23 +99,23 @@ lemma det_on_connected_component {Λ Λ' : lorentzGroup} (h : Λ' ∈ connected
(@IsPreconnected.subsingleton ℤ₂ _ _ _ (isPreconnected_range f.2)) (@IsPreconnected.subsingleton ℤ₂ _ _ _ (isPreconnected_range f.2))
(Set.mem_range_self ⟨Λ, hs.2⟩) (Set.mem_range_self ⟨Λ', hΛ'⟩) (Set.mem_range_self ⟨Λ, hs.2⟩) (Set.mem_range_self ⟨Λ', hΛ'⟩)
lemma detRep_on_connected_component {Λ Λ' : lorentzGroup} (h : Λ' ∈ connectedComponent Λ) : lemma detRep_on_connected_component {Λ Λ' : LorentzGroup} (h : Λ' ∈ connectedComponent Λ) :
detRep Λ = detRep Λ' := by detRep Λ = detRep Λ' := by
simp [detRep_apply, detRep_apply, detContinuous] simp [detRep_apply, detRep_apply, detContinuous]
rw [det_on_connected_component h] rw [det_on_connected_component h]
lemma det_of_joined {Λ Λ' : lorentzGroup} (h : Joined Λ Λ') : Λ.1.det = Λ'.1.det := lemma det_of_joined {Λ Λ' : LorentzGroup} (h : Joined Λ Λ') : Λ.1.det = Λ'.1.det :=
det_on_connected_component $ pathComponent_subset_component _ h det_on_connected_component $ pathComponent_subset_component _ h
/-- A Lorentz Matrix is proper if its determinant is 1. -/ /-- A Lorentz Matrix is proper if its determinant is 1. -/
@[simp] @[simp]
def IsProper (Λ : lorentzGroup) : Prop := Λ.1.det = 1 def IsProper (Λ : LorentzGroup) : Prop := Λ.1.det = 1
instance : DecidablePred IsProper := by instance : DecidablePred IsProper := by
intro Λ intro Λ
apply Real.decidableEq apply Real.decidableEq
lemma IsProper_iff (Λ : lorentzGroup) : IsProper Λ ↔ detRep Λ = 1 := by lemma IsProper_iff (Λ : LorentzGroup) : IsProper Λ ↔ detRep Λ = 1 := by
rw [show 1 = detRep 1 from Eq.symm (MonoidHom.map_one detRep)] rw [show 1 = detRep 1 from Eq.symm (MonoidHom.map_one detRep)]
rw [detRep_apply, detRep_apply, detContinuous_eq_iff_det_eq] rw [detRep_apply, detRep_apply, detContinuous_eq_iff_det_eq]
simp only [IsProper, lorentzGroupIsGroup_one_coe, det_one] simp only [IsProper, lorentzGroupIsGroup_one_coe, det_one]
@ -123,12 +123,12 @@ lemma IsProper_iff (Λ : lorentzGroup) : IsProper Λ ↔ detRep Λ = 1 := by
lemma id_IsProper : IsProper 1 := by lemma id_IsProper : IsProper 1 := by
simp [IsProper] simp [IsProper]
lemma isProper_on_connected_component {Λ Λ' : lorentzGroup} (h : Λ' ∈ connectedComponent Λ) : lemma isProper_on_connected_component {Λ Λ' : LorentzGroup} (h : Λ' ∈ connectedComponent Λ) :
IsProper Λ ↔ IsProper Λ' := by IsProper Λ ↔ IsProper Λ' := by
simp [detRep_apply, detRep_apply, detContinuous] simp [detRep_apply, detRep_apply, detContinuous]
rw [det_on_connected_component h] rw [det_on_connected_component h]
end lorentzGroup end LorentzGroup
end spaceTime end SpaceTime
end end

View file

@ -11,7 +11,7 @@ import Mathlib.Topology.Constructions
-/ -/
noncomputable section noncomputable section
namespace spaceTime namespace SpaceTime
namespace lorentzGroup namespace lorentzGroup
open GroupTheory open GroupTheory
@ -54,5 +54,5 @@ def SO3ToLorentz : SO(3) →* 𝓛 where
end lorentzGroup end lorentzGroup
end spaceTime end SpaceTime
end end

View file

@ -15,7 +15,7 @@ This file introduces the metric on spacetime in the (+, -, -, -) signature.
noncomputable section noncomputable section
namespace spaceTime namespace SpaceTime
open Manifold open Manifold
open Matrix open Matrix
@ -108,7 +108,7 @@ lemma η_sq : η * η = 1 := by
lemma η_diag_mul_self (μ : Fin 4) : η μ μ * η μ μ = 1 := by lemma η_diag_mul_self (μ : Fin 4) : η μ μ * η μ μ = 1 := by
fin_cases μ <;> simp [η_explicit] fin_cases μ <;> simp [η_explicit]
lemma η_mulVec (x : spaceTime) : η *ᵥ x = ![x 0, -x 1, -x 2, -x 3] := by lemma η_mulVec (x : SpaceTime) : η *ᵥ x = ![x 0, -x 1, -x 2, -x 3] := by
rw [explicit x, η_explicit] rw [explicit x, η_explicit]
funext i funext i
fin_cases i <;> fin_cases i <;>
@ -143,7 +143,7 @@ lemma η_contract_self' (μ ν : Fin 4) : ∑ x, (η^[x]_[μ] * η_[ν]_[x]) =
/-- Given a point in spaceTime `x` the linear map `y → x ⬝ᵥ (η *ᵥ y)`. -/ /-- Given a point in spaceTime `x` the linear map `y → x ⬝ᵥ (η *ᵥ y)`. -/
@[simps!] @[simps!]
def linearMapForSpaceTime (x : spaceTime) : spaceTime →ₗ[] where def linearMapForSpaceTime (x : SpaceTime) : SpaceTime →ₗ[] where
toFun y := x ⬝ᵥ (η *ᵥ y) toFun y := x ⬝ᵥ (η *ᵥ y)
map_add' y z := by map_add' y z := by
simp only simp only
@ -154,7 +154,7 @@ def linearMapForSpaceTime (x : spaceTime) : spaceTime →ₗ[] where
rfl rfl
/-- The metric as a bilinear map from `spaceTime` to ``. -/ /-- The metric as a bilinear map from `spaceTime` to ``. -/
def ηLin : LinearMap.BilinForm spaceTime where def ηLin : LinearMap.BilinForm SpaceTime where
toFun x := linearMapForSpaceTime x toFun x := linearMapForSpaceTime x
map_add' x y := by map_add' x y := by
apply LinearMap.ext apply LinearMap.ext
@ -168,7 +168,7 @@ def ηLin : LinearMap.BilinForm spaceTime where
rw [smul_dotProduct] rw [smul_dotProduct]
rfl rfl
lemma ηLin_expand (x y : spaceTime) : ηLin x y = x 0 * y 0 - x 1 * y 1 - x 2 * y 2 - x 3 * y 3 := by lemma ηLin_expand (x y : SpaceTime) : ηLin x y = x 0 * y 0 - x 1 * y 1 - x 2 * y 2 - x 3 * y 3 := by
rw [ηLin] rw [ηLin]
simp only [LinearMap.coe_mk, AddHom.coe_mk, linearMapForSpaceTime_apply, Fin.isValue] simp only [LinearMap.coe_mk, AddHom.coe_mk, linearMapForSpaceTime_apply, Fin.isValue]
erw [η_mulVec] erw [η_mulVec]
@ -177,40 +177,40 @@ lemma ηLin_expand (x y : spaceTime) : ηLin x y = x 0 * y 0 - x 1 * y 1 - x 2 *
cons_val_zero, cons_val_one, head_cons, mul_neg, cons_val_two, tail_cons, cons_val_three] cons_val_zero, cons_val_one, head_cons, mul_neg, cons_val_two, tail_cons, cons_val_three]
ring ring
lemma ηLin_expand_self (x : spaceTime) : ηLin x x = x 0 ^ 2 - ‖x.space‖ ^ 2 := by lemma ηLin_expand_self (x : SpaceTime) : ηLin x x = x 0 ^ 2 - ‖x.space‖ ^ 2 := by
rw [← @real_inner_self_eq_norm_sq, @PiLp.inner_apply, Fin.sum_univ_three, ηLin_expand] rw [← @real_inner_self_eq_norm_sq, @PiLp.inner_apply, Fin.sum_univ_three, ηLin_expand]
noncomm_ring noncomm_ring
lemma time_elm_sq_of_ηLin (x : spaceTime) : x 0 ^ 2 = ηLin x x + ‖x.space‖ ^ 2 := by lemma time_elm_sq_of_ηLin (x : SpaceTime) : x 0 ^ 2 = ηLin x x + ‖x.space‖ ^ 2 := by
rw [ηLin_expand_self] rw [ηLin_expand_self]
ring ring
lemma ηLin_leq_time_sq (x : spaceTime) : ηLin x x ≤ x 0 ^ 2 := by lemma ηLin_leq_time_sq (x : SpaceTime) : ηLin x x ≤ x 0 ^ 2 := by
rw [time_elm_sq_of_ηLin] rw [time_elm_sq_of_ηLin]
exact (le_add_iff_nonneg_right _).mpr $ sq_nonneg ‖x.space‖ exact (le_add_iff_nonneg_right _).mpr $ sq_nonneg ‖x.space‖
lemma ηLin_space_inner_product (x y : spaceTime) : lemma ηLin_space_inner_product (x y : SpaceTime) :
ηLin x y = x 0 * y 0 - ⟪x.space, y.space⟫_ := by ηLin x y = x 0 * y 0 - ⟪x.space, y.space⟫_ := by
rw [ηLin_expand, @PiLp.inner_apply, Fin.sum_univ_three] rw [ηLin_expand, @PiLp.inner_apply, Fin.sum_univ_three]
noncomm_ring noncomm_ring
lemma ηLin_ge_abs_inner_product (x y : spaceTime) : lemma ηLin_ge_abs_inner_product (x y : SpaceTime) :
x 0 * y 0 - ‖⟪x.space, y.space⟫_‖ ≤ (ηLin x y) := by x 0 * y 0 - ‖⟪x.space, y.space⟫_‖ ≤ (ηLin x y) := by
rw [ηLin_space_inner_product, sub_le_sub_iff_left] rw [ηLin_space_inner_product, sub_le_sub_iff_left]
exact Real.le_norm_self ⟪x.space, y.space⟫_ exact Real.le_norm_self ⟪x.space, y.space⟫_
lemma ηLin_ge_sub_norm (x y : spaceTime) : lemma ηLin_ge_sub_norm (x y : SpaceTime) :
x 0 * y 0 - ‖x.space‖ * ‖y.space‖ ≤ (ηLin x y) := by x 0 * y 0 - ‖x.space‖ * ‖y.space‖ ≤ (ηLin x y) := by
apply le_trans ?_ (ηLin_ge_abs_inner_product x y) apply le_trans ?_ (ηLin_ge_abs_inner_product x y)
rw [sub_le_sub_iff_left] rw [sub_le_sub_iff_left]
exact norm_inner_le_norm x.space y.space exact norm_inner_le_norm x.space y.space
lemma ηLin_symm (x y : spaceTime) : ηLin x y = ηLin y x := by lemma ηLin_symm (x y : SpaceTime) : ηLin x y = ηLin y x := by
rw [ηLin_expand, ηLin_expand] rw [ηLin_expand, ηLin_expand]
ring ring
lemma ηLin_stdBasis_apply (μ : Fin 4) (x : spaceTime) : ηLin (stdBasis μ) x = η μ μ * x μ := by lemma ηLin_stdBasis_apply (μ : Fin 4) (x : SpaceTime) : ηLin (stdBasis μ) x = η μ μ * x μ := by
rw [ηLin_expand] rw [ηLin_expand]
fin_cases μ fin_cases μ
<;> simp [stdBasis_0, stdBasis_1, stdBasis_2, stdBasis_3, η_explicit] <;> simp [stdBasis_0, stdBasis_1, stdBasis_2, stdBasis_3, η_explicit]
@ -227,13 +227,13 @@ lemma ηLin_η_stdBasis (μ ν : Fin 4) : ηLin (stdBasis μ) (stdBasis ν) = η
exact fun a ↦ h (id a.symm) exact fun a ↦ h (id a.symm)
set_option maxHeartbeats 0 set_option maxHeartbeats 0
lemma ηLin_mulVec_left (x y : spaceTime) (Λ : Matrix (Fin 4) (Fin 4) ) : lemma ηLin_mulVec_left (x y : SpaceTime) (Λ : Matrix (Fin 4) (Fin 4) ) :
ηLin (Λ *ᵥ x) y = ηLin x ((η * Λᵀ * η) *ᵥ y) := by ηLin (Λ *ᵥ x) y = ηLin x ((η * Λᵀ * η) *ᵥ y) := by
simp [ηLin, LinearMap.coe_mk, AddHom.coe_mk, linearMapForSpaceTime_apply, simp [ηLin, LinearMap.coe_mk, AddHom.coe_mk, linearMapForSpaceTime_apply,
mulVec_mulVec, (vecMul_transpose Λ x).symm, ← dotProduct_mulVec, mulVec_mulVec, mulVec_mulVec, (vecMul_transpose Λ x).symm, ← dotProduct_mulVec, mulVec_mulVec,
← mul_assoc, ← mul_assoc, η_sq, one_mul] ← mul_assoc, ← mul_assoc, η_sq, one_mul]
lemma ηLin_mulVec_right (x y : spaceTime) (Λ : Matrix (Fin 4) (Fin 4) ) : lemma ηLin_mulVec_right (x y : SpaceTime) (Λ : Matrix (Fin 4) (Fin 4) ) :
ηLin x (Λ *ᵥ y) = ηLin ((η * Λᵀ * η) *ᵥ x) y := by ηLin x (Λ *ᵥ y) = ηLin ((η * Λᵀ * η) *ᵥ x) y := by
rw [ηLin_symm, ηLin_symm ((η * Λᵀ * η) *ᵥ x) _ ] rw [ηLin_symm, ηLin_symm ((η * Λᵀ * η) *ᵥ x) _ ]
exact ηLin_mulVec_left y x Λ exact ηLin_mulVec_left y x Λ
@ -247,7 +247,7 @@ lemma ηLin_matrix_stdBasis' (μ ν : Fin 4) (Λ : Matrix (Fin 4) (Fin 4) ) :
rw [ηLin_matrix_stdBasis, ← mul_assoc, η_diag_mul_self, one_mul] rw [ηLin_matrix_stdBasis, ← mul_assoc, η_diag_mul_self, one_mul]
lemma ηLin_matrix_eq_identity_iff (Λ : Matrix (Fin 4) (Fin 4) ) : lemma ηLin_matrix_eq_identity_iff (Λ : Matrix (Fin 4) (Fin 4) ) :
Λ = 1 ↔ ∀ (x y : spaceTime), ηLin x y = ηLin x (Λ *ᵥ y) := by Λ = 1 ↔ ∀ (x y : SpaceTime), ηLin x y = ηLin x (Λ *ᵥ y) := by
apply Iff.intro apply Iff.intro
· intro h · intro h
subst h subst h
@ -260,8 +260,8 @@ lemma ηLin_matrix_eq_identity_iff (Λ : Matrix (Fin 4) (Fin 4) ) :
simp_all [η_explicit, Fin.mk_one, Fin.mk_one, vecHead, vecTail] simp_all [η_explicit, Fin.mk_one, Fin.mk_one, vecHead, vecTail]
/-- The metric as a quadratic form on `spaceTime`. -/ /-- The metric as a quadratic form on `spaceTime`. -/
def quadraticForm : QuadraticForm spaceTime := ηLin.toQuadraticForm def quadraticForm : QuadraticForm SpaceTime := ηLin.toQuadraticForm
end spaceTime end SpaceTime
end end

View file

@ -11,7 +11,7 @@ import Mathlib.RepresentationTheory.Basic
The aim of this file is to give the relationship between `SL(2, )` and the Lorentz group. The aim of this file is to give the relationship between `SL(2, )` and the Lorentz group.
-/ -/
namespace spaceTime namespace SpaceTime
open Matrix open Matrix
open MatrixGroups open MatrixGroups
@ -19,7 +19,7 @@ open Complex
namespace SL2C namespace SL2C
open spaceTime open SpaceTime
noncomputable section noncomputable section
@ -64,7 +64,7 @@ def repSelfAdjointMatrix : Representation SL(2, ) $ selfAdjoint (Matrix (
/-- The representation of `SL(2, )` on `spaceTime` obtained from `toSelfAdjointMatrix` and /-- The representation of `SL(2, )` on `spaceTime` obtained from `toSelfAdjointMatrix` and
`repSelfAdjointMatrix`. -/ `repSelfAdjointMatrix`. -/
def repSpaceTime : Representation SL(2, ) spaceTime where def repSpaceTime : Representation SL(2, ) SpaceTime where
toFun M := toSelfAdjointMatrix.symm.comp ((repSelfAdjointMatrix M).comp toFun M := toSelfAdjointMatrix.symm.comp ((repSelfAdjointMatrix M).comp
toSelfAdjointMatrix.toLinearMap) toSelfAdjointMatrix.toLinearMap)
map_one' := by map_one' := by
@ -120,4 +120,4 @@ Complete this section.
end end
end SL2C end SL2C
end spaceTime end SpaceTime

View file

@ -26,7 +26,7 @@ open Complex
open ComplexConjugate open ComplexConjugate
/-- The global gauge group of the standard model. TODO: Generalize to quotient. -/ /-- The global gauge group of the standard model. TODO: Generalize to quotient. -/
abbrev gaugeGroup : Type := abbrev GaugeGroup : Type :=
specialUnitaryGroup (Fin 3) × specialUnitaryGroup (Fin 2) × unitary specialUnitaryGroup (Fin 3) × specialUnitaryGroup (Fin 2) × unitary
end StandardModel end StandardModel

View file

@ -34,85 +34,85 @@ open Manifold
open Matrix open Matrix
open Complex open Complex
open ComplexConjugate open ComplexConjugate
open spaceTime open SpaceTime
/-- The trivial vector bundle 𝓡² × ℂ². (TODO: Make associated bundle.) -/ /-- The trivial vector bundle 𝓡² × ℂ². (TODO: Make associated bundle.) -/
abbrev higgsBundle := Bundle.Trivial spaceTime higgsVec abbrev HiggsBundle := Bundle.Trivial SpaceTime HiggsVec
instance : SmoothVectorBundle higgsVec higgsBundle spaceTime.asSmoothManifold := instance : SmoothVectorBundle HiggsVec HiggsBundle SpaceTime.asSmoothManifold :=
Bundle.Trivial.smoothVectorBundle higgsVec 𝓘(, spaceTime) Bundle.Trivial.smoothVectorBundle HiggsVec 𝓘(, SpaceTime)
/-- A higgs field is a smooth section of the higgs bundle. -/ /-- A higgs field is a smooth section of the higgs bundle. -/
abbrev higgsField : Type := SmoothSection spaceTime.asSmoothManifold higgsVec higgsBundle abbrev HiggsField : Type := SmoothSection SpaceTime.asSmoothManifold HiggsVec HiggsBundle
instance : NormedAddCommGroup (Fin 2 → ) := by instance : NormedAddCommGroup (Fin 2 → ) := by
exact Pi.normedAddCommGroup exact Pi.normedAddCommGroup
/-- Given a vector `ℂ²` the constant higgs field with value equal to that /-- Given a vector `ℂ²` the constant higgs field with value equal to that
section. -/ section. -/
noncomputable def higgsVec.toField (φ : higgsVec) : higgsField where noncomputable def HiggsVec.toField (φ : HiggsVec) : HiggsField where
toFun := fun _ ↦ φ toFun := fun _ ↦ φ
contMDiff_toFun := by contMDiff_toFun := by
intro x intro x
rw [Bundle.contMDiffAt_section] rw [Bundle.contMDiffAt_section]
exact smoothAt_const exact smoothAt_const
namespace higgsField namespace HiggsField
open Complex Real open Complex Real
/-- Given a `higgsField`, the corresponding map from `spaceTime` to `higgsVec`. -/ /-- Given a `higgsField`, the corresponding map from `spaceTime` to `higgsVec`. -/
def toHiggsVec (φ : higgsField) : spaceTime → higgsVec := φ def toHiggsVec (φ : HiggsField) : SpaceTime → HiggsVec := φ
lemma toHiggsVec_smooth (φ : higgsField) : Smooth 𝓘(, spaceTime) 𝓘(, higgsVec) φ.toHiggsVec := by lemma toHiggsVec_smooth (φ : HiggsField) : Smooth 𝓘(, SpaceTime) 𝓘(, HiggsVec) φ.toHiggsVec := by
intro x0 intro x0
have h1 := φ.contMDiff x0 have h1 := φ.contMDiff x0
rw [Bundle.contMDiffAt_section] at h1 rw [Bundle.contMDiffAt_section] at h1
have h2 : have h2 :
(fun x => ((trivializationAt higgsVec (Bundle.Trivial spaceTime higgsVec) x0) (fun x => ((trivializationAt HiggsVec (Bundle.Trivial SpaceTime HiggsVec) x0)
{ proj := x, snd := φ x }).2) = φ := by { proj := x, snd := φ x }).2) = φ := by
rfl rfl
simp only [h2] at h1 simp only [h2] at h1
exact h1 exact h1
lemma toField_toHiggsVec_apply (φ : higgsField) (x : spaceTime) : lemma toField_toHiggsVec_apply (φ : HiggsField) (x : SpaceTime) :
(φ.toHiggsVec x).toField x = φ x := rfl (φ.toHiggsVec x).toField x = φ x := rfl
lemma higgsVecToFin2_toHiggsVec (φ : higgsField) : lemma higgsVecToFin2_toHiggsVec (φ : HiggsField) :
higgsVecToFin2 ∘ φ.toHiggsVec = φ := rfl higgsVecToFin2 ∘ φ.toHiggsVec = φ := rfl
lemma toVec_smooth (φ : higgsField) : Smooth 𝓘(, spaceTime) 𝓘(, Fin 2 → ) φ := lemma toVec_smooth (φ : HiggsField) : Smooth 𝓘(, SpaceTime) 𝓘(, Fin 2 → ) φ :=
smooth_higgsVecToFin2.comp φ.toHiggsVec_smooth smooth_higgsVecToFin2.comp φ.toHiggsVec_smooth
lemma apply_smooth (φ : higgsField) : lemma apply_smooth (φ : HiggsField) :
∀ i, Smooth 𝓘(, spaceTime) 𝓘(, ) (fun (x : spaceTime) => (φ x i)) := ∀ i, Smooth 𝓘(, SpaceTime) 𝓘(, ) (fun (x : SpaceTime) => (φ x i)) :=
(smooth_pi_space).mp (φ.toVec_smooth) (smooth_pi_space).mp (φ.toVec_smooth)
lemma apply_re_smooth (φ : higgsField) (i : Fin 2): lemma apply_re_smooth (φ : HiggsField) (i : Fin 2):
Smooth 𝓘(, spaceTime) 𝓘(, ) (reCLM ∘ (fun (x : spaceTime) => (φ x i))) := Smooth 𝓘(, SpaceTime) 𝓘(, ) (reCLM ∘ (fun (x : SpaceTime) => (φ x i))) :=
(ContinuousLinearMap.smooth reCLM).comp (φ.apply_smooth i) (ContinuousLinearMap.smooth reCLM).comp (φ.apply_smooth i)
lemma apply_im_smooth (φ : higgsField) (i : Fin 2): lemma apply_im_smooth (φ : HiggsField) (i : Fin 2):
Smooth 𝓘(, spaceTime) 𝓘(, ) (imCLM ∘ (fun (x : spaceTime) => (φ x i))) := Smooth 𝓘(, SpaceTime) 𝓘(, ) (imCLM ∘ (fun (x : SpaceTime) => (φ x i))) :=
(ContinuousLinearMap.smooth imCLM).comp (φ.apply_smooth i) (ContinuousLinearMap.smooth imCLM).comp (φ.apply_smooth i)
/-- Given two `higgsField`, the map `spaceTime → ` obtained by taking their inner product. -/ /-- Given two `higgsField`, the map `spaceTime → ` obtained by taking their inner product. -/
def innerProd (φ1 φ2 : higgsField) : spaceTime → := fun x => ⟪φ1 x, φ2 x⟫_ def innerProd (φ1 φ2 : HiggsField) : SpaceTime → := fun x => ⟪φ1 x, φ2 x⟫_
/-- Given a `higgsField`, the map `spaceTime → ` obtained by taking the square norm of the /-- Given a `higgsField`, the map `spaceTime → ` obtained by taking the square norm of the
higgs vector. -/ higgs vector. -/
@[simp] @[simp]
def normSq (φ : higgsField) : spaceTime → := fun x => ( ‖φ x‖ ^ 2) def normSq (φ : HiggsField) : SpaceTime → := fun x => ( ‖φ x‖ ^ 2)
lemma toHiggsVec_norm (φ : higgsField) (x : spaceTime) : lemma toHiggsVec_norm (φ : HiggsField) (x : SpaceTime) :
‖φ x‖ = ‖φ.toHiggsVec x‖ := rfl ‖φ x‖ = ‖φ.toHiggsVec x‖ := rfl
lemma normSq_expand (φ : higgsField) : lemma normSq_expand (φ : HiggsField) :
φ.normSq = fun x => (conj (φ x 0) * (φ x 0) + conj (φ x 1) * (φ x 1) ).re := by φ.normSq = fun x => (conj (φ x 0) * (φ x 0) + conj (φ x 1) * (φ x 1) ).re := by
funext x funext x
simp [normSq, add_re, mul_re, conj_re, conj_im, neg_mul, sub_neg_eq_add, @norm_sq_eq_inner ] simp [normSq, add_re, mul_re, conj_re, conj_im, neg_mul, sub_neg_eq_add, @norm_sq_eq_inner ]
lemma normSq_smooth (φ : higgsField) : Smooth 𝓘(, spaceTime) 𝓘(, ) φ.normSq := by lemma normSq_smooth (φ : HiggsField) : Smooth 𝓘(, SpaceTime) 𝓘(, ) φ.normSq := by
rw [normSq_expand] rw [normSq_expand]
refine Smooth.add ?_ ?_ refine Smooth.add ?_ ?_
simp only [mul_re, conj_re, conj_im, neg_mul, sub_neg_eq_add] simp only [mul_re, conj_re, conj_im, neg_mul, sub_neg_eq_add]
@ -132,50 +132,50 @@ lemma normSq_smooth (φ : higgsField) : Smooth 𝓘(, spaceTime) 𝓘(,
exact φ.apply_im_smooth 1 exact φ.apply_im_smooth 1
exact φ.apply_im_smooth 1 exact φ.apply_im_smooth 1
lemma normSq_nonneg (φ : higgsField) (x : spaceTime) : 0 ≤ φ.normSq x := by lemma normSq_nonneg (φ : HiggsField) (x : SpaceTime) : 0 ≤ φ.normSq x := by
simp [normSq, ge_iff_le, norm_nonneg, pow_nonneg] simp [normSq, ge_iff_le, norm_nonneg, pow_nonneg]
lemma normSq_zero (φ : higgsField) (x : spaceTime) : φ.normSq x = 0 ↔ φ x = 0 := by lemma normSq_zero (φ : HiggsField) (x : SpaceTime) : φ.normSq x = 0 ↔ φ x = 0 := by
simp [normSq, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, norm_eq_zero] simp [normSq, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, norm_eq_zero]
/-- The Higgs potential of the form `- μ² * |φ|² + λ * |φ|⁴`. -/ /-- The Higgs potential of the form `- μ² * |φ|² + λ * |φ|⁴`. -/
@[simp] @[simp]
def potential (φ : higgsField) (μSq lambda : ) (x : spaceTime) : := def potential (φ : HiggsField) (μSq lambda : ) (x : SpaceTime) : :=
- μSq * φ.normSq x + lambda * φ.normSq x * φ.normSq x - μSq * φ.normSq x + lambda * φ.normSq x * φ.normSq x
lemma potential_smooth (φ : higgsField) (μSq lambda : ) : lemma potential_smooth (φ : HiggsField) (μSq lambda : ) :
Smooth 𝓘(, spaceTime) 𝓘(, ) (fun x => φ.potential μSq lambda x) := by Smooth 𝓘(, SpaceTime) 𝓘(, ) (fun x => φ.potential μSq lambda x) := by
simp only [potential, normSq, neg_mul] simp only [potential, normSq, neg_mul]
exact (smooth_const.smul φ.normSq_smooth).neg.add exact (smooth_const.smul φ.normSq_smooth).neg.add
((smooth_const.smul φ.normSq_smooth).smul φ.normSq_smooth) ((smooth_const.smul φ.normSq_smooth).smul φ.normSq_smooth)
lemma potential_apply (φ : higgsField) (μSq lambda : ) (x : spaceTime) : lemma potential_apply (φ : HiggsField) (μSq lambda : ) (x : SpaceTime) :
(φ.potential μSq lambda) x = higgsVec.potential μSq lambda (φ.toHiggsVec x) := by (φ.potential μSq lambda) x = HiggsVec.potential μSq lambda (φ.toHiggsVec x) := by
simp [higgsVec.potential, toHiggsVec_norm] simp [HiggsVec.potential, toHiggsVec_norm]
ring ring
/-- A higgs field is constant if it is equal for all `x` `y` in `spaceTime`. -/ /-- A higgs field is constant if it is equal for all `x` `y` in `spaceTime`. -/
def isConst (Φ : higgsField) : Prop := ∀ x y, Φ x = Φ y def IsConst (Φ : HiggsField) : Prop := ∀ x y, Φ x = Φ y
lemma isConst_of_higgsVec (φ : higgsVec) : φ.toField.isConst := by lemma isConst_of_higgsVec (φ : HiggsVec) : φ.toField.IsConst := by
intro x _ intro x _
simp [higgsVec.toField] simp [HiggsVec.toField]
lemma isConst_iff_of_higgsVec (Φ : higgsField) : Φ.isConst ↔ ∃ (φ : higgsVec), Φ = φ.toField := lemma isConst_iff_of_higgsVec (Φ : HiggsField) : Φ.IsConst ↔ ∃ (φ : HiggsVec), Φ = φ.toField :=
Iff.intro (fun h ↦ ⟨Φ 0, by ext x y; rw [← h x 0]; rfl⟩) (fun ⟨φ, hφ⟩ x y ↦ by subst hφ; rfl) Iff.intro (fun h ↦ ⟨Φ 0, by ext x y; rw [← h x 0]; rfl⟩) (fun ⟨φ, hφ⟩ x y ↦ by subst hφ; rfl)
lemma normSq_of_higgsVec (φ : higgsVec) : φ.toField.normSq = fun x => (norm φ) ^ 2 := by lemma normSq_of_higgsVec (φ : HiggsVec) : φ.toField.normSq = fun x => (norm φ) ^ 2 := by
funext x funext x
simp [normSq, higgsVec.toField] simp [normSq, HiggsVec.toField]
lemma potential_of_higgsVec (φ : higgsVec) (μSq lambda : ) : lemma potential_of_higgsVec (φ : HiggsVec) (μSq lambda : ) :
φ.toField.potential μSq lambda = fun _ => higgsVec.potential μSq lambda φ := by φ.toField.potential μSq lambda = fun _ => HiggsVec.potential μSq lambda φ := by
simp [higgsVec.potential] simp [HiggsVec.potential]
unfold potential unfold potential
rw [normSq_of_higgsVec] rw [normSq_of_higgsVec]
ring_nf ring_nf
end higgsField end HiggsField
end end
end StandardModel end StandardModel

View file

@ -37,25 +37,24 @@ open Complex
open ComplexConjugate open ComplexConjugate
/-- The complex vector space in which the Higgs field takes values. -/ /-- The complex vector space in which the Higgs field takes values. -/
abbrev higgsVec := EuclideanSpace (Fin 2) abbrev HiggsVec := EuclideanSpace (Fin 2)
section higgsVec
/-- The continuous linear map from the vector space `higgsVec` to `(Fin 2 → )` achieved by /-- The continuous linear map from the vector space `higgsVec` to `(Fin 2 → )` achieved by
casting vectors. -/ casting vectors. -/
def higgsVecToFin2 : higgsVec →L[] (Fin 2 → ) where def higgsVecToFin2 : HiggsVec →L[] (Fin 2 → ) where
toFun x := x toFun x := x
map_add' x y := by simp map_add' x y := by simp
map_smul' a x := by simp map_smul' a x := by simp
lemma smooth_higgsVecToFin2 : Smooth 𝓘(, higgsVec) 𝓘(, Fin 2 → ) higgsVecToFin2 := lemma smooth_higgsVecToFin2 : Smooth 𝓘(, HiggsVec) 𝓘(, Fin 2 → ) higgsVecToFin2 :=
ContinuousLinearMap.smooth higgsVecToFin2 ContinuousLinearMap.smooth higgsVecToFin2
namespace higgsVec namespace HiggsVec
/-- The Higgs representation as a homomorphism from the gauge group to unitary `2×2`-matrices. -/ /-- The Higgs representation as a homomorphism from the gauge group to unitary `2×2`-matrices. -/
@[simps!] @[simps!]
noncomputable def higgsRepUnitary : gaugeGroup →* unitaryGroup (Fin 2) where noncomputable def higgsRepUnitary : GaugeGroup →* unitaryGroup (Fin 2) where
toFun g := repU1 g.2.2 * fundamentalSU2 g.2.1 toFun g := repU1 g.2.2 * fundamentalSU2 g.2.1
map_mul' := by map_mul' := by
intro ⟨_, a2, a3⟩ ⟨_, b2, b3⟩ intro ⟨_, a2, a3⟩ ⟨_, b2, b3⟩
@ -66,11 +65,11 @@ noncomputable def higgsRepUnitary : gaugeGroup →* unitaryGroup (Fin 2) whe
map_one' := by simp map_one' := by simp
/-- An orthonormal basis of higgsVec. -/ /-- An orthonormal basis of higgsVec. -/
noncomputable def orthonormBasis : OrthonormalBasis (Fin 2) higgsVec := noncomputable def orthonormBasis : OrthonormalBasis (Fin 2) HiggsVec :=
EuclideanSpace.basisFun (Fin 2) EuclideanSpace.basisFun (Fin 2)
/-- Takes in a `2×2`-matrix and returns a linear map of `higgsVec`. -/ /-- Takes in a `2×2`-matrix and returns a linear map of `higgsVec`. -/
noncomputable def matrixToLin : Matrix (Fin 2) (Fin 2) →* (higgsVec →L[] higgsVec) where noncomputable def matrixToLin : Matrix (Fin 2) (Fin 2) →* (HiggsVec →L[] HiggsVec) where
toFun g := LinearMap.toContinuousLinearMap toFun g := LinearMap.toContinuousLinearMap
$ Matrix.toLin orthonormBasis.toBasis orthonormBasis.toBasis g $ Matrix.toLin orthonormBasis.toBasis orthonormBasis.toBasis g
map_mul' g h := ContinuousLinearMap.coe_inj.mp $ map_mul' g h := ContinuousLinearMap.coe_inj.mp $
@ -82,36 +81,36 @@ lemma matrixToLin_star (g : Matrix (Fin 2) (Fin 2) ) :
ContinuousLinearMap.coe_inj.mp $ Matrix.toLin_conjTranspose orthonormBasis orthonormBasis g ContinuousLinearMap.coe_inj.mp $ Matrix.toLin_conjTranspose orthonormBasis orthonormBasis g
lemma matrixToLin_unitary (g : unitaryGroup (Fin 2) ) : lemma matrixToLin_unitary (g : unitaryGroup (Fin 2) ) :
matrixToLin g ∈ unitary (higgsVec →L[] higgsVec) := by matrixToLin g ∈ unitary (HiggsVec →L[] HiggsVec) := by
rw [@unitary.mem_iff, ← matrixToLin_star, ← matrixToLin.map_mul, ← matrixToLin.map_mul, rw [@unitary.mem_iff, ← matrixToLin_star, ← matrixToLin.map_mul, ← matrixToLin.map_mul,
mem_unitaryGroup_iff.mp g.prop, mem_unitaryGroup_iff'.mp g.prop, matrixToLin.map_one] mem_unitaryGroup_iff.mp g.prop, mem_unitaryGroup_iff'.mp g.prop, matrixToLin.map_one]
simp simp
/-- The natural homomorphism from unitary `2×2` complex matrices to unitary transformations /-- The natural homomorphism from unitary `2×2` complex matrices to unitary transformations
of `higgsVec`. -/ of `higgsVec`. -/
noncomputable def unitaryToLin : unitaryGroup (Fin 2) →* unitary (higgsVec →L[] higgsVec) where noncomputable def unitaryToLin : unitaryGroup (Fin 2) →* unitary (HiggsVec →L[] HiggsVec) where
toFun g := ⟨matrixToLin g, matrixToLin_unitary g⟩ toFun g := ⟨matrixToLin g, matrixToLin_unitary g⟩
map_mul' g h := by simp map_mul' g h := by simp
map_one' := by simp map_one' := by simp
/-- The inclusion of unitary transformations on `higgsVec` into all linear transformations. -/ /-- The inclusion of unitary transformations on `higgsVec` into all linear transformations. -/
@[simps!] @[simps!]
def unitToLinear : unitary (higgsVec →L[] higgsVec) →* higgsVec →ₗ[] higgsVec := def unitToLinear : unitary (HiggsVec →L[] HiggsVec) →* HiggsVec →ₗ[] HiggsVec :=
DistribMulAction.toModuleEnd higgsVec DistribMulAction.toModuleEnd HiggsVec
/-- The representation of the gauge group acting on `higgsVec`. -/ /-- The representation of the gauge group acting on `higgsVec`. -/
@[simps!] @[simps!]
def rep : Representation gaugeGroup higgsVec := def rep : Representation GaugeGroup HiggsVec :=
unitToLinear.comp (unitaryToLin.comp higgsRepUnitary) unitToLinear.comp (unitaryToLin.comp higgsRepUnitary)
lemma higgsRepUnitary_mul (g : gaugeGroup) (φ : higgsVec) : lemma higgsRepUnitary_mul (g : GaugeGroup) (φ : HiggsVec) :
(higgsRepUnitary g).1 *ᵥ φ = g.2.2 ^ 3 • (g.2.1.1 *ᵥ φ) := by (higgsRepUnitary g).1 *ᵥ φ = g.2.2 ^ 3 • (g.2.1.1 *ᵥ φ) := by
simp [higgsRepUnitary_apply_coe, smul_mulVec_assoc] simp [higgsRepUnitary_apply_coe, smul_mulVec_assoc]
lemma rep_apply (g : gaugeGroup) (φ : higgsVec) : rep g φ = g.2.2 ^ 3 • (g.2.1.1 *ᵥ φ) := lemma rep_apply (g : GaugeGroup) (φ : HiggsVec) : rep g φ = g.2.2 ^ 3 • (g.2.1.1 *ᵥ φ) :=
higgsRepUnitary_mul g φ higgsRepUnitary_mul g φ
lemma norm_invariant (g : gaugeGroup) (φ : higgsVec) : ‖rep g φ‖ = ‖φ‖ := lemma norm_invariant (g : GaugeGroup) (φ : HiggsVec) : ‖rep g φ‖ = ‖φ‖ :=
ContinuousLinearMap.norm_map_of_mem_unitary (unitaryToLin (higgsRepUnitary g)).2 φ ContinuousLinearMap.norm_map_of_mem_unitary (unitaryToLin (higgsRepUnitary g)).2 φ
section potentialDefn section potentialDefn
@ -120,13 +119,13 @@ variable (μSq lambda : )
local notation "λ" => lambda local notation "λ" => lambda
/-- The higgs potential for `higgsVec`, i.e. for constant higgs fields. -/ /-- The higgs potential for `higgsVec`, i.e. for constant higgs fields. -/
def potential (φ : higgsVec) : := - μSq * ‖φ‖ ^ 2 + λ * ‖φ‖ ^ 4 def potential (φ : HiggsVec) : := - μSq * ‖φ‖ ^ 2 + λ * ‖φ‖ ^ 4
lemma potential_invariant (φ : higgsVec) (g : gaugeGroup) : lemma potential_invariant (φ : HiggsVec) (g : GaugeGroup) :
potential μSq (λ) (rep g φ) = potential μSq (λ) φ := by potential μSq (λ) (rep g φ) = potential μSq (λ) φ := by
simp only [potential, neg_mul, norm_invariant] simp only [potential, neg_mul, norm_invariant]
lemma potential_as_quad (φ : higgsVec) : lemma potential_as_quad (φ : HiggsVec) :
λ * ‖φ‖ ^ 2 * ‖φ‖ ^ 2 + (- μSq ) * ‖φ‖ ^ 2 + (- potential μSq (λ) φ) = 0 := by λ * ‖φ‖ ^ 2 * ‖φ‖ ^ 2 + (- μSq ) * ‖φ‖ ^ 2 + (- potential μSq (λ) φ) = 0 := by
simp [potential]; ring simp [potential]; ring
@ -138,7 +137,7 @@ variable (μSq : )
variable (hLam : 0 < lambda) variable (hLam : 0 < lambda)
local notation "λ" => lambda local notation "λ" => lambda
lemma potential_snd_term_nonneg (φ : higgsVec) : lemma potential_snd_term_nonneg (φ : HiggsVec) :
0 ≤ λ * ‖φ‖ ^ 4 := by 0 ≤ λ * ‖φ‖ ^ 4 := by
rw [mul_nonneg_iff] rw [mul_nonneg_iff]
apply Or.inl apply Or.inl
@ -146,7 +145,7 @@ lemma potential_snd_term_nonneg (φ : higgsVec) :
exact le_of_lt hLam exact le_of_lt hLam
lemma zero_le_potential_discrim (φ : higgsVec) : lemma zero_le_potential_discrim (φ : HiggsVec) :
0 ≤ discrim (λ) (- μSq ) (- potential μSq (λ) φ) := by 0 ≤ discrim (λ) (- μSq ) (- potential μSq (λ) φ) := by
have h1 := potential_as_quad μSq (λ) φ have h1 := potential_as_quad μSq (λ) φ
rw [quadratic_eq_zero_iff_discrim_eq_sq] at h1 rw [quadratic_eq_zero_iff_discrim_eq_sq] at h1
@ -154,7 +153,7 @@ lemma zero_le_potential_discrim (φ : higgsVec) :
exact sq_nonneg (2 * lambda * ‖φ‖ ^ 2 + -μSq) exact sq_nonneg (2 * lambda * ‖φ‖ ^ 2 + -μSq)
· exact ne_of_gt hLam · exact ne_of_gt hLam
lemma potential_eq_zero_sol (φ : higgsVec) lemma potential_eq_zero_sol (φ : HiggsVec)
(hV : potential μSq (λ) φ = 0) : φ = 0 ‖φ‖ ^ 2 = μSq / λ := by (hV : potential μSq (λ) φ = 0) : φ = 0 ‖φ‖ ^ 2 = μSq / λ := by
have h1 := potential_as_quad μSq (λ) φ have h1 := potential_as_quad μSq (λ) φ
rw [hV] at h1 rw [hV] at h1
@ -169,7 +168,7 @@ lemma potential_eq_zero_sol (φ : higgsVec)
linear_combination h2 linear_combination h2
lemma potential_eq_zero_sol_of_μSq_nonpos (hμSq : μSq ≤ 0) lemma potential_eq_zero_sol_of_μSq_nonpos (hμSq : μSq ≤ 0)
(φ : higgsVec) (hV : potential μSq (λ) φ = 0) : φ = 0 := by (φ : HiggsVec) (hV : potential μSq (λ) φ = 0) : φ = 0 := by
cases' (potential_eq_zero_sol μSq hLam φ hV) with h1 h1 cases' (potential_eq_zero_sol μSq hLam φ hV) with h1 h1
exact h1 exact h1
by_cases hμSqZ : μSq = 0 by_cases hμSqZ : μSq = 0
@ -181,7 +180,7 @@ lemma potential_eq_zero_sol_of_μSq_nonpos (hμSq : μSq ≤ 0)
· rw [← h1] · rw [← h1]
exact sq_nonneg ‖φ‖ exact sq_nonneg ‖φ‖
lemma potential_bounded_below (φ : higgsVec) : lemma potential_bounded_below (φ : HiggsVec) :
- μSq ^ 2 / (4 * (λ)) ≤ potential μSq (λ) φ := by - μSq ^ 2 / (4 * (λ)) ≤ potential μSq (λ) φ := by
have h1 := zero_le_potential_discrim μSq hLam φ have h1 := zero_le_potential_discrim μSq hLam φ
simp [discrim] at h1 simp [discrim] at h1
@ -195,17 +194,17 @@ lemma potential_bounded_below (φ : higgsVec) :
exact h2 exact h2
lemma potential_bounded_below_of_μSq_nonpos {μSq : } lemma potential_bounded_below_of_μSq_nonpos {μSq : }
(hμSq : μSq ≤ 0) (φ : higgsVec) : 0 ≤ potential μSq (λ) φ := by (hμSq : μSq ≤ 0) (φ : HiggsVec) : 0 ≤ potential μSq (λ) φ := by
refine add_nonneg ?_ (potential_snd_term_nonneg hLam φ) refine add_nonneg ?_ (potential_snd_term_nonneg hLam φ)
field_simp [mul_nonpos_iff] field_simp [mul_nonpos_iff]
simp_all [ge_iff_le, norm_nonneg, pow_nonneg, and_self, or_true] simp_all [ge_iff_le, norm_nonneg, pow_nonneg, and_self, or_true]
lemma potential_eq_bound_discrim_zero (φ : higgsVec) lemma potential_eq_bound_discrim_zero (φ : HiggsVec)
(hV : potential μSq (λ) φ = - μSq ^ 2 / (4 * λ)) : (hV : potential μSq (λ) φ = - μSq ^ 2 / (4 * λ)) :
discrim (λ) (- μSq) (- potential μSq (λ) φ) = 0 := by discrim (λ) (- μSq) (- potential μSq (λ) φ) = 0 := by
field_simp [discrim, hV] field_simp [discrim, hV]
lemma potential_eq_bound_higgsVec_sq (φ : higgsVec) lemma potential_eq_bound_higgsVec_sq (φ : HiggsVec)
(hV : potential μSq (λ) φ = - μSq ^ 2 / (4 * (λ))) : (hV : potential μSq (λ) φ = - μSq ^ 2 / (4 * (λ))) :
‖φ‖ ^ 2 = μSq / (2 * λ) := by ‖φ‖ ^ 2 = μSq / (2 * λ) := by
have h1 := potential_as_quad μSq (λ) φ have h1 := potential_as_quad μSq (λ) φ
@ -214,7 +213,7 @@ lemma potential_eq_bound_higgsVec_sq (φ : higgsVec)
simp_rw [h1, neg_neg] simp_rw [h1, neg_neg]
exact ne_of_gt hLam exact ne_of_gt hLam
lemma potential_eq_bound_iff (φ : higgsVec) : lemma potential_eq_bound_iff (φ : HiggsVec) :
potential μSq (λ) φ = - μSq ^ 2 / (4 * (λ)) ↔ ‖φ‖ ^ 2 = μSq / (2 * (λ)) := potential μSq (λ) φ = - μSq ^ 2 / (4 * (λ)) ↔ ‖φ‖ ^ 2 = μSq / (2 * (λ)) :=
Iff.intro (potential_eq_bound_higgsVec_sq μSq hLam φ) Iff.intro (potential_eq_bound_higgsVec_sq μSq hLam φ)
(fun h ↦ by (fun h ↦ by
@ -223,24 +222,24 @@ lemma potential_eq_bound_iff (φ : higgsVec) :
ring_nf) ring_nf)
lemma potential_eq_bound_iff_of_μSq_nonpos {μSq : } lemma potential_eq_bound_iff_of_μSq_nonpos {μSq : }
(hμSq : μSq ≤ 0) (φ : higgsVec) : potential μSq (λ) φ = 0 ↔ φ = 0 := (hμSq : μSq ≤ 0) (φ : HiggsVec) : potential μSq (λ) φ = 0 ↔ φ = 0 :=
Iff.intro (fun h ↦ potential_eq_zero_sol_of_μSq_nonpos μSq hLam hμSq φ h) Iff.intro (fun h ↦ potential_eq_zero_sol_of_μSq_nonpos μSq hLam hμSq φ h)
(fun h ↦ by simp [potential, h]) (fun h ↦ by simp [potential, h])
lemma potential_eq_bound_IsMinOn (φ : higgsVec) lemma potential_eq_bound_IsMinOn (φ : HiggsVec)
(hv : potential μSq lambda φ = - μSq ^ 2 / (4 * lambda)) : (hv : potential μSq lambda φ = - μSq ^ 2 / (4 * lambda)) :
IsMinOn (potential μSq lambda) Set.univ φ := by IsMinOn (potential μSq lambda) Set.univ φ := by
rw [isMinOn_univ_iff, hv] rw [isMinOn_univ_iff, hv]
exact fun x ↦ potential_bounded_below μSq hLam x exact fun x ↦ potential_bounded_below μSq hLam x
lemma potential_eq_bound_IsMinOn_of_μSq_nonpos {μSq : } lemma potential_eq_bound_IsMinOn_of_μSq_nonpos {μSq : }
(hμSq : μSq ≤ 0) (φ : higgsVec) (hv : potential μSq lambda φ = 0) : (hμSq : μSq ≤ 0) (φ : HiggsVec) (hv : potential μSq lambda φ = 0) :
IsMinOn (potential μSq lambda) Set.univ φ := by IsMinOn (potential μSq lambda) Set.univ φ := by
rw [isMinOn_univ_iff, hv] rw [isMinOn_univ_iff, hv]
exact fun x ↦ potential_bounded_below_of_μSq_nonpos hLam hμSq x exact fun x ↦ potential_bounded_below_of_μSq_nonpos hLam hμSq x
lemma potential_bound_reached_of_μSq_nonneg {μSq : } (hμSq : 0 ≤ μSq) : lemma potential_bound_reached_of_μSq_nonneg {μSq : } (hμSq : 0 ≤ μSq) :
∃ (φ : higgsVec), potential μSq lambda φ = - μSq ^ 2 / (4 * lambda) := by ∃ (φ : HiggsVec), potential μSq lambda φ = - μSq ^ 2 / (4 * lambda) := by
use ![√(μSq/(2 * lambda)), 0] use ![√(μSq/(2 * lambda)), 0]
refine (potential_eq_bound_iff μSq hLam _).mpr ?_ refine (potential_eq_bound_iff μSq hLam _).mpr ?_
simp [PiLp.norm_sq_eq_of_L2] simp [PiLp.norm_sq_eq_of_L2]
@ -269,24 +268,24 @@ lemma IsMinOn_potential_iff_of_μSq_nonpos {μSq : } (hμSq : μSq ≤ 0) :
end potentialProp end potentialProp
/-- Given a Higgs vector, a rotation matrix which puts the first component of the /-- Given a Higgs vector, a rotation matrix which puts the first component of the
vector to zero, and the second component to a real -/ vector to zero, and the second component to a real -/
def rotateMatrix (φ : higgsVec) : Matrix (Fin 2) (Fin 2) := def rotateMatrix (φ : HiggsVec) : Matrix (Fin 2) (Fin 2) :=
![![φ 1 /‖φ‖ , - φ 0 /‖φ‖], ![conj (φ 0) / ‖φ‖ , conj (φ 1) / ‖φ‖] ] ![![φ 1 /‖φ‖ , - φ 0 /‖φ‖], ![conj (φ 0) / ‖φ‖ , conj (φ 1) / ‖φ‖] ]
lemma rotateMatrix_star (φ : higgsVec) : lemma rotateMatrix_star (φ : HiggsVec) :
star φ.rotateMatrix = star φ.rotateMatrix =
![![conj (φ 1) /‖φ‖ , φ 0 /‖φ‖], ![- conj (φ 0) / ‖φ‖ , φ 1 / ‖φ‖] ] := by ![![conj (φ 1) /‖φ‖ , φ 0 /‖φ‖], ![- conj (φ 0) / ‖φ‖ , φ 1 / ‖φ‖] ] := by
simp_rw [star, rotateMatrix, conjTranspose] simp_rw [star, rotateMatrix, conjTranspose]
ext i j ext i j
fin_cases i <;> fin_cases j <;> simp [conj_ofReal] fin_cases i <;> fin_cases j <;> simp [conj_ofReal]
lemma rotateMatrix_det {φ : higgsVec} (hφ : φ ≠ 0) : (rotateMatrix φ).det = 1 := by lemma rotateMatrix_det {φ : HiggsVec} (hφ : φ ≠ 0) : (rotateMatrix φ).det = 1 := by
have h1 : (‖φ‖ : ) ≠ 0 := ofReal_inj.mp.mt (norm_ne_zero_iff.mpr hφ) have h1 : (‖φ‖ : ) ≠ 0 := ofReal_inj.mp.mt (norm_ne_zero_iff.mpr hφ)
field_simp [rotateMatrix, det_fin_two] field_simp [rotateMatrix, det_fin_two]
rw [← ofReal_mul, ← sq, ← @real_inner_self_eq_norm_sq] rw [← ofReal_mul, ← sq, ← @real_inner_self_eq_norm_sq]
simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add, simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm, add_comm] Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm, add_comm]
lemma rotateMatrix_unitary {φ : higgsVec} (hφ : φ ≠ 0) : lemma rotateMatrix_unitary {φ : HiggsVec} (hφ : φ ≠ 0) :
(rotateMatrix φ) ∈ unitaryGroup (Fin 2) := by (rotateMatrix φ) ∈ unitaryGroup (Fin 2) := by
rw [mem_unitaryGroup_iff', rotateMatrix_star, rotateMatrix] rw [mem_unitaryGroup_iff', rotateMatrix_star, rotateMatrix]
erw [mul_fin_two, one_fin_two] erw [mul_fin_two, one_fin_two]
@ -301,16 +300,16 @@ lemma rotateMatrix_unitary {φ : higgsVec} (hφ : φ ≠ 0) :
· simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add, · simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm] Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm]
lemma rotateMatrix_specialUnitary {φ : higgsVec} (hφ : φ ≠ 0) : lemma rotateMatrix_specialUnitary {φ : HiggsVec} (hφ : φ ≠ 0) :
(rotateMatrix φ) ∈ specialUnitaryGroup (Fin 2) := (rotateMatrix φ) ∈ specialUnitaryGroup (Fin 2) :=
mem_specialUnitaryGroup_iff.mpr ⟨rotateMatrix_unitary hφ, rotateMatrix_det hφ⟩ mem_specialUnitaryGroup_iff.mpr ⟨rotateMatrix_unitary hφ, rotateMatrix_det hφ⟩
/-- Given a Higgs vector, an element of the gauge group which puts the first component of the /-- Given a Higgs vector, an element of the gauge group which puts the first component of the
vector to zero, and the second component to a real -/ vector to zero, and the second component to a real -/
def rotateGuageGroup {φ : higgsVec} (hφ : φ ≠ 0) : gaugeGroup := def rotateGuageGroup {φ : HiggsVec} (hφ : φ ≠ 0) : GaugeGroup :=
⟨1, ⟨(rotateMatrix φ), rotateMatrix_specialUnitary hφ⟩, 1⟩ ⟨1, ⟨(rotateMatrix φ), rotateMatrix_specialUnitary hφ⟩, 1⟩
lemma rotateGuageGroup_apply {φ : higgsVec} (hφ : φ ≠ 0) : lemma rotateGuageGroup_apply {φ : HiggsVec} (hφ : φ ≠ 0) :
rep (rotateGuageGroup hφ) φ = ![0, ofReal ‖φ‖] := by rep (rotateGuageGroup hφ) φ = ![0, ofReal ‖φ‖] := by
rw [rep_apply] rw [rep_apply]
simp only [rotateGuageGroup, rotateMatrix, one_pow, one_smul, simp only [rotateGuageGroup, rotateMatrix, one_pow, one_smul,
@ -330,8 +329,8 @@ lemma rotateGuageGroup_apply {φ : higgsVec} (hφ : φ ≠ 0) :
simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add, simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm] Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm]
theorem rotate_fst_zero_snd_real (φ : higgsVec) : theorem rotate_fst_zero_snd_real (φ : HiggsVec) :
∃ (g : gaugeGroup), rep g φ = ![0, ofReal ‖φ‖] := by ∃ (g : GaugeGroup), rep g φ = ![0, ofReal ‖φ‖] := by
by_cases h : φ = 0 by_cases h : φ = 0
· use ⟨1, 1, 1⟩ · use ⟨1, 1, 1⟩
simp [h] simp [h]
@ -340,8 +339,7 @@ theorem rotate_fst_zero_snd_real (φ : higgsVec) :
· use rotateGuageGroup h · use rotateGuageGroup h
exact rotateGuageGroup_apply h exact rotateGuageGroup_apply h
end higgsVec end HiggsVec
end higgsVec
end end
end StandardModel end StandardModel

View file

@ -27,6 +27,12 @@ name = "HepLean"
#git = "https://github.com/lean-dojo/LeanCopilot.git" #git = "https://github.com/lean-dojo/LeanCopilot.git"
#rev = "v1.2.2" #rev = "v1.2.2"
# lean_exe commands defined specifically for HepLean.
[[lean_exe]] [[lean_exe]]
name = "check_file_imports" name = "check_file_imports"
srcDir = "scripts" srcDir = "scripts"
[[lean_exe]]
name = "type_former_lint"
srcDir = "scripts"

View file

@ -0,0 +1,50 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
import Batteries.Lean.HashSet
import Lean
/-!
# Type former Lint
This file checks that all definitions which form types start with a capital letter
(or some non-latin character). This linter is not currently strictly enforced
in HepLean.
-/
open Lean System Meta
/-- A rough definition of Upper Camal, checking that if a string starts with a
latin letter then that letter is captial. -/
def IsUpperCamal (s : String) : Bool :=
let parts := s.splitOn "."
let lastPart := parts.get! (parts.length - 1)
lastPart = "noConfusionType"
(¬ (lastPart.get 0 ≥ 'A' ∧ lastPart.get 0 ≤ 'z')) (lastPart.get 0 ≥ 'A' ∧ lastPart.get 0 ≤ 'Z')
def main (_ : List String) : IO UInt32 := do
initSearchPath (← findSysroot)
let mods : Name := `HepLean
let imp : Import := {module := mods}
let mFile ← findOLean imp.module
unless (← mFile.pathExists) do
throw <| IO.userError s!"object file '{mFile}' of module {imp.module} does not exist"
let (hepLeanMod, _) ← readModuleData mFile
let mut warned := false
for imp in hepLeanMod.imports do
let mFile ← findOLean imp.module
let (modData, _) ← readModuleData mFile
let cons := modData.constants
for c in cons do
if c.hasValue ∧ ¬ Lean.Name.isInternal c.name ∧ Lean.Compiler.LCNF.isTypeFormerType c.type
∧ ¬ IsUpperCamal c.name.toString then
warned := true
IO.println s!"name: #{c.name}"
if warned then
throw <| IO.userError s!"There are type former definitions not in upper camal style."
pure 0