refactor: Lint

This commit is contained in:
jstoobysmith 2024-04-19 10:08:56 -04:00
parent e710c9278e
commit 12a568b45f
5 changed files with 44 additions and 13 deletions

View file

@ -6,7 +6,13 @@ Authors: Joseph Tooby-Smith
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic
import HepLean.AnomalyCancellation.SMNu.PlusU1.FamilyMaps
import HepLean.AnomalyCancellation.SMNu.PlusU1.PlaneNonSols
/-!
# Bound on plane dimension
We place an upper bound on the dimension of a plane of charges on which every point is a solution.
The upper bound is 7, proven in the theorem `plane_exists_dim_le_7`.
-/
universe v u
namespace SMRHN
@ -16,6 +22,8 @@ open SMνCharges
open SMνACCs
open BigOperators
/-- 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. -/
def existsPlane (n : ) : Prop := ∃ (B : Fin n → (PlusU1 3).charges),
LinearIndependent B ∧ ∀ (f : Fin n → ), (PlusU1 3).isSolution (∑ i, f i • B i)
@ -33,7 +41,7 @@ lemma exists_plane_exists_basis {n : } (hE : existsPlane n) :
have h1 : ∑ x : Fin n, -(g (Sum.inr x) • Y (Sum.inr x)) =
∑ x : Fin n, (-g (Sum.inr x)) • Y (Sum.inr x):= by
apply Finset.sum_congr
simp
simp only
intro i _
simp
rw [h1] at hg

View file

@ -26,28 +26,40 @@ open BigOperators
namespace ElevenPlane
/-- 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]
/-- 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]
/-- 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]
/-- 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]
/-- 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]
/-- 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]
/-- 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]
/-- 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]
/-- 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]
/-- 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]
/-- 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]
/-- The charge assignment forming a basis of the plane. -/
def B : Fin 11 → (PlusU1 3).charges := fun i =>
match i with
| 0 => B₀
@ -76,7 +88,7 @@ lemma Bi_sum_quad (i : Fin 11) (f : Fin 11 → ) :
rw [quadBiLin.map_smul₂, Bi_Bj_quad hij.symm]
simp
/-- The coefficents of the quadratic equation in our basis. -/
@[simp]
def quadCoeff : Fin 11 → := ![1, 1, 1, 1, 1, 1, 1, 1, 1, 0, 0]
@ -102,7 +114,7 @@ lemma isSolution_quadCoeff_f_sq_zero (f : Fin 11 → ) (hS : (PlusU1 3).isSol
rw [Fintype.sum_eq_zero_iff_of_nonneg] at hQ
exact congrFun hQ k
intro i
simp
simp only [Pi.zero_apply, quadCoeff]
rw [mul_nonneg_iff]
apply Or.inl
apply And.intro