Skip to content

Commit

Permalink
test: test case for #4751
Browse files Browse the repository at this point in the history
and tracing for `IndPredBelow.backwardsChaining`.
  • Loading branch information
nomeata committed Jul 24, 2024
1 parent 6d97182 commit a8f09ff
Show file tree
Hide file tree
Showing 2 changed files with 83 additions and 15 deletions.
37 changes: 22 additions & 15 deletions src/Lean/Meta/IndPredBelow.lean
Original file line number Diff line number Diff line change
Expand Up @@ -230,22 +230,28 @@ def mkBelowDecl (ctx : Context) : MetaM Declaration := do
ctx.typeInfos[0]!.isUnsafe

partial def backwardsChaining (m : MVarId) (depth : Nat) : MetaM Bool := do
if depth = 0 then return false
else
m.withContext do
let lctx ← getLCtx
m.withContext do
let mTy ← m.getType
lctx.anyM fun localDecl =>
if localDecl.isAuxDecl then
return false
else
commitWhen do
let (mvars, _, t) ← forallMetaTelescope localDecl.type
if ←isDefEq mTy t then
m.assign (mkAppN localDecl.toExpr mvars)
mvars.allM fun v =>
v.mvarId!.isAssigned <||> backwardsChaining v.mvarId! (depth - 1)
else return false
if depth = 0 then
trace[Meta.IndPredBelow.search] "searching for {mTy}: ran out of max depth"
return false
else
let lctx ← getLCtx
let r ← lctx.anyM fun localDecl =>
if localDecl.isAuxDecl then
return false
else
commitWhen do
let (mvars, _, t) ← forallMetaTelescope localDecl.type
if (← isDefEq mTy t) then
trace[Meta.IndPredBelow.search] "searching for {mTy}: trying {mkFVar localDecl.fvarId} : {localDecl.type}"
m.assign (mkAppN localDecl.toExpr mvars)
mvars.allM fun v =>
v.mvarId!.isAssigned <||> backwardsChaining v.mvarId! (depth - 1)
else return false
unless r do
trace[Meta.IndPredBelow.search] "searching for {mTy} failed"
return r

partial def proveBrecOn (ctx : Context) (indVal : InductiveVal) (type : Expr) : MetaM Expr := do
let main ← mkFreshExprSyntheticOpaqueMVar type
Expand Down Expand Up @@ -596,5 +602,6 @@ def mkBelow (declName : Name) : MetaM Unit := do
builtin_initialize
registerTraceClass `Meta.IndPredBelow
registerTraceClass `Meta.IndPredBelow.match
registerTraceClass `Meta.IndPredBelow.search

end Lean.Meta.IndPredBelow
61 changes: 61 additions & 0 deletions tests/lean/run/4751.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,61 @@
inductive F: Prop where
| base
| step (fn: Nat → F)

-- set_option trace.Meta.IndPredBelow.search true
set_option pp.proofs true

/--
error: failed to infer structural recursion:
Cannot use parameter #1:
could not solve using backwards chaining x✝¹ : F
x✝ : x✝¹.below
f : Nat → F
a✝¹ : ∀ (a : Nat), (f a).below
a✝ : Nat → True
⊢ True
-/
#guard_msgs in
def F.asdf1 : (f : F) → True
| base => trivial
| step f => F.asdf1 (f 0)
termination_by structural f => f


def TTrue (_f : F) := True

/--
error: failed to infer structural recursion:
Cannot use parameter #1:
could not solve using backwards chaining x✝¹ : F
x✝ : x✝¹.below
f : Nat → F
a✝¹ : ∀ (a : Nat), (f a).below
a✝ : ∀ (a : Nat), TTrue (f a)
⊢ TTrue (f 0)
-/
#guard_msgs in
def F.asdf2 : (f : F) → TTrue f
| base => trivial
| step f => F.asdf2 (f 0)
termination_by structural f => f



inductive ITrue (f : F) : Prop where | trivial

/--
error: failed to infer structural recursion:
Cannot use parameter #1:
could not solve using backwards chaining x✝¹ : F
x✝ : x✝¹.below
f : Nat → F
a✝¹ : ∀ (a : Nat), (f a).below
a✝ : ∀ (a : Nat), ITrue (f a)
⊢ ITrue (f 0)
-/
#guard_msgs in
def F.asdf3 : (f : F) → ITrue f
| base => .trivial
| step f => F.asdf3 (f 0)
termination_by structural f => f

0 comments on commit a8f09ff

Please sign in to comment.