Skip to content

Commit 89ba3be

Browse files
committed
use (and modify!) logLintIf
1 parent 9a030d9 commit 89ba3be

File tree

4 files changed

+18
-16
lines changed

4 files changed

+18
-16
lines changed

Diff for: src/Lean/Elab/Declaration.lean

+2-4
Original file line numberDiff line numberDiff line change
@@ -169,10 +169,8 @@ private def inductiveSyntaxToView (modifiers : Modifiers) (decl : Syntax) : Comm
169169
let classes ← liftCoreM <| getOptDerivingClasses decl[6]
170170
if decl[3][0].isToken ":=" then
171171
-- https://github.com/leanprover/lean4/issues/5236
172-
if Linter.getLinterValue Linter.linter.deprecated (← getOptions) then
173-
withRef decl[0] <| withRef decl[3] <| logWarning <| .tagged ``Linter.deprecatedAttr "\
174-
'inductive ... :=' has been deprecated in favor of 'inductive ... where'.\n\
175-
You can disable this warning with 'set_option linter.deprecated false'."
172+
withRef decl[0] <| Linter.logLintIf Linter.linter.deprecated decl[3]
173+
"'inductive ... :=' has been deprecated in favor of 'inductive ... where'."
176174
return {
177175
ref := decl
178176
shortDeclName := name

Diff for: src/Lean/Elab/Structure.lean

+3-5
Original file line numberDiff line numberDiff line change
@@ -140,11 +140,9 @@ def structFields := leading_parser many (structExplicitBinder <|> struct
140140
private def expandFields (structStx : Syntax) (structModifiers : Modifiers) (structDeclName : Name) : TermElabM (Array StructFieldView) := do
141141
if structStx[5][0].isToken ":=" then
142142
-- https://github.com/leanprover/lean4/issues/5236
143-
if Linter.getLinterValue Linter.linter.deprecated (← getOptions) then
144-
let cmd := if structStx[0].getKind == ``Parser.Command.classTk then "class" else "structure"
145-
withRef structStx[0] <| withRef structStx[5][0] <| logWarning <| .tagged ``Linter.deprecatedAttr s!"\
146-
'{cmd} ... :=' has been deprecated in favor of '{cmd} ... where'.\n\
147-
You can disable this warning with 'set_option linter.deprecated false'."
143+
let cmd := if structStx[0].getKind == ``Parser.Command.classTk then "class" else "structure"
144+
withRef structStx[0] <| Linter.logLintIf Linter.linter.deprecated structStx[5][0]
145+
s!"{cmd} ... :=' has been deprecated in favor of '{cmd} ... where'."
148146
let fieldBinders := if structStx[5].isNone then #[] else structStx[5][2][0].getArgs
149147
fieldBinders.foldlM (init := #[]) fun (views : Array StructFieldView) fieldBinder => withRef fieldBinder do
150148
let mut fieldBinder := fieldBinder

Diff for: src/Lean/Linter/Basic.lean

+8-2
Original file line numberDiff line numberDiff line change
@@ -23,8 +23,14 @@ def logLint [Monad m] [MonadLog m] [AddMessageContext m] [MonadOptions m]
2323
let disable := m!"note: this linter can be disabled with `set_option {linterOption.name} false`"
2424
logWarningAt stx (.tagged linterOption.name m!"{msg}\n{disable}")
2525

26-
/-- If `linterOption` is true, print a linter warning message at the position determined by `stx`.
26+
/--
27+
If `linterOption` is enabled, print a linter warning message at the position determined by `stx`.
28+
29+
Whether a linter option is enabled or not is determined by the following sequence:
30+
1. If it is set, then the value determines whether or not it is enabled.
31+
2. Otherwise, if `linter.all` is set, then its value determines whether or not the option is enabled.
32+
3. Otherwise, the default value determines whether or not it is enabled.
2733
-/
2834
def logLintIf [Monad m] [MonadLog m] [AddMessageContext m] [MonadOptions m]
2935
(linterOption : Lean.Option Bool) (stx : Syntax) (msg : MessageData) : m Unit := do
30-
if linterOption.get (← getOptions) then logLint linterOption stx msg
36+
if getLinterValue linterOption (← getOptions) then logLint linterOption stx msg

Diff for: tests/lean/run/5236.lean

+5-5
Original file line numberDiff line numberDiff line change
@@ -6,24 +6,24 @@ set_option linter.deprecated true
66

77
/--
88
warning: 'inductive ... :=' has been deprecated in favor of 'inductive ... where'.
9-
You can disable this warning with 'set_option linter.deprecated false'.
9+
note: this linter can be disabled with `set_option linter.deprecated false`
1010
-/
1111
#guard_msgs in
1212
inductive DogSound' :=
1313
| woof
1414
| grr
1515

1616
/--
17-
warning: 'structure ... :=' has been deprecated in favor of 'structure ... where'.
18-
You can disable this warning with 'set_option linter.deprecated false'.
17+
warning: structure ... :=' has been deprecated in favor of 'structure ... where'.
18+
note: this linter can be disabled with `set_option linter.deprecated false`
1919
-/
2020
#guard_msgs in
2121
structure S :=
2222
(n : Nat)
2323

2424
/--
25-
warning: 'class ... :=' has been deprecated in favor of 'class ... where'.
26-
You can disable this warning with 'set_option linter.deprecated false'.
25+
warning: class ... :=' has been deprecated in favor of 'class ... where'.
26+
note: this linter can be disabled with `set_option linter.deprecated false`
2727
-/
2828
#guard_msgs in
2929
class C :=

0 commit comments

Comments
 (0)