Merge pull request #92 from HEPLean/Update-versions

refactor: Basic lint
This commit is contained in:
Joseph Tooby-Smith 2024-07-18 17:00:13 -04:00 committed by GitHub
commit ea45922df8
No known key found for this signature in database
GPG key ID: B5690EEEBB952194
27 changed files with 53 additions and 57 deletions

View file

@ -187,8 +187,7 @@ instance quadSolsMulAction (χ : ACCSystemQuad) : MulAction χ.QuadSols wher
intro i intro i
erw [(χ.quadraticACCs i).map_smul] erw [(χ.quadraticACCs i).map_smul]
rw [S.quadSol i] rw [S.quadSol i]
simp only [mul_zero] simp only [mul_zero]⟩
mul_smul a b S := by mul_smul a b S := by
apply QuadSols.ext apply QuadSols.ext
exact mul_smul _ _ _ exact mul_smul _ _ _

View file

@ -44,8 +44,7 @@ def linSolMap {χ : ACCSystem} (G : ACCSystemGroupAction χ) (g : G.group) :
χ.LinSols →ₗ[] χ.LinSols where χ.LinSols →ₗ[] χ.LinSols where
toFun S := ⟨G.rep g S.val, by toFun S := ⟨G.rep g S.val, by
intro i intro i
rw [G.linearInvariant, S.linearSol] rw [G.linearInvariant, S.linearSol]⟩
map_add' S T := by map_add' S T := by
apply ACCSystemLinear.LinSols.ext apply ACCSystemLinear.LinSols.ext
exact (G.rep g).map_add' _ _ exact (G.rep g).map_add' _ _
@ -85,8 +84,7 @@ instance quadSolAction {χ : ACCSystem} (G : ACCSystemGroupAction χ) :
smul f S := ⟨G.linSolRep f S.1, by smul f S := ⟨G.linSolRep f S.1, by
intro i intro i
simp only [linSolRep_apply_apply_val] simp only [linSolRep_apply_apply_val]
rw [G.quadInvariant, S.quadSol] rw [G.quadInvariant, S.quadSol]⟩
mul_smul f1 f2 S := by mul_smul f1 f2 S := by
apply ACCSystemQuad.QuadSols.ext apply ACCSystemQuad.QuadSols.ext
change (G.rep.toFun (f1 * f2)) S.val = _ change (G.rep.toFun (f1 * f2)) S.val = _

View file

@ -473,8 +473,7 @@ def AnomalyFreeMk (S : MSSMACC.Charges) (hg : accGrav S = 0)
intro i intro i
simp at i simp at i
match i with match i with
| 0 => exact hquad | 0 => exact hquad⟩, 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)
@ -490,8 +489,7 @@ def AnomalyFreeQuadMk' (S : MSSMACC.LinSols) (hquad : accQuad S.val = 0) :
intro i intro i
simp at i simp at i
match i with match i with
| 0 => exact hquad | 0 => exact hquad⟩
/-- A `Sol` from a `LinSol` satisfying the quadratic and cubic ACCs. -/ /-- A `Sol` from a `LinSol` satisfying the quadratic and cubic ACCs. -/
@[simp] @[simp]
@ -501,13 +499,12 @@ def AnomalyFreeMk' (S : MSSMACC.LinSols) (hquad : accQuad S.val = 0)
intro i intro i
simp at i simp at i
match i with match i with
| 0 => exact hquad | 0 => exact hquad⟩, hcube⟩
⟩ , by exact hcube ⟩
/-- A `Sol` from a `QuadSol` satisfying the cubic ACCs. -/ /-- A `Sol` from a `QuadSol` satisfying the cubic ACCs. -/
@[simp] @[simp]
def AnomalyFreeMk'' (S : MSSMACC.QuadSols) (hcube : accCube S.val = 0) : MSSMACC.Sols := def AnomalyFreeMk'' (S : MSSMACC.QuadSols) (hcube : accCube S.val = 0) : MSSMACC.Sols :=
⟨S , by exact hcube ⟨S, hcube⟩
lemma AnomalyFreeMk''_val (S : MSSMACC.QuadSols) lemma AnomalyFreeMk''_val (S : MSSMACC.QuadSols)
(hcube : accCube S.val = 0) : (hcube : accCube S.val = 0) :

View file

@ -84,7 +84,9 @@ def hepLeanLintFile (path : FilePath) : IO Bool := do
(Array.map (fun (e, n, c) ↦ HepLeanErrorContext.mk e n c path)) (lint lines))) (Array.map (fun (e, n, c) ↦ HepLeanErrorContext.mk e n c path)) (lint lines)))
#[doubleEmptyLineLinter, doubleSpaceLinter, substringLinter ".-/", substringLinter " )", #[doubleEmptyLineLinter, doubleSpaceLinter, substringLinter ".-/", substringLinter " )",
substringLinter "( ", substringLinter "=by", substringLinter " def ", substringLinter "( ", substringLinter "=by", substringLinter " def ",
substringLinter "/-- We "] substringLinter "/-- We ", substringLinter "[ ", substringLinter " ]", substringLinter " ,"
, substringLinter "by exact ",
substringLinter "⟨ ", substringLinter " ⟩"]
let errors := allOutput.flatten let errors := allOutput.flatten
printErrors errors printErrors errors
return errors.size > 0 return errors.size > 0