From b74f85accd0347383faeefd62c49d4f47d556084 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 17 Sep 2024 10:57:35 +0100 Subject: [PATCH] fix: do not ban `..` with a `.` on the next line (#4768) Without this change, ```lean example : True := by refine' trivial .. . trivial ``` is a parse error. --- src/Lean/Parser/Term.lean | 2 +- tests/lean/run/4768.lean | 10 ++++++++++ 2 files changed, 11 insertions(+), 1 deletion(-) create mode 100644 tests/lean/run/4768.lean diff --git a/src/Lean/Parser/Term.lean b/src/Lean/Parser/Term.lean index 5a88595b7ec2..e9980bea794f 100644 --- a/src/Lean/Parser/Term.lean +++ b/src/Lean/Parser/Term.lean @@ -763,7 +763,7 @@ We use them to implement `macro_rules` and `elab_rules` def namedArgument := leading_parser (withAnonymousAntiquot := false) atomic ("(" >> ident >> " := ") >> withoutPosition termParser >> ")" def ellipsis := leading_parser (withAnonymousAntiquot := false) - ".." >> notFollowedBy "." "`.` immediately after `..`" + ".." >> notFollowedBy (checkNoWsBefore >> ".") "`.` immediately after `..`" def argument := checkWsBefore "expected space" >> checkColGt "expected to be indented" >> diff --git a/tests/lean/run/4768.lean b/tests/lean/run/4768.lean new file mode 100644 index 000000000000..520f25caed2e --- /dev/null +++ b/tests/lean/run/4768.lean @@ -0,0 +1,10 @@ +/-! +Test that `..` tokens do not break nearby `.`s. + +Note that this tests specifically for issues with `.` that are not present with `·`. +-/ + +example : True ∧ True := by + constructor + refine trivial .. + . trivial -- this has to be . not · for this test to be useful