Skip to content

Commit 4da40f0

Browse files
committed
fix: improve error message when decide +kernel fails
This PR improves the error message when `decide +kernel` fails in the kernel, but not the elaborator. Fixes #10766.
1 parent 14ff08d commit 4da40f0

File tree

2 files changed

+101
-57
lines changed

2 files changed

+101
-57
lines changed

src/Lean/Elab/Tactic/ElabTerm.lean

Lines changed: 62 additions & 57 deletions
Original file line numberDiff line numberDiff line change
@@ -482,7 +482,10 @@ where
482482
-- efficient term. The kernel handles the unification `e =?= true` specially.
483483
return pf
484484
else
485-
diagnose expectedType s r
485+
-- Diagnose the failure, lazily so that there is no performance impact if `decide` isn't being used interactively.
486+
throwError MessageData.ofLazyM (es := #[expectedType]) do
487+
diagnose expectedType s r
488+
486489
doKernel (expectedType : Expr) : TacticM Expr := do
487490
let pf ← mkDecideProof expectedType
488491
-- Get instance from `pf`
@@ -498,64 +501,66 @@ where
498501
let lemmaName ← withOptions (Elab.async.set · false) do
499502
mkAuxLemma lemmaLevels expectedType pf
500503
return mkConst lemmaName (lemmaLevels.map .param)
501-
catch _ =>
502-
diagnose expectedType s none
503-
diagnose {α : Type} (expectedType s : Expr) (r? : Option Expr) : TacticM α :=
504-
-- Diagnose the failure, lazily so that there is no performance impact if `decide` isn't being used interactively.
505-
throwError MessageData.ofLazyM (es := #[expectedType]) do
506-
let r ← r?.getDM (withAtLeastTransparency .default <| whnf s)
507-
if r.isAppOf ``isTrue then
508-
return m!"\
509-
Tactic `{tacticName}` failed. Internal error: The elaborator is able to reduce the \
510-
`{.ofConstName ``Decidable}` instance, but the kernel is not able to"
511-
else if r.isAppOf ``isFalse then
512-
return m!"\
513-
Tactic `{tacticName}` proved that the proposition\
514-
{indentExpr expectedType}\n\
515-
is false"
516-
-- Re-reduce the instance and collect diagnostics, to get all unfolded Decidable instances
517-
let (reason, unfoldedInsts) ← withoutModifyingState <| withOptions (fun opt => diagnostics.set opt true) do
518-
modifyDiag (fun _ => {})
519-
let reason ← withAtLeastTransparency .default <| blameDecideReductionFailure s
520-
let unfolded := (← get).diag.unfoldCounter.foldl (init := #[]) fun cs n _ => cs.push n
521-
let unfoldedInsts ← unfolded |>.qsort Name.lt |>.filterMapM fun n => do
522-
let e ← mkConstWithLevelParams n
523-
if (← Meta.isClass? (← inferType e)) == ``Decidable then
524-
return m!"`{.ofConst e}`"
525-
else
526-
return none
527-
return (reason, unfoldedInsts)
528-
let stuckMsg :=
529-
if unfoldedInsts.isEmpty then
530-
m!"Reduction got stuck at the `{.ofConstName ``Decidable}` instance{indentExpr reason}"
531-
else
532-
let instances := if unfoldedInsts.size == 1 then "instance" else "instances"
533-
m!"After unfolding the {instances} {.andList unfoldedInsts.toList}, \
534-
reduction got stuck at the `{.ofConstName ``Decidable}` instance{indentExpr reason}"
535-
let hint :=
536-
if reason.isAppOf ``Eq.rec then
537-
.hint' m!"Reduction got stuck on `▸` ({.ofConstName ``Eq.rec}), \
538-
which suggests that one of the `{.ofConstName ``Decidable}` instances is defined using tactics such as `rw` or `simp`. \
539-
To avoid tactics, make use of functions such as \
540-
`{.ofConstName ``inferInstanceAs}` or `{.ofConstName ``decidable_of_decidable_of_iff}` \
541-
to alter a proposition."
542-
else if reason.isAppOf ``Classical.choice then
543-
.hint' m!"Reduction got stuck on `{.ofConstName ``Classical.choice}`, \
544-
which indicates that a `{.ofConstName ``Decidable}` instance \
545-
is defined using classical reasoning, proving an instance exists rather than giving a concrete construction. \
546-
The `{tacticName}` tactic works by evaluating a decision procedure via reduction, \
547-
and it cannot make progress with such instances. \
548-
This can occur due to the `open scoped Classical` command, which enables the instance \
549-
`{.ofConstName ``Classical.propDecidable}`."
550-
else
551-
MessageData.nil
504+
catch ex =>
505+
-- Diagnose the failure, lazily so that there is no performance impact if `decide` isn't being used interactively.
506+
throwError MessageData.ofLazyM (es := #[expectedType]) do
507+
let r ← withAtLeastTransparency .default <| whnf s
508+
if r.isAppOf ``isTrue then
509+
return m!"\
510+
Tactic `{tacticName}` failed. The elaborator is able to reduce the \
511+
`{.ofConstName ``Decidable}` instance, but the kernel fails with:\n\
512+
{indentD ex.toMessageData}"
513+
diagnose expectedType s r
514+
515+
diagnose (expectedType s : Expr) (r : Expr) : MetaM MessageData := do
516+
if r.isAppOf ``isFalse then
552517
return m!"\
553-
Tactic `{tacticName}` failed for proposition\
518+
Tactic `{tacticName}` proved that the proposition\
554519
{indentExpr expectedType}\n\
555-
because its `{.ofConstName ``Decidable}` instance\
556-
{indentExpr s}\n\
557-
did not reduce to `{.ofConstName ``isTrue}` or `{.ofConstName ``isFalse}`.\n\n\
558-
{stuckMsg}{hint}"
520+
is false"
521+
-- Re-reduce the instance and collect diagnostics, to get all unfolded Decidable instances
522+
let (reason, unfoldedInsts) ← withoutModifyingState <| withOptions (fun opt => diagnostics.set opt true) do
523+
modifyDiag (fun _ => {})
524+
let reason ← withAtLeastTransparency .default <| blameDecideReductionFailure s
525+
let unfolded := (← get).diag.unfoldCounter.foldl (init := #[]) fun cs n _ => cs.push n
526+
let unfoldedInsts ← unfolded |>.qsort Name.lt |>.filterMapM fun n => do
527+
let e ← mkConstWithLevelParams n
528+
if (← Meta.isClass? (← inferType e)) == ``Decidable then
529+
return m!"`{.ofConst e}`"
530+
else
531+
return none
532+
return (reason, unfoldedInsts)
533+
let stuckMsg :=
534+
if unfoldedInsts.isEmpty then
535+
m!"Reduction got stuck at the `{.ofConstName ``Decidable}` instance{indentExpr reason}"
536+
else
537+
let instances := if unfoldedInsts.size == 1 then "instance" else "instances"
538+
m!"After unfolding the {instances} {.andList unfoldedInsts.toList}, \
539+
reduction got stuck at the `{.ofConstName ``Decidable}` instance{indentExpr reason}"
540+
let hint :=
541+
if reason.isAppOf ``Eq.rec then
542+
.hint' m!"Reduction got stuck on `▸` ({.ofConstName ``Eq.rec}), \
543+
which suggests that one of the `{.ofConstName ``Decidable}` instances is defined using tactics such as `rw` or `simp`. \
544+
To avoid tactics, make use of functions such as \
545+
`{.ofConstName ``inferInstanceAs}` or `{.ofConstName ``decidable_of_decidable_of_iff}` \
546+
to alter a proposition."
547+
else if reason.isAppOf ``Classical.choice then
548+
.hint' m!"Reduction got stuck on `{.ofConstName ``Classical.choice}`, \
549+
which indicates that a `{.ofConstName ``Decidable}` instance \
550+
is defined using classical reasoning, proving an instance exists rather than giving a concrete construction. \
551+
The `{tacticName}` tactic works by evaluating a decision procedure via reduction, \
552+
and it cannot make progress with such instances. \
553+
This can occur due to the `open scoped Classical` command, which enables the instance \
554+
`{.ofConstName ``Classical.propDecidable}`."
555+
else
556+
MessageData.nil
557+
return m!"\
558+
Tactic `{tacticName}` failed for proposition\
559+
{indentExpr expectedType}\n\
560+
because its `{.ofConstName ``Decidable}` instance\
561+
{indentExpr s}\n\
562+
did not reduce to `{.ofConstName ``isTrue}` or `{.ofConstName ``isFalse}`.\n\n\
563+
{stuckMsg}{hint}"
559564

560565
declare_config_elab elabDecideConfig Parser.Tactic.DecideConfig
561566

tests/lean/run/issue10766.lean

Lines changed: 39 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,39 @@
1+
noncomputable def powMod (a b n : Nat) : Nat :=
2+
aux (b + 1) (a.mod n) b 1
3+
where
4+
aux : Nat → ((a b c : Nat) → Nat) :=
5+
Nat.rec (fun _ _ _ => 0)
6+
(fun _ r a b c =>
7+
(b.beq 0).rec
8+
(((b.mod 2).beq 0).rec
9+
(r ((a.mul a).mod n) (b.div 2) ((a.mul c).mod n))
10+
(r ((a.mul a).mod n) (b.div 2) c))
11+
(c.mod n))
12+
13+
def x : Nat := 127780414391497973212171930170926986757577048484820926201064729783485263494817422495127775983679039078116803697168137524940219819335799478153348592755198599590903607242050230924443865709697486743641039970666450337071378658828331722728467720393963808366917988956767802913905167890490075236068196363700359481304279948916896583006686025357237170212018946813663108217900835975808683160984117514866915965161953626338070145596982334808959718966160701183250747572515090867613655044807172211728519357721287835503689517292364425608325467094686443862517374850243698013720305871319056887431952190952721719757200172695537054790570648290887720009455171821568413052107356003828041937567129362866696549587422369864562815134637684140271767482353107080370450890024342225936273158281477009232714640818424893445193089479459814572594522258577931514012256573162006292678354475638319009668319255772179069845291474717503333030909793536116894869761453687330048252587304656806182949368202671739705463406846852567720022377005763291104588535681445561286808586673846016527511475331939430139687698419185010117348285933672139833826832898565919546377321517928825162277951756632134321102813522053716838646284289
14+
15+
set_option maxRecDepth 20000
16+
17+
/--
18+
error: Tactic `decide` failed. The elaborator is able to reduce the `Decidable` instance, but the kernel fails with:
19+
20+
(kernel) deep recursion detected
21+
-/
22+
#guard_msgs in
23+
example : powMod 2 (x - 1) x = 1 := by decide +kernel
24+
25+
/-- error: (kernel) deep recursion detected -/
26+
#guard_msgs in
27+
example : powMod 2 (x - 1) x = 1 := by decide
28+
29+
/-- error: (kernel) deep recursion detected -/
30+
#guard_msgs in
31+
example : powMod 2 (x - 1) x = 1 := by rfl
32+
33+
set_option debug.skipKernelTC true
34+
35+
#guard_msgs in
36+
example : powMod 2 (x - 1) x = 1 := by decide
37+
38+
#guard_msgs in
39+
example : powMod 2 (x - 1) x = 1 := by rfl

0 commit comments

Comments
 (0)