File tree 4 files changed +18
-16
lines changed
4 files changed +18
-16
lines changed Original file line number Diff line number Diff line change @@ -169,10 +169,8 @@ private def inductiveSyntaxToView (modifiers : Modifiers) (decl : Syntax) : Comm
169
169
let classes ← liftCoreM <| getOptDerivingClasses decl[6 ]
170
170
if decl[3 ][0 ].isToken ":=" then
171
171
-- 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'."
176
174
return {
177
175
ref := decl
178
176
shortDeclName := name
Original file line number Diff line number Diff line change @@ -140,11 +140,9 @@ def structFields := leading_parser many (structExplicitBinder <|> struct
140
140
private def expandFields (structStx : Syntax) (structModifiers : Modifiers) (structDeclName : Name) : TermElabM (Array StructFieldView) := do
141
141
if structStx[5 ][0 ].isToken ":=" then
142
142
-- 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'."
148
146
let fieldBinders := if structStx[5 ].isNone then #[] else structStx[5 ][2 ][0 ].getArgs
149
147
fieldBinders.foldlM (init := #[]) fun (views : Array StructFieldView) fieldBinder => withRef fieldBinder do
150
148
let mut fieldBinder := fieldBinder
Original file line number Diff line number Diff line change @@ -23,8 +23,14 @@ def logLint [Monad m] [MonadLog m] [AddMessageContext m] [MonadOptions m]
23
23
let disable := m!"note: this linter can be disabled with `set_option {linterOption.name} false`"
24
24
logWarningAt stx (.tagged linterOption.name m!"{msg}\n {disable}" )
25
25
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.
27
33
-/
28
34
def logLintIf [Monad m] [MonadLog m] [AddMessageContext m] [MonadOptions m]
29
35
(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
Original file line number Diff line number Diff line change @@ -6,24 +6,24 @@ set_option linter.deprecated true
6
6
7
7
/--
8
8
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`
10
10
-/
11
11
#guard_msgs in
12
12
inductive DogSound' :=
13
13
| woof
14
14
| grr
15
15
16
16
/--
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`
19
19
-/
20
20
#guard_msgs in
21
21
structure S :=
22
22
(n : Nat)
23
23
24
24
/--
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`
27
27
-/
28
28
#guard_msgs in
29
29
class C :=
You can’t perform that action at this time.
0 commit comments