Skip to content

Commit

Permalink
fix: .. in patterns should not make use of opt-params
Browse files Browse the repository at this point in the history
In patterns, ellipsis should always be an implicit argument, even if it is an optparam or autoparam. This prevents examples such as the one in #4555 from failing:
```lean
match e with
| .internal .. => sorry
| .error .. => sorry
```
The `internal` constructor has an optparam (`| internal (id : InternalExceptionId) (extra : KVMap := {})`).

We may consider having ellipsis suppress optparams and autoparams in general. We avoid doing so for now since it's possible to opt-out of them individually (for example with `.internal (extra := _) ..`) but it's not possible to opt-in. While we can opt-out like this for patterns as well, with patterns it is suprising how strongly the optparams are inserted.

Closes #4555
  • Loading branch information
kmill committed Nov 3, 2024
1 parent 0de925e commit 832ab07
Show file tree
Hide file tree
Showing 2 changed files with 32 additions and 0 deletions.
16 changes: 16 additions & 0 deletions src/Lean/Elab/App.lean
Original file line number Diff line number Diff line change
Expand Up @@ -595,6 +595,22 @@ mutual
elabAndAddNewArg argName arg
main
| _ =>
if (← read).ellipsis && (← readThe Term.Context).inPattern then
/-
In patterns, ellipsis should always be an implicit argument, even if it is an optparam or autoparam.
This prevents examples such as the one in #4555 from failing:
```lean
match e with
| .internal .. => sorry
| .error .. => sorry
```
The `internal` has an optparam (`| internal (id : InternalExceptionId) (extra : KVMap := {})`).
We may consider having ellipsis suppress optparams and autoparams in general.
We avoid doing so for now since it's possible to opt-out of them (for example with `.internal (extra := _) ..`)
but it's not possible to opt-in.
-/
return ← addImplicitArg argName
let argType ← getArgExpectedType
match (← read).explicit, argType.getOptParamDefault?, argType.getAutoParamTactic? with
| false, some defVal, _ => addNewArg argName defVal; main
Expand Down
16 changes: 16 additions & 0 deletions tests/lean/run/4555.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,16 @@
import Lean.Exception
/-!
# In pattern, `..` shouldn't make use of opt-params
https://github.com/leanprover/lean4/issues/4555
-/

open Lean

/-!
The `.internal` constructor has an opt-param, which caused this to be a non-exhaustive match.
-/
example (e : Exception) : Nat :=
match e with
| .internal .. => 0
| .error .. => 1

0 comments on commit 832ab07

Please sign in to comment.