Skip to content

Commit

Permalink
chore(Linter): upstream nolint attribute
Browse files Browse the repository at this point in the history
  • Loading branch information
jcommelin committed Nov 25, 2024
1 parent 0eca3bd commit 578e749
Show file tree
Hide file tree
Showing 3 changed files with 32 additions and 0 deletions.
1 change: 1 addition & 0 deletions src/Lean/Linter.lean
Original file line number Diff line number Diff line change
Expand Up @@ -8,6 +8,7 @@ import Lean.Linter.Util
import Lean.Linter.Builtin
import Lean.Linter.ConstructorAsVariable
import Lean.Linter.Deprecated
import Lean.Linter.Environment
import Lean.Linter.UnusedVariables
import Lean.Linter.MissingDocs
import Lean.Linter.Omit
1 change: 1 addition & 0 deletions src/Lean/Linter/Environment.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1 @@
import Lean.Linter.Environment.Nolint
30 changes: 30 additions & 0 deletions src/Lean/Linter/Environment/Nolint.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,30 @@
/-
Copyright (c) 2020 Floris van Doorn. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Floris van Doorn, Robert Y. Lewis, Gabriel Ebner
-/
prelude

/-- `@[nolint linterName]` omits the tagged declaration from being checked by
the linter with name `linterName`. -/
syntax (name := nolint) "nolint" (ppSpace ident)+ : attr

/-- Defines the user attribute `nolint` for skipping `#lint` -/
initialize nolintAttr : ParametricAttribute (Array Name) ←
registerParametricAttribute {
name := `nolint
descr := "Do not report this declaration in any of the tests of `#lint`"
getParam := fun _ => fun
| `(attr| nolint $[$ids]*) => ids.mapM fun id => withRef id <| do
let shortName := id.getId.eraseMacroScopes
let some (declName, _) := (batteriesLinterExt.getState (← getEnv)).find? shortName
| throwError "linter '{shortName}' not found"
Elab.addConstInfo id declName
pure shortName
| _ => Elab.throwUnsupportedSyntax
}

/-- Returns true if `decl` should be checked
using `linter`, i.e., if there is no `nolint` attribute. -/
def shouldBeLinted [Monad m] [MonadEnv m] (linter : Name) (decl : Name) : m Bool :=
return !((nolintAttr.getParam? (← getEnv) decl).getD #[]).contains linter

0 comments on commit 578e749

Please sign in to comment.