Skip to content

Commit

Permalink
fix: GuessLex: delaborate unused parameters as _
Browse files Browse the repository at this point in the history
fixes #4230
  • Loading branch information
nomeata committed Jun 3, 2024
1 parent b53a74d commit 7f71f01
Show file tree
Hide file tree
Showing 2 changed files with 33 additions and 9 deletions.
20 changes: 11 additions & 9 deletions src/Lean/Elab/PreDefinition/WF/TerminationArgument.lean
Original file line number Diff line number Diff line change
Expand Up @@ -86,21 +86,23 @@ def TerminationArgument.delab (termArg : TerminationArgument) : MetaM (TSyntax `
let e ← mkLambdaFVars ys[termArg.arity - termArg.extraParams:] e -- undo overshooting by lambdaTelescope
pure (← delabCore e (delab := go termArg.extraParams #[])).1
where
go : Nat → TSyntaxArray [`ident, `Lean.Parser.Term.hole] → DelabM (TSyntax ``terminationBy)
go : Nat → TSyntaxArray `ident → DelabM (TSyntax ``terminationBy)
| 0, vars => do
let stxBody ← Delaborator.delab
let hole : TSyntax `Lean.Parser.Term.hole ← `(hole|_)

-- any variable not mentioned syntatically (it may appear in the `Expr`, so do not just use
-- `e.bindingBody!.hasLooseBVar`) should be delaborated as a hole.
let vars : TSyntaxArray [`ident, `Lean.Parser.Term.hole] :=
Array.map (fun (i : Ident) => if hasIdent i.getId stxBody then i else hole) vars
-- drop trailing underscores
let mut vars := vars
while ! vars.isEmpty && vars.back.raw.isOfKind ``hole do vars := vars.pop
if vars.isEmpty then
`(terminationBy|termination_by $(← Delaborator.delab))
`(terminationBy|termination_by $stxBody)
else
`(terminationBy|termination_by $vars* => $(← Delaborator.delab))
`(terminationBy|termination_by $vars* => $stxBody)
| i+1, vars => do
let e ← getExpr
unless e.isLambda do return ← go 0 vars -- should not happen

-- Delaborate unused parameters with `_`
if e.bindingBody!.hasLooseBVar 0 then
withBindingBodyUnusedName fun n => go i (vars.push ⟨n⟩)
else
descend e.bindingBody! 1 (go i (vars.push (← `(hole|_))))
withBindingBodyUnusedName fun n => go i (vars.push ⟨n⟩)
22 changes: 22 additions & 0 deletions tests/lean/run/4230.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,22 @@
def copy (curr : Nat) (input : Array Nat) (output : Array Nat) : Array Nat :=
if hcurr:curr < input.size then
copy (curr + 1) input (output.push (input[curr]'hcurr))
else
output
termination_by input.size - curr

/-- info: Try this: termination_by input.size - curr -/
#guard_msgs(drop warning, info) in
theorem foo (curr : Nat) (input : Array Nat) (output : Array Nat)
: ∀ (idx : Nat) (hidx1 : idx < curr),
(copy curr input output)[idx]'sorry
=
output[idx]'sorry := by
intro idx hidx
unfold copy
split
. rw [foo]
. rw [Array.get_push_lt]
. omega
. rfl
termination_by?

0 comments on commit 7f71f01

Please sign in to comment.