refactor: Lint
This commit is contained in:
parent
52e591fa7a
commit
9f27a3a9fd
46 changed files with 176 additions and 168 deletions
|
@ -5,6 +5,7 @@ Authors: Joseph Tooby-Smith
|
|||
-/
|
||||
import Lean
|
||||
import Mathlib.Tactic.Linter.TextBased
|
||||
import Batteries.Data.Array.Merge
|
||||
/-!
|
||||
|
||||
# HepLean style linter
|
||||
|
@ -60,7 +61,17 @@ def substringLinter (s : String) : HepLeanTextLinter := fun lines ↦ Id.run do
|
|||
let col := match k with
|
||||
| none => 1
|
||||
| some k => String.offsetOfPos l k.stopPos
|
||||
some (s!" Found instance of substring {s}.", lno, col)
|
||||
some (s!" Found instance of substring `{s}`.", lno, col)
|
||||
else none)
|
||||
errors.toArray
|
||||
|
||||
/-- Number of space at new line must be even. -/
|
||||
def numInitialSpacesEven : HepLeanTextLinter := fun lines ↦ Id.run do
|
||||
let enumLines := (lines.toList.enumFrom 1)
|
||||
let errors := enumLines.filterMap (fun (lno, l) ↦
|
||||
let numSpaces := (l.takeWhile (· == ' ')).length
|
||||
if numSpaces % 2 != 0 then
|
||||
some (s!"Number of initial spaces is not even.", lno, 1)
|
||||
else none)
|
||||
errors.toArray
|
||||
|
||||
|
@ -78,18 +89,18 @@ def printErrors (errors : Array HepLeanErrorContext) : IO Unit := do
|
|||
for e in errors do
|
||||
IO.println (s!"error: {e.path}:{e.lineNumber}:{e.columnNumber}: {e.error}")
|
||||
|
||||
def hepLeanLintFile (path : FilePath) : IO Bool := do
|
||||
def hepLeanLintFile (path : FilePath) : IO (Array HepLeanErrorContext) := do
|
||||
let lines ← IO.FS.lines path
|
||||
let allOutput := (Array.map (fun lint ↦
|
||||
(Array.map (fun (e, n, c) ↦ HepLeanErrorContext.mk e n c path)) (lint lines)))
|
||||
#[doubleEmptyLineLinter, doubleSpaceLinter, substringLinter ".-/", substringLinter " )",
|
||||
#[doubleEmptyLineLinter, doubleSpaceLinter, numInitialSpacesEven,
|
||||
substringLinter ".-/", substringLinter " )",
|
||||
substringLinter "( ", substringLinter "=by", substringLinter " def ",
|
||||
substringLinter "/-- We ", substringLinter "[ ", substringLinter " ]", substringLinter " ,"
|
||||
, substringLinter "by exact ",
|
||||
substringLinter "⟨ ", substringLinter " ⟩"]
|
||||
substringLinter "⟨ ", substringLinter " ⟩", substringLinter "):"]
|
||||
let errors := allOutput.flatten
|
||||
printErrors errors
|
||||
return errors.size > 0
|
||||
return errors
|
||||
|
||||
def main (_ : List String) : IO UInt32 := do
|
||||
initSearchPath (← findSysroot)
|
||||
|
@ -99,13 +110,19 @@ def main (_ : List String) : IO UInt32 := do
|
|||
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
|
||||
if imp.module == `Init then continue
|
||||
let filePath := (mkFilePath (imp.module.toString.split (· == '.'))).addExtension "lean"
|
||||
if ← hepLeanLintFile filePath then
|
||||
warned := true
|
||||
if warned then
|
||||
let filePaths := hepLeanMod.imports.filterMap (fun imp ↦
|
||||
if imp.module == `Init then
|
||||
none
|
||||
else
|
||||
some ((mkFilePath (imp.module.toString.split (· == '.'))).addExtension "lean"))
|
||||
let errors := (← filePaths.mapM hepLeanLintFile).flatten
|
||||
let errorMessagesPresent := (errors.map (fun e => e.error)).sortDedup
|
||||
for eM in errorMessagesPresent do
|
||||
IO.println s!"\n\nError: {eM}"
|
||||
for e in errors do
|
||||
if e.error == eM then
|
||||
IO.println s!"{e.path}:{e.lineNumber}:{e.columnNumber}: {e.error}"
|
||||
if errors.size > 0 then
|
||||
throw <| IO.userError s!"Errors found."
|
||||
else
|
||||
IO.println "No linting issues found."
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue