Skip to content

Commit a1e7068

Browse files
committed
feat: make loose docstrings into elaboration errors rather than parse errors
Now loose docstrings will not yield a confusing parser error and instead log an "unexpected docstring" elaboration error and continue processing. Furthermore, if the docstring is preceding the `in` command combinator, it will add a hint to the error: ``` unexpected doc string Hint: the docstring must come after '... in', immediately before the declaration that accepts a docstring. ``` Closes #3135
1 parent e5e5778 commit a1e7068

File tree

4 files changed

+79
-0
lines changed

4 files changed

+79
-0
lines changed

src/Lean/Elab/BuiltinCommand.lean

Lines changed: 11 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -23,6 +23,17 @@ namespace Lean.Elab.Command
2323
modifyEnv fun env => addMainModuleDoc env ⟨doc, range⟩
2424
| _ => throwErrorAt stx "unexpected module doc string{indentD stx[1]}"
2525

26+
@[builtin_command_elab looseDocComment] def elabLooseDocComment : CommandElab := fun stx => do
27+
let trailingKind := stx[1][0].getKind
28+
let addlHint :=
29+
if trailingKind == ``Lean.Parser.Command.in then
30+
"\n\nHint: the docstring must come after '... in', immediately before the declaration that accepts a docstring."
31+
else
32+
""
33+
logError m!"unexpected doc string{addlHint}"
34+
if stx[1].getNumArgs == 1 then
35+
elabCommand stx[1]
36+
2637
private def addScope (isNewNamespace : Bool) (isNoncomputable : Bool) (header : String) (newNamespace : Name) : CommandElabM Unit := do
2738
modify fun s => { s with
2839
env := s.env.registerNamespace newNamespace,

src/Lean/Parser/Command.lean

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -53,6 +53,13 @@ multiple times in the same file.
5353
def moduleDoc := leading_parser ppDedent <|
5454
"/-!" >> commentBody >> ppLine
5555

56+
-- A low-priority parser to be able to log an error that it's not attached to anything.
57+
-- This is more friendly than a parser error, and it saves needing to add `docComment` before
58+
-- every command that doesn't accept a doc comment.
59+
-- It is trailed by an optional command so that the error can give a hint.
60+
@[builtin_command_parser low]
61+
def looseDocComment := leading_parser ppDedent (docComment >> optional commandParser)
62+
5663
def namedPrio := leading_parser
5764
atomic (" (" >> nonReservedSymbol "priority") >> " := " >> withoutPosition priorityParser >> ")"
5865
def optNamedPrio := optional namedPrio

tests/lean/interactive/3135.lean

Lines changed: 29 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,29 @@
1+
/-!
2+
# Loose docstrings
3+
-/
4+
5+
/-!
6+
Basic test
7+
-/
8+
9+
section
10+
/-- This is a loose docstring -/
11+
--^ collectDiagnostics
12+
end
13+
14+
section
15+
/-- This is a loose docstring before an 'in' command. -/
16+
--^ collectDiagnostics
17+
set_option pp.all true in
18+
def x := 0
19+
20+
-- Still elaborates the `def`
21+
#check x
22+
--^ textDocument/hover
23+
end
24+
25+
section
26+
set_option pp.all true in
27+
/-- This is a docstring in its correct position. -/
28+
def y := 0
29+
end
Lines changed: 32 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,32 @@
1+
{"version": 1,
2+
"uri": "file:///3135.lean",
3+
"diagnostics":
4+
[{"source": "Lean 4",
5+
"severity": 1,
6+
"range":
7+
{"start": {"line": 9, "character": 0}, "end": {"line": 10, "character": 0}},
8+
"message": "unexpected doc string",
9+
"fullRange":
10+
{"start": {"line": 9, "character": 0}, "end": {"line": 11, "character": 3}}},
11+
{"source": "Lean 4",
12+
"severity": 1,
13+
"range":
14+
{"start": {"line": 14, "character": 0}, "end": {"line": 15, "character": 0}},
15+
"message":
16+
"unexpected doc string\n\nHint: the docstring must come after '... in', immediately before the declaration that accepts a docstring.",
17+
"fullRange":
18+
{"start": {"line": 14, "character": 0},
19+
"end": {"line": 17, "character": 10}}},
20+
{"source": "Lean 4",
21+
"severity": 3,
22+
"range":
23+
{"start": {"line": 20, "character": 0}, "end": {"line": 20, "character": 6}},
24+
"message": "x : Nat",
25+
"fullRange":
26+
{"start": {"line": 20, "character": 0},
27+
"end": {"line": 20, "character": 6}}}]}
28+
{"textDocument": {"uri": "file:///3135.lean"},
29+
"position": {"line": 20, "character": 7}}
30+
{"range":
31+
{"start": {"line": 20, "character": 7}, "end": {"line": 20, "character": 8}},
32+
"contents": {"value": "```lean\nx : Nat\n```", "kind": "markdown"}}

0 commit comments

Comments
 (0)