Skip to content

Commit 2152a3a

Browse files
committed
feat: rewrite the trailing whitespace linter in Lean
There is an in-progress PR batteries to the same effect: when that PR lands, this linter can be removed. Both implement basically the same logic. This linter is consciously not rewritte as a syntax linter, as a fair number of mathlib contributors have their editors configured to remove trailing whitespace when saving: in effect, the linter would only fire intermittently (such as, between keypresses), which is not helpful. Running this as a text-based linter in CI, which is auto-fixable, is far more helpful.
1 parent 2900855 commit 2152a3a

File tree

2 files changed

+23
-3
lines changed

2 files changed

+23
-3
lines changed

Mathlib/Tactic/Linter/TextBased.lean

+22-1
Original file line numberDiff line numberDiff line change
@@ -71,6 +71,8 @@ inductive StyleError where
7171
| fileTooLong (numberLines : ℕ) (newSizeLimit : ℕ) (previousLimit : Option ℕ) : StyleError
7272
/-- A line ends with windows line endings (\r\n) instead of unix ones (\n). -/
7373
| windowsLineEnding
74+
/-- A line contains trailing whitespace -/
75+
| trailingWhitespace
7476
deriving BEq
7577

7678
/-- How to format style errors -/
@@ -111,6 +113,7 @@ def StyleError.errorMessage (err : StyleError) (style : ErrorFormat) : String :=
111113
| ErrorFormat.humanReadable => s!"file contains {currentSize} lines, try to split it up"
112114
| windowsLineEnding => "This line ends with a windows line ending (\r\n): please use Unix line\
113115
endings (\n) instead"
116+
| trailingWhitespace => "This line ends with some whitespace: please remove this"
114117

115118
/-- The error code for a given style error. Keep this in sync with `parse?_errorContext` below! -/
116119
-- FUTURE: we're matching the old codes in `lint-style.py` for compatibility;
@@ -122,6 +125,7 @@ def StyleError.errorCode (err : StyleError) : String := match err with
122125
| StyleError.broadImport _ => "ERR_IMP"
123126
| StyleError.fileTooLong _ _ _ => "ERR_NUM_LIN"
124127
| StyleError.windowsLineEnding => "ERR_WIN"
128+
| StyleError.trailingWhitespace => "ERR_TWS"
125129

126130
/-- Context for a style error: the actual error, the line number in the file we're reading
127131
and the path to the file. -/
@@ -216,6 +220,7 @@ def parse?_errorContext (line : String) : Option ErrorContext := Id.run do
216220
| "ERR_COP" => some (StyleError.copyright none)
217221
| "ERR_AUT" => some (StyleError.authors)
218222
| "ERR_ADN" => some (StyleError.adaptationNote)
223+
| "ERR_TWS" => some (StyleError.trailingWhitespace)
219224
| "ERR_WIN" => some (StyleError.windowsLineEnding)
220225
| "ERR_IMP" =>
221226
-- XXX tweak exceptions messages to ease parsing?
@@ -363,6 +368,21 @@ def windowsLineEndingLinter : TextbasedLinter := fun lines ↦ Id.run do
363368
lineNumber := lineNumber + 1
364369
return (errors, fixedLines)
365370

371+
372+
/-- Lint a collection of input strings if one of them contains trailing whitespace. -/
373+
def trailingWhitespaceLinter : TextbasedLinter := fun lines ↦ Id.run do
374+
let mut errors := Array.mkEmpty 0
375+
let mut fixedLines := lines
376+
-- invariant: this equals the current index of 'line' in 'lines',
377+
-- hence starts at 0 and is incremented *at the end* of the loop
378+
let mut lineNumber := 0
379+
for line in lines do
380+
if line.back == ' ' then
381+
errors := errors.push (StyleError.trailingWhitespace, lineNumber)
382+
fixedLines := fixedLines.set! 0 (line.dropRightWhile (· == ' '))
383+
lineNumber := lineNumber + 1
384+
return (errors, fixedLines)
385+
366386
/-- Whether a collection of lines consists *only* of imports, blank lines and single-line comments.
367387
In practice, this means it's an imports-only file and exempt from almost all linting. -/
368388
def isImportsOnlyFile (lines : Array String) : Bool :=
@@ -395,7 +415,8 @@ end
395415

396416
/-- All text-based linters registered in this file. -/
397417
def allLinters : Array TextbasedLinter := #[
398-
copyrightHeaderLinter, adaptationNoteLinter, broadImportsLinter, windowsLineEndingLinter
418+
copyrightHeaderLinter, adaptationNoteLinter, broadImportsLinter,
419+
windowsLineEndingLinter, trailingWhitespaceLinter
399420
]
400421

401422
/-- Controls what kind of output this programme produces. -/

scripts/lint-style.py

+1-2
Original file line numberDiff line numberDiff line change
@@ -42,7 +42,6 @@
4242
ERR_IWH = 22 # isolated where
4343
ERR_DOT = 12 # isolated or low focusing dot
4444
ERR_SEM = 13 # the substring " ;"
45-
ERR_TWS = 15 # trailing whitespace
4645
ERR_CLN = 16 # line starts with a colon
4746
ERR_IND = 17 # second line not correctly indented
4847
ERR_ARR = 18 # space after "←"
@@ -62,7 +61,7 @@
6261
exceptions += [(ERR_MOD, path, None)]
6362
elif errno == "ERR_IND":
6463
exceptions += [(ERR_IND, path, None)]
65-
elif errno in ["ERR_COP", "ERR_LIN", "ERR_ADN", "ERR_NUM_LIN"]:
64+
elif errno in ["ERR_COP", "ERR_LIN", "ERR_ADN", "ERR_WIN", "ERR_TWS", "ERR_NUM_LIN"]:
6665
pass # maintained by the Lean style linter now
6766
else:
6867
print(f"Error: unexpected errno in style-exceptions.txt: {errno}")

0 commit comments

Comments
 (0)