PhysLean/scripts/hepLean_style_lint.lean

87 lines
2.9 KiB
Text
Raw Normal View History

2024-07-03 07:41:06 -04:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
import Lean
import Mathlib.Tactic.Linter.TextBased
/-!
# Double line Lint
This linter double empty lines in files.
## Note
Parts of this file are adapted from `Mathlib.Tactic.Linter.TextBased`,
authored by Michael Rothgang.
-/
open Lean System Meta
/-- Given a list of lines, outputs an error message and a line number. -/
2024-07-12 10:23:59 -04:00
def HepLeanTextLinter : Type := Array String → Array (String × × )
2024-07-03 07:41:06 -04:00
/-- Checks if there are two consecutive empty lines. -/
def doubleEmptyLineLinter : HepLeanTextLinter := fun lines ↦ Id.run do
let enumLines := (lines.toList.enumFrom 1)
let pairLines := List.zip enumLines (List.tail! enumLines)
let errors := pairLines.filterMap (fun ((lno1, l1), _, l2) ↦
if l1.length == 0 && l2.length == 0 then
2024-07-12 10:23:59 -04:00
some (s!" Double empty line. ", lno1, 1)
2024-07-03 07:41:06 -04:00
else none)
errors.toArray
2024-07-12 09:58:40 -04:00
/-- Checks if there is a souble space in the line, which is not at the start. -/
def doubleSpaceLinter : HepLeanTextLinter := fun lines ↦ Id.run do
let enumLines := (lines.toList.enumFrom 1)
let errors := enumLines.filterMap (fun (lno, l) ↦
if String.containsSubstr l.trimLeft " " then
2024-07-12 10:23:59 -04:00
let k := (Substring.findAllSubstr l " ").toList.getLast?
let col := match k with
| none => 1
| some k => String.offsetOfPos l k.stopPos
2024-07-12 10:23:59 -04:00
some (s!" Non-initial double space in line.", lno, col)
2024-07-12 09:58:40 -04:00
else none)
errors.toArray
2024-07-03 07:41:06 -04:00
structure HepLeanErrorContext where
/-- The underlying `message`. -/
error : String
/-- The line number -/
lineNumber :
2024-07-12 10:23:59 -04:00
/-- The column number -/
columnNumber :
2024-07-03 07:41:06 -04:00
/-- The file path -/
path : FilePath
def printErrors (errors : Array HepLeanErrorContext) : IO Unit := do
for e in errors do
2024-07-12 10:23:59 -04:00
IO.println (s!"error: {e.path}:{e.lineNumber}:{e.columnNumber}: {e.error}")
2024-07-03 07:41:06 -04:00
def hepLeanLintFile (path : FilePath) : IO Bool := do
let lines ← IO.FS.lines path
let allOutput := (Array.map (fun lint ↦
2024-07-12 10:23:59 -04:00
(Array.map (fun (e, n, c) ↦ HepLeanErrorContext.mk e n c path)) (lint lines)))
2024-07-12 09:58:40 -04:00
#[doubleEmptyLineLinter, doubleSpaceLinter]
2024-07-03 07:41:06 -04:00
let errors := allOutput.flatten
printErrors errors
return errors.size > 0
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
if imp.module == `Init then continue
let filePath := (mkFilePath (imp.module.toString.split (· == '.'))).addExtension "lean"
if ← hepLeanLintFile filePath then
warned := true
if warned then
throw <| IO.userError s!"Errors found."
return 0