2024-04-19 09:57:30 -04:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license.
|
|
|
|
|
Authors: Joseph Tooby-Smith
|
|
|
|
|
-/
|
|
|
|
|
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic
|
|
|
|
|
/-!
|
|
|
|
|
# Plane of non-solutions
|
|
|
|
|
|
|
|
|
|
Working in the three family case, we show that there exists an eleven dimensional plane in the
|
|
|
|
|
vector space of charges on which there are no solutions.
|
|
|
|
|
|
|
|
|
|
The main result of this file is `eleven_dim_plane_of_no_sols_exists`, which states that
|
|
|
|
|
an 11 dimensional plane of charges exists on which there are no solutions except the origin.
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
universe v u
|
|
|
|
|
|
|
|
|
|
namespace SMRHN
|
|
|
|
|
namespace PlusU1
|
|
|
|
|
|
|
|
|
|
open SMνCharges
|
|
|
|
|
open SMνACCs
|
|
|
|
|
open BigOperators
|
|
|
|
|
|
|
|
|
|
namespace ElevenPlane
|
|
|
|
|
|
2024-04-19 10:08:56 -04:00
|
|
|
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def B₀ : (PlusU1 3).Charges := ![1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0]
|
2024-04-19 09:57:30 -04:00
|
|
|
|
|
2024-04-19 10:08:56 -04:00
|
|
|
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def B₁ : (PlusU1 3).Charges := ![0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0]
|
2024-04-19 09:57:30 -04:00
|
|
|
|
|
2024-04-19 10:08:56 -04:00
|
|
|
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def B₂ : (PlusU1 3).Charges := ![0, 0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0]
|
2024-04-19 09:57:30 -04:00
|
|
|
|
|
2024-04-19 10:08:56 -04:00
|
|
|
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def B₃ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0]
|
2024-04-19 09:57:30 -04:00
|
|
|
|
|
2024-04-19 10:08:56 -04:00
|
|
|
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def B₄ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0]
|
2024-04-19 09:57:30 -04:00
|
|
|
|
|
2024-04-19 10:08:56 -04:00
|
|
|
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def B₅ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0]
|
2024-04-19 09:57:30 -04:00
|
|
|
|
|
2024-04-19 10:08:56 -04:00
|
|
|
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def B₆ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0, 0]
|
2024-04-19 09:57:30 -04:00
|
|
|
|
|
2024-04-19 10:08:56 -04:00
|
|
|
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def B₇ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0]
|
2024-04-19 09:57:30 -04:00
|
|
|
|
|
2024-04-19 10:08:56 -04:00
|
|
|
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def B₈ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0]
|
2024-04-19 09:57:30 -04:00
|
|
|
|
|
2024-04-19 10:08:56 -04:00
|
|
|
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def B₉ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 2, 0]
|
2024-04-19 09:57:30 -04:00
|
|
|
|
|
2024-04-19 10:08:56 -04:00
|
|
|
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def B₁₀ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1]
|
2024-04-19 09:57:30 -04:00
|
|
|
|
|
2024-04-19 10:08:56 -04:00
|
|
|
|
/-- The charge assignment forming a basis of the plane. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def B : Fin 11 → (PlusU1 3).Charges := fun i =>
|
2024-04-19 09:57:30 -04:00
|
|
|
|
match i with
|
|
|
|
|
| 0 => B₀
|
|
|
|
|
| 1 => B₁
|
|
|
|
|
| 2 => B₂
|
|
|
|
|
| 3 => B₃
|
|
|
|
|
| 4 => B₄
|
|
|
|
|
| 5 => B₅
|
|
|
|
|
| 6 => B₆
|
|
|
|
|
| 7 => B₇
|
|
|
|
|
| 8 => B₈
|
|
|
|
|
| 9 => B₉
|
|
|
|
|
| 10 => B₁₀
|
|
|
|
|
|
2024-04-22 08:41:50 -04:00
|
|
|
|
lemma Bi_Bj_quad {i j : Fin 11} (hi : i ≠ j) : quadBiLin (B i) (B j) = 0 := by
|
2024-04-19 09:57:30 -04:00
|
|
|
|
fin_cases i <;> fin_cases j
|
|
|
|
|
any_goals rfl
|
|
|
|
|
all_goals simp at hi
|
|
|
|
|
|
|
|
|
|
lemma Bi_sum_quad (i : Fin 11) (f : Fin 11 → ℚ) :
|
2024-04-22 08:41:50 -04:00
|
|
|
|
quadBiLin (B i) (∑ k, f k • B k) = f i * quadBiLin (B i) (B i) := by
|
2024-04-19 09:57:30 -04:00
|
|
|
|
rw [quadBiLin.map_sum₂]
|
|
|
|
|
rw [Fintype.sum_eq_single i]
|
|
|
|
|
rw [quadBiLin.map_smul₂]
|
|
|
|
|
intro k hij
|
|
|
|
|
rw [quadBiLin.map_smul₂, Bi_Bj_quad hij.symm]
|
|
|
|
|
simp
|
|
|
|
|
|
2024-04-19 10:08:56 -04:00
|
|
|
|
/-- The coefficents of the quadratic equation in our basis. -/
|
2024-04-19 09:57:30 -04:00
|
|
|
|
@[simp]
|
|
|
|
|
def quadCoeff : Fin 11 → ℚ := ![1, 1, 1, 1, 1, 1, 1, 1, 1, 0, 0]
|
|
|
|
|
|
2024-04-22 08:41:50 -04:00
|
|
|
|
lemma quadCoeff_eq_bilinear (i : Fin 11) : quadCoeff i = quadBiLin (B i) (B i) := by
|
2024-04-19 09:57:30 -04:00
|
|
|
|
fin_cases i
|
|
|
|
|
all_goals rfl
|
|
|
|
|
|
|
|
|
|
lemma on_accQuad (f : Fin 11 → ℚ) :
|
|
|
|
|
accQuad (∑ i, f i • B i) = ∑ i, quadCoeff i * (f i)^2 := by
|
2024-04-22 08:41:50 -04:00
|
|
|
|
change quadBiLin _ _ = _
|
2024-04-19 09:57:30 -04:00
|
|
|
|
rw [quadBiLin.map_sum₁]
|
|
|
|
|
apply Fintype.sum_congr
|
|
|
|
|
intro i
|
|
|
|
|
rw [quadBiLin.map_smul₁, Bi_sum_quad, quadCoeff_eq_bilinear]
|
|
|
|
|
ring
|
|
|
|
|
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isSolution_quadCoeff_f_sq_zero (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i))
|
2024-04-19 09:57:30 -04:00
|
|
|
|
(k : Fin 11) : quadCoeff k * (f k)^2 = 0 := by
|
|
|
|
|
obtain ⟨S, hS⟩ := hS
|
|
|
|
|
have hQ := quadSol S.1
|
|
|
|
|
rw [hS, on_accQuad] at hQ
|
|
|
|
|
rw [Fintype.sum_eq_zero_iff_of_nonneg] at hQ
|
|
|
|
|
exact congrFun hQ k
|
|
|
|
|
intro i
|
2024-04-19 10:08:56 -04:00
|
|
|
|
simp only [Pi.zero_apply, quadCoeff]
|
2024-04-19 09:57:30 -04:00
|
|
|
|
rw [mul_nonneg_iff]
|
|
|
|
|
apply Or.inl
|
|
|
|
|
apply And.intro
|
|
|
|
|
fin_cases i <;> rfl
|
|
|
|
|
exact sq_nonneg (f i)
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isSolution_f0 (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 0 = 0 := by
|
2024-04-19 09:57:30 -04:00
|
|
|
|
simpa using (isSolution_quadCoeff_f_sq_zero f hS 0)
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isSolution_f1 (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 1 = 0 := by
|
2024-04-19 09:57:30 -04:00
|
|
|
|
simpa using (isSolution_quadCoeff_f_sq_zero f hS 1)
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isSolution_f2 (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 2 = 0 := by
|
2024-04-19 09:57:30 -04:00
|
|
|
|
simpa using (isSolution_quadCoeff_f_sq_zero f hS 2)
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isSolution_f3 (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 3 = 0 := by
|
2024-04-19 09:57:30 -04:00
|
|
|
|
simpa using (isSolution_quadCoeff_f_sq_zero f hS 3)
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isSolution_f4 (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 4 = 0 := by
|
2024-04-19 09:57:30 -04:00
|
|
|
|
simpa using (isSolution_quadCoeff_f_sq_zero f hS 4)
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isSolution_f5 (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 5 = 0 := by
|
2024-04-19 09:57:30 -04:00
|
|
|
|
have h := isSolution_quadCoeff_f_sq_zero f hS 5
|
|
|
|
|
rw [mul_eq_zero] at h
|
|
|
|
|
change 1 = 0 ∨ _ = _ at h
|
|
|
|
|
simpa using h
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isSolution_f6 (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 6 = 0 := by
|
2024-04-19 09:57:30 -04:00
|
|
|
|
have h := isSolution_quadCoeff_f_sq_zero f hS 6
|
|
|
|
|
rw [mul_eq_zero] at h
|
|
|
|
|
change 1 = 0 ∨ _ = _ at h
|
|
|
|
|
simpa using h
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isSolution_f7 (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 7 = 0 := by
|
2024-04-19 09:57:30 -04:00
|
|
|
|
have h := isSolution_quadCoeff_f_sq_zero f hS 7
|
|
|
|
|
rw [mul_eq_zero] at h
|
|
|
|
|
change 1 = 0 ∨ _ = _ at h
|
|
|
|
|
simpa using h
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isSolution_f8 (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 8 = 0 := by
|
2024-04-19 09:57:30 -04:00
|
|
|
|
have h := isSolution_quadCoeff_f_sq_zero f hS 8
|
|
|
|
|
rw [mul_eq_zero] at h
|
|
|
|
|
change 1 = 0 ∨ _ = _ at h
|
|
|
|
|
simpa using h
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isSolution_sum_part (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
|
2024-04-19 09:57:30 -04:00
|
|
|
|
∑ 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]
|
|
|
|
|
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 +
|
|
|
|
|
f 7 • B 7 + f 8 • B 8 + f 9 • B 9 + f 10 • B 10 = f 9 • B₉ + f 10 • B₁₀
|
|
|
|
|
rw [isSolution_f0 f hS, isSolution_f1 f hS, isSolution_f2 f hS, isSolution_f3 f hS,
|
|
|
|
|
isSolution_f4 f hS, isSolution_f5 f hS,
|
|
|
|
|
isSolution_f6 f hS, isSolution_f7 f hS, isSolution_f8 f hS]
|
|
|
|
|
simp
|
|
|
|
|
rfl
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isSolution_grav (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
|
2024-04-19 09:57:30 -04:00
|
|
|
|
f 10 = - 3 * f 9 := by
|
|
|
|
|
have hx := isSolution_sum_part f hS
|
|
|
|
|
obtain ⟨S, hS'⟩ := hS
|
|
|
|
|
have hg := gravSol S.toLinSols
|
|
|
|
|
rw [hS'] at hg
|
|
|
|
|
rw [hx] at hg
|
|
|
|
|
rw [accGrav.map_add, accGrav.map_smul, accGrav.map_smul] at hg
|
|
|
|
|
rw [show accGrav B₉ = 3 by rfl] at hg
|
|
|
|
|
rw [show accGrav B₁₀ = 1 by rfl] at hg
|
|
|
|
|
simp at hg
|
|
|
|
|
linear_combination hg
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isSolution_sum_part' (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
|
2024-04-19 09:57:30 -04:00
|
|
|
|
∑ i, f i • B i = f 9 • B₉ + (- 3 * f 9) • B₁₀ := by
|
|
|
|
|
rw [isSolution_sum_part f hS]
|
|
|
|
|
rw [isSolution_grav f hS]
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isSolution_f9 (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
|
2024-04-19 09:57:30 -04:00
|
|
|
|
f 9 = 0 := by
|
|
|
|
|
have hx := isSolution_sum_part' f hS
|
|
|
|
|
obtain ⟨S, hS'⟩ := hS
|
|
|
|
|
have hc := cubeSol S
|
|
|
|
|
rw [hS'] at hc
|
|
|
|
|
rw [hx] at hc
|
|
|
|
|
change cubeTriLin.toCubic _ = _ at hc
|
|
|
|
|
rw [cubeTriLin.toCubic_add] at hc
|
|
|
|
|
erw [accCube.map_smul] at hc
|
|
|
|
|
erw [accCube.map_smul (- 3 * f 9) B₁₀] at hc
|
|
|
|
|
rw [cubeTriLin.map_smul₁, cubeTriLin.map_smul₁, cubeTriLin.map_smul₂, cubeTriLin.map_smul₂,
|
|
|
|
|
cubeTriLin.map_smul₃, cubeTriLin.map_smul₃] at hc
|
|
|
|
|
rw [show accCube B₉ = 9 by rfl] at hc
|
|
|
|
|
rw [show accCube B₁₀ = 1 by rfl] at hc
|
2024-04-22 09:48:44 -04:00
|
|
|
|
rw [show cubeTriLin B₉ B₉ B₁₀ = 0 by rfl] at hc
|
|
|
|
|
rw [show cubeTriLin B₁₀ B₁₀ B₉ = 0 by rfl] at hc
|
2024-04-19 09:57:30 -04:00
|
|
|
|
simp at hc
|
|
|
|
|
have h1 : f 9 ^ 3 * 9 + (-(3 * f 9)) ^ 3 = - 18 * f 9 ^ 3 := by
|
|
|
|
|
ring
|
|
|
|
|
rw [h1] at hc
|
|
|
|
|
simpa using hc
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isSolution_f10 (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
|
2024-04-19 09:57:30 -04:00
|
|
|
|
f 10 = 0 := by
|
|
|
|
|
rw [isSolution_grav f hS, isSolution_f9 f hS]
|
|
|
|
|
simp
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isSolution_f_zero (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i))
|
2024-04-19 09:57:30 -04:00
|
|
|
|
(k : Fin 11) : f k = 0 := by
|
|
|
|
|
fin_cases k
|
|
|
|
|
exact isSolution_f0 f hS
|
|
|
|
|
exact isSolution_f1 f hS
|
|
|
|
|
exact isSolution_f2 f hS
|
|
|
|
|
exact isSolution_f3 f hS
|
|
|
|
|
exact isSolution_f4 f hS
|
|
|
|
|
exact isSolution_f5 f hS
|
|
|
|
|
exact isSolution_f6 f hS
|
|
|
|
|
exact isSolution_f7 f hS
|
|
|
|
|
exact isSolution_f8 f hS
|
|
|
|
|
exact isSolution_f9 f hS
|
|
|
|
|
exact isSolution_f10 f hS
|
|
|
|
|
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma isSolution_only_if_zero (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
|
2024-04-19 09:57:30 -04:00
|
|
|
|
∑ i, f i • B i = 0 := by
|
|
|
|
|
rw [isSolution_sum_part f hS]
|
|
|
|
|
rw [isSolution_grav f hS]
|
|
|
|
|
rw [isSolution_f9 f hS]
|
|
|
|
|
simp
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
theorem basis_linear_independent : LinearIndependent ℚ B := by
|
|
|
|
|
apply Fintype.linearIndependent_iff.mpr
|
|
|
|
|
intro f h
|
|
|
|
|
let X : (PlusU1 3).Sols := chargeToAF 0 (by rfl) (by rfl) (by rfl) (by rfl) (by rfl) (by rfl)
|
2024-06-26 11:54:02 -04:00
|
|
|
|
have hS : (PlusU1 3).IsSolution (∑ i, f i • B i) := by
|
2024-04-19 09:57:30 -04:00
|
|
|
|
use X
|
|
|
|
|
rw [h]
|
|
|
|
|
rfl
|
|
|
|
|
exact isSolution_f_zero f hS
|
|
|
|
|
|
|
|
|
|
end ElevenPlane
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
theorem eleven_dim_plane_of_no_sols_exists : ∃ (B : Fin 11 → (PlusU1 3).Charges),
|
2024-04-19 09:57:30 -04:00
|
|
|
|
LinearIndependent ℚ B ∧
|
2024-06-26 11:54:02 -04:00
|
|
|
|
∀ (f : Fin 11 → ℚ), (PlusU1 3).IsSolution (∑ i, f i • B i) → ∑ i, f i • B i = 0 := by
|
2024-04-19 09:57:30 -04:00
|
|
|
|
use ElevenPlane.B
|
|
|
|
|
apply And.intro
|
|
|
|
|
exact ElevenPlane.basis_linear_independent
|
|
|
|
|
exact ElevenPlane.isSolution_only_if_zero
|
|
|
|
|
|
|
|
|
|
end PlusU1
|
|
|
|
|
end SMRHN
|