summaryrefslogtreecommitdiff
path: root/backends/lean/Base/Progress/Progress.lean
diff options
context:
space:
mode:
Diffstat (limited to 'backends/lean/Base/Progress/Progress.lean')
-rw-r--r--backends/lean/Base/Progress/Progress.lean35
1 files changed, 26 insertions, 9 deletions
diff --git a/backends/lean/Base/Progress/Progress.lean b/backends/lean/Base/Progress/Progress.lean
index 6a4729dc..8b0759c5 100644
--- a/backends/lean/Base/Progress/Progress.lean
+++ b/backends/lean/Base/Progress/Progress.lean
@@ -110,8 +110,9 @@ def progressWith (fExpr : Expr) (th : TheoremOrLocal)
-- then continue splitting the post-condition
splitEqAndPost fun hEq hPost ids => do
trace[Progress] "eq and post:\n{hEq} : {← inferType hEq}\n{hPost}"
- simpAt [] [``Primitives.bind_tc_ret, ``Primitives.bind_tc_fail, ``Primitives.bind_tc_div]
- [hEq.fvarId!] (.targets #[] true)
+ tryTac (
+ simpAt [] [``Primitives.bind_tc_ret, ``Primitives.bind_tc_fail, ``Primitives.bind_tc_div]
+ [hEq.fvarId!] (.targets #[] true))
-- Clear the equality, unless the user requests not to do so
let mgoal ← do
if keep.isSome then getMainGoal
@@ -242,21 +243,26 @@ def progressAsmsOrLookupTheorem (keep : Option Name) (withTh : Option TheoremOrL
tryLookupApply keep ids splitPost asmTac fExpr "pspec theorem" pspec do
-- It failed: try to lookup a *class* expr spec theorem (those are more
-- specific than class spec theorems)
+ trace[Progress] "Failed using a pspec theorem: trying to lookup a pspec class expr theorem"
let pspecClassExpr ← do
match getFirstArg args with
| none => pure none
| some arg => do
+ trace[Progress] "Using: f:{fName}, arg: {arg}"
let thName ← pspecClassExprAttr.find? fName arg
pure (thName.map fun th => .Theorem th)
tryLookupApply keep ids splitPost asmTac fExpr "pspec class expr theorem" pspecClassExpr do
-- It failed: try to lookup a *class* spec theorem
+ trace[Progress] "Failed using a pspec class expr theorem: trying to lookup a pspec class theorem"
let pspecClass ← do
match ← getFirstArgAppName args with
| none => pure none
| some argName => do
+ trace[Progress] "Using: f: {fName}, arg: {argName}"
let thName ← pspecClassAttr.find? fName argName
pure (thName.map fun th => .Theorem th)
tryLookupApply keep ids splitPost asmTac fExpr "pspec class theorem" pspecClass do
+ trace[Progress] "Failed using a pspec class theorem: trying to use a recursive assumption"
-- Try a recursive call - we try the assumptions of kind "auxDecl"
let ctx ← Lean.MonadLCtx.getLCtx
let decls ← ctx.getAllDecls
@@ -314,12 +320,14 @@ def evalProgress (args : TSyntax `Progress.progressArgs) : TacticM Unit := do
else pure none
let ids :=
let args := asArgs.getArgs
- let args := (args.get! 2).getSepArgs
- args.map (λ s => if s.isIdent then some s.getId else none)
+ if args.size > 2 then
+ let args := (args.get! 2).getSepArgs
+ args.map (λ s => if s.isIdent then some s.getId else none)
+ else #[]
trace[Progress] "User-provided ids: {ids}"
let splitPost : Bool :=
let args := asArgs.getArgs
- (args.get! 3).getArgs.size > 0
+ args.size > 3 ∧ (args.get! 3).getArgs.size > 0
trace[Progress] "Split post: {splitPost}"
/- For scalarTac we have a fast track: if the goal is not a linear
arithmetic goal, we skip (note that otherwise, scalarTac would try
@@ -343,11 +351,14 @@ elab "progress" args:progressArgs : tactic =>
namespace Test
open Primitives Result
- set_option trace.Progress true
- set_option pp.rawOnError true
+ -- Show the traces
+ -- set_option trace.Progress true
+ -- set_option pp.rawOnError true
- #eval showStoredPSpec
- #eval showStoredPSpecClass
+ -- The following commands display the databases of theorems
+ -- #eval showStoredPSpec
+ -- #eval showStoredPSpecClass
+ -- #eval showStoredPSpecExprClass
example {ty} {x y : Scalar ty}
(hmin : Scalar.min ty ≤ x.val + y.val)
@@ -363,6 +374,12 @@ namespace Test
progress keep h with Scalar.add_spec as ⟨ z ⟩
simp [*, h]
+ example {x y : U32}
+ (hmax : x.val + y.val ≤ U32.max) :
+ ∃ z, x + y = ret z ∧ z.val = x.val + y.val := by
+ progress keep _ as ⟨ z, h1 .. ⟩
+ simp [*, h1]
+
/- Checking that universe instantiation works: the original spec uses
`α : Type u` where u is quantified, while here we use `α : Type 0` -/
example {α : Type} (v: Vec α) (i: Usize) (x : α)