2023-05-21 17:41:39 -07:00
|
|
|
|
import Lean
|
|
|
|
|
|
2023-10-15 17:15:23 -07:00
|
|
|
|
import Pantograph.Symbol
|
2023-05-21 17:41:39 -07:00
|
|
|
|
|
2023-05-24 00:54:48 -07:00
|
|
|
|
def Lean.MessageLog.getErrorMessages (log : MessageLog) : MessageLog :=
|
2023-08-14 17:07:53 -07:00
|
|
|
|
{
|
|
|
|
|
msgs := log.msgs.filter fun m => match m.severity with | MessageSeverity.error => true | _ => false
|
|
|
|
|
}
|
2023-05-21 17:41:39 -07:00
|
|
|
|
|
|
|
|
|
|
2023-05-24 00:54:48 -07:00
|
|
|
|
namespace Pantograph
|
|
|
|
|
open Lean
|
2023-05-21 17:41:39 -07:00
|
|
|
|
|
2023-08-27 19:53:09 -07:00
|
|
|
|
structure GoalState where
|
2023-05-24 00:54:48 -07:00
|
|
|
|
savedState : Elab.Tactic.SavedState
|
2023-05-21 17:41:39 -07:00
|
|
|
|
|
2023-10-15 17:15:23 -07:00
|
|
|
|
-- The root hole which is the search target
|
|
|
|
|
root: MVarId
|
|
|
|
|
-- New metavariables acquired in this state
|
|
|
|
|
newMVars: SSet MVarId
|
|
|
|
|
|
2023-10-26 22:47:42 -07:00
|
|
|
|
-- The id of the goal in the parent
|
|
|
|
|
parentGoalId: Nat := 0
|
|
|
|
|
|
2023-05-24 00:54:48 -07:00
|
|
|
|
abbrev M := Elab.TermElabM
|
2023-05-21 17:41:39 -07:00
|
|
|
|
|
2023-10-15 17:15:23 -07:00
|
|
|
|
protected def GoalState.create (expr: Expr): M GoalState := do
|
|
|
|
|
-- May be necessary to immediately synthesise all metavariables if we need to leave the elaboration context.
|
2023-10-15 12:31:22 -07:00
|
|
|
|
-- See https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/Unknown.20universe.20metavariable/near/360130070
|
2023-10-15 17:15:23 -07:00
|
|
|
|
|
2023-10-15 12:31:22 -07:00
|
|
|
|
--Elab.Term.synthesizeSyntheticMVarsNoPostponing
|
2023-10-15 17:15:23 -07:00
|
|
|
|
--let expr ← instantiateMVars expr
|
2023-10-15 12:31:22 -07:00
|
|
|
|
let goal := (← Meta.mkFreshExprMVar expr (kind := MetavarKind.synthetic) (userName := .anonymous))
|
2023-05-24 00:54:48 -07:00
|
|
|
|
let savedStateMonad: Elab.Tactic.TacticM Elab.Tactic.SavedState := MonadBacktrack.saveState
|
2023-10-25 16:03:45 -07:00
|
|
|
|
let root := goal.mvarId!
|
|
|
|
|
let savedState ← savedStateMonad { elaborator := .anonymous } |>.run' { goals := [root]}
|
2023-05-23 05:12:46 -07:00
|
|
|
|
return {
|
2023-10-15 12:31:22 -07:00
|
|
|
|
savedState,
|
2023-10-25 16:03:45 -07:00
|
|
|
|
root,
|
|
|
|
|
newMVars := SSet.insert .empty root,
|
2023-05-21 17:41:39 -07:00
|
|
|
|
}
|
2023-10-26 11:22:02 -07:00
|
|
|
|
protected def GoalState.goals (state: GoalState): List MVarId := state.savedState.tactic.goals
|
2023-05-21 17:41:39 -07:00
|
|
|
|
|
2023-10-26 11:22:02 -07:00
|
|
|
|
protected def GoalState.runM {α: Type} (state: GoalState) (m: Elab.TermElabM α) : M α := do
|
|
|
|
|
state.savedState.term.restore
|
|
|
|
|
m
|
|
|
|
|
|
|
|
|
|
protected def GoalState.mctx (state: GoalState): MetavarContext :=
|
2023-10-25 16:03:45 -07:00
|
|
|
|
state.savedState.term.meta.meta.mctx
|
2023-10-27 15:15:22 -07:00
|
|
|
|
protected def GoalState.env (state: GoalState): Environment :=
|
|
|
|
|
state.savedState.term.meta.core.env
|
2023-10-25 16:03:45 -07:00
|
|
|
|
private def GoalState.mvars (state: GoalState): SSet MVarId :=
|
|
|
|
|
state.mctx.decls.foldl (init := .empty) fun acc k _ => acc.insert k
|
|
|
|
|
|
2023-10-26 22:47:42 -07:00
|
|
|
|
/-- Inner function for executing tactic on goal state -/
|
2023-10-15 17:15:23 -07:00
|
|
|
|
def executeTactic (state: Elab.Tactic.SavedState) (goal: MVarId) (tactic: Syntax) :
|
2023-10-27 15:15:22 -07:00
|
|
|
|
M (Except (Array String) Elab.Tactic.SavedState):= do
|
|
|
|
|
let tacticM (stx: Syntax): Elab.Tactic.TacticM (Except (Array String) Elab.Tactic.SavedState) := do
|
2023-05-23 05:12:46 -07:00
|
|
|
|
state.restore
|
2023-05-24 00:54:48 -07:00
|
|
|
|
Elab.Tactic.setGoals [goal]
|
2023-05-21 17:41:39 -07:00
|
|
|
|
try
|
2023-05-24 00:54:48 -07:00
|
|
|
|
Elab.Tactic.evalTactic stx
|
|
|
|
|
if (← getThe Core.State).messages.hasErrors then
|
|
|
|
|
let messages := (← getThe Core.State).messages.getErrorMessages |>.toList.toArray
|
|
|
|
|
let errors ← (messages.map Message.data).mapM fun md => md.toString
|
2023-05-23 05:12:46 -07:00
|
|
|
|
return .error errors
|
|
|
|
|
else
|
2023-10-27 15:15:22 -07:00
|
|
|
|
return .ok (← MonadBacktrack.saveState)
|
2023-05-23 05:12:46 -07:00
|
|
|
|
catch exception =>
|
|
|
|
|
return .error #[← exception.toMessageData.toString]
|
2023-10-15 17:15:23 -07:00
|
|
|
|
tacticM tactic { elaborator := .anonymous } |>.run' state.tactic
|
2023-05-21 17:41:39 -07:00
|
|
|
|
|
|
|
|
|
/-- Response for executing a tactic -/
|
|
|
|
|
inductive TacticResult where
|
|
|
|
|
-- Goes to next state
|
2023-10-26 22:47:42 -07:00
|
|
|
|
| success (state: GoalState)
|
2023-10-15 17:15:23 -07:00
|
|
|
|
-- Tactic failed with messages
|
2023-05-21 17:41:39 -07:00
|
|
|
|
| failure (messages: Array String)
|
2023-10-15 17:15:23 -07:00
|
|
|
|
-- Could not parse tactic
|
|
|
|
|
| parseError (message: String)
|
|
|
|
|
-- The goal index is out of bounds
|
|
|
|
|
| indexError (goalId: Nat)
|
2023-08-27 19:53:09 -07:00
|
|
|
|
|
2023-05-21 17:41:39 -07:00
|
|
|
|
/-- Execute tactic on given state -/
|
2023-10-15 17:15:23 -07:00
|
|
|
|
protected def GoalState.execute (state: GoalState) (goalId: Nat) (tactic: String):
|
2023-10-26 22:47:42 -07:00
|
|
|
|
M TacticResult := do
|
2023-10-15 17:15:23 -07:00
|
|
|
|
let goal ← match state.savedState.tactic.goals.get? goalId with
|
|
|
|
|
| .some goal => pure $ goal
|
|
|
|
|
| .none => return .indexError goalId
|
|
|
|
|
let tactic ← match Parser.runParserCategory
|
|
|
|
|
(env := ← MonadEnv.getEnv)
|
|
|
|
|
(catName := `tactic)
|
|
|
|
|
(input := tactic)
|
|
|
|
|
(fileName := "<stdin>") with
|
|
|
|
|
| .ok stx => pure $ stx
|
|
|
|
|
| .error error => return .parseError error
|
|
|
|
|
match (← executeTactic (state := state.savedState) (goal := goal) (tactic := tactic)) with
|
2023-08-27 19:53:09 -07:00
|
|
|
|
| .error errors =>
|
|
|
|
|
return .failure errors
|
2023-10-27 15:15:22 -07:00
|
|
|
|
| .ok nextSavedState =>
|
2023-10-15 17:15:23 -07:00
|
|
|
|
-- Assert that the definition of metavariables are the same
|
|
|
|
|
let nextMCtx := nextSavedState.term.meta.meta.mctx
|
|
|
|
|
let prevMCtx := state.savedState.term.meta.meta.mctx
|
|
|
|
|
-- Generate a list of mvarIds that exist in the parent state; Also test the
|
|
|
|
|
-- assertion that the types have not changed on any mvars.
|
2023-10-25 16:03:45 -07:00
|
|
|
|
let newMVars ← nextMCtx.decls.foldlM (fun acc mvarId mvarDecl => do
|
2023-10-15 17:15:23 -07:00
|
|
|
|
if let .some prevMVarDecl := prevMCtx.decls.find? mvarId then
|
|
|
|
|
assert! prevMVarDecl.type == mvarDecl.type
|
|
|
|
|
return acc
|
|
|
|
|
else
|
2023-10-25 16:03:45 -07:00
|
|
|
|
return acc.insert mvarId
|
|
|
|
|
) SSet.empty
|
2023-10-26 22:47:42 -07:00
|
|
|
|
return .success {
|
2023-10-27 15:15:22 -07:00
|
|
|
|
state with
|
2023-10-15 17:15:23 -07:00
|
|
|
|
savedState := nextSavedState
|
|
|
|
|
newMVars,
|
2023-10-26 22:47:42 -07:00
|
|
|
|
parentGoalId := goalId,
|
2023-10-15 17:15:23 -07:00
|
|
|
|
}
|
|
|
|
|
|
2023-10-27 15:15:22 -07:00
|
|
|
|
protected def GoalState.tryAssign (state: GoalState) (goalId: Nat) (expr: String): M TacticResult := do
|
|
|
|
|
let goal ← match state.savedState.tactic.goals.get? goalId with
|
|
|
|
|
| .some goal => pure goal
|
|
|
|
|
| .none => return .indexError goalId
|
|
|
|
|
let expr ← match Parser.runParserCategory
|
|
|
|
|
(env := state.env)
|
|
|
|
|
(catName := `term)
|
|
|
|
|
(input := expr)
|
|
|
|
|
(fileName := "<stdin>") with
|
|
|
|
|
| .ok syn => pure syn
|
|
|
|
|
| .error error => return .parseError error
|
|
|
|
|
let tacticM: Elab.Tactic.TacticM TacticResult := do
|
|
|
|
|
state.savedState.restore
|
|
|
|
|
Elab.Tactic.setGoals [goal]
|
|
|
|
|
try
|
|
|
|
|
let expr ← Elab.Term.elabTerm (stx := expr) (expectedType? := .none)
|
|
|
|
|
-- Attempt to unify the expression
|
|
|
|
|
let goalType ← goal.getType
|
|
|
|
|
let exprType ← Meta.inferType expr
|
|
|
|
|
if !(← Meta.isDefEq goalType exprType) then
|
|
|
|
|
return .failure #["Type unification failed", toString (← Meta.ppExpr goalType), toString (← Meta.ppExpr exprType)]
|
|
|
|
|
goal.checkNotAssigned `GoalState.tryAssign
|
|
|
|
|
goal.assign expr
|
|
|
|
|
if (← getThe Core.State).messages.hasErrors then
|
|
|
|
|
let messages := (← getThe Core.State).messages.getErrorMessages |>.toList.toArray
|
|
|
|
|
let errors ← (messages.map Message.data).mapM fun md => md.toString
|
|
|
|
|
return .failure errors
|
|
|
|
|
else
|
|
|
|
|
let prevMCtx := state.savedState.term.meta.meta.mctx
|
|
|
|
|
let nextMCtx ← getMCtx
|
|
|
|
|
-- Generate a list of mvarIds that exist in the parent state; Also test the
|
|
|
|
|
-- assertion that the types have not changed on any mvars.
|
|
|
|
|
let newMVars ← nextMCtx.decls.foldlM (fun acc mvarId mvarDecl => do
|
|
|
|
|
if let .some prevMVarDecl := prevMCtx.decls.find? mvarId then
|
|
|
|
|
assert! prevMVarDecl.type == mvarDecl.type
|
|
|
|
|
return acc
|
|
|
|
|
else
|
|
|
|
|
return mvarId :: acc
|
|
|
|
|
) []
|
|
|
|
|
-- The new goals are the newMVars that lack an assignment
|
|
|
|
|
Elab.Tactic.setGoals (← newMVars.filterM (λ mvar => do pure !(← mvar.isAssigned)))
|
|
|
|
|
let nextSavedState ← MonadBacktrack.saveState
|
|
|
|
|
return .success {
|
|
|
|
|
state with
|
|
|
|
|
savedState := nextSavedState,
|
|
|
|
|
newMVars := newMVars.toSSet,
|
|
|
|
|
parentGoalId := goalId,
|
|
|
|
|
}
|
|
|
|
|
catch exception =>
|
|
|
|
|
return .failure #[← exception.toMessageData.toString]
|
|
|
|
|
tacticM { elaborator := .anonymous } |>.run' state.savedState.tactic
|
|
|
|
|
|
2023-11-04 15:33:53 -07:00
|
|
|
|
/--
|
|
|
|
|
Brings into scope a list of goals
|
|
|
|
|
-/
|
|
|
|
|
protected def GoalState.resume (state: GoalState) (goals: List MVarId): Except String GoalState :=
|
|
|
|
|
if ¬ (goals.all (λ goal => state.mvars.contains goal)) then
|
|
|
|
|
.error s!"Goals not in scope"
|
2023-10-25 16:03:45 -07:00
|
|
|
|
else
|
|
|
|
|
-- Set goals to the goals that have not been assigned yet, similar to the `focus` tactic.
|
2023-11-04 15:00:51 -07:00
|
|
|
|
let unassigned := goals.filter (λ goal =>
|
2023-11-04 15:33:53 -07:00
|
|
|
|
let mctx := state.mctx
|
2023-10-25 16:03:45 -07:00
|
|
|
|
¬(mctx.eAssignment.contains goal || mctx.dAssignment.contains goal))
|
|
|
|
|
.ok {
|
2023-11-04 15:33:53 -07:00
|
|
|
|
state with
|
2023-10-25 16:03:45 -07:00
|
|
|
|
savedState := {
|
2023-11-04 15:33:53 -07:00
|
|
|
|
term := state.savedState.term,
|
2023-10-25 16:03:45 -07:00
|
|
|
|
tactic := { goals := unassigned },
|
|
|
|
|
},
|
|
|
|
|
}
|
|
|
|
|
|
2023-11-04 15:33:53 -07:00
|
|
|
|
/--
|
|
|
|
|
Brings into scope all goals from `branch`
|
|
|
|
|
-/
|
|
|
|
|
protected def GoalState.continue (target: GoalState) (branch: GoalState): Except String GoalState :=
|
2023-11-04 15:53:57 -07:00
|
|
|
|
if !target.goals.isEmpty then
|
|
|
|
|
.error s!"Target state has unresolved goals"
|
|
|
|
|
else if target.root != branch.root then
|
2023-11-04 15:33:53 -07:00
|
|
|
|
.error s!"Roots of two continued goal states do not match: {target.root.name} != {branch.root.name}"
|
|
|
|
|
else
|
|
|
|
|
target.resume (goals := branch.goals)
|
|
|
|
|
|
2023-10-27 15:41:12 -07:00
|
|
|
|
protected def GoalState.rootExpr? (goalState: GoalState): Option Expr :=
|
2023-10-25 22:18:59 -07:00
|
|
|
|
let expr := goalState.mctx.eAssignment.find! goalState.root
|
|
|
|
|
let (expr, _) := instantiateMVarsCore (mctx := goalState.mctx) (e := expr)
|
|
|
|
|
if expr.hasMVar then
|
2023-10-26 22:47:42 -07:00
|
|
|
|
-- Must not assert that the goal state is empty here. We could be in a branch goal.
|
|
|
|
|
--assert! ¬goalState.goals.isEmpty
|
2023-10-25 22:18:59 -07:00
|
|
|
|
.none
|
|
|
|
|
else
|
2023-10-26 22:47:42 -07:00
|
|
|
|
assert! goalState.goals.isEmpty
|
2023-10-25 22:18:59 -07:00
|
|
|
|
.some expr
|
2023-10-25 16:03:45 -07:00
|
|
|
|
|
2023-05-24 00:54:48 -07:00
|
|
|
|
end Pantograph
|