Compare commits
3 Commits
dev
...
frontend/t
Author | SHA1 | Date |
---|---|---|
|
53ea808c90 | |
|
1a9e450066 | |
|
d7001c1b28 |
|
@ -454,6 +454,7 @@ def serializeExpression (options: @&Protocol.Options) (e: Expr): MetaM Protocol.
|
||||||
dependentMVars?,
|
dependentMVars?,
|
||||||
}
|
}
|
||||||
|
|
||||||
|
|
||||||
/-- Adapted from ppGoal -/
|
/-- Adapted from ppGoal -/
|
||||||
def serializeGoal (options: @&Protocol.Options) (goal: MVarId) (mvarDecl: MetavarDecl) (parentDecl?: Option MetavarDecl := .none)
|
def serializeGoal (options: @&Protocol.Options) (goal: MVarId) (mvarDecl: MetavarDecl) (parentDecl?: Option MetavarDecl := .none)
|
||||||
: MetaM Protocol.Goal := do
|
: MetaM Protocol.Goal := do
|
||||||
|
@ -519,6 +520,7 @@ def serializeGoal (options: @&Protocol.Options) (goal: MVarId) (mvarDecl: Metava
|
||||||
return {
|
return {
|
||||||
name := goal.name.toString,
|
name := goal.name.toString,
|
||||||
userName? := if mvarDecl.userName == .anonymous then .none else .some (ofName mvarDecl.userName),
|
userName? := if mvarDecl.userName == .anonymous then .none else .some (ofName mvarDecl.userName),
|
||||||
|
isConversion := isLHSGoal? mvarDecl.type |>.isSome,
|
||||||
target := (← serializeExpression options (← instantiate mvarDecl.type)),
|
target := (← serializeExpression options (← instantiate mvarDecl.type)),
|
||||||
vars := vars.reverse.toArray
|
vars := vars.reverse.toArray
|
||||||
}
|
}
|
||||||
|
@ -528,20 +530,17 @@ def serializeGoal (options: @&Protocol.Options) (goal: MVarId) (mvarDecl: Metava
|
||||||
|
|
||||||
protected def GoalState.serializeGoals
|
protected def GoalState.serializeGoals
|
||||||
(state: GoalState)
|
(state: GoalState)
|
||||||
|
(parent: Option GoalState := .none)
|
||||||
(options: @&Protocol.Options := {}):
|
(options: @&Protocol.Options := {}):
|
||||||
MetaM (Array Protocol.Goal):= do
|
MetaM (Array Protocol.Goal):= do
|
||||||
state.restoreMetaM
|
state.restoreMetaM
|
||||||
let goals := state.goals.toArray
|
let goals := state.goals.toArray
|
||||||
|
let parentDecl? := parent.bind (λ parentState => parentState.mctx.findDecl? state.parentMVar?.get!)
|
||||||
goals.mapM fun goal => do
|
goals.mapM fun goal => do
|
||||||
let fragment := match state.fragments[goal]? with
|
|
||||||
| .none => .tactic
|
|
||||||
| .some $ .calc .. => .calc
|
|
||||||
| .some $ .conv .. => .conv
|
|
||||||
| .some $ .convSentinel .. => .conv
|
|
||||||
match state.mctx.findDecl? goal with
|
match state.mctx.findDecl? goal with
|
||||||
| .some mvarDecl =>
|
| .some mvarDecl =>
|
||||||
let serializedGoal ← serializeGoal options goal mvarDecl (parentDecl? := .none)
|
let serializedGoal ← serializeGoal options goal mvarDecl (parentDecl? := parentDecl?)
|
||||||
pure { serializedGoal with fragment }
|
pure serializedGoal
|
||||||
| .none => throwError s!"Metavariable does not exist in context {goal.name}"
|
| .none => throwError s!"Metavariable does not exist in context {goal.name}"
|
||||||
|
|
||||||
/-- Print the metavariables in a readable format -/
|
/-- Print the metavariables in a readable format -/
|
||||||
|
@ -607,9 +606,7 @@ protected def GoalState.diag (goalState: GoalState) (parent?: Option GoalState :
|
||||||
userNameToString : Name → String
|
userNameToString : Name → String
|
||||||
| .anonymous => ""
|
| .anonymous => ""
|
||||||
| other => s!"[{other}]"
|
| other => s!"[{other}]"
|
||||||
parentHasMVar (mvarId: MVarId): Bool := match parent? with
|
parentHasMVar (mvarId: MVarId): Bool := parent?.map (λ state => state.mctx.decls.contains mvarId) |>.getD true
|
||||||
| .some state => state.mctx.decls.contains mvarId
|
|
||||||
| .none => true
|
|
||||||
|
|
||||||
initialize
|
initialize
|
||||||
registerTraceClass `Pantograph.Delate
|
registerTraceClass `Pantograph.Delate
|
||||||
|
|
|
@ -10,7 +10,7 @@ namespace Lean.Elab
|
||||||
private def elaboratorToString : Name → String
|
private def elaboratorToString : Name → String
|
||||||
| .anonymous => ""
|
| .anonymous => ""
|
||||||
| n => s!"⟨{n}⟩ "
|
| n => s!"⟨{n}⟩ "
|
||||||
private def indent (s : String) : String := "\n".intercalate $ s.splitOn "\n" |>.map ("\t" ++ .)
|
private def indent (s : String) : String := "\n".intercalate $ s.splitOn "\n" |>.map (" " ++ .)
|
||||||
|
|
||||||
/-- The `Syntax` for a `Lean.Elab.Info`, if there is one. -/
|
/-- The `Syntax` for a `Lean.Elab.Info`, if there is one. -/
|
||||||
protected def Info.stx? : Info → Option Syntax
|
protected def Info.stx? : Info → Option Syntax
|
||||||
|
|
|
@ -10,69 +10,22 @@ import Lean
|
||||||
namespace Pantograph
|
namespace Pantograph
|
||||||
open Lean
|
open Lean
|
||||||
|
|
||||||
/-- The acting area of a tactic -/
|
|
||||||
inductive Site where
|
|
||||||
-- Dormant all other goals
|
|
||||||
| focus (goal : MVarId)
|
|
||||||
-- Move the goal to the first in the list
|
|
||||||
| prefer (goal : MVarId)
|
|
||||||
-- Execute as-is, no goals go dormant
|
|
||||||
| unfocus
|
|
||||||
deriving BEq, Inhabited
|
|
||||||
|
|
||||||
instance : Coe MVarId Site where
|
|
||||||
coe := .focus
|
|
||||||
instance : ToString Site where
|
|
||||||
toString
|
|
||||||
| .focus { name } => s!"[{name}]"
|
|
||||||
| .prefer { name } => s!"[{name},...]"
|
|
||||||
| .unfocus => "[*]"
|
|
||||||
|
|
||||||
/-- Executes a `TacticM` on a site and return affected goals -/
|
|
||||||
protected def Site.runTacticM (site : Site)
|
|
||||||
{ m } [Monad m] [MonadLiftT Elab.Tactic.TacticM m] [MonadControlT Elab.Tactic.TacticM m] [MonadMCtx m] [MonadError m]
|
|
||||||
(f : m α) : m (α × List MVarId) :=
|
|
||||||
match site with
|
|
||||||
| .focus goal => do
|
|
||||||
Elab.Tactic.setGoals [goal]
|
|
||||||
let a ← f
|
|
||||||
return (a, [goal])
|
|
||||||
| .prefer goal => do
|
|
||||||
let before ← Elab.Tactic.getUnsolvedGoals
|
|
||||||
let otherGoals := before.filter (· != goal)
|
|
||||||
Elab.Tactic.setGoals (goal :: otherGoals)
|
|
||||||
let a ← f
|
|
||||||
let after ← Elab.Tactic.getUnsolvedGoals
|
|
||||||
let parents := before.filter (¬ after.contains ·)
|
|
||||||
Elab.Tactic.pruneSolvedGoals
|
|
||||||
return (a, parents)
|
|
||||||
| .unfocus => do
|
|
||||||
let before ← Elab.Tactic.getUnsolvedGoals
|
|
||||||
let a ← f
|
|
||||||
let after ← Elab.Tactic.getUnsolvedGoals
|
|
||||||
let parents := before.filter (¬ after.contains ·)
|
|
||||||
Elab.Tactic.pruneSolvedGoals
|
|
||||||
return (a, parents)
|
|
||||||
|
|
||||||
/--
|
/--
|
||||||
Kernel view of the state of a proof
|
Represents an interconnected set of metavariables, or a state in proof search
|
||||||
-/
|
-/
|
||||||
structure GoalState where
|
structure GoalState where
|
||||||
-- Captured `TacticM` state
|
|
||||||
savedState : Elab.Tactic.SavedState
|
savedState : Elab.Tactic.SavedState
|
||||||
|
|
||||||
-- The root goal which is the search target
|
-- The root hole which is the search target
|
||||||
root: MVarId
|
root: MVarId
|
||||||
|
|
||||||
-- Parent goals assigned to produce this state
|
-- Parent state metavariable source
|
||||||
parentMVars : List MVarId := []
|
parentMVar?: Option MVarId := .none
|
||||||
|
|
||||||
-- Any goal associated with a fragment has a partial tactic which has not
|
-- Any goal associated with a fragment has a partial tactic which has not
|
||||||
-- finished executing.
|
-- finished executing.
|
||||||
fragments : FragmentMap := .empty
|
fragments : FragmentMap := .empty
|
||||||
|
|
||||||
def throwNoGoals { m α } [Monad m] [MonadError m] : m α := throwError "no goals to be solved"
|
|
||||||
|
|
||||||
@[export pantograph_goal_state_create_m]
|
@[export pantograph_goal_state_create_m]
|
||||||
protected def GoalState.create (expr: Expr): Elab.TermElabM GoalState := do
|
protected def GoalState.create (expr: Expr): Elab.TermElabM GoalState := do
|
||||||
-- May be necessary to immediately synthesise all metavariables if we need to leave the elaboration context.
|
-- May be necessary to immediately synthesise all metavariables if we need to leave the elaboration context.
|
||||||
|
@ -86,6 +39,7 @@ protected def GoalState.create (expr: Expr): Elab.TermElabM GoalState := do
|
||||||
return {
|
return {
|
||||||
root := root.mvarId!,
|
root := root.mvarId!,
|
||||||
savedState,
|
savedState,
|
||||||
|
parentMVar? := .none,
|
||||||
}
|
}
|
||||||
@[export pantograph_goal_state_create_from_mvars_m]
|
@[export pantograph_goal_state_create_from_mvars_m]
|
||||||
protected def GoalState.createFromMVars (goals: List MVarId) (root: MVarId): MetaM GoalState := do
|
protected def GoalState.createFromMVars (goals: List MVarId) (root: MVarId): MetaM GoalState := do
|
||||||
|
@ -94,19 +48,10 @@ protected def GoalState.createFromMVars (goals: List MVarId) (root: MVarId): Met
|
||||||
return {
|
return {
|
||||||
root,
|
root,
|
||||||
savedState,
|
savedState,
|
||||||
|
parentMVar? := .none,
|
||||||
}
|
}
|
||||||
@[always_inline]
|
|
||||||
protected def GoalState.goals (state: GoalState): List MVarId :=
|
protected def GoalState.goals (state: GoalState): List MVarId :=
|
||||||
state.savedState.tactic.goals
|
state.savedState.tactic.goals
|
||||||
@[always_inline]
|
|
||||||
protected def GoalState.mainGoal? (state : GoalState) : Option MVarId :=
|
|
||||||
state.goals.head?
|
|
||||||
@[always_inline]
|
|
||||||
protected def GoalState.actingGoal? (state : GoalState) (site : Site) : Option MVarId := do
|
|
||||||
match site with
|
|
||||||
| .focus goal | .prefer goal => return goal
|
|
||||||
| .unfocus => state.mainGoal?
|
|
||||||
|
|
||||||
@[export pantograph_goal_state_goals]
|
@[export pantograph_goal_state_goals]
|
||||||
protected def GoalState.goalsArray (state: GoalState): Array MVarId := state.goals.toArray
|
protected def GoalState.goalsArray (state: GoalState): Array MVarId := state.goals.toArray
|
||||||
protected def GoalState.mctx (state: GoalState): MetavarContext :=
|
protected def GoalState.mctx (state: GoalState): MetavarContext :=
|
||||||
|
@ -118,10 +63,8 @@ protected def GoalState.env (state: GoalState): Environment :=
|
||||||
protected def GoalState.metaContextOfGoal (state: GoalState) (mvarId: MVarId): Option Meta.Context := do
|
protected def GoalState.metaContextOfGoal (state: GoalState) (mvarId: MVarId): Option Meta.Context := do
|
||||||
let mvarDecl ← state.mctx.findDecl? mvarId
|
let mvarDecl ← state.mctx.findDecl? mvarId
|
||||||
return { lctx := mvarDecl.lctx, localInstances := mvarDecl.localInstances }
|
return { lctx := mvarDecl.lctx, localInstances := mvarDecl.localInstances }
|
||||||
@[always_inline]
|
|
||||||
protected def GoalState.metaState (state: GoalState): Meta.State :=
|
protected def GoalState.metaState (state: GoalState): Meta.State :=
|
||||||
state.savedState.term.meta.meta
|
state.savedState.term.meta.meta
|
||||||
@[always_inline]
|
|
||||||
protected def GoalState.coreState (state: GoalState): Core.SavedState :=
|
protected def GoalState.coreState (state: GoalState): Core.SavedState :=
|
||||||
state.savedState.term.meta.core
|
state.savedState.term.meta.core
|
||||||
|
|
||||||
|
@ -129,18 +72,18 @@ protected def GoalState.withContext' (state: GoalState) (mvarId: MVarId) (m: Met
|
||||||
mvarId.withContext m |>.run' (← read) state.metaState
|
mvarId.withContext m |>.run' (← read) state.metaState
|
||||||
protected def GoalState.withContext { m } [MonadControlT MetaM m] [Monad m] (state: GoalState) (mvarId: MVarId) : m α → m α :=
|
protected def GoalState.withContext { m } [MonadControlT MetaM m] [Monad m] (state: GoalState) (mvarId: MVarId) : m α → m α :=
|
||||||
Meta.mapMetaM <| state.withContext' mvarId
|
Meta.mapMetaM <| state.withContext' mvarId
|
||||||
/-- Uses context of the first parent -/
|
|
||||||
protected def GoalState.withParentContext { n } [MonadControlT MetaM n] [Monad n] (state: GoalState): n α → n α :=
|
protected def GoalState.withParentContext { n } [MonadControlT MetaM n] [Monad n] (state: GoalState): n α → n α :=
|
||||||
Meta.mapMetaM <| state.withContext' state.parentMVars[0]!
|
Meta.mapMetaM <| state.withContext' state.parentMVar?.get!
|
||||||
protected def GoalState.withRootContext { n } [MonadControlT MetaM n] [Monad n] (state: GoalState): n α → n α :=
|
protected def GoalState.withRootContext { n } [MonadControlT MetaM n] [Monad n] (state: GoalState): n α → n α :=
|
||||||
Meta.mapMetaM <| state.withContext' state.root
|
Meta.mapMetaM <| state.withContext' state.root
|
||||||
|
|
||||||
private def restoreCoreMExtra (state : Core.SavedState) : CoreM Unit :=
|
private def GoalState.mvars (state: GoalState): SSet MVarId :=
|
||||||
modifyGetThe Core.State (fun st => ((),
|
state.mctx.decls.foldl (init := .empty) fun acc k _ => acc.insert k
|
||||||
{ st with nextMacroScope := state.nextMacroScope, ngen := state.ngen }))
|
|
||||||
-- Restore the name generator and macro scopes of the core state
|
-- Restore the name generator and macro scopes of the core state
|
||||||
protected def GoalState.restoreCoreMExtra (state: GoalState): CoreM Unit :=
|
protected def GoalState.restoreCoreMExtra (state: GoalState): CoreM Unit := do
|
||||||
restoreCoreMExtra state.coreState
|
let savedCore := state.coreState
|
||||||
|
modifyGetThe Core.State (fun st => ((),
|
||||||
|
{ st with nextMacroScope := savedCore.nextMacroScope, ngen := savedCore.ngen }))
|
||||||
protected def GoalState.restoreMetaM (state: GoalState): MetaM Unit := do
|
protected def GoalState.restoreMetaM (state: GoalState): MetaM Unit := do
|
||||||
state.restoreCoreMExtra
|
state.restoreCoreMExtra
|
||||||
state.savedState.term.meta.restore
|
state.savedState.term.meta.restore
|
||||||
|
@ -151,13 +94,40 @@ private def GoalState.restoreTacticM (state: GoalState) (goal: MVarId): Elab.Tac
|
||||||
state.restoreElabM
|
state.restoreElabM
|
||||||
Elab.Tactic.setGoals [goal]
|
Elab.Tactic.setGoals [goal]
|
||||||
|
|
||||||
|
@[export pantograph_goal_state_focus]
|
||||||
|
protected def GoalState.focus (state : GoalState) (goal : MVarId): Option GoalState := do
|
||||||
|
return {
|
||||||
|
state with
|
||||||
|
savedState := {
|
||||||
|
state.savedState with
|
||||||
|
tactic := { goals := [goal] },
|
||||||
|
},
|
||||||
|
}
|
||||||
|
|
||||||
|
/-- Immediately bring all parent goals back into scope. Used in automatic mode -/
|
||||||
|
@[export pantograph_goal_state_immediate_resume]
|
||||||
|
protected def GoalState.immediateResume (state: GoalState) (parent: GoalState): GoalState :=
|
||||||
|
-- Prune parents solved goals
|
||||||
|
let mctx := state.mctx
|
||||||
|
let parentGoals := parent.goals.filter λ goal =>
|
||||||
|
let isDuplicate := state.goals.contains goal
|
||||||
|
let isSolved := mctx.eAssignment.contains goal || mctx.dAssignment.contains goal
|
||||||
|
(¬ isDuplicate) && (¬ isSolved)
|
||||||
|
{
|
||||||
|
state with
|
||||||
|
savedState := {
|
||||||
|
state.savedState with
|
||||||
|
tactic := { goals := state.goals ++ parentGoals },
|
||||||
|
},
|
||||||
|
}
|
||||||
|
|
||||||
/--
|
/--
|
||||||
Brings into scope a list of goals. User must ensure `goals` are distinct.
|
Brings into scope a list of goals. User must ensure `goals` are distinct.
|
||||||
-/
|
-/
|
||||||
@[export pantograph_goal_state_resume]
|
@[export pantograph_goal_state_resume]
|
||||||
protected def GoalState.resume (state : GoalState) (goals : List MVarId) : Except String GoalState := do
|
protected def GoalState.resume (state : GoalState) (goals : List MVarId) : Except String GoalState := do
|
||||||
if ¬ (goals.all (state.mctx.decls.contains ·)) then
|
if ¬ (goals.all (λ goal => state.mvars.contains goal)) then
|
||||||
let invalid_goals := goals.filter (λ goal => ¬ state.mctx.decls.contains goal) |>.map (·.name.toString)
|
let invalid_goals := goals.filter (λ goal => ¬ state.mvars.contains goal) |>.map (·.name.toString)
|
||||||
.error s!"Goals {invalid_goals} are not in scope"
|
.error s!"Goals {invalid_goals} are not in scope"
|
||||||
-- Set goals to the goals that have not been assigned yet, similar to the `focus` tactic.
|
-- Set goals to the goals that have not been assigned yet, similar to the `focus` tactic.
|
||||||
let unassigned := goals.filter λ goal =>
|
let unassigned := goals.filter λ goal =>
|
||||||
|
@ -195,21 +165,17 @@ protected def GoalState.isSolved (goalState : GoalState) : Bool :=
|
||||||
| .some e => ¬ e.hasExprMVar
|
| .some e => ¬ e.hasExprMVar
|
||||||
| .none => true
|
| .none => true
|
||||||
goalState.goals.isEmpty && solvedRoot
|
goalState.goals.isEmpty && solvedRoot
|
||||||
|
@[export pantograph_goal_state_parent_expr]
|
||||||
|
protected def GoalState.parentExpr? (goalState: GoalState): Option Expr := do
|
||||||
|
let parent ← goalState.parentMVar?
|
||||||
|
let expr := goalState.mctx.eAssignment.find! parent
|
||||||
|
let (expr, _) := instantiateMVarsCore (mctx := goalState.mctx) (e := expr)
|
||||||
|
return expr
|
||||||
@[export pantograph_goal_state_get_mvar_e_assignment]
|
@[export pantograph_goal_state_get_mvar_e_assignment]
|
||||||
protected def GoalState.getMVarEAssignment (goalState: GoalState) (mvarId: MVarId): Option Expr := do
|
protected def GoalState.getMVarEAssignment (goalState: GoalState) (mvarId: MVarId): Option Expr := do
|
||||||
let expr ← goalState.mctx.eAssignment.find? mvarId
|
let expr ← goalState.mctx.eAssignment.find? mvarId
|
||||||
let (expr, _) := instantiateMVarsCore (mctx := goalState.mctx) (e := expr)
|
let (expr, _) := instantiateMVarsCore (mctx := goalState.mctx) (e := expr)
|
||||||
return expr
|
return expr
|
||||||
@[export pantograph_goal_state_parent_exprs]
|
|
||||||
protected def GoalState.parentExprs (state : GoalState) : List Expr :=
|
|
||||||
state.parentMVars.map λ goal => state.getMVarEAssignment goal |>.get!
|
|
||||||
@[always_inline]
|
|
||||||
protected def GoalState.hasUniqueParent (state : GoalState) : Bool :=
|
|
||||||
state.parentMVars.length == 1
|
|
||||||
@[always_inline]
|
|
||||||
protected def GoalState.parentExpr! (state : GoalState) : Expr :=
|
|
||||||
assert! state.parentMVars.length == 1
|
|
||||||
(state.getMVarEAssignment state.parentMVars[0]!).get!
|
|
||||||
|
|
||||||
deriving instance BEq for DelayedMetavarAssignment
|
deriving instance BEq for DelayedMetavarAssignment
|
||||||
|
|
||||||
|
@ -221,8 +187,7 @@ if conflicting assignments exist. We also assume the monotonicity property: In a
|
||||||
chain of descending goal states, a mvar cannot be unassigned, and once assigned
|
chain of descending goal states, a mvar cannot be unassigned, and once assigned
|
||||||
its assignment cannot change. -/
|
its assignment cannot change. -/
|
||||||
@[export pantograph_goal_state_replay_m]
|
@[export pantograph_goal_state_replay_m]
|
||||||
protected def GoalState.replay (dst : GoalState) (src src' : GoalState) : CoreM GoalState :=
|
protected def GoalState.replay (dst : GoalState) (src src' : GoalState) : CoreM GoalState := do
|
||||||
withTraceNode `Pantograph.GoalState.replay (fun _ => return m!"replay") do
|
|
||||||
let srcNGen := src.coreState.ngen
|
let srcNGen := src.coreState.ngen
|
||||||
let srcNGen' := src'.coreState.ngen
|
let srcNGen' := src'.coreState.ngen
|
||||||
let dstNGen := dst.coreState.ngen
|
let dstNGen := dst.coreState.ngen
|
||||||
|
@ -246,8 +211,6 @@ protected def GoalState.replay (dst : GoalState) (src src' : GoalState) : CoreM
|
||||||
else
|
else
|
||||||
id
|
id
|
||||||
| id => id
|
| id => id
|
||||||
let mapMVar : MVarId → MVarId
|
|
||||||
| { name } => ⟨mapId name⟩
|
|
||||||
let rec mapLevel : Level → Level
|
let rec mapLevel : Level → Level
|
||||||
| .succ x => .succ (mapLevel x)
|
| .succ x => .succ (mapLevel x)
|
||||||
| .max l1 l2 => .max (mapLevel l1) (mapLevel l2)
|
| .max l1 l2 => .max (mapLevel l1) (mapLevel l2)
|
||||||
|
@ -261,7 +224,7 @@ protected def GoalState.replay (dst : GoalState) (src src' : GoalState) : CoreM
|
||||||
let mapDelayedAssignment (d : DelayedMetavarAssignment) : CoreM DelayedMetavarAssignment := do
|
let mapDelayedAssignment (d : DelayedMetavarAssignment) : CoreM DelayedMetavarAssignment := do
|
||||||
let { mvarIdPending, fvars } := d
|
let { mvarIdPending, fvars } := d
|
||||||
return {
|
return {
|
||||||
mvarIdPending := mapMVar mvarIdPending,
|
mvarIdPending := ⟨mapId mvarIdPending.name⟩,
|
||||||
fvars := ← fvars.mapM mapExpr,
|
fvars := ← fvars.mapM mapExpr,
|
||||||
}
|
}
|
||||||
let mapLocalDecl (ldecl : LocalDecl) : CoreM LocalDecl := do
|
let mapLocalDecl (ldecl : LocalDecl) : CoreM LocalDecl := do
|
||||||
|
@ -342,7 +305,6 @@ protected def GoalState.replay (dst : GoalState) (src src' : GoalState) : CoreM
|
||||||
}
|
}
|
||||||
}
|
}
|
||||||
let m : MetaM Meta.SavedState := Meta.withMCtx mctx do
|
let m : MetaM Meta.SavedState := Meta.withMCtx mctx do
|
||||||
restoreCoreMExtra savedMeta.core
|
|
||||||
savedMeta.restore
|
savedMeta.restore
|
||||||
|
|
||||||
for (lmvarId, l') in src'.mctx.lAssignment do
|
for (lmvarId, l') in src'.mctx.lAssignment do
|
||||||
|
@ -374,7 +336,7 @@ protected def GoalState.replay (dst : GoalState) (src src' : GoalState) : CoreM
|
||||||
throwError "Conflicting assignment of expr metavariable (d != d) {mvarId.name}"
|
throwError "Conflicting assignment of expr metavariable (d != d) {mvarId.name}"
|
||||||
|
|
||||||
Meta.saveState
|
Meta.saveState
|
||||||
let goals := dst.savedState.tactic.goals ++
|
let goals :=dst.savedState.tactic.goals ++
|
||||||
src'.savedState.tactic.goals.map (⟨mapId ·.name⟩)
|
src'.savedState.tactic.goals.map (⟨mapId ·.name⟩)
|
||||||
let fragments ← src'.fragments.foldM (init := dst.fragments) λ acc mvarId' fragment' => do
|
let fragments ← src'.fragments.foldM (init := dst.fragments) λ acc mvarId' fragment' => do
|
||||||
let mvarId := ⟨mapId mvarId'.name⟩
|
let mvarId := ⟨mapId mvarId'.name⟩
|
||||||
|
@ -393,21 +355,18 @@ protected def GoalState.replay (dst : GoalState) (src src' : GoalState) : CoreM
|
||||||
meta := ← m.run',
|
meta := ← m.run',
|
||||||
},
|
},
|
||||||
},
|
},
|
||||||
parentMVars := dst.parentMVars ++ src.parentMVars.map mapMVar,
|
|
||||||
fragments,
|
fragments,
|
||||||
}
|
}
|
||||||
|
|
||||||
--- Tactic execution functions ---
|
--- Tactic execution functions ---
|
||||||
|
|
||||||
/--
|
-- Mimics `Elab.Term.logUnassignedUsingErrorInfos`
|
||||||
These descendants serve as "seed" mvars. If a MVarError's mvar is related to one
|
|
||||||
of these seed mvars, it means an error has occurred when a tactic was executing
|
|
||||||
on `src`. `evalTactic`, will not capture these mvars, so we need to manually
|
|
||||||
find them and save them into the goal list. See the rationales document for the
|
|
||||||
inspiration of this function.
|
|
||||||
-/
|
|
||||||
private def collectAllErroredMVars (src : MVarId) : Elab.TermElabM (List MVarId) := do
|
private def collectAllErroredMVars (src : MVarId) : Elab.TermElabM (List MVarId) := do
|
||||||
-- Mimics `Elab.Term.logUnassignedUsingErrorInfos`
|
-- These descendants serve as "seed" mvars. If a MVarError's mvar is related
|
||||||
|
-- to one of these seed mvars, it means an error has occurred when a tactic
|
||||||
|
-- was executing on `src`. `evalTactic`, will not capture these mvars, so we
|
||||||
|
-- need to manually find them and save them into the goal list.
|
||||||
|
|
||||||
let descendants ← Meta.getMVars (.mvar src)
|
let descendants ← Meta.getMVars (.mvar src)
|
||||||
--let _ ← Elab.Term.logUnassignedUsingErrorInfos descendants
|
--let _ ← Elab.Term.logUnassignedUsingErrorInfos descendants
|
||||||
let mut alreadyVisited : MVarIdSet := {}
|
let mut alreadyVisited : MVarIdSet := {}
|
||||||
|
@ -422,7 +381,6 @@ private def collectAllErroredMVars (src : MVarId) : Elab.TermElabM (List MVarId)
|
||||||
result := mvarDeps.foldl (·.insert ·) result
|
result := mvarDeps.foldl (·.insert ·) result
|
||||||
return result.toList
|
return result.toList
|
||||||
|
|
||||||
/-- Merger of two unique lists -/
|
|
||||||
private def mergeMVarLists (li1 li2 : List MVarId) : List MVarId :=
|
private def mergeMVarLists (li1 li2 : List MVarId) : List MVarId :=
|
||||||
let li2' := li2.filter (¬ li1.contains ·)
|
let li2' := li2.filter (¬ li1.contains ·)
|
||||||
li1 ++ li2'
|
li1 ++ li2'
|
||||||
|
@ -432,29 +390,28 @@ Set `guardMVarErrors` to true to capture mvar errors. Lean will not
|
||||||
automatically collect mvars from text tactics (vide
|
automatically collect mvars from text tactics (vide
|
||||||
`test_tactic_failure_synthesize_placeholder`)
|
`test_tactic_failure_synthesize_placeholder`)
|
||||||
-/
|
-/
|
||||||
protected def GoalState.step' { α } (state : GoalState) (site : Site) (tacticM : Elab.Tactic.TacticM α) (guardMVarErrors : Bool := false)
|
protected def GoalState.step' { α } (state: GoalState) (goal: MVarId) (tacticM: Elab.Tactic.TacticM α) (guardMVarErrors : Bool := false)
|
||||||
: Elab.TermElabM (α × GoalState) := do
|
: Elab.TermElabM (α × GoalState) := do
|
||||||
let ((a, parentMVars), { goals }) ← site.runTacticM tacticM
|
unless (← getMCtx).decls.contains goal do
|
||||||
|>.run { elaborator := .anonymous }
|
throwError s!"Goal is not in context: {goal.name}"
|
||||||
|>.run state.savedState.tactic
|
goal.checkNotAssigned `GoalState.step
|
||||||
|
let (a, { goals }) ← tacticM { elaborator := .anonymous } |>.run { goals := [goal] }
|
||||||
let nextElabState ← MonadBacktrack.saveState
|
let nextElabState ← MonadBacktrack.saveState
|
||||||
--Elab.Term.synthesizeSyntheticMVarsNoPostponing
|
--Elab.Term.synthesizeSyntheticMVarsNoPostponing
|
||||||
|
|
||||||
let goals ← if guardMVarErrors then
|
let goals ← if guardMVarErrors then
|
||||||
parentMVars.foldlM (init := goals) λ goals parent => do
|
pure $ mergeMVarLists goals (← collectAllErroredMVars goal)
|
||||||
let errors ← collectAllErroredMVars parent
|
|
||||||
return mergeMVarLists goals errors
|
|
||||||
else
|
else
|
||||||
pure goals
|
pure goals
|
||||||
let state' := {
|
let state' := {
|
||||||
state with
|
state with
|
||||||
savedState := { term := nextElabState, tactic := { goals }, },
|
savedState := { term := nextElabState, tactic := { goals }, },
|
||||||
parentMVars,
|
parentMVar? := .some goal,
|
||||||
}
|
}
|
||||||
return (a, state')
|
return (a, state')
|
||||||
protected def GoalState.step (state : GoalState) (site : Site) (tacticM : Elab.Tactic.TacticM Unit) (guardMVarErrors : Bool := false)
|
protected def GoalState.step (state: GoalState) (goal: MVarId) (tacticM: Elab.Tactic.TacticM Unit) (guardMVarErrors : Bool := false)
|
||||||
: Elab.TermElabM GoalState :=
|
: Elab.TermElabM GoalState :=
|
||||||
Prod.snd <$> GoalState.step' state site tacticM guardMVarErrors
|
Prod.snd <$> GoalState.step' state goal tacticM guardMVarErrors
|
||||||
|
|
||||||
/-- Response for executing a tactic -/
|
/-- Response for executing a tactic -/
|
||||||
inductive TacticResult where
|
inductive TacticResult where
|
||||||
|
@ -467,21 +424,21 @@ inductive TacticResult where
|
||||||
-- The given action cannot be executed in the state
|
-- The given action cannot be executed in the state
|
||||||
| invalidAction (message : String)
|
| invalidAction (message : String)
|
||||||
|
|
||||||
private def dumpMessageLog (prevMessageLength : Nat := 0) : CoreM (Bool × Array String) := do
|
private def dumpMessageLog (prevMessageLength : Nat) : CoreM (Bool × Array String) := do
|
||||||
let newMessages := (← Core.getMessageLog).toList.drop prevMessageLength
|
let newMessages := (← Core.getMessageLog).toList.drop prevMessageLength
|
||||||
let hasErrors := newMessages.any (·.severity == .error)
|
let hasErrors := newMessages.any (·.severity == .error)
|
||||||
let newMessages ← newMessages.mapM λ m => m.toString
|
let newMessages ← newMessages.mapM λ m => m.toString
|
||||||
Core.resetMessageLog
|
Core.resetMessageLog
|
||||||
return (hasErrors, newMessages.toArray)
|
return (hasErrors, newMessages.toArray)
|
||||||
|
|
||||||
/-- Execute a `TermElabM` producing a goal state, capturing the error and turn it into a `TacticResult` -/
|
|
||||||
def withCapturingError (elabM : Elab.Term.TermElabM GoalState) : Elab.TermElabM TacticResult := do
|
def withCapturingError (elabM : Elab.Term.TermElabM GoalState) : Elab.TermElabM TacticResult := do
|
||||||
assert! (← Core.getMessageLog).toList.isEmpty
|
-- FIXME: Maybe message log should be empty
|
||||||
|
let prevMessageLength := (← Core.getMessageLog).toList.length
|
||||||
try
|
try
|
||||||
let state ← elabM
|
let state ← elabM
|
||||||
|
|
||||||
-- Check if error messages have been generated in the core.
|
-- Check if error messages have been generated in the core.
|
||||||
let (hasError, newMessages) ← dumpMessageLog
|
let (hasError, newMessages) ← dumpMessageLog prevMessageLength
|
||||||
if hasError then
|
if hasError then
|
||||||
return .failure newMessages
|
return .failure newMessages
|
||||||
else
|
else
|
||||||
|
@ -489,58 +446,62 @@ def withCapturingError (elabM : Elab.Term.TermElabM GoalState) : Elab.TermElabM
|
||||||
catch exception =>
|
catch exception =>
|
||||||
match exception with
|
match exception with
|
||||||
| .internal _ =>
|
| .internal _ =>
|
||||||
let (_, messages) ← dumpMessageLog
|
let (_, messages) ← dumpMessageLog prevMessageLength
|
||||||
return .failure messages
|
return .failure messages
|
||||||
| _ => return .failure #[← exception.toMessageData.toString]
|
| _ => return .failure #[← exception.toMessageData.toString]
|
||||||
|
|
||||||
/-- Executes a `TacticM` monad on this `GoalState`, collecting the errors as necessary -/
|
/-- Executes a `TacticM` monad on this `GoalState`, collecting the errors as necessary -/
|
||||||
protected def GoalState.tryTacticM
|
protected def GoalState.tryTacticM
|
||||||
(state: GoalState) (site : Site)
|
(state: GoalState) (goal: MVarId) (tacticM: Elab.Tactic.TacticM Unit)
|
||||||
(tacticM: Elab.Tactic.TacticM Unit)
|
|
||||||
(guardMVarErrors : Bool := false)
|
(guardMVarErrors : Bool := false)
|
||||||
: Elab.TermElabM TacticResult :=
|
: Elab.TermElabM TacticResult := do
|
||||||
withCapturingError do
|
withCapturingError do
|
||||||
state.step site tacticM guardMVarErrors
|
state.step goal tacticM guardMVarErrors
|
||||||
|
|
||||||
/-- Execute a string tactic on given state. Restores TermElabM -/
|
/-- Execute a string tactic on given state. Restores TermElabM -/
|
||||||
@[export pantograph_goal_state_try_tactic_m]
|
@[export pantograph_goal_state_try_tactic_m]
|
||||||
protected def GoalState.tryTactic (state: GoalState) (site : Site) (tactic: String):
|
protected def GoalState.tryTactic (state: GoalState) (goal: MVarId) (tactic: String):
|
||||||
Elab.TermElabM TacticResult := do
|
Elab.TermElabM TacticResult := do
|
||||||
state.restoreElabM
|
state.restoreElabM
|
||||||
let .some goal := state.actingGoal? site | throwNoGoals
|
/-
|
||||||
if let .some fragment := state.fragments[goal]? then
|
if let .some fragment := state.fragments[goal]? then
|
||||||
return ← withCapturingError do
|
return ← withCapturingError do
|
||||||
let (fragments, state') ← state.step' site do
|
let (moreFragments, state') ← state.step' goal (fragment.step goal tactic)
|
||||||
fragment.step goal tactic $ state.fragments.erase goal
|
let fragments := moreFragments.fold (init := state.fragments.erase goal) λ acc mvarId f =>
|
||||||
return { state' with fragments }
|
acc.insert mvarId f
|
||||||
|
return { state' with fragments } -/
|
||||||
|
|
||||||
|
let catName := match isLHSGoal? (← goal.getType) with
|
||||||
|
| .some _ => `conv
|
||||||
|
| .none => `tactic
|
||||||
-- Normal tactic without fragment
|
-- Normal tactic without fragment
|
||||||
let tactic ← match Parser.runParserCategory
|
let tactic ← match Parser.runParserCategory
|
||||||
(env := ← getEnv)
|
(env := ← MonadEnv.getEnv)
|
||||||
(catName := `tactic)
|
(catName := catName)
|
||||||
(input := tactic)
|
(input := tactic)
|
||||||
(fileName := ← getFileName) with
|
(fileName := ← getFileName) with
|
||||||
| .ok stx => pure $ stx
|
| .ok stx => pure $ stx
|
||||||
| .error error => return .parseError error
|
| .error error => return .parseError error
|
||||||
let tacticM := Elab.Tactic.evalTactic tactic
|
let tacticM := Elab.Tactic.evalTactic tactic
|
||||||
withCapturingError do
|
withCapturingError do
|
||||||
state.step site tacticM (guardMVarErrors := true)
|
state.step goal tacticM (guardMVarErrors := true)
|
||||||
|
|
||||||
-- Specialized Tactics
|
-- Specialized Tactics
|
||||||
|
|
||||||
protected def GoalState.tryAssign (state : GoalState) (site : Site) (expr : String)
|
protected def GoalState.tryAssign (state: GoalState) (goal: MVarId) (expr: String):
|
||||||
: Elab.TermElabM TacticResult := do
|
Elab.TermElabM TacticResult := do
|
||||||
state.restoreElabM
|
state.restoreElabM
|
||||||
let expr ← match Parser.runParserCategory
|
let expr ← match Parser.runParserCategory
|
||||||
(env := ← getEnv)
|
(env := ← MonadEnv.getEnv)
|
||||||
(catName := `term)
|
(catName := `term)
|
||||||
(input := expr)
|
(input := expr)
|
||||||
(fileName := ← getFileName) with
|
(fileName := ← getFileName) with
|
||||||
| .ok syn => pure syn
|
| .ok syn => pure syn
|
||||||
| .error error => return .parseError error
|
| .error error => return .parseError error
|
||||||
state.tryTacticM site $ Tactic.evalAssign expr
|
state.tryTacticM goal $ Tactic.evalAssign expr
|
||||||
|
|
||||||
protected def GoalState.tryLet (state : GoalState) (site : Site) (binderName : String) (type : String)
|
protected def GoalState.tryLet (state: GoalState) (goal: MVarId) (binderName: String) (type: String):
|
||||||
: Elab.TermElabM TacticResult := do
|
Elab.TermElabM TacticResult := do
|
||||||
state.restoreElabM
|
state.restoreElabM
|
||||||
let type ← match Parser.runParserCategory
|
let type ← match Parser.runParserCategory
|
||||||
(env := ← MonadEnv.getEnv)
|
(env := ← MonadEnv.getEnv)
|
||||||
|
@ -549,30 +510,28 @@ protected def GoalState.tryLet (state : GoalState) (site : Site) (binderName : S
|
||||||
(fileName := ← getFileName) with
|
(fileName := ← getFileName) with
|
||||||
| .ok syn => pure syn
|
| .ok syn => pure syn
|
||||||
| .error error => return .parseError error
|
| .error error => return .parseError error
|
||||||
state.tryTacticM site $ Tactic.evalLet binderName.toName type
|
state.tryTacticM goal $ Tactic.evalLet binderName.toName type
|
||||||
|
|
||||||
/-- Enter conv tactic mode -/
|
/-- Enter conv tactic mode -/
|
||||||
@[export pantograph_goal_state_conv_enter_m]
|
@[export pantograph_goal_state_conv_enter_m]
|
||||||
protected def GoalState.convEnter (state : GoalState) (site : Site) :
|
protected def GoalState.conv (state : GoalState) (goal : MVarId) :
|
||||||
Elab.TermElabM TacticResult := do
|
Elab.TermElabM TacticResult := do
|
||||||
let .some goal := state.actingGoal? site | throwNoGoals
|
|
||||||
if let .some (.conv ..) := state.fragments[goal]? then
|
if let .some (.conv ..) := state.fragments[goal]? then
|
||||||
return .invalidAction "Already in conv state"
|
return .invalidAction "Already in conv state"
|
||||||
|
goal.checkNotAssigned `GoalState.conv
|
||||||
withCapturingError do
|
withCapturingError do
|
||||||
let (fragments, state') ← state.step' site Fragment.enterConv
|
let (fragment, state') ← state.step' goal Fragment.enterConv
|
||||||
return {
|
return {
|
||||||
state' with
|
state' with
|
||||||
fragments := fragments.fold (init := state'.fragments) λ acc goal fragment =>
|
fragments := state'.fragments.insert goal fragment
|
||||||
acc.insert goal fragment
|
|
||||||
}
|
}
|
||||||
|
|
||||||
/-- Exit from a tactic fragment. -/
|
/-- Exit from `conv` mode. Resumes all goals before the mode starts and applys the conv -/
|
||||||
@[export pantograph_goal_state_fragment_exit_m]
|
@[export pantograph_goal_state_conv_exit_m]
|
||||||
protected def GoalState.fragmentExit (state : GoalState) (site : Site):
|
protected def GoalState.convExit (state : GoalState) (goal : MVarId):
|
||||||
Elab.TermElabM TacticResult := do
|
Elab.TermElabM TacticResult := do
|
||||||
let .some goal := state.actingGoal? site | throwNoGoals
|
let .some fragment@(.conv ..) := state.fragments[goal]? |
|
||||||
let .some fragment := state.fragments[goal]? |
|
return .invalidAction "Not in conv state"
|
||||||
return .invalidAction "Goal does not have a fragment"
|
|
||||||
withCapturingError do
|
withCapturingError do
|
||||||
let (fragments, state') ← state.step' goal (fragment.exit goal state.fragments)
|
let (fragments, state') ← state.step' goal (fragment.exit goal state.fragments)
|
||||||
return {
|
return {
|
||||||
|
@ -584,17 +543,18 @@ protected def GoalState.calcPrevRhsOf? (state : GoalState) (goal : MVarId) : Opt
|
||||||
let .some (.calc prevRhs?) := state.fragments[goal]? | .none
|
let .some (.calc prevRhs?) := state.fragments[goal]? | .none
|
||||||
prevRhs?
|
prevRhs?
|
||||||
|
|
||||||
@[export pantograph_goal_state_calc_enter_m]
|
@[export pantograph_goal_state_try_calc_m]
|
||||||
protected def GoalState.calcEnter (state : GoalState) (site : Site)
|
protected def GoalState.tryCalc (state : GoalState) (goal : MVarId) (pred : String)
|
||||||
: Elab.TermElabM TacticResult := do
|
: Elab.TermElabM TacticResult := do
|
||||||
let .some goal := state.actingGoal? site | throwNoGoals
|
let prevRhs? := state.calcPrevRhsOf? goal
|
||||||
if let .some _ := state.fragments[goal]? then
|
|
||||||
return .invalidAction "Goal already has a fragment"
|
|
||||||
withCapturingError do
|
withCapturingError do
|
||||||
let fragment := Fragment.enterCalc
|
let (moreFragments, state') ← state.step' goal do
|
||||||
let fragments := state.fragments.insert goal fragment
|
let fragment := Fragment.calc prevRhs?
|
||||||
|
fragment.step goal pred
|
||||||
|
let fragments := moreFragments.fold (init := state.fragments.erase goal) λ acc mvarId f =>
|
||||||
|
acc.insert mvarId f
|
||||||
return {
|
return {
|
||||||
state with
|
state' with
|
||||||
fragments,
|
fragments,
|
||||||
}
|
}
|
||||||
|
|
||||||
|
|
|
@ -117,10 +117,10 @@ def goalStartExpr (expr: String) : Protocol.FallibleT Elab.TermElabM GoalState :
|
||||||
|
|
||||||
@[export pantograph_goal_serialize_m]
|
@[export pantograph_goal_serialize_m]
|
||||||
def goalSerialize (state: GoalState) (options: @&Protocol.Options): CoreM (Array Protocol.Goal) :=
|
def goalSerialize (state: GoalState) (options: @&Protocol.Options): CoreM (Array Protocol.Goal) :=
|
||||||
runMetaM <| state.serializeGoals options
|
runMetaM <| state.serializeGoals (parent := .none) options
|
||||||
|
|
||||||
@[export pantograph_goal_print_m]
|
@[export pantograph_goal_print_m]
|
||||||
def goalPrint (state: GoalState) (rootExpr: Bool) (parentExprs: Bool) (goals: Bool) (extraMVars : Array String) (options: @&Protocol.Options)
|
def goalPrint (state: GoalState) (rootExpr: Bool) (parentExpr: Bool) (goals: Bool) (extraMVars : Array String) (options: @&Protocol.Options)
|
||||||
: CoreM Protocol.GoalPrintResult := runMetaM do
|
: CoreM Protocol.GoalPrintResult := runMetaM do
|
||||||
state.restoreMetaM
|
state.restoreMetaM
|
||||||
|
|
||||||
|
@ -130,10 +130,9 @@ def goalPrint (state: GoalState) (rootExpr: Bool) (parentExprs: Bool) (goals: Bo
|
||||||
serializeExpression options (← instantiateAll expr)
|
serializeExpression options (← instantiateAll expr)
|
||||||
else
|
else
|
||||||
pure .none
|
pure .none
|
||||||
let parentExprs? ← if parentExprs then
|
let parent? ← if parentExpr then
|
||||||
.some <$> state.parentMVars.mapM λ parent => parent.withContext do
|
state.parentExpr?.mapM λ expr => state.withParentContext do
|
||||||
let val := state.getMVarEAssignment parent |>.get!
|
serializeExpression options (← instantiateAll expr)
|
||||||
serializeExpression options (← instantiateAll val)
|
|
||||||
else
|
else
|
||||||
pure .none
|
pure .none
|
||||||
let goals ← if goals then
|
let goals ← if goals then
|
||||||
|
@ -149,7 +148,7 @@ def goalPrint (state: GoalState) (rootExpr: Bool) (parentExprs: Bool) (goals: Bo
|
||||||
let env ← getEnv
|
let env ← getEnv
|
||||||
return {
|
return {
|
||||||
root?,
|
root?,
|
||||||
parentExprs?,
|
parent?,
|
||||||
goals,
|
goals,
|
||||||
extraMVars,
|
extraMVars,
|
||||||
rootHasSorry := rootExpr?.map (·.hasSorry) |>.getD false,
|
rootHasSorry := rootExpr?.map (·.hasSorry) |>.getD false,
|
||||||
|
@ -158,26 +157,26 @@ def goalPrint (state: GoalState) (rootExpr: Bool) (parentExprs: Bool) (goals: Bo
|
||||||
}
|
}
|
||||||
|
|
||||||
@[export pantograph_goal_have_m]
|
@[export pantograph_goal_have_m]
|
||||||
protected def GoalState.tryHave (state: GoalState) (site : Site) (binderName: String) (type: String): Elab.TermElabM TacticResult := do
|
protected def GoalState.tryHave (state: GoalState) (goal: MVarId) (binderName: String) (type: String): Elab.TermElabM TacticResult := do
|
||||||
let type ← match (← parseTermM type) with
|
let type ← match (← parseTermM type) with
|
||||||
| .ok syn => pure syn
|
| .ok syn => pure syn
|
||||||
| .error error => return .parseError error
|
| .error error => return .parseError error
|
||||||
state.restoreElabM
|
state.restoreElabM
|
||||||
state.tryTacticM site $ Tactic.evalHave binderName.toName type
|
state.tryTacticM goal $ Tactic.evalHave binderName.toName type
|
||||||
@[export pantograph_goal_try_define_m]
|
@[export pantograph_goal_try_define_m]
|
||||||
protected def GoalState.tryDefine (state: GoalState) (site : Site) (binderName: String) (expr: String): Elab.TermElabM TacticResult := do
|
protected def GoalState.tryDefine (state: GoalState) (goal: MVarId) (binderName: String) (expr: String): Elab.TermElabM TacticResult := do
|
||||||
let expr ← match (← parseTermM expr) with
|
let expr ← match (← parseTermM expr) with
|
||||||
| .ok syn => pure syn
|
| .ok syn => pure syn
|
||||||
| .error error => return .parseError error
|
| .error error => return .parseError error
|
||||||
state.restoreElabM
|
state.restoreElabM
|
||||||
state.tryTacticM site $ Tactic.evalDefine binderName.toName expr
|
state.tryTacticM goal (Tactic.evalDefine binderName.toName expr)
|
||||||
@[export pantograph_goal_try_draft_m]
|
@[export pantograph_goal_try_draft_m]
|
||||||
protected def GoalState.tryDraft (state: GoalState) (site : Site) (expr: String): Elab.TermElabM TacticResult := do
|
protected def GoalState.tryDraft (state: GoalState) (goal: MVarId) (expr: String): Elab.TermElabM TacticResult := do
|
||||||
let expr ← match (← parseTermM expr) with
|
let expr ← match (← parseTermM expr) with
|
||||||
| .ok syn => pure syn
|
| .ok syn => pure syn
|
||||||
| .error error => return .parseError error
|
| .error error => return .parseError error
|
||||||
state.restoreElabM
|
state.restoreElabM
|
||||||
state.tryTacticM site $ Tactic.evalDraft expr
|
state.tryTacticM goal (Tactic.evalDraft expr)
|
||||||
|
|
||||||
-- Cancel the token after a timeout.
|
-- Cancel the token after a timeout.
|
||||||
@[export pantograph_run_cancel_token_with_timeout_m]
|
@[export pantograph_run_cancel_token_with_timeout_m]
|
||||||
|
|
|
@ -60,21 +60,16 @@ structure Variable where
|
||||||
type?: Option Expression := .none
|
type?: Option Expression := .none
|
||||||
value?: Option Expression := .none
|
value?: Option Expression := .none
|
||||||
deriving Lean.ToJson
|
deriving Lean.ToJson
|
||||||
inductive Fragment where
|
|
||||||
| tactic
|
|
||||||
| conv
|
|
||||||
| calc
|
|
||||||
deriving BEq, DecidableEq, Repr, Lean.ToJson
|
|
||||||
structure Goal where
|
structure Goal where
|
||||||
|
name: String := ""
|
||||||
/-- Name of the metavariable -/
|
/-- Name of the metavariable -/
|
||||||
name : String := ""
|
userName?: Option String := .none
|
||||||
/-- User-facing name -/
|
/-- Is the goal in conversion mode -/
|
||||||
userName? : Option String := .none
|
isConversion: Bool := false
|
||||||
fragment : Fragment := .tactic
|
|
||||||
/-- target expression type -/
|
/-- target expression type -/
|
||||||
target : Expression
|
target: Expression
|
||||||
/-- Variables -/
|
/-- Variables -/
|
||||||
vars : Array Variable := #[]
|
vars: Array Variable := #[]
|
||||||
deriving Lean.ToJson
|
deriving Lean.ToJson
|
||||||
|
|
||||||
|
|
||||||
|
@ -92,7 +87,6 @@ structure InteractionError where
|
||||||
deriving Lean.ToJson
|
deriving Lean.ToJson
|
||||||
|
|
||||||
def errorIndex (desc: String): InteractionError := { error := "index", desc }
|
def errorIndex (desc: String): InteractionError := { error := "index", desc }
|
||||||
def errorOperation (desc: String): InteractionError := { error := "operation", desc }
|
|
||||||
def errorExpr (desc: String): InteractionError := { error := "expr", desc }
|
def errorExpr (desc: String): InteractionError := { error := "expr", desc }
|
||||||
|
|
||||||
|
|
||||||
|
@ -254,17 +248,17 @@ structure GoalStartResult where
|
||||||
root: String
|
root: String
|
||||||
deriving Lean.ToJson
|
deriving Lean.ToJson
|
||||||
structure GoalTactic where
|
structure GoalTactic where
|
||||||
|
-- Identifiers for tree, state, and goal
|
||||||
stateId: Nat
|
stateId: Nat
|
||||||
-- If omitted, act on the first goal
|
goalId: Nat := 0
|
||||||
goalId?: Option Nat := .none
|
|
||||||
-- If set to true, goal will not go dormant. Defaults to `automaticMode`
|
|
||||||
autoResume?: Option Bool := .none
|
|
||||||
-- One of the fields here must be filled
|
-- One of the fields here must be filled
|
||||||
tactic?: Option String := .none
|
tactic?: Option String := .none
|
||||||
mode?: Option String := .none -- Changes the current category to {"tactic", "calc", "conv"}
|
|
||||||
expr?: Option String := .none
|
expr?: Option String := .none
|
||||||
have?: Option String := .none
|
have?: Option String := .none
|
||||||
let?: Option String := .none
|
let?: Option String := .none
|
||||||
|
calc?: Option String := .none
|
||||||
|
-- true to enter `conv`, `false` to exit. In case of exit the `goalId` is ignored.
|
||||||
|
conv?: Option Bool := .none
|
||||||
draft?: Option String := .none
|
draft?: Option String := .none
|
||||||
|
|
||||||
-- In case of the `have` tactic, the new free variable name is provided here
|
-- In case of the `have` tactic, the new free variable name is provided here
|
||||||
|
@ -314,8 +308,8 @@ structure GoalPrint where
|
||||||
|
|
||||||
-- Print root?
|
-- Print root?
|
||||||
rootExpr?: Option Bool := .some False
|
rootExpr?: Option Bool := .some False
|
||||||
-- Print the parent expressions
|
-- Print the parent expr?
|
||||||
parentExprs?: Option Bool := .some False
|
parentExpr?: Option Bool := .some False
|
||||||
-- Print goals?
|
-- Print goals?
|
||||||
goals?: Option Bool := .some False
|
goals?: Option Bool := .some False
|
||||||
-- Print values of extra mvars?
|
-- Print values of extra mvars?
|
||||||
|
@ -325,7 +319,7 @@ structure GoalPrintResult where
|
||||||
-- The root expression
|
-- The root expression
|
||||||
root?: Option Expression := .none
|
root?: Option Expression := .none
|
||||||
-- The filling expression of the parent goal
|
-- The filling expression of the parent goal
|
||||||
parentExprs?: Option (List Expression) := .none
|
parent?: Option Expression := .none
|
||||||
goals: Array Goal := #[]
|
goals: Array Goal := #[]
|
||||||
extraMVars: Array Expression := #[]
|
extraMVars: Array Expression := #[]
|
||||||
|
|
||||||
|
|
|
@ -103,7 +103,7 @@ def goalStatePickle (goalState : GoalState) (path : System.FilePath) : IO Unit :
|
||||||
tactic
|
tactic
|
||||||
}
|
}
|
||||||
root,
|
root,
|
||||||
parentMVars,
|
parentMVar?,
|
||||||
fragments,
|
fragments,
|
||||||
} := goalState
|
} := goalState
|
||||||
Pantograph.pickle path (
|
Pantograph.pickle path (
|
||||||
|
@ -115,7 +115,7 @@ def goalStatePickle (goalState : GoalState) (path : System.FilePath) : IO Unit :
|
||||||
tactic,
|
tactic,
|
||||||
|
|
||||||
root,
|
root,
|
||||||
parentMVars,
|
parentMVar?,
|
||||||
fragments,
|
fragments,
|
||||||
)
|
)
|
||||||
|
|
||||||
|
@ -131,7 +131,7 @@ def goalStateUnpickle (path : System.FilePath) (env : Environment)
|
||||||
tactic,
|
tactic,
|
||||||
|
|
||||||
root,
|
root,
|
||||||
parentMVars,
|
parentMVar?,
|
||||||
fragments,
|
fragments,
|
||||||
), region) ← Pantograph.unpickle (
|
), region) ← Pantograph.unpickle (
|
||||||
PHashMap Name ConstantInfo ×
|
PHashMap Name ConstantInfo ×
|
||||||
|
@ -142,7 +142,7 @@ def goalStateUnpickle (path : System.FilePath) (env : Environment)
|
||||||
Elab.Tactic.State ×
|
Elab.Tactic.State ×
|
||||||
|
|
||||||
MVarId ×
|
MVarId ×
|
||||||
List MVarId ×
|
Option MVarId ×
|
||||||
FragmentMap
|
FragmentMap
|
||||||
) path
|
) path
|
||||||
let env ← env.replay (Std.HashMap.ofList map₂.toList)
|
let env ← env.replay (Std.HashMap.ofList map₂.toList)
|
||||||
|
@ -162,7 +162,7 @@ def goalStateUnpickle (path : System.FilePath) (env : Environment)
|
||||||
tactic,
|
tactic,
|
||||||
},
|
},
|
||||||
root,
|
root,
|
||||||
parentMVars,
|
parentMVar?,
|
||||||
fragments,
|
fragments,
|
||||||
}
|
}
|
||||||
return (goalState, region)
|
return (goalState, region)
|
||||||
|
|
|
@ -4,13 +4,6 @@ Here, a unified system handles all fragments.
|
||||||
|
|
||||||
Inside a tactic fragment, the parser category may be different. An incomplete
|
Inside a tactic fragment, the parser category may be different. An incomplete
|
||||||
fragmented tactic may not be elaboratable..
|
fragmented tactic may not be elaboratable..
|
||||||
|
|
||||||
In line with continuation/resumption paradigms, the exit function of a fragment
|
|
||||||
tactic is responsible for resuming incomplete goals with fragments. For example,
|
|
||||||
when a conversion tactic finishes, the sentinels should resume the root of the
|
|
||||||
conversion tactic goal. The user cannot be expected to execute this resumption,
|
|
||||||
since the root is automatically dormanted at the entry of the conversion tactic
|
|
||||||
mode.
|
|
||||||
-/
|
-/
|
||||||
import Lean.Meta
|
import Lean.Meta
|
||||||
import Lean.Elab
|
import Lean.Elab
|
||||||
|
@ -21,79 +14,59 @@ namespace Pantograph
|
||||||
|
|
||||||
inductive Fragment where
|
inductive Fragment where
|
||||||
| calc (prevRhs? : Option Expr)
|
| calc (prevRhs? : Option Expr)
|
||||||
| conv (rhs : MVarId)
|
| conv (rhs : MVarId) (dormant : List MVarId)
|
||||||
-- This goal is spawned from a `conv`
|
deriving BEq
|
||||||
| convSentinel (parent : MVarId)
|
|
||||||
deriving BEq, Inhabited
|
|
||||||
|
|
||||||
abbrev FragmentMap := Std.HashMap MVarId Fragment
|
abbrev FragmentMap := Std.HashMap MVarId Fragment
|
||||||
def FragmentMap.empty : FragmentMap := Std.HashMap.emptyWithCapacity 2
|
def FragmentMap.empty : FragmentMap := Std.HashMap.emptyWithCapacity 2
|
||||||
protected def FragmentMap.filter (map : FragmentMap) (pred : MVarId → Fragment → Bool) : FragmentMap :=
|
|
||||||
map.fold (init := FragmentMap.empty) λ acc mvarId fragment =>
|
|
||||||
if pred mvarId fragment then
|
|
||||||
acc.insert mvarId fragment
|
|
||||||
else
|
|
||||||
acc
|
|
||||||
|
|
||||||
protected def Fragment.map (fragment : Fragment) (mapExpr : Expr → CoreM Expr) : CoreM Fragment :=
|
protected def Fragment.map (fragment : Fragment) (mapExpr : Expr → CoreM Expr) : CoreM Fragment :=
|
||||||
let mapMVar (mvarId : MVarId) : CoreM MVarId :=
|
let mapMVarId (mvarId : MVarId) : CoreM MVarId :=
|
||||||
return (← mapExpr (.mvar mvarId)) |>.mvarId!
|
return (← mapExpr (.mvar mvarId)) |>.mvarId!
|
||||||
match fragment with
|
match fragment with
|
||||||
| .calc prevRhs? => return .calc (← prevRhs?.mapM mapExpr)
|
| .calc prevRhs? => return .calc (← prevRhs?.mapM mapExpr)
|
||||||
| .conv rhs => do
|
| .conv rhs dormant => do
|
||||||
let rhs' ← mapMVar rhs
|
let rhs' ← mapMVarId rhs
|
||||||
return .conv rhs'
|
let dormant' ← dormant.mapM mapMVarId
|
||||||
| .convSentinel parent => do
|
return .conv rhs' dormant'
|
||||||
return .convSentinel (← mapMVar parent)
|
|
||||||
|
|
||||||
protected def Fragment.enterCalc : Fragment := .calc .none
|
protected def Fragment.enterCalc : Elab.Tactic.TacticM Fragment := do
|
||||||
protected def Fragment.enterConv : Elab.Tactic.TacticM FragmentMap := do
|
return .calc .none
|
||||||
|
protected def Fragment.enterConv : Elab.Tactic.TacticM Fragment := do
|
||||||
let goal ← Elab.Tactic.getMainGoal
|
let goal ← Elab.Tactic.getMainGoal
|
||||||
goal.checkNotAssigned `GoalState.conv
|
let rhs ← goal.withContext do
|
||||||
let (rhs, newGoal) ← goal.withContext do
|
let (rhs, newGoal) ← Elab.Tactic.Conv.mkConvGoalFor (← Elab.Tactic.getMainTarget)
|
||||||
let target ← instantiateMVars (← goal.getType)
|
Elab.Tactic.replaceMainGoal [newGoal.mvarId!]
|
||||||
let (rhs, newGoal) ← Elab.Tactic.Conv.mkConvGoalFor target
|
pure rhs.mvarId!
|
||||||
pure (rhs.mvarId!, newGoal.mvarId!)
|
let otherGoals := (← Elab.Tactic.getGoals).filter (· != goal)
|
||||||
Elab.Tactic.replaceMainGoal [newGoal]
|
return conv rhs otherGoals
|
||||||
return FragmentMap.empty
|
|
||||||
|>.insert goal (.conv rhs)
|
|
||||||
|>.insert newGoal (.convSentinel goal)
|
|
||||||
|
|
||||||
protected partial def Fragment.exit (fragment : Fragment) (goal : MVarId) (fragments : FragmentMap)
|
protected def Fragment.exit (fragment : Fragment) (goal : MVarId) (fragments : FragmentMap)
|
||||||
: Elab.Tactic.TacticM FragmentMap :=
|
: Elab.Tactic.TacticM FragmentMap :=
|
||||||
match fragment with
|
match fragment with
|
||||||
| .calc .. => do
|
| .calc .. => do
|
||||||
Elab.Tactic.setGoals [goal]
|
Elab.Tactic.setGoals [goal]
|
||||||
return fragments.erase goal
|
return fragments.erase goal
|
||||||
| .conv rhs => do
|
| .conv rhs otherGoals => do
|
||||||
let goals := (← Elab.Tactic.getGoals).filter λ descendant =>
|
-- FIXME: Only process the goals that are descendants of `goal`
|
||||||
match fragments[descendant]? with
|
let goals := (← Elab.Tactic.getGoals).filter λ goal => true
|
||||||
| .some s => (.convSentinel goal) == s
|
--match fragments[goal]? with
|
||||||
| _ => false -- Not a conv goal from this
|
--| .some f => fragment == f
|
||||||
|
--| .none => false -- Not a conv goal from this
|
||||||
-- Close all existing goals with `refl`
|
-- Close all existing goals with `refl`
|
||||||
for mvarId in goals do
|
for mvarId in goals do
|
||||||
liftM <| mvarId.refl <|> mvarId.inferInstance <|> pure ()
|
liftM <| mvarId.refl <|> mvarId.inferInstance <|> pure ()
|
||||||
|
Elab.Tactic.pruneSolvedGoals
|
||||||
unless (← goals.filterM (·.isAssignedOrDelayedAssigned)).isEmpty do
|
unless (← goals.filterM (·.isAssignedOrDelayedAssigned)).isEmpty do
|
||||||
throwError "convert tactic failed, there are unsolved goals\n{Elab.goalsToMessageData (goals)}"
|
throwError "convert tactic failed, there are unsolved goals\n{Elab.goalsToMessageData (goals)}"
|
||||||
|
|
||||||
-- Ensure the meta tactic runs on `goal` even if its dormant by forcing resumption
|
Elab.Tactic.setGoals $ [goal] ++ otherGoals
|
||||||
Elab.Tactic.setGoals $ goal :: (← Elab.Tactic.getGoals)
|
|
||||||
let targetNew ← instantiateMVars (.mvar rhs)
|
let targetNew ← instantiateMVars (.mvar rhs)
|
||||||
let proof ← instantiateMVars (.mvar goal)
|
let proof ← instantiateMVars (.mvar goal)
|
||||||
|
|
||||||
Elab.Tactic.liftMetaTactic1 (·.replaceTargetEq targetNew proof)
|
Elab.Tactic.liftMetaTactic1 (·.replaceTargetEq targetNew proof)
|
||||||
|
return fragments.erase goal
|
||||||
|
|
||||||
-- Try to solve maiinline by rfl
|
protected def Fragment.step (fragment : Fragment) (goal : MVarId) (s : String)
|
||||||
let mvarId ← Elab.Tactic.getMainGoal
|
|
||||||
liftM <| mvarId.refl <|> mvarId.inferInstance <|> pure ()
|
|
||||||
Elab.Tactic.pruneSolvedGoals
|
|
||||||
return fragments.filter λ mvarId fragment =>
|
|
||||||
!(mvarId == goal || fragment == .convSentinel goal)
|
|
||||||
| .convSentinel parent =>
|
|
||||||
let parentFragment := fragments[parent]!
|
|
||||||
parentFragment.exit parent (fragments.erase goal)
|
|
||||||
|
|
||||||
protected def Fragment.step (fragment : Fragment) (goal : MVarId) (s : String) (map : FragmentMap)
|
|
||||||
: Elab.Tactic.TacticM FragmentMap := goal.withContext do
|
: Elab.Tactic.TacticM FragmentMap := goal.withContext do
|
||||||
assert! ¬ (← goal.isAssigned)
|
assert! ¬ (← goal.isAssigned)
|
||||||
match fragment with
|
match fragment with
|
||||||
|
@ -148,41 +121,9 @@ protected def Fragment.step (fragment : Fragment) (goal : MVarId) (s : String) (
|
||||||
let goals := [ mvarBranch ] ++ remainder?.toList
|
let goals := [ mvarBranch ] ++ remainder?.toList
|
||||||
Elab.Tactic.setGoals goals
|
Elab.Tactic.setGoals goals
|
||||||
match remainder? with
|
match remainder? with
|
||||||
| .some goal => return map.erase goal |>.insert goal $ .calc (.some rhs)
|
| .some goal => return FragmentMap.empty.insert goal $ .calc (.some rhs)
|
||||||
| .none => return map
|
| .none => return .empty
|
||||||
| .conv .. => do
|
| .conv .. => do
|
||||||
throwError "Direct operation on conversion tactic parent goal is not allowed"
|
throwError "Direct operation on conversion tactic parent goal is not allowed"
|
||||||
| fragment@(.convSentinel parent) => do
|
|
||||||
assert! isLHSGoal? (← goal.getType) |>.isSome
|
|
||||||
let tactic ← match Parser.runParserCategory
|
|
||||||
(env := ← MonadEnv.getEnv)
|
|
||||||
(catName := `conv)
|
|
||||||
(input := s)
|
|
||||||
(fileName := ← getFileName) with
|
|
||||||
| .ok stx => pure $ stx
|
|
||||||
| .error error => throwError error
|
|
||||||
let oldGoals ← Elab.Tactic.getGoals
|
|
||||||
-- Label newly generated goals as conv sentinels
|
|
||||||
Elab.Tactic.evalTactic tactic
|
|
||||||
let newConvGoals ← (← Elab.Tactic.getUnsolvedGoals).filterM λ g => do
|
|
||||||
-- conv tactic might generate non-conv goals
|
|
||||||
if oldGoals.contains g then
|
|
||||||
return false
|
|
||||||
return isLHSGoal? (← g.getType) |>.isSome
|
|
||||||
-- Conclude the conv by exiting the parent fragment if new goals is empty
|
|
||||||
if newConvGoals.isEmpty then
|
|
||||||
let hasSiblingFragment := map.fold (init := false) λ flag _ fragment =>
|
|
||||||
if flag then
|
|
||||||
true
|
|
||||||
else match fragment with
|
|
||||||
| .convSentinel parent' => parent == parent'
|
|
||||||
| _ => false
|
|
||||||
if ¬ hasSiblingFragment then
|
|
||||||
-- This fragment must exist since we have conv goals
|
|
||||||
let parentFragment := map[parent]!
|
|
||||||
-- All descendants exhausted. Exit from the parent conv.
|
|
||||||
return ← parentFragment.exit parent map
|
|
||||||
return newConvGoals.foldl (init := map) λ acc g =>
|
|
||||||
acc.insert g fragment
|
|
||||||
|
|
||||||
end Pantograph
|
end Pantograph
|
||||||
|
|
140
Repl.lean
140
Repl.lean
|
@ -97,77 +97,6 @@ def liftTermElabM { α } (termElabM : Elab.TermElabM α) (levelNames : List Name
|
||||||
}
|
}
|
||||||
runCoreM $ termElabM.run' context state |>.run'
|
runCoreM $ termElabM.run' context state |>.run'
|
||||||
|
|
||||||
section Goal
|
|
||||||
|
|
||||||
def goal_tactic (args: Protocol.GoalTactic): EMainM Protocol.GoalTacticResult := do
|
|
||||||
let state ← getMainState
|
|
||||||
let .some goalState := state.goalStates[args.stateId]? |
|
|
||||||
Protocol.throw $ Protocol.errorIndex s!"Invalid state index {args.stateId}"
|
|
||||||
let unshielded := args.autoResume?.getD state.options.automaticMode
|
|
||||||
let site ← match args.goalId?, unshielded with
|
|
||||||
| .some goalId, true => do
|
|
||||||
let .some goal := goalState.goals[goalId]? |
|
|
||||||
Protocol.throw $ Protocol.errorIndex s!"Invalid goal index {goalId}"
|
|
||||||
pure (.prefer goal)
|
|
||||||
| .some goalId, false => do
|
|
||||||
let .some goal := goalState.goals[goalId]? |
|
|
||||||
Protocol.throw $ Protocol.errorIndex s!"Invalid goal index {goalId}"
|
|
||||||
pure (.focus goal)
|
|
||||||
| .none, true => pure .unfocus
|
|
||||||
| .none, false => do
|
|
||||||
let .some goal := goalState.mainGoal? |
|
|
||||||
Protocol.throw $ Protocol.errorIndex s!"No goals to be solved"
|
|
||||||
pure (.focus goal)
|
|
||||||
let nextGoalState?: Except _ TacticResult ← liftTermElabM do
|
|
||||||
-- NOTE: Should probably use a macro to handle this...
|
|
||||||
match args.tactic?, args.mode?, args.expr?, args.have?, args.let?, args.draft? with
|
|
||||||
| .some tactic, .none, .none, .none, .none, .none => do
|
|
||||||
pure $ Except.ok $ ← goalState.tryTactic site tactic
|
|
||||||
| .none, .some mode, .none, .none, .none, .none => match mode with
|
|
||||||
| "tactic" => do -- Exit from the current fragment
|
|
||||||
pure $ Except.ok $ ← goalState.fragmentExit site
|
|
||||||
| "conv" => do
|
|
||||||
pure $ Except.ok $ ← goalState.convEnter site
|
|
||||||
| "calc" => do
|
|
||||||
pure $ Except.ok $ ← goalState.calcEnter site
|
|
||||||
| _ => pure $ .error $ Protocol.errorOperation s!"Invalid mode {mode}"
|
|
||||||
| .none, .none, .some expr, .none, .none, .none => do
|
|
||||||
pure $ Except.ok $ ← goalState.tryAssign site expr
|
|
||||||
| .none, .none, .none, .some type, .none, .none => do
|
|
||||||
let binderName := args.binderName?.getD ""
|
|
||||||
pure $ Except.ok $ ← goalState.tryHave site binderName type
|
|
||||||
| .none, .none, .none, .none, .some type, .none => do
|
|
||||||
let binderName := args.binderName?.getD ""
|
|
||||||
pure $ Except.ok $ ← goalState.tryLet site binderName type
|
|
||||||
| .none, .none, .none, .none, .none, .some draft => do
|
|
||||||
pure $ Except.ok $ ← goalState.tryDraft site draft
|
|
||||||
| _, _, _, _, _, _ =>
|
|
||||||
pure $ .error $ Protocol.errorOperation
|
|
||||||
"Exactly one of {tactic, mode, expr, have, let, draft} must be supplied"
|
|
||||||
match nextGoalState? with
|
|
||||||
| .error error => Protocol.throw error
|
|
||||||
| .ok (.success nextGoalState messages) => do
|
|
||||||
let nextStateId ← newGoalState nextGoalState
|
|
||||||
let parentExprs := nextGoalState.parentExprs
|
|
||||||
let hasSorry := parentExprs.any (·.hasSorry)
|
|
||||||
let hasUnsafe := parentExprs.any ((← getEnv).hasUnsafe ·)
|
|
||||||
let goals ← runCoreM $ nextGoalState.serializeGoals (options := state.options) |>.run'
|
|
||||||
return {
|
|
||||||
nextStateId? := .some nextStateId,
|
|
||||||
goals? := .some goals,
|
|
||||||
messages? := .some messages,
|
|
||||||
hasSorry,
|
|
||||||
hasUnsafe,
|
|
||||||
}
|
|
||||||
| .ok (.parseError message) =>
|
|
||||||
return { messages? := .none, parseError? := .some message }
|
|
||||||
| .ok (.invalidAction message) =>
|
|
||||||
Protocol.throw $ errorI "invalid" message
|
|
||||||
| .ok (.failure messages) =>
|
|
||||||
return { messages? := .some messages }
|
|
||||||
|
|
||||||
end Goal
|
|
||||||
|
|
||||||
section Frontend
|
section Frontend
|
||||||
|
|
||||||
structure CompilationUnit where
|
structure CompilationUnit where
|
||||||
|
@ -296,6 +225,7 @@ def execute (command: Protocol.Command): MainM Json := do
|
||||||
return toJson error
|
return toJson error
|
||||||
where
|
where
|
||||||
errorCommand := errorI "command"
|
errorCommand := errorI "command"
|
||||||
|
errorIndex := errorI "index"
|
||||||
errorIO := errorI "io"
|
errorIO := errorI "io"
|
||||||
-- Command Functions
|
-- Command Functions
|
||||||
reset (_: Protocol.Reset): EMainM Protocol.StatResult := do
|
reset (_: Protocol.Reset): EMainM Protocol.StatResult := do
|
||||||
|
@ -360,7 +290,7 @@ def execute (command: Protocol.Command): MainM Json := do
|
||||||
| .some expr, .none => goalStartExpr expr |>.run
|
| .some expr, .none => goalStartExpr expr |>.run
|
||||||
| .none, .some copyFrom => do
|
| .none, .some copyFrom => do
|
||||||
(match (← getEnv).find? <| copyFrom.toName with
|
(match (← getEnv).find? <| copyFrom.toName with
|
||||||
| .none => return .error <| Protocol.errorIndex s!"Symbol not found: {copyFrom}"
|
| .none => return .error <| errorIndex s!"Symbol not found: {copyFrom}"
|
||||||
| .some cInfo => return .ok (← GoalState.create cInfo.type))
|
| .some cInfo => return .ok (← GoalState.create cInfo.type))
|
||||||
| _, _ =>
|
| _, _ =>
|
||||||
return .error <| errorI "arguments" "Exactly one of {expr, copyFrom} must be supplied"
|
return .error <| errorI "arguments" "Exactly one of {expr, copyFrom} must be supplied"
|
||||||
|
@ -369,14 +299,70 @@ def execute (command: Protocol.Command): MainM Json := do
|
||||||
| .ok goalState =>
|
| .ok goalState =>
|
||||||
let stateId ← newGoalState goalState
|
let stateId ← newGoalState goalState
|
||||||
return { stateId, root := goalState.root.name.toString }
|
return { stateId, root := goalState.root.name.toString }
|
||||||
|
goal_tactic (args: Protocol.GoalTactic): EMainM Protocol.GoalTacticResult := do
|
||||||
|
let state ← getMainState
|
||||||
|
let .some goalState := state.goalStates[args.stateId]? |
|
||||||
|
Protocol.throw $ errorIndex s!"Invalid state index {args.stateId}"
|
||||||
|
-- FIXME: Allow proper conversion tactic exit even in automatic mode
|
||||||
|
let .some goal := goalState.goals[args.goalId]? |
|
||||||
|
Protocol.throw $ errorIndex s!"Invalid goal index {args.goalId}"
|
||||||
|
let nextGoalState?: Except _ TacticResult ← liftTermElabM do
|
||||||
|
-- NOTE: Should probably use a macro to handle this...
|
||||||
|
match args.tactic?, args.expr?, args.have?, args.let?, args.calc?, args.conv?, args.draft? with
|
||||||
|
| .some tactic, .none, .none, .none, .none, .none, .none => do
|
||||||
|
pure <| Except.ok <| ← goalState.tryTactic goal tactic
|
||||||
|
| .none, .some expr, .none, .none, .none, .none, .none => do
|
||||||
|
pure <| Except.ok <| ← goalState.tryAssign goal expr
|
||||||
|
| .none, .none, .some type, .none, .none, .none, .none => do
|
||||||
|
let binderName := args.binderName?.getD ""
|
||||||
|
pure <| Except.ok <| ← goalState.tryHave goal binderName type
|
||||||
|
| .none, .none, .none, .some type, .none, .none, .none => do
|
||||||
|
let binderName := args.binderName?.getD ""
|
||||||
|
pure <| Except.ok <| ← goalState.tryLet goal binderName type
|
||||||
|
| .none, .none, .none, .none, .some pred, .none, .none => do
|
||||||
|
pure <| Except.ok <| ← goalState.tryCalc goal pred
|
||||||
|
| .none, .none, .none, .none, .none, .some true, .none => do
|
||||||
|
pure <| Except.ok <| ← goalState.conv goal
|
||||||
|
| .none, .none, .none, .none, .none, .some false, .none => do
|
||||||
|
pure <| Except.ok <| ← goalState.convExit goal
|
||||||
|
| .none, .none, .none, .none, .none, .none, .some draft => do
|
||||||
|
pure <| Except.ok <| ← goalState.tryDraft goal draft
|
||||||
|
| _, _, _, _, _, _, _ =>
|
||||||
|
let error := errorI "arguments" "Exactly one of {tactic, expr, have, let, calc, conv, draft} must be supplied"
|
||||||
|
pure $ .error error
|
||||||
|
match nextGoalState? with
|
||||||
|
| .error error => Protocol.throw error
|
||||||
|
| .ok (.success nextGoalState messages) => do
|
||||||
|
let nextGoalState ← match state.options.automaticMode, args.conv? with
|
||||||
|
| true, .none => do
|
||||||
|
pure $ nextGoalState.immediateResume goalState
|
||||||
|
| true, .some true => pure nextGoalState
|
||||||
|
| true, .some false => pure nextGoalState
|
||||||
|
| false, _ => pure nextGoalState
|
||||||
|
let nextStateId ← newGoalState nextGoalState
|
||||||
|
let parentExpr := nextGoalState.parentExpr?.get!
|
||||||
|
let goals ← runCoreM $ nextGoalState.serializeGoals (parent := .some goalState) (options := state.options) |>.run'
|
||||||
|
return {
|
||||||
|
nextStateId? := .some nextStateId,
|
||||||
|
goals? := .some goals,
|
||||||
|
messages? := .some messages,
|
||||||
|
hasSorry := parentExpr.hasSorry,
|
||||||
|
hasUnsafe := (← getEnv).hasUnsafe parentExpr,
|
||||||
|
}
|
||||||
|
| .ok (.parseError message) =>
|
||||||
|
return { messages? := .none, parseError? := .some message }
|
||||||
|
| .ok (.invalidAction message) =>
|
||||||
|
Protocol.throw $ errorI "invalid" message
|
||||||
|
| .ok (.failure messages) =>
|
||||||
|
return { messages? := .some messages }
|
||||||
goal_continue (args: Protocol.GoalContinue): EMainM Protocol.GoalContinueResult := do
|
goal_continue (args: Protocol.GoalContinue): EMainM Protocol.GoalContinueResult := do
|
||||||
let state ← getMainState
|
let state ← getMainState
|
||||||
let .some target := state.goalStates[args.target]? |
|
let .some target := state.goalStates[args.target]? |
|
||||||
Protocol.throw $ Protocol.errorIndex s!"Invalid state index {args.target}"
|
Protocol.throw $ errorIndex s!"Invalid state index {args.target}"
|
||||||
let nextGoalState? : GoalState ← match args.branch?, args.goals? with
|
let nextGoalState? : GoalState ← match args.branch?, args.goals? with
|
||||||
| .some branchId, .none => do
|
| .some branchId, .none => do
|
||||||
match state.goalStates[branchId]? with
|
match state.goalStates[branchId]? with
|
||||||
| .none => Protocol.throw $ Protocol.errorIndex s!"Invalid state index {branchId}"
|
| .none => Protocol.throw $ errorIndex s!"Invalid state index {branchId}"
|
||||||
| .some branch => pure $ target.continue branch
|
| .some branch => pure $ target.continue branch
|
||||||
| .none, .some goals =>
|
| .none, .some goals =>
|
||||||
let goals := goals.toList.map (λ n => { name := n.toName })
|
let goals := goals.toList.map (λ n => { name := n.toName })
|
||||||
|
@ -399,11 +385,11 @@ def execute (command: Protocol.Command): MainM Json := do
|
||||||
goal_print (args: Protocol.GoalPrint): EMainM Protocol.GoalPrintResult := do
|
goal_print (args: Protocol.GoalPrint): EMainM Protocol.GoalPrintResult := do
|
||||||
let state ← getMainState
|
let state ← getMainState
|
||||||
let .some goalState := state.goalStates[args.stateId]? |
|
let .some goalState := state.goalStates[args.stateId]? |
|
||||||
Protocol.throw $ Protocol.errorIndex s!"Invalid state index {args.stateId}"
|
Protocol.throw $ errorIndex s!"Invalid state index {args.stateId}"
|
||||||
let result ← liftMetaM <| goalPrint
|
let result ← liftMetaM <| goalPrint
|
||||||
goalState
|
goalState
|
||||||
(rootExpr := args.rootExpr?.getD False)
|
(rootExpr := args.rootExpr?.getD False)
|
||||||
(parentExprs := args.parentExprs?.getD False)
|
(parentExpr := args.parentExpr?.getD False)
|
||||||
(goals := args.goals?.getD False)
|
(goals := args.goals?.getD False)
|
||||||
(extraMVars := args.extraMVars?.getD #[])
|
(extraMVars := args.extraMVars?.getD #[])
|
||||||
(options := state.options)
|
(options := state.options)
|
||||||
|
@ -411,7 +397,7 @@ def execute (command: Protocol.Command): MainM Json := do
|
||||||
goal_save (args: Protocol.GoalSave): EMainM Protocol.GoalSaveResult := do
|
goal_save (args: Protocol.GoalSave): EMainM Protocol.GoalSaveResult := do
|
||||||
let state ← getMainState
|
let state ← getMainState
|
||||||
let .some goalState := state.goalStates[args.id]? |
|
let .some goalState := state.goalStates[args.id]? |
|
||||||
Protocol.throw $ Protocol.errorIndex s!"Invalid state index {args.id}"
|
Protocol.throw $ errorIndex s!"Invalid state index {args.id}"
|
||||||
goalStatePickle goalState args.path
|
goalStatePickle goalState args.path
|
||||||
return {}
|
return {}
|
||||||
goal_load (args: Protocol.GoalLoad): EMainM Protocol.GoalLoadResult := do
|
goal_load (args: Protocol.GoalLoad): EMainM Protocol.GoalLoadResult := do
|
||||||
|
|
|
@ -88,10 +88,10 @@ def test_tactic : Test := do
|
||||||
({ stateId := 0, root := "_uniq.9" }: Protocol.GoalStartResult)
|
({ stateId := 0, root := "_uniq.9" }: Protocol.GoalStartResult)
|
||||||
step "goal.tactic" ({ stateId := 0, tactic? := .some "intro x" }: Protocol.GoalTactic)
|
step "goal.tactic" ({ stateId := 0, tactic? := .some "intro x" }: Protocol.GoalTactic)
|
||||||
({ nextStateId? := .some 1, goals? := #[goal1], }: Protocol.GoalTacticResult)
|
({ nextStateId? := .some 1, goals? := #[goal1], }: Protocol.GoalTacticResult)
|
||||||
step "goal.print" ({ stateId := 1, parentExprs? := .some true, rootExpr? := .some true }: Protocol.GoalPrint)
|
step "goal.print" ({ stateId := 1, parentExpr? := .some true, rootExpr? := .some true }: Protocol.GoalPrint)
|
||||||
({
|
({
|
||||||
root? := .some { pp? := "fun x => ?m.11"},
|
root? := .some { pp? := "fun x => ?m.11"},
|
||||||
parentExprs? := .some [{ pp? := .some "fun x => ?m.11" }],
|
parent? := .some { pp? := .some "fun x => ?m.11" },
|
||||||
}: Protocol.GoalPrintResult)
|
}: Protocol.GoalPrintResult)
|
||||||
step "goal.tactic" ({ stateId := 1, tactic? := .some "intro y" }: Protocol.GoalTactic)
|
step "goal.tactic" ({ stateId := 1, tactic? := .some "intro y" }: Protocol.GoalTactic)
|
||||||
({ nextStateId? := .some 2, goals? := #[goal2], }: Protocol.GoalTacticResult)
|
({ nextStateId? := .some 2, goals? := #[goal2], }: Protocol.GoalTacticResult)
|
||||||
|
|
|
@ -55,7 +55,6 @@ def main (args: List String) := do
|
||||||
("Delate", Delate.suite),
|
("Delate", Delate.suite),
|
||||||
("Serial", Serial.suite),
|
("Serial", Serial.suite),
|
||||||
("Tactic/Assign", Tactic.Assign.suite),
|
("Tactic/Assign", Tactic.Assign.suite),
|
||||||
("Tactic/Fragment", Tactic.Fragment.suite),
|
|
||||||
("Tactic/Prograde", Tactic.Prograde.suite),
|
("Tactic/Prograde", Tactic.Prograde.suite),
|
||||||
("Tactic/Special", Tactic.Special.suite),
|
("Tactic/Special", Tactic.Special.suite),
|
||||||
]
|
]
|
||||||
|
|
250
Test/Proofs.lean
250
Test/Proofs.lean
|
@ -32,7 +32,7 @@ def startProof (start: Start): TestM (Option GoalState) := do
|
||||||
let expr ← parseSentence expr
|
let expr ← parseSentence expr
|
||||||
return .some $ ← GoalState.create (expr := expr)
|
return .some $ ← GoalState.create (expr := expr)
|
||||||
|
|
||||||
private def buildNamedGoal (name: String) (nameType: List (String × String)) (target: String)
|
def buildNamedGoal (name: String) (nameType: List (String × String)) (target: String)
|
||||||
(userName?: Option String := .none): Protocol.Goal :=
|
(userName?: Option String := .none): Protocol.Goal :=
|
||||||
{
|
{
|
||||||
name,
|
name,
|
||||||
|
@ -43,7 +43,7 @@ private def buildNamedGoal (name: String) (nameType: List (String × String)) (t
|
||||||
type? := .some { pp? := .some x.snd },
|
type? := .some { pp? := .some x.snd },
|
||||||
})).toArray
|
})).toArray
|
||||||
}
|
}
|
||||||
private def buildGoal (nameType: List (String × String)) (target: String) (userName?: Option String := .none):
|
def buildGoal (nameType: List (String × String)) (target: String) (userName?: Option String := .none):
|
||||||
Protocol.Goal :=
|
Protocol.Goal :=
|
||||||
{
|
{
|
||||||
userName?,
|
userName?,
|
||||||
|
@ -53,7 +53,7 @@ private def buildGoal (nameType: List (String × String)) (target: String) (user
|
||||||
type? := .some { pp? := .some x.snd },
|
type? := .some { pp? := .some x.snd },
|
||||||
})).toArray
|
})).toArray
|
||||||
}
|
}
|
||||||
private def proofRunner (env: Lean.Environment) (tests: TestM Unit): IO LSpec.TestSeq := do
|
def proofRunner (env: Lean.Environment) (tests: TestM Unit): IO LSpec.TestSeq := do
|
||||||
let termElabM := tests.run LSpec.TestSeq.done |>.run {} -- with default options
|
let termElabM := tests.run LSpec.TestSeq.done |>.run {} -- with default options
|
||||||
|
|
||||||
let coreContext: Lean.Core.Context ← createCoreContext #[]
|
let coreContext: Lean.Core.Context ← createCoreContext #[]
|
||||||
|
@ -77,7 +77,7 @@ def test_identity: TestM Unit := do
|
||||||
addTest $ LSpec.check "intro" ((← state1.serializeGoals (options := ← read)).map (·.name) =
|
addTest $ LSpec.check "intro" ((← state1.serializeGoals (options := ← read)).map (·.name) =
|
||||||
#[inner])
|
#[inner])
|
||||||
let state1parent ← state1.withParentContext do
|
let state1parent ← state1.withParentContext do
|
||||||
serializeExpressionSexp (← instantiateAll state1.parentExpr!)
|
serializeExpressionSexp (← instantiateAll state1.parentExpr?.get!)
|
||||||
addTest $ LSpec.test "(1 parent)" (state1parent == s!"(:lambda p (:sort 0) (:lambda h 0 (:subst (:mv {inner}) 1 0)))")
|
addTest $ LSpec.test "(1 parent)" (state1parent == s!"(:lambda p (:sort 0) (:lambda h 0 (:subst (:mv {inner}) 1 0)))")
|
||||||
|
|
||||||
-- Individual test cases
|
-- Individual test cases
|
||||||
|
@ -117,6 +117,41 @@ def test_nat_add_comm (manual: Bool): TestM Unit := do
|
||||||
addTest $ LSpec.test "rw [Nat.add_comm]" state2.goals.isEmpty
|
addTest $ LSpec.test "rw [Nat.add_comm]" state2.goals.isEmpty
|
||||||
|
|
||||||
return ()
|
return ()
|
||||||
|
def test_delta_variable: TestM Unit := do
|
||||||
|
let options: Protocol.Options := { noRepeat := true }
|
||||||
|
let state? ← startProof <| .expr "∀ (a b: Nat), a + b = b + a"
|
||||||
|
addTest $ LSpec.check "Start goal" state?.isSome
|
||||||
|
let state0 ← match state? with
|
||||||
|
| .some state => pure state
|
||||||
|
| .none => do
|
||||||
|
addTest $ assertUnreachable "Goal could not parse"
|
||||||
|
return ()
|
||||||
|
|
||||||
|
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := "intro n") with
|
||||||
|
| .success state _ => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check "intro n" ((← state1.serializeGoals (parent := state0) options).map (·.devolatilize) =
|
||||||
|
#[buildGoalSelective [("n", .some "Nat")] "∀ (b : Nat), n + b = b + n"])
|
||||||
|
let state2 ← match ← state1.tacticOn (goalId := 0) (tactic := "intro m") with
|
||||||
|
| .success state _ => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check "intro m" ((← state2.serializeGoals (parent := state1) options).map (·.devolatilize) =
|
||||||
|
#[buildGoalSelective [("n", .none), ("m", .some "Nat")] "n + m = m + n"])
|
||||||
|
return ()
|
||||||
|
where
|
||||||
|
-- Like `buildGoal` but allow certain variables to be elided.
|
||||||
|
buildGoalSelective (nameType: List (String × Option String)) (target: String): Protocol.Goal :=
|
||||||
|
{
|
||||||
|
target := { pp? := .some target},
|
||||||
|
vars := (nameType.map fun x => ({
|
||||||
|
userName := x.fst,
|
||||||
|
type? := x.snd.map (λ type => { pp? := type }),
|
||||||
|
})).toArray
|
||||||
|
}
|
||||||
|
|
||||||
example (w x y z : Nat) (p : Nat → Prop)
|
example (w x y z : Nat) (p : Nat → Prop)
|
||||||
(h : p (x * y + z * w * x)) : p (x * w * z + y * x) := by
|
(h : p (x * y + z * w * x)) : p (x * w * z + y * x) := by
|
||||||
|
@ -177,8 +212,8 @@ def test_or_comm: TestM Unit := do
|
||||||
| .none => do
|
| .none => do
|
||||||
addTest $ assertUnreachable "Goal could not parse"
|
addTest $ assertUnreachable "Goal could not parse"
|
||||||
return ()
|
return ()
|
||||||
checkTrue "(0 parent)" state0.parentMVars.isEmpty
|
addTest $ LSpec.check "(0 parent)" state0.parentExpr?.isNone
|
||||||
checkTrue "(0 root)" state0.rootExpr?.isNone
|
addTest $ LSpec.check "(0 root)" state0.rootExpr?.isNone
|
||||||
|
|
||||||
let tactic := "intro p q h"
|
let tactic := "intro p q h"
|
||||||
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
|
@ -202,11 +237,11 @@ def test_or_comm: TestM Unit := do
|
||||||
{ name := fvH, userName := "h", type? := .some { pp? := .some "p ∨ q" } }
|
{ name := fvH, userName := "h", type? := .some { pp? := .some "p ∨ q" } }
|
||||||
]
|
]
|
||||||
}])
|
}])
|
||||||
checkTrue "(1 parent)" state1.hasUniqueParent
|
checkTrue "(1 parent)" state1.parentExpr?.isSome
|
||||||
checkTrue "(1 root)" $ ¬ state1.isSolved
|
checkTrue "(1 root)" $ ¬ state1.isSolved
|
||||||
|
|
||||||
let state1parent ← state1.withParentContext do
|
let state1parent ← state1.withParentContext do
|
||||||
serializeExpressionSexp (← instantiateAll state1.parentExpr!)
|
serializeExpressionSexp (← instantiateAll state1.parentExpr?.get!)
|
||||||
addTest $ LSpec.test "(1 parent)" (state1parent == s!"(:lambda p (:sort 0) (:lambda q (:sort 0) (:lambda h ((:c Or) 1 0) (:subst (:mv {state1g0}) 2 1 0))))")
|
addTest $ LSpec.test "(1 parent)" (state1parent == s!"(:lambda p (:sort 0) (:lambda q (:sort 0) (:lambda h ((:c Or) 1 0) (:subst (:mv {state1g0}) 2 1 0))))")
|
||||||
let tactic := "cases h"
|
let tactic := "cases h"
|
||||||
let state2 ← match ← state1.tacticOn (goalId := 0) (tactic := tactic) with
|
let state2 ← match ← state1.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
|
@ -221,11 +256,11 @@ def test_or_comm: TestM Unit := do
|
||||||
let (caseL, caseR) := (state2g0.name.toString, state2g1.name.toString)
|
let (caseL, caseR) := (state2g0.name.toString, state2g1.name.toString)
|
||||||
addTest $ LSpec.check tactic ((← state2.serializeGoals (options := ← read)).map (·.name) =
|
addTest $ LSpec.check tactic ((← state2.serializeGoals (options := ← read)).map (·.name) =
|
||||||
#[caseL, caseR])
|
#[caseL, caseR])
|
||||||
checkTrue "(2 parent exists)" state2.hasUniqueParent
|
checkTrue "(2 parent exists)" state2.parentExpr?.isSome
|
||||||
checkTrue "(2 root)" $ ¬ state2.isSolved
|
checkTrue "(2 root)" $ ¬ state2.isSolved
|
||||||
|
|
||||||
let state2parent ← state2.withParentContext do
|
let state2parent ← state2.withParentContext do
|
||||||
serializeExpressionSexp (← instantiateAll state2.parentExpr!)
|
serializeExpressionSexp (← instantiateAll state2.parentExpr?.get!)
|
||||||
let orPQ := s!"((:c Or) (:fv {fvP}) (:fv {fvQ}))"
|
let orPQ := s!"((:c Or) (:fv {fvP}) (:fv {fvQ}))"
|
||||||
let orQP := s!"((:c Or) (:fv {fvQ}) (:fv {fvP}))"
|
let orQP := s!"((:c Or) (:fv {fvQ}) (:fv {fvP}))"
|
||||||
let motive := s!"(:lambda t {orPQ} (:forall h ((:c Eq) ((:c Or) (:fv {fvP}) (:fv {fvQ})) (:fv {fvH}) 0) {orQP}))"
|
let motive := s!"(:lambda t {orPQ} (:forall h ((:c Eq) ((:c Or) (:fv {fvP}) (:fv {fvQ})) (:fv {fvH}) 0) {orQP}))"
|
||||||
|
@ -241,7 +276,7 @@ def test_or_comm: TestM Unit := do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
let state3_1parent ← state3_1.withParentContext do
|
let state3_1parent ← state3_1.withParentContext do
|
||||||
serializeExpressionSexp (← instantiateAll state3_1.parentExpr!)
|
serializeExpressionSexp (← instantiateAll state3_1.parentExpr?.get!)
|
||||||
let [state3_1goal0] := state3_1.goals | fail "Should have 1 goal"
|
let [state3_1goal0] := state3_1.goals | fail "Should have 1 goal"
|
||||||
addTest $ LSpec.test "(3_1 parent)" (state3_1parent == s!"((:c Or.inr) (:fv {fvQ}) (:fv {fvP}) (:mv {state3_1goal0}))")
|
addTest $ LSpec.test "(3_1 parent)" (state3_1parent == s!"((:c Or.inr) (:fv {fvQ}) (:fv {fvP}) (:mv {state3_1goal0}))")
|
||||||
addTest $ LSpec.check "· apply Or.inr" (state3_1.goals.length = 1)
|
addTest $ LSpec.check "· apply Or.inr" (state3_1.goals.length = 1)
|
||||||
|
@ -251,7 +286,7 @@ def test_or_comm: TestM Unit := do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
addTest $ LSpec.check " assumption" state4_1.goals.isEmpty
|
addTest $ LSpec.check " assumption" state4_1.goals.isEmpty
|
||||||
let state4_1parent ← instantiateAll state4_1.parentExpr!
|
let state4_1parent ← instantiateAll state4_1.parentExpr?.get!
|
||||||
addTest $ LSpec.test "(4_1 parent)" state4_1parent.isFVar
|
addTest $ LSpec.test "(4_1 parent)" state4_1parent.isFVar
|
||||||
checkTrue "(4_1 root)" $ ¬ state4_1.isSolved
|
checkTrue "(4_1 root)" $ ¬ state4_1.isSolved
|
||||||
let state3_2 ← match ← state2.tacticOn (goalId := 1) (tactic := "apply Or.inl") with
|
let state3_2 ← match ← state2.tacticOn (goalId := 1) (tactic := "apply Or.inl") with
|
||||||
|
@ -301,6 +336,194 @@ def test_or_comm: TestM Unit := do
|
||||||
]
|
]
|
||||||
}
|
}
|
||||||
|
|
||||||
|
example : ∀ (a b c1 c2: Nat), (b + a) + c1 = (b + a) + c2 → (a + b) + c1 = (b + a) + c2 := by
|
||||||
|
intro a b c1 c2 h
|
||||||
|
conv =>
|
||||||
|
lhs
|
||||||
|
congr
|
||||||
|
. rw [Nat.add_comm]
|
||||||
|
. rfl
|
||||||
|
exact h
|
||||||
|
|
||||||
|
def test_conv: TestM Unit := do
|
||||||
|
let state? ← startProof (.expr "∀ (a b c1 c2: Nat), (b + a) + c1 = (b + a) + c2 → (a + b) + c1 = (b + a) + c2")
|
||||||
|
let state0 ← match state? with
|
||||||
|
| .some state => pure state
|
||||||
|
| .none => do
|
||||||
|
addTest $ assertUnreachable "Goal could not parse"
|
||||||
|
return ()
|
||||||
|
|
||||||
|
let tactic := "intro a b c1 c2 h"
|
||||||
|
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
|
| .success state _ => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check tactic ((← state1.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[interiorGoal [] "a + b + c1 = b + a + c2"])
|
||||||
|
|
||||||
|
let goalConv := state1.goals[0]!
|
||||||
|
let state2 ← match ← state1.conv (state1.get! 0) with
|
||||||
|
| .success state _ => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check "conv => ..." ((← state2.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[{ interiorGoal [] "a + b + c1 = b + a + c2" with isConversion := true }])
|
||||||
|
|
||||||
|
let convTactic := "rhs"
|
||||||
|
let state3R ← match ← state2.tacticOn (goalId := 0) convTactic with
|
||||||
|
| .success state _ => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check s!" {convTactic} (discard)" ((← state3R.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[{ interiorGoal [] "b + a + c2" with isConversion := true }])
|
||||||
|
|
||||||
|
let convTactic := "lhs"
|
||||||
|
let state3L ← match ← state2.tacticOn (goalId := 0) convTactic with
|
||||||
|
| .success state _ => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check s!" {convTactic}" ((← state3L.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[{ interiorGoal [] "a + b + c1" with isConversion := true }])
|
||||||
|
|
||||||
|
let convTactic := "congr"
|
||||||
|
let state4 ← match ← state3L.tacticOn (goalId := 0) convTactic with
|
||||||
|
| .success state _ => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check s!" {convTactic}" ((← state4.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[
|
||||||
|
{ interiorGoal [] "a + b" with isConversion := true, userName? := .some "a" },
|
||||||
|
{ interiorGoal [] "c1" with isConversion := true, userName? := .some "a" }
|
||||||
|
])
|
||||||
|
|
||||||
|
let convTactic := "rw [Nat.add_comm]"
|
||||||
|
let state5_1 ← match ← state4.tacticOn (goalId := 0) convTactic with
|
||||||
|
| .success state _ => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check s!" · {convTactic}" ((← state5_1.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[{ interiorGoal [] "b + a" with isConversion := true, userName? := .some "a" }])
|
||||||
|
|
||||||
|
let convTactic := "rfl"
|
||||||
|
let state6_1 ← match ← state5_1.tacticOn (goalId := 0) convTactic with
|
||||||
|
| .success state _ => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check s!" {convTactic}" ((← state6_1.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[])
|
||||||
|
|
||||||
|
let state4_1 ← match state6_1.continue state4 with
|
||||||
|
| .ok state => pure state
|
||||||
|
| .error e => do
|
||||||
|
addTest $ expectationFailure "continue" e
|
||||||
|
return ()
|
||||||
|
|
||||||
|
let convTactic := "rfl"
|
||||||
|
let state6 ← match ← state4_1.tacticOn (goalId := 0) convTactic with
|
||||||
|
| .success state _ => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check s!" · {convTactic}" ((← state6.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[])
|
||||||
|
|
||||||
|
let state1_1 ← match ← state6.convExit goalConv with
|
||||||
|
| .success state _ => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
|
||||||
|
let tactic := "exact h"
|
||||||
|
let stateF ← match ← state1_1.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
|
| .success state _ => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check tactic ((← stateF.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[])
|
||||||
|
|
||||||
|
where
|
||||||
|
h := "b + a + c1 = b + a + c2"
|
||||||
|
interiorGoal (free: List (String × String)) (target: String) :=
|
||||||
|
let free := [("a", "Nat"), ("b", "Nat"), ("c1", "Nat"), ("c2", "Nat"), ("h", h)] ++ free
|
||||||
|
buildGoal free target
|
||||||
|
|
||||||
|
example : ∀ (a b c d: Nat), a + b = b + c → b + c = c + d → a + b = c + d := by
|
||||||
|
intro a b c d h1 h2
|
||||||
|
calc a + b = b + c := by apply h1
|
||||||
|
_ = c + d := by apply h2
|
||||||
|
|
||||||
|
def test_calc: TestM Unit := do
|
||||||
|
let state? ← startProof (.expr "∀ (a b c d: Nat), a + b = b + c → b + c = c + d → a + b = c + d")
|
||||||
|
let state0 ← match state? with
|
||||||
|
| .some state => pure state
|
||||||
|
| .none => do
|
||||||
|
addTest $ assertUnreachable "Goal could not parse"
|
||||||
|
return ()
|
||||||
|
let tactic := "intro a b c d h1 h2"
|
||||||
|
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
|
| .success state _ => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check tactic ((← state1.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[interiorGoal [] "a + b = c + d"])
|
||||||
|
let pred := "a + b = b + c"
|
||||||
|
let state2 ← match ← state1.tryCalc (state1.get! 0) (pred := pred) with
|
||||||
|
| .success state _ => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check s!"calc {pred} := _" ((← state2.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[
|
||||||
|
interiorGoal [] "a + b = b + c" (.some "calc"),
|
||||||
|
interiorGoal [] "b + c = c + d"
|
||||||
|
])
|
||||||
|
addTest $ LSpec.test "(2.0 prev rhs)" (state2.calcPrevRhsOf? (state2.get! 0) |>.isNone)
|
||||||
|
addTest $ LSpec.test "(2.1 prev rhs)" (state2.calcPrevRhsOf? (state2.get! 1) |>.isSome)
|
||||||
|
|
||||||
|
let tactic := "apply h1"
|
||||||
|
let state2m ← match ← state2.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
|
| .success state _ => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
let state3 ← match state2m.continue state2 with
|
||||||
|
| .ok state => pure state
|
||||||
|
| .error e => do
|
||||||
|
addTest $ expectationFailure "continue" e
|
||||||
|
return ()
|
||||||
|
let pred := "_ = c + d"
|
||||||
|
let state4 ← match ← state3.tryCalc (state3.get! 0) (pred := pred) with
|
||||||
|
| .success state _ => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check s!"calc {pred} := _" ((← state4.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[
|
||||||
|
interiorGoal [] "b + c = c + d" (.some "calc")
|
||||||
|
])
|
||||||
|
addTest $ LSpec.test "(4.0 prev rhs)" (state4.calcPrevRhsOf? (state4.get! 0) |>.isNone)
|
||||||
|
let tactic := "apply h2"
|
||||||
|
let state4m ← match ← state4.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
|
| .success state _ => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.test "(4m root)" state4m.rootExpr?.isSome
|
||||||
|
where
|
||||||
|
interiorGoal (free: List (String × String)) (target: String) (userName?: Option String := .none) :=
|
||||||
|
let free := [("a", "Nat"), ("b", "Nat"), ("c", "Nat"), ("d", "Nat"),
|
||||||
|
("h1", "a + b = b + c"), ("h2", "b + c = c + d")] ++ free
|
||||||
|
buildGoal free target userName?
|
||||||
|
|
||||||
def test_tactic_failure_unresolved_goals : TestM Unit := do
|
def test_tactic_failure_unresolved_goals : TestM Unit := do
|
||||||
let state? ← startProof (.expr "∀ (p : Nat → Prop), ∃ (x : Nat), p (0 + x + 0)")
|
let state? ← startProof (.expr "∀ (p : Nat → Prop), ∃ (x : Nat), p (0 + x + 0)")
|
||||||
let state0 ← match state? with
|
let state0 ← match state? with
|
||||||
|
@ -377,8 +600,11 @@ def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
|
||||||
("identity", test_identity),
|
("identity", test_identity),
|
||||||
("Nat.add_comm", test_nat_add_comm false),
|
("Nat.add_comm", test_nat_add_comm false),
|
||||||
("Nat.add_comm manual", test_nat_add_comm true),
|
("Nat.add_comm manual", test_nat_add_comm true),
|
||||||
|
("Nat.add_comm delta", test_delta_variable),
|
||||||
("arithmetic", test_arith),
|
("arithmetic", test_arith),
|
||||||
("Or.comm", test_or_comm),
|
("Or.comm", test_or_comm),
|
||||||
|
("conv", test_conv),
|
||||||
|
("calc", test_calc),
|
||||||
("tactic failure with unresolved goals", test_tactic_failure_unresolved_goals),
|
("tactic failure with unresolved goals", test_tactic_failure_unresolved_goals),
|
||||||
("tactic failure with synthesize placeholder", test_tactic_failure_synthesize_placeholder),
|
("tactic failure with synthesize placeholder", test_tactic_failure_synthesize_placeholder),
|
||||||
("deconstruct", test_deconstruct),
|
("deconstruct", test_deconstruct),
|
||||||
|
|
|
@ -87,12 +87,12 @@ def test_pickling_env_extensions : TestM Unit := do
|
||||||
let .ok value ← elabTerm (← `(term|sorry)) (.some type) | unreachable!
|
let .ok value ← elabTerm (← `(term|sorry)) (.some type) | unreachable!
|
||||||
pure (type, value)
|
pure (type, value)
|
||||||
let .success state1 _ ← state.tryTacticM goal (Tactic.assignWithAuxLemma type value) | unreachable!
|
let .success state1 _ ← state.tryTacticM goal (Tactic.assignWithAuxLemma type value) | unreachable!
|
||||||
let parentExpr := state1.parentExpr!
|
let parentExpr := state1.parentExpr?.get!
|
||||||
checkTrue "src has aux lemma" $ parentExpr.getUsedConstants.any λ name => name.isAuxLemma
|
checkTrue "src has aux lemma" $ parentExpr.getUsedConstants.any λ name => name.isAuxLemma
|
||||||
goalStatePickle state1 statePath
|
goalStatePickle state1 statePath
|
||||||
let ((), _) ← runCoreM coreDst $ transformTestT runTermElabMInCore do
|
let ((), _) ← runCoreM coreDst $ transformTestT runTermElabMInCore do
|
||||||
let (state1, _) ← goalStateUnpickle statePath (← getEnv)
|
let (state1, _) ← goalStateUnpickle statePath (← getEnv)
|
||||||
let parentExpr := state1.parentExpr!
|
let parentExpr := state1.parentExpr?.get!
|
||||||
checkTrue "dst has aux lemma" $ parentExpr.getUsedConstants.any λ name => name.isAuxLemma
|
checkTrue "dst has aux lemma" $ parentExpr.getUsedConstants.any λ name => name.isAuxLemma
|
||||||
|
|
||||||
return ()
|
return ()
|
||||||
|
|
|
@ -1,4 +1,3 @@
|
||||||
import Test.Tactic.Assign
|
import Test.Tactic.Assign
|
||||||
import Test.Tactic.Fragment
|
|
||||||
import Test.Tactic.Prograde
|
import Test.Tactic.Prograde
|
||||||
import Test.Tactic.Special
|
import Test.Tactic.Special
|
||||||
|
|
|
@ -1,312 +0,0 @@
|
||||||
import Pantograph.Goal
|
|
||||||
import Test.Common
|
|
||||||
|
|
||||||
open Lean
|
|
||||||
|
|
||||||
namespace Pantograph.Test.Tactic.Fragment
|
|
||||||
|
|
||||||
private def buildGoal (nameType: List (String × String)) (target: String):
|
|
||||||
Protocol.Goal :=
|
|
||||||
{
|
|
||||||
target := { pp? := .some target},
|
|
||||||
vars := (nameType.map fun x => ({
|
|
||||||
userName := x.fst,
|
|
||||||
type? := .some { pp? := .some x.snd },
|
|
||||||
})).toArray
|
|
||||||
}
|
|
||||||
|
|
||||||
abbrev TestM := TestT $ Elab.TermElabM
|
|
||||||
|
|
||||||
example : ∀ (a b c1 c2: Nat), (b + a) + c1 = (b + a) + c2 → (a + b) + c1 = (b + a) + c2 := by
|
|
||||||
intro a b c1 c2 h
|
|
||||||
conv =>
|
|
||||||
lhs
|
|
||||||
congr
|
|
||||||
. rw [Nat.add_comm]
|
|
||||||
. rfl
|
|
||||||
exact h
|
|
||||||
|
|
||||||
def test_conv_simple: TestM Unit := do
|
|
||||||
let rootTarget ← parseSentence "∀ (a b c1 c2: Nat), (b + a) + c1 = (b + a) + c2 → (a + b) + c1 = (b + a) + c2"
|
|
||||||
let state0 ← GoalState.create rootTarget
|
|
||||||
|
|
||||||
let tactic := "intro a b c1 c2 h"
|
|
||||||
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check tactic ((← state1.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[interiorGoal [] "a + b + c1 = b + a + c2"])
|
|
||||||
|
|
||||||
let goalConv := state1.goals[0]!
|
|
||||||
let state2 ← match ← state1.convEnter (state1.get! 0) with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check "conv => ..." ((← state2.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[{ interiorGoal [] "a + b + c1 = b + a + c2" with fragment := .conv }])
|
|
||||||
|
|
||||||
let convTactic := "rhs"
|
|
||||||
let state3R ← match ← state2.tacticOn (goalId := 0) convTactic with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check s!" {convTactic} (discard)" ((← state3R.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[{ interiorGoal [] "b + a + c2" with fragment := .conv }])
|
|
||||||
|
|
||||||
let convTactic := "lhs"
|
|
||||||
let state3L ← match ← state2.tacticOn (goalId := 0) convTactic with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check s!" {convTactic}" ((← state3L.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[{ interiorGoal [] "a + b + c1" with fragment := .conv }])
|
|
||||||
|
|
||||||
let convTactic := "congr"
|
|
||||||
let state4 ← match ← state3L.tacticOn (goalId := 0) convTactic with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check s!" {convTactic}" ((← state4.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[
|
|
||||||
{ interiorGoal [] "a + b" with fragment := .conv, userName? := .some "a" },
|
|
||||||
{ interiorGoal [] "c1" with fragment := .conv, userName? := .some "a" }
|
|
||||||
])
|
|
||||||
|
|
||||||
let convTactic := "rw [Nat.add_comm]"
|
|
||||||
let state5_1 ← match ← state4.tacticOn (goalId := 0) convTactic with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check s!" · {convTactic}" ((← state5_1.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[{ interiorGoal [] "b + a" with fragment := .conv, userName? := .some "a" }])
|
|
||||||
|
|
||||||
let convTactic := "rfl"
|
|
||||||
let state6_1 ← match ← state5_1.tacticOn (goalId := 0) convTactic with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check s!" {convTactic}" ((← state6_1.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[])
|
|
||||||
|
|
||||||
let state4_1 ← match state6_1.continue state4 with
|
|
||||||
| .ok state => pure state
|
|
||||||
| .error e => do
|
|
||||||
addTest $ expectationFailure "continue" e
|
|
||||||
return ()
|
|
||||||
|
|
||||||
let convTactic := "rfl"
|
|
||||||
let state1_1 ← match ← state4_1.tacticOn (goalId := 0) convTactic with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check s!" · {convTactic}" ((← state1_1.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[interiorGoal [] "b + a + c1 = b + a + c2"])
|
|
||||||
checkEq "(fragments)" state1_1.fragments.size 0
|
|
||||||
|
|
||||||
/-
|
|
||||||
let state1_1 ← match ← state6.fragmentExit goalConv with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
-/
|
|
||||||
|
|
||||||
let tactic := "exact h"
|
|
||||||
let stateF ← match ← state1_1.tacticOn (goalId := 0) (tactic := tactic) with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check tactic ((← stateF.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[])
|
|
||||||
|
|
||||||
where
|
|
||||||
h := "b + a + c1 = b + a + c2"
|
|
||||||
interiorGoal (free: List (String × String)) (target: String) :=
|
|
||||||
let free := [("a", "Nat"), ("b", "Nat"), ("c1", "Nat"), ("c2", "Nat"), ("h", h)] ++ free
|
|
||||||
buildGoal free target
|
|
||||||
|
|
||||||
example (p : Prop) (x y z : Nat) : p → (p → x = y) → x + z = y + z ∧ p := by
|
|
||||||
intro hp hi
|
|
||||||
apply And.intro
|
|
||||||
conv =>
|
|
||||||
rhs
|
|
||||||
arg 1
|
|
||||||
rw [←hi]
|
|
||||||
rfl
|
|
||||||
tactic => exact hp
|
|
||||||
exact hp
|
|
||||||
|
|
||||||
def test_conv_unshielded : TestM Unit := do
|
|
||||||
let rootTarget ← parseSentence "∀ (p : Prop) (x y z : Nat), p → (p → x = y) → x + z = y + z ∧ p"
|
|
||||||
let state ← GoalState.create rootTarget
|
|
||||||
let tactic := "intro p x y z hp hi"
|
|
||||||
let .success state _ ← state.tacticOn 0 tactic | fail "intro failed"
|
|
||||||
let tactic := "apply And.intro"
|
|
||||||
let .success state _ ← state.tacticOn 0 tactic | fail "apply failed"
|
|
||||||
let .success state _ ← state.convEnter (.prefer state.goals[0]!) | fail "Cannot enter conversion tactic mode"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus "rhs" | fail "rhs failed"
|
|
||||||
let tactic := "arg 1"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize))
|
|
||||||
#[
|
|
||||||
{ interiorGoal [] "y" with fragment := .conv },
|
|
||||||
{ interiorGoal [] "p" with userName? := "right", },
|
|
||||||
]
|
|
||||||
let tactic := "rw [←hi]"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
checkEq s!" {tactic}" state.goals.length 3
|
|
||||||
let tactic := "rfl"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize))
|
|
||||||
#[
|
|
||||||
interiorGoal [] "p",
|
|
||||||
{ interiorGoal [] "p" with userName? := "right", },
|
|
||||||
]
|
|
||||||
checkEq "(n goals)" state.goals.length 2
|
|
||||||
checkEq "(fragments)" state.fragments.size 0
|
|
||||||
let tactic := "exact hp"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
let tactic := "exact hp"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
let root? := state.rootExpr?
|
|
||||||
checkTrue "root" root?.isSome
|
|
||||||
where
|
|
||||||
interiorGoal (free: List (String × String)) (target: String) :=
|
|
||||||
let free := [("p", "Prop"), ("x", "Nat"), ("y", "Nat"), ("z", "Nat"), ("hp", "p"), ("hi", "p → x = y")] ++ free
|
|
||||||
buildGoal free target
|
|
||||||
|
|
||||||
example : ∀ (x y z w : Nat), y = z → x + z = w → x + y = w := by
|
|
||||||
intro x y z w hyz hxzw
|
|
||||||
conv =>
|
|
||||||
lhs
|
|
||||||
arg 2
|
|
||||||
rw [hyz]
|
|
||||||
rfl
|
|
||||||
exact hxzw
|
|
||||||
|
|
||||||
def test_conv_unfinished : TestM Unit := do
|
|
||||||
let rootTarget ← parseSentence "∀ (x y z w : Nat), y = z → x + z = w → x + y = w"
|
|
||||||
let state ← GoalState.create rootTarget
|
|
||||||
let tactic := "intro x y z w hyz hxzw"
|
|
||||||
let .success state _ ← state.tacticOn 0 tactic | fail "intro failed"
|
|
||||||
let convParent := state.goals[0]!
|
|
||||||
let .success state _ ← state.convEnter (.prefer convParent) | fail "Cannot enter conversion tactic mode"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus "lhs" | fail "rhs failed"
|
|
||||||
let tactic := "arg 2"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize))
|
|
||||||
#[
|
|
||||||
{ interiorGoal [] "y" with fragment := .conv },
|
|
||||||
]
|
|
||||||
let tactic := "rw [hyz]"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize))
|
|
||||||
#[
|
|
||||||
{ interiorGoal [] "z" with fragment := .conv },
|
|
||||||
]
|
|
||||||
checkTrue " (fragment)" $ state.fragments.contains state.mainGoal?.get!
|
|
||||||
checkTrue " (fragment parent)" $ state.fragments.contains convParent
|
|
||||||
checkTrue " (main goal)" state.mainGoal?.isSome
|
|
||||||
let tactic := "rfl"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize))
|
|
||||||
#[
|
|
||||||
interiorGoal [] "x + z = w",
|
|
||||||
]
|
|
||||||
checkEq "(fragments)" state.fragments.size 0
|
|
||||||
checkEq s!" {tactic}" state.goals.length 1
|
|
||||||
let tactic := "exact hxzw"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
let root? := state.rootExpr?
|
|
||||||
checkTrue "root" root?.isSome
|
|
||||||
where
|
|
||||||
interiorGoal (free: List (String × String)) (target: String) :=
|
|
||||||
let free := [("x", "Nat"), ("y", "Nat"), ("z", "Nat"), ("w", "Nat"), ("hyz", "y = z"), ("hxzw", "x + z = w")] ++ free
|
|
||||||
buildGoal free target
|
|
||||||
|
|
||||||
example : ∀ (a b c d: Nat), a + b = b + c → b + c = c + d → a + b = c + d := by
|
|
||||||
intro a b c d h1 h2
|
|
||||||
calc a + b = b + c := by apply h1
|
|
||||||
_ = c + d := by apply h2
|
|
||||||
|
|
||||||
def test_calc: TestM Unit := do
|
|
||||||
let rootTarget ← parseSentence "∀ (a b c d: Nat), a + b = b + c → b + c = c + d → a + b = c + d"
|
|
||||||
let state0 ← GoalState.create rootTarget
|
|
||||||
let tactic := "intro a b c d h1 h2"
|
|
||||||
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check tactic ((← state1.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[interiorGoal [] "a + b = c + d"])
|
|
||||||
let pred := "a + b = b + c"
|
|
||||||
let .success state1 _ ← state1.calcEnter state1.mainGoal?.get! | fail "Could not enter calc"
|
|
||||||
let state2 ← match ← state1.tacticOn 0 pred with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check s!"calc {pred} := _" ((← state2.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[
|
|
||||||
{ interiorGoal [] "a + b = b + c" with userName? := .some "calc" },
|
|
||||||
{ interiorGoal [] "b + c = c + d" with fragment := .calc },
|
|
||||||
])
|
|
||||||
addTest $ LSpec.test "(2.0 prev rhs)" (state2.calcPrevRhsOf? (state2.get! 0) |>.isNone)
|
|
||||||
addTest $ LSpec.test "(2.1 prev rhs)" (state2.calcPrevRhsOf? (state2.get! 1) |>.isSome)
|
|
||||||
|
|
||||||
let tactic := "apply h1"
|
|
||||||
let state2m ← match ← state2.tacticOn (goalId := 0) (tactic := tactic) with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
let state3 ← match state2m.continue state2 with
|
|
||||||
| .ok state => pure state
|
|
||||||
| .error e => do
|
|
||||||
addTest $ expectationFailure "continue" e
|
|
||||||
return ()
|
|
||||||
let pred := "_ = c + d"
|
|
||||||
let state4 ← match ← state3.tacticOn 0 pred with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check s!"calc {pred} := _" ((← state4.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[
|
|
||||||
{ interiorGoal [] "b + c = c + d" with userName? := .some "calc" },
|
|
||||||
])
|
|
||||||
addTest $ LSpec.test "(4.0 prev rhs)" (state4.calcPrevRhsOf? (state4.get! 0) |>.isNone)
|
|
||||||
let tactic := "apply h2"
|
|
||||||
let state4m ← match ← state4.tacticOn (goalId := 0) (tactic := tactic) with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
checkEq "(fragments)" state4m.fragments.size 0
|
|
||||||
addTest $ LSpec.test "(4m root)" state4m.rootExpr?.isSome
|
|
||||||
where
|
|
||||||
interiorGoal (free: List (String × String)) (target: String) :=
|
|
||||||
let free := [("a", "Nat"), ("b", "Nat"), ("c", "Nat"), ("d", "Nat"),
|
|
||||||
("h1", "a + b = b + c"), ("h2", "b + c = c + d")] ++ free
|
|
||||||
buildGoal free target
|
|
||||||
|
|
||||||
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
|
|
||||||
[
|
|
||||||
("conv simple", test_conv_simple),
|
|
||||||
("conv unshielded", test_conv_unshielded),
|
|
||||||
("conv unfinished", test_conv_unfinished),
|
|
||||||
("calc", test_calc),
|
|
||||||
] |>.map (λ (name, t) => (name, runTestTermElabM env t))
|
|
||||||
|
|
||||||
end Pantograph.Test.Tactic.Fragment
|
|
|
@ -0,0 +1,29 @@
|
||||||
|
import Pantograph.Frontend
|
||||||
|
|
||||||
|
open Lean
|
||||||
|
|
||||||
|
namespace Pantograph
|
||||||
|
|
||||||
|
def fail (s : String) : IO Unit := do
|
||||||
|
IO.eprintln s
|
||||||
|
|
||||||
|
def dissect (args: List String): IO Unit := do
|
||||||
|
let fileName :: _args := args | fail s!"Must supply a file name"
|
||||||
|
let file ← IO.FS.readFile fileName
|
||||||
|
let (context, state) ← do Frontend.createContextStateFromFile file fileName (env? := .none) {}
|
||||||
|
let frontendM: Elab.Frontend.FrontendM _ :=
|
||||||
|
Frontend.mapCompilationSteps λ step => do
|
||||||
|
for tree in step.trees do
|
||||||
|
IO.println s!"{← tree.toString}"
|
||||||
|
let (_, _) ← frontendM.run context |>.run state
|
||||||
|
return ()
|
||||||
|
|
||||||
|
end Pantograph
|
||||||
|
|
||||||
|
open Pantograph
|
||||||
|
|
||||||
|
def main (args : List String) : IO Unit := do
|
||||||
|
let command :: args := args | IO.eprintln "Must supply a command"
|
||||||
|
match command with
|
||||||
|
| "dissect" => dissect args
|
||||||
|
| _ => IO.eprintln s!"Unknown command {command}"
|
21
doc/repl.md
21
doc/repl.md
|
@ -13,7 +13,7 @@ See `Pantograph/Protocol.lean` for a description of the parameters and return va
|
||||||
only the values of definitions are printed.
|
only the values of definitions are printed.
|
||||||
* `env.save { "path": <fileName> }`, `env.load { "path": <fileName> }`: Save/Load the
|
* `env.save { "path": <fileName> }`, `env.load { "path": <fileName> }`: Save/Load the
|
||||||
current environment to/from a file
|
current environment to/from a file
|
||||||
* `env.module_read { "module": <name> }`: Reads a list of symbols from a module
|
* `env.module_read { "module": <name }`: Reads a list of symbols from a module
|
||||||
* `env.describe {}`: Describes the imports and modules in the current environment
|
* `env.describe {}`: Describes the imports and modules in the current environment
|
||||||
* `options.set { key: value, ... }`: Set one or more options (not Lean options; those
|
* `options.set { key: value, ... }`: Set one or more options (not Lean options; those
|
||||||
have to be set via command line arguments.), for options, see `Pantograph/Protocol.lean`
|
have to be set via command line arguments.), for options, see `Pantograph/Protocol.lean`
|
||||||
|
@ -28,19 +28,16 @@ See `Pantograph/Protocol.lean` for a description of the parameters and return va
|
||||||
* `options.print`: Display the current set of options
|
* `options.print`: Display the current set of options
|
||||||
* `goal.start {["name": <name>], ["expr": <expr>], ["levels": [<levels>]], ["copyFrom": <symbol>]}`:
|
* `goal.start {["name": <name>], ["expr": <expr>], ["levels": [<levels>]], ["copyFrom": <symbol>]}`:
|
||||||
Start a new proof from a given expression or symbol
|
Start a new proof from a given expression or symbol
|
||||||
* `goal.tactic {"stateId": <id>, ["goalId": <id>], ["autoResume": <bool>], ...}`:
|
* `goal.tactic {"stateId": <id>, "goalId": <id>, ...}`: Execute a tactic string on a
|
||||||
Execute a tactic string on a given goal site. The tactic is supplied as additional
|
given goal. The tactic is supplied as additional key-value pairs in one of the following formats:
|
||||||
key-value pairs in one of the following formats:
|
- `{ "tactic": <tactic> }`: Execute an ordinary tactic
|
||||||
- `{ "tactic": <tactic> }`: Executes a tactic in the current mode
|
|
||||||
- `{ "mode": <mode> }`: Enter a different tactic mode. The permitted values
|
|
||||||
are `tactic` (default), `conv`, `calc`. In case of `calc`, each step must
|
|
||||||
be of the form `lhs op rhs`. An `lhs` of `_` indicates that it should be set
|
|
||||||
to the previous `rhs`.
|
|
||||||
- `{ "expr": <expr> }`: Assign the given proof term to the current goal
|
- `{ "expr": <expr> }`: Assign the given proof term to the current goal
|
||||||
- `{ "have": <expr>, "binderName": <name> }`: Execute `have` and creates a branch goal
|
- `{ "have": <expr>, "binderName": <name> }`: Execute `have` and creates a branch goal
|
||||||
- `{ "let": <expr>, "binderName": <name> }`: Execute `let` and creates a branch goal
|
- `{ "calc": <expr> }`: Execute one step of a `calc` tactic. Each step must
|
||||||
- `{ "draft": <expr> }`: Draft an expression with `sorry`s, turning them into
|
be of the form `lhs op rhs`. An `lhs` of `_` indicates that it should be set
|
||||||
goals. Coupling is not allowed.
|
to the previous `rhs`.
|
||||||
|
- `{ "conv": <bool> }`: Enter or exit conversion tactic mode.
|
||||||
|
- `{ "draft": <expr> }`: Draft an expression with `sorry`s, turning them into goals. Coupling is not allowed.
|
||||||
If the `goals` field does not exist, the tactic execution has failed. Read
|
If the `goals` field does not exist, the tactic execution has failed. Read
|
||||||
`messages` to find the reason.
|
`messages` to find the reason.
|
||||||
* `goal.continue {"stateId": <id>, ["branch": <id>], ["goals": <names>]}`:
|
* `goal.continue {"stateId": <id>, ["branch": <id>], ["goals": <names>]}`:
|
||||||
|
|
|
@ -10,6 +10,7 @@ lean_lib Pantograph {
|
||||||
|
|
||||||
lean_lib Repl {
|
lean_lib Repl {
|
||||||
}
|
}
|
||||||
|
|
||||||
@[default_target]
|
@[default_target]
|
||||||
lean_exe repl {
|
lean_exe repl {
|
||||||
root := `Main
|
root := `Main
|
||||||
|
@ -17,6 +18,13 @@ lean_exe repl {
|
||||||
supportInterpreter := true
|
supportInterpreter := true
|
||||||
}
|
}
|
||||||
|
|
||||||
|
@[default_target]
|
||||||
|
lean_exe tomograph {
|
||||||
|
root := `Tomograph
|
||||||
|
-- Solves the native symbol not found problem
|
||||||
|
supportInterpreter := true
|
||||||
|
}
|
||||||
|
|
||||||
require LSpec from git
|
require LSpec from git
|
||||||
"https://github.com/argumentcomputer/LSpec.git" @ "a6652a48b5c67b0d8dd3930fad6390a97d127e8d"
|
"https://github.com/argumentcomputer/LSpec.git" @ "a6652a48b5c67b0d8dd3930fad6390a97d127e8d"
|
||||||
lean_lib Test {
|
lean_lib Test {
|
||||||
|
|
Loading…
Reference in New Issue