Merge branch 'dev' into misc/version

This commit is contained in:
Leni Aniva 2024-12-11 20:51:17 -08:00
commit 13e01b9e62
Signed by: aniva
GPG Key ID: 4D9B1C8D10EA4C50
21 changed files with 541 additions and 183 deletions

View File

@ -1,4 +1,4 @@
/- Adapted from lean-training-data by semorrison -/
import Pantograph.Frontend.Basic import Pantograph.Frontend.Basic
import Pantograph.Frontend.Elab import Pantograph.Frontend.Elab
import Pantograph.Frontend.InfoTree
import Pantograph.Frontend.MetaTranslate import Pantograph.Frontend.MetaTranslate

View File

@ -30,6 +30,13 @@ end Lean.PersistentArray
namespace Pantograph.Frontend namespace Pantograph.Frontend
@[export pantograph_frontend_stx_byte_range]
def stxByteRange (stx : Syntax) : String.Pos × String.Pos :=
let pos := stx.getPos?.getD 0
let endPos := stx.getTailPos?.getD 0
(pos, endPos)
abbrev FrontendM := Elab.Frontend.FrontendM abbrev FrontendM := Elab.Frontend.FrontendM
structure CompilationStep where structure CompilationStep where

View File

@ -1,87 +1,21 @@
/- Adapted from https://github.com/semorrison/lean-training-data -/
import Lean.Elab.Import import Lean.Elab.Import
import Lean.Elab.Command import Lean.Elab.Command
import Lean.Elab.InfoTree import Lean.Elab.InfoTree
import Lean.DeclarationRange
import Pantograph.Frontend.Basic import Pantograph.Frontend.Basic
import Pantograph.Frontend.MetaTranslate import Pantograph.Frontend.MetaTranslate
import Pantograph.Goal import Pantograph.Goal
import Pantograph.Protocol import Pantograph.Protocol
import Pantograph.Frontend.InfoTree
open Lean open Lean
namespace Lean.Elab.Info
/-- The `Syntax` for a `Lean.Elab.Info`, if there is one. -/
protected def stx? : Info → Option Syntax
| .ofTacticInfo info => info.stx
| .ofTermInfo info => info.stx
| .ofCommandInfo info => info.stx
| .ofMacroExpansionInfo info => info.stx
| .ofOptionInfo info => info.stx
| .ofFieldInfo info => info.stx
| .ofCompletionInfo info => info.stx
| .ofUserWidgetInfo info => info.stx
| .ofCustomInfo info => info.stx
| .ofFVarAliasInfo _ => none
| .ofFieldRedeclInfo info => info.stx
| .ofOmissionInfo info => info.stx
/-- Is the `Syntax` for this `Lean.Elab.Info` original, or synthetic? -/
protected def isOriginal (i : Info) : Bool :=
match i.stx? with
| none => true -- Somewhat unclear what to do with `FVarAliasInfo`, so be conservative.
| some stx => match stx.getHeadInfo with
| .original .. => true
| _ => false
end Lean.Elab.Info
namespace Lean.Elab.TacticInfo
/-- Find the name for the outermost `Syntax` in this `TacticInfo`. -/
def name? (t : TacticInfo) : Option Name :=
match t.stx with
| Syntax.node _ n _ => some n
| _ => none
/-- Decide whether a tactic is "substantive",
or is merely a tactic combinator (e.g. `by`, `;`, multiline tactics, parenthesized tactics). -/
def isSubstantive (t : TacticInfo) : Bool :=
match t.name? with
| none => false
| some `null => false
| some ``cdot => false
| some ``cdotTk => false
| some ``Lean.Parser.Term.byTactic => false
| some ``Lean.Parser.Tactic.tacticSeq => false
| some ``Lean.Parser.Tactic.tacticSeq1Indented => false
| some ``Lean.Parser.Tactic.«tactic_<;>_» => false
| some ``Lean.Parser.Tactic.paren => false
| _ => true
end Lean.Elab.TacticInfo
namespace Lean.Elab.InfoTree
/--
Keep `.node` nodes and `.hole` nodes satisfying predicates.
Returns a `List InfoTree`, although in most situations this will be a singleton.
-/
partial def filter (p : Info → Bool) (m : MVarId → Bool := fun _ => false) :
InfoTree → List InfoTree
| .context ctx tree => tree.filter p m |>.map (.context ctx)
| .node info children =>
if p info then
[.node info (children.toList.map (filter p m)).flatten.toPArray']
else
(children.toList.map (filter p m)).flatten
| .hole mvar => if m mvar then [.hole mvar] else []
end Lean.Elab.InfoTree
namespace Pantograph.Frontend namespace Pantograph.Frontend
-- Info tree filtering functions -- Info tree filtering functions
/- Adapted from lean-training-data -/
structure TacticInvocation where structure TacticInvocation where
info : Elab.TacticInfo info : Elab.TacticInfo
ctx : Elab.ContextInfo ctx : Elab.ContextInfo
@ -131,20 +65,10 @@ protected def usedConstants (t: TacticInvocation) : NameSet :=
end TacticInvocation end TacticInvocation
/-- Analogue of `Lean.Elab.InfoTree.findInfo?`, but that returns a list of all results. -/
partial def findAllInfo (t : Elab.InfoTree) (context?: Option Elab.ContextInfo) (pred : Elab.Info → Bool) :
List (Elab.Info × Option Elab.ContextInfo × PersistentArray Elab.InfoTree) :=
match t with
| .context inner t => findAllInfo t (inner.mergeIntoOuter? context?) pred
| .node i children =>
(if pred i then [(i, context?, children)] else []) ++
children.toList.flatMap (fun t => findAllInfo t context? pred)
| _ => []
/-- Return all `TacticInfo` nodes in an `InfoTree` corresponding to tactics, /-- Return all `TacticInfo` nodes in an `InfoTree` corresponding to tactics,
each equipped with its relevant `ContextInfo`, and any children info trees. -/ each equipped with its relevant `ContextInfo`, and any children info trees. -/
private def collectTacticNodes (t : Elab.InfoTree) : List TacticInvocation := private def collectTacticNodes (t : Elab.InfoTree) : List TacticInvocation :=
let infos := findAllInfo t none fun i => match i with let infos := t.findAllInfo none false fun i => match i with
| .ofTacticInfo _ => true | .ofTacticInfo _ => true
| _ => false | _ => false
infos.filterMap fun p => match p with infos.filterMap fun p => match p with
@ -163,9 +87,11 @@ def collectTacticsFromCompilationStep (step : CompilationStep) : IO (List Protoc
tactics.mapM λ invocation => do tactics.mapM λ invocation => do
let goalBefore := (Format.joinSep (← invocation.goalState) "\n").pretty let goalBefore := (Format.joinSep (← invocation.goalState) "\n").pretty
let goalAfter := (Format.joinSep (← invocation.goalStateAfter) "\n").pretty let goalAfter := (Format.joinSep (← invocation.goalStateAfter) "\n").pretty
let tactic ← invocation.ctx.runMetaM {} do let tactic ← invocation.ctx.runMetaM {} <| Meta.withMCtx invocation.info.mctxBefore do
let t ← PrettyPrinter.ppTactic ⟨invocation.info.stx⟩ return (← invocation.ctx.ppSyntax {} invocation.info.stx).pretty
return t.pretty -- FIXME: Why does this not work? There are problems with `term.pseudo.antiquot`
--PrettyPrinter.ppTactic ⟨invocation.info.stx⟩
--return t.pretty
let usedConstants := invocation.usedConstants.toArray.map λ n => n.toString let usedConstants := invocation.usedConstants.toArray.map λ n => n.toString
return { return {
goalBefore, goalBefore,
@ -178,47 +104,79 @@ structure InfoWithContext where
info: Elab.Info info: Elab.Info
context?: Option Elab.ContextInfo := .none context?: Option Elab.ContextInfo := .none
private def collectSorrysInTree (t : Elab.InfoTree) : List InfoWithContext := private def collectSorrysInTree (t : Elab.InfoTree) : IO (List InfoWithContext) := do
let infos := findAllInfo t none fun i => match i with let infos ← t.findAllInfoM none fun i ctx? => match i with
| .ofTermInfo { expectedType?, expr, stx, .. } => | .ofTermInfo { expectedType?, expr, stx, lctx, .. } => do
expr.isSorry ∧ expectedType?.isSome ∧ stx.isOfKind `Lean.Parser.Term.sorry let .some ctx := ctx? | return (false, true)
if expr.isSorry ∧ stx.isOfKind `Lean.Parser.Term.sorry then
if expectedType?.isNone then
throw $ .userError "Sorry of indeterminant type is not allowed"
return (true, false)
let .some expectedType := expectedType? | return (false, true)
let typeMatch ← ctx.runMetaM lctx do
let type ← Meta.inferType expr
Meta.isExprDefEqGuarded type expectedType
return match typeMatch, expr.hasSorry with
| false, true => (true, false) -- Types mismatch but has sorry -> collect, halt
| false, false => (true, false) -- Types mistmatch but no sorry -> collect, halt
| true, true => (false, true) -- Types match but has sorry -> continue
| true, false => (false, false) -- Types match but no sorries -> halt
| .ofTacticInfo { stx, goalsBefore, .. } => | .ofTacticInfo { stx, goalsBefore, .. } =>
-- The `sorry` term is distinct from the `sorry` tactic -- The `sorry` term is distinct from the `sorry` tactic
let isSorry := stx.isOfKind `Lean.Parser.Tactic.tacticSorry let isSorry := stx.isOfKind `Lean.Parser.Tactic.tacticSorry
isSorry ∧ !goalsBefore.isEmpty return (isSorry ∧ !goalsBefore.isEmpty, ¬ isSorry)
| _ => false | _ => return (false, true)
infos.map fun (info, context?, _) => { info, context? } return infos.map fun (info, context?, _) => { info, context? }
-- NOTE: Plural deliberately not spelled "sorries" -- NOTE: Plural deliberately not spelled "sorries"
@[export pantograph_frontend_collect_sorrys_m] @[export pantograph_frontend_collect_sorrys_m]
def collectSorrys (step: CompilationStep) : List InfoWithContext := def collectSorrys (step: CompilationStep) : IO (List InfoWithContext) := do
step.trees.flatMap collectSorrysInTree return (← step.trees.mapM collectSorrysInTree).flatten
structure AnnotatedGoalState where
state : GoalState
srcBoundaries : List (String.Pos × String.Pos)
/-- /--
Since we cannot directly merge `MetavarContext`s, we have to get creative. This Since we cannot directly merge `MetavarContext`s, we have to get creative. This
function duplicates frozen mvars in term and tactic info nodes, and add them to function duplicates frozen mvars in term and tactic info nodes, and add them to
the current `MetavarContext`. the current `MetavarContext`.
-/ -/
@[export pantograph_frontend_sorrys_to_goal_state] @[export pantograph_frontend_sorrys_to_goal_state_m]
def sorrysToGoalState (sorrys : List InfoWithContext) : MetaM GoalState := do def sorrysToGoalState (sorrys : List InfoWithContext) : MetaM AnnotatedGoalState := do
assert! !sorrys.isEmpty assert! !sorrys.isEmpty
let goalsM := sorrys.mapM λ i => do let goalsM := sorrys.mapM λ i => do
match i.info with match i.info with
| .ofTermInfo termInfo => do | .ofTermInfo termInfo => do
let mvarId ← MetaTranslate.translateMVarFromTermInfo termInfo i.context? let mvarId ← MetaTranslate.translateMVarFromTermInfo termInfo i.context?
return [mvarId] return [(mvarId, stxByteRange termInfo.stx)]
| .ofTacticInfo tacticInfo => do | .ofTacticInfo tacticInfo => do
MetaTranslate.translateMVarFromTacticInfoBefore tacticInfo i.context? let mvarIds ← MetaTranslate.translateMVarFromTacticInfoBefore tacticInfo i.context?
let range := stxByteRange tacticInfo.stx
return mvarIds.map (·, range)
| _ => panic! "Invalid info" | _ => panic! "Invalid info"
let goals := List.flatten (← goalsM.run {} |>.run' {}) let annotatedGoals := List.flatten (← goalsM.run {} |>.run' {})
let goals := annotatedGoals.map Prod.fst
let srcBoundaries := annotatedGoals.map Prod.snd
let root := match goals with let root := match goals with
| [] => panic! "No MVars generated" | [] => panic! "No MVars generated"
| [g] => g | [g] => g
| _ => { name := .anonymous } | _ => { name := .anonymous }
GoalState.createFromMVars goals root let state ← GoalState.createFromMVars goals root
return { state, srcBoundaries }
@[export pantograph_frontend_collect_new_defined_constants_m]
def collectNewDefinedConstants (step : CompilationStep) : IO (List Name) := do
step.after.constants.map₂.foldlM (λ acc name _ => do
if step.before.contains name then
return acc
let coreM : CoreM Bool := Option.isSome <$> findDeclarationRanges? name
let hasRange ← coreM.run' { fileName := step.fileName, fileMap := step.fileMap } { env := step.after } |>.toBaseIO
match hasRange with
| .ok true => return name :: acc
| .ok false => return acc
| .error e => throw $ IO.userError (← e.toMessageData.toString)
) []
end Pantograph.Frontend end Pantograph.Frontend

View File

@ -0,0 +1,153 @@
/- Adapted from lean-training-data -/
import Lean.Elab.InfoTree
import Lean.Parser.Term
import Lean.PrettyPrinter
open Lean
namespace Lean.Elab
private def elaboratorToString : Name → String
| .anonymous => ""
| n => s!"⟨{n}⟩ "
private def indent (s : String) : String := "\n".intercalate $ s.splitOn "\n" |>.map ("\t" ++ .)
/-- The `Syntax` for a `Lean.Elab.Info`, if there is one. -/
protected def Info.stx? : Info → Option Syntax
| .ofTacticInfo info => info.stx
| .ofTermInfo info => info.stx
| .ofCommandInfo info => info.stx
| .ofMacroExpansionInfo info => info.stx
| .ofOptionInfo info => info.stx
| .ofFieldInfo info => info.stx
| .ofCompletionInfo info => info.stx
| .ofUserWidgetInfo info => info.stx
| .ofCustomInfo info => info.stx
| .ofFVarAliasInfo _ => none
| .ofFieldRedeclInfo info => info.stx
| .ofOmissionInfo info => info.stx
/-- Is the `Syntax` for this `Lean.Elab.Info` original, or synthetic? -/
protected def Info.isOriginal (i : Info) : Bool :=
match i.stx? with
| none => true -- Somewhat unclear what to do with `FVarAliasInfo`, so be conservative.
| some stx => match stx.getHeadInfo with
| .original .. => true
| _ => false
def ContextInfo.ppExpr (ctx : ContextInfo) (lctx : LocalContext) (e : Expr) : IO Format :=
ctx.runMetaM lctx (do Meta.ppExpr (← instantiateMVars e))
def CommandInfo.toString (info : CommandInfo) (ctx : ContextInfo) : IO String := do
let stx := (← ctx.ppSyntax {} info.stx).pretty
return s!"{elaboratorToString info.elaborator}\n{stx}"
def TermInfo.toString (info : TermInfo) (ctx : ContextInfo) : IO String := do
let stx := (← ctx.ppSyntax info.lctx info.stx).pretty
let expectedType := (← info.expectedType?.mapM fun ty => do
pure s!": {(← ctx.ppExpr info.lctx ty).pretty}").getD ""
let expr := (← ctx.ppExpr info.lctx info.expr).pretty
return s!"{elaboratorToString info.elaborator}{expr}{expectedType}\n{stx}"
/-- Find the name for the outermost `Syntax` in this `TacticInfo`. -/
def TacticInfo.name? (t : TacticInfo) : Option Name :=
match t.stx with
| Syntax.node _ n _ => some n
| _ => none
/-- Decide whether a tactic is "substantive",
or is merely a tactic combinator (e.g. `by`, `;`, multiline tactics, parenthesized tactics). -/
def TacticInfo.isSubstantive (t : TacticInfo) : Bool :=
match t.name? with
| none => false
| some `null => false
| some ``cdot => false
| some ``cdotTk => false
| some ``Lean.Parser.Term.byTactic => false
| some ``Lean.Parser.Tactic.tacticSeq => false
| some ``Lean.Parser.Tactic.tacticSeq1Indented => false
| some ``Lean.Parser.Tactic.«tactic_<;>_» => false
| some ``Lean.Parser.Tactic.paren => false
| _ => true
def TacticInfo.pp (info : TacticInfo) (ctx : ContextInfo) : IO Format :=
ctx.runMetaM {} try
Lean.PrettyPrinter.ppTactic ⟨info.stx⟩
catch _ =>
pure "<failed to pretty print>"
def TacticInfo.toString (i : TacticInfo) (ctx : ContextInfo) : IO String := do
let name := i.name?
let stx := Format.pretty (← i.pp ctx)
return s!"{name}\n{stx}"
/--
Keep `.node` nodes and `.hole` nodes satisfying predicates.
Returns a `List InfoTree`, although in most situations this will be a singleton.
-/
partial def InfoTree.filter (p : Info → Bool) (m : MVarId → Bool := fun _ => false) :
InfoTree → List InfoTree
| .context ctx tree => tree.filter p m |>.map (.context ctx)
| .node info children =>
if p info then
[.node info (children.toList.map (filter p m)).flatten.toPArray']
else
(children.toList.map (filter p m)).flatten
| .hole mvar => if m mvar then [.hole mvar] else []
/-- Analogue of `Lean.Elab.InfoTree.findInfo?`, but that returns a list of all results. -/
partial def InfoTree.findAllInfo
(t : InfoTree)
(context?: Option Elab.ContextInfo)
(haltOnMatch : Bool := false)
(pred : Elab.Info → Bool)
: List (Elab.Info × Option Elab.ContextInfo × PersistentArray Elab.InfoTree) :=
match t with
| .context inner t => findAllInfo t (inner.mergeIntoOuter? context?) haltOnMatch pred
| .node i children =>
let head := if pred i then [(i, context?, children)] else []
let tail := if haltOnMatch ∧ !head.isEmpty then [] else children.toList.flatMap (fun t => findAllInfo t context? haltOnMatch pred)
head ++ tail
| _ => []
/-- Monadic analogue of `findAllInfo`, but predicate controls whether to recurse. -/
partial def InfoTree.findAllInfoM [Monad m]
(t : InfoTree)
(context?: Option Elab.ContextInfo)
(pred : Elab.Info → Option Elab.ContextInfo → m (Bool × Bool))
: m (List (Elab.Info × Option Elab.ContextInfo × PersistentArray Elab.InfoTree)) := do
match t with
| .context inner t => t.findAllInfoM (inner.mergeIntoOuter? context?) pred
| .node i children =>
let (flagCollect, flagRecurse) ← pred i context?
let head := if flagCollect then [(i, context?, children)] else []
let tail := if ¬ flagRecurse then pure [] else children.toList.mapM (fun t => t.findAllInfoM context? pred)
return head ++ (← tail).flatten
| _ => return []
@[export pantograph_infotree_to_string_m]
partial def InfoTree.toString (t : InfoTree) (ctx?: Option Elab.ContextInfo := .none) : IO String := do
match t with
| .context ctx t => t.toString (ctx.mergeIntoOuter? ctx?)
| .node info children =>
if let some ctx := ctx? then
let node : String ← match info with
| .ofTermInfo info => pure s!"[term] {(← info.toString ctx)}"
| .ofCommandInfo info => pure s!"[command] {(← info.toString ctx)}"
| .ofTacticInfo info => pure s!"[tactic] {(← info.toString ctx)}"
| .ofMacroExpansionInfo _ => pure "[macro_exp]"
| .ofOptionInfo _ => pure "[option]"
| .ofFieldInfo _ => pure "[field]"
| .ofCompletionInfo _ => pure "[completion]"
| .ofUserWidgetInfo _ => pure "[user_widget]"
| .ofCustomInfo _ => pure "[custom]"
| .ofFVarAliasInfo _ => pure "[fvar]"
| .ofFieldRedeclInfo _ => pure "[field_redecl]"
| .ofOmissionInfo _ => pure "[omission]"
let children := "\n".intercalate (← children.toList.mapM λ t' => do pure $ indent $ ← t'.toString ctx)
return s!"{node}\n{children}"
else throw <| IO.userError "No `ContextInfo` available."
| .hole mvarId =>
if let some ctx := ctx? then
let payload := (← ctx.runMetaM {} (do Meta.ppGoal mvarId)).pretty
return s!"[hole] {payload}"
else throw <| IO.userError "No `ContextInfo` available."
end Lean.Elab

View File

@ -10,8 +10,6 @@ import Lean
namespace Pantograph namespace Pantograph
open Lean open Lean
def filename: String := "<pantograph>"
/-- /--
Represents an interconnected set of metavariables, or a state in proof search Represents an interconnected set of metavariables, or a state in proof search
-/ -/
@ -73,6 +71,8 @@ protected def GoalState.metaContextOfGoal (state: GoalState) (mvarId: MVarId): O
return { lctx := mvarDecl.lctx, localInstances := mvarDecl.localInstances } return { lctx := mvarDecl.lctx, localInstances := mvarDecl.localInstances }
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
protected def GoalState.coreState (state: GoalState): Core.SavedState :=
state.savedState.term.meta.core
protected def GoalState.withContext (state: GoalState) (mvarId: MVarId) (m: MetaM α): MetaM α := do protected def GoalState.withContext (state: GoalState) (mvarId: MVarId) (m: MetaM α): MetaM α := do
mvarId.withContext m |>.run' (← read) state.metaState mvarId.withContext m |>.run' (← read) state.metaState
@ -177,16 +177,51 @@ protected def GoalState.getMVarEAssignment (goalState: GoalState) (mvarId: MVarI
--- Tactic execution functions --- --- Tactic execution functions ---
protected def GoalState.step (state: GoalState) (goal: MVarId) (tacticM: Elab.Tactic.TacticM Unit) -- Mimics `Elab.Term.logUnassignedUsingErrorInfos`
private def collectAllErroredMVars (src : MVarId) : Elab.TermElabM (List MVarId) := do
-- 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 $ ← instantiateMVars (.mvar src)
--let _ ← Elab.Term.logUnassignedUsingErrorInfos descendants
let mut alreadyVisited : MVarIdSet := {}
let mut result : MVarIdSet := {}
for { mvarId, .. } in (← get).mvarErrorInfos do
unless alreadyVisited.contains mvarId do
alreadyVisited := alreadyVisited.insert mvarId
/- The metavariable `mvarErrorInfo.mvarId` may have been assigned or
delayed assigned to another metavariable that is unassigned. -/
let mvarDeps ← Meta.getMVars (.mvar mvarId)
if mvarDeps.any descendants.contains then do
result := mvarDeps.foldl (·.insert ·) result
return result.toList
private def mergeMVarLists (li1 li2 : List MVarId) : List MVarId :=
let li2' := li2.filter (¬ li1.contains ·)
li1 ++ li2'
/--
Set `guardMVarErrors` to true to capture mvar errors. Lean will not
automatically collect mvars from text tactics (vide
`test_tactic_failure_synthesize_placeholder`)
-/
protected def GoalState.step (state: GoalState) (goal: MVarId) (tacticM: Elab.Tactic.TacticM Unit) (guardMVarErrors : Bool := false)
: Elab.TermElabM GoalState := do : Elab.TermElabM GoalState := do
unless (← getMCtx).decls.contains goal do unless (← getMCtx).decls.contains goal do
throwError s!"Goal is not in context: {goal.name}" throwError s!"Goal is not in context: {goal.name}"
goal.checkNotAssigned `GoalState.step goal.checkNotAssigned `GoalState.step
let (_, newGoals) ← tacticM { elaborator := .anonymous } |>.run { goals := [goal] } let (_, { goals }) ← tacticM { elaborator := .anonymous } |>.run { goals := [goal] }
let nextElabState ← MonadBacktrack.saveState let nextElabState ← MonadBacktrack.saveState
Elab.Term.synthesizeSyntheticMVarsNoPostponing
let goals ← if guardMVarErrors then
pure $ mergeMVarLists goals (← collectAllErroredMVars goal)
else
pure goals
return { return {
state with state with
savedState := { term := nextElabState, tactic := newGoals }, savedState := { term := nextElabState, tactic := { goals }, },
parentMVar? := .some goal, parentMVar? := .some goal,
calcPrevRhs? := .none, calcPrevRhs? := .none,
} }
@ -202,16 +237,28 @@ 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)
/-- Executes a `TacticM` monads on this `GoalState`, collecting the errors as necessary -/ /-- Executes a `TacticM` monad on this `GoalState`, collecting the errors as necessary -/
protected def GoalState.tryTacticM (state: GoalState) (goal: MVarId) (tacticM: Elab.Tactic.TacticM Unit): protected def GoalState.tryTacticM (state: GoalState) (goal: MVarId) (tacticM: Elab.Tactic.TacticM Unit) (guardMVarErrors : Bool := false):
Elab.TermElabM TacticResult := do Elab.TermElabM TacticResult := do
try try
let nextState ← state.step goal tacticM let nextState ← state.step goal tacticM guardMVarErrors
-- Check if error messages have been generated in the core.
let newMessages ← (← Core.getMessageLog).toList.drop state.coreState.messages.toList.length
|>.filterMapM λ m => do
if m.severity == .error then
return .some $ ← m.toString
else
return .none
Core.resetMessageLog
if ¬ newMessages.isEmpty then
return .failure newMessages.toArray
return .success nextState return .success nextState
catch exception => catch exception =>
return .failure #[← exception.toMessageData.toString] return .failure #[← exception.toMessageData.toString]
/-- 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]
protected def GoalState.tryTactic (state: GoalState) (goal: MVarId) (tactic: String): protected def GoalState.tryTactic (state: GoalState) (goal: MVarId) (tactic: String):
Elab.TermElabM TacticResult := do Elab.TermElabM TacticResult := do
state.restoreElabM state.restoreElabM
@ -219,10 +266,10 @@ protected def GoalState.tryTactic (state: GoalState) (goal: MVarId) (tactic: Str
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
(catName := if state.isConv then `conv else `tactic) (catName := if state.isConv then `conv else `tactic)
(input := tactic) (input := tactic)
(fileName := filename) with (fileName := ← getFileName) with
| .ok stx => pure $ stx | .ok stx => pure $ stx
| .error error => return .parseError error | .error error => return .parseError error
state.tryTacticM goal $ Elab.Tactic.evalTactic tactic state.tryTacticM goal (Elab.Tactic.evalTactic tactic) true
protected def GoalState.tryAssign (state: GoalState) (goal: MVarId) (expr: String): protected def GoalState.tryAssign (state: GoalState) (goal: MVarId) (expr: String):
Elab.TermElabM TacticResult := do Elab.TermElabM TacticResult := do
@ -231,7 +278,7 @@ protected def GoalState.tryAssign (state: GoalState) (goal: MVarId) (expr: Strin
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
(catName := `term) (catName := `term)
(input := expr) (input := expr)
(fileName := filename) with (fileName := ← getFileName) with
| .ok syn => pure syn | .ok syn => pure syn
| .error error => return .parseError error | .error error => return .parseError error
state.tryTacticM goal $ Tactic.evalAssign expr state.tryTacticM goal $ Tactic.evalAssign expr
@ -245,7 +292,7 @@ protected def GoalState.tryLet (state: GoalState) (goal: MVarId) (binderName: St
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
(catName := `term) (catName := `term)
(input := type) (input := type)
(fileName := filename) with (fileName := ← getFileName) with
| .ok syn => pure syn | .ok syn => pure syn
| .error error => return .parseError error | .error error => return .parseError error
state.tryTacticM goal $ Tactic.evalLet binderName.toName type state.tryTacticM goal $ Tactic.evalLet binderName.toName type
@ -332,7 +379,7 @@ protected def GoalState.tryCalc (state: GoalState) (goal: MVarId) (pred: String)
(env := state.env) (env := state.env)
(catName := `term) (catName := `term)
(input := pred) (input := pred)
(fileName := filename) with (fileName := ← getFileName) with
| .ok syn => pure syn | .ok syn => pure syn
| .error error => return .parseError error | .error error => return .parseError error
goal.checkNotAssigned `GoalState.tryCalc goal.checkNotAssigned `GoalState.tryCalc
@ -353,7 +400,7 @@ protected def GoalState.tryCalc (state: GoalState) (goal: MVarId) (pred: String)
throwErrorAt pred "invalid 'calc' step, relation expected{indentExpr step}" throwErrorAt pred "invalid 'calc' step, relation expected{indentExpr step}"
if let some prevRhs := calcPrevRhs? then if let some prevRhs := calcPrevRhs? then
unless ← Meta.isDefEqGuarded lhs prevRhs do unless ← Meta.isDefEqGuarded lhs prevRhs do
throwErrorAt pred "invalid 'calc' step, left-hand-side is{indentD m!"{lhs} : {← Meta.inferType lhs}"}\nprevious right-hand-side is{indentD m!"{prevRhs} : {← Meta.inferType prevRhs}"}" -- " throwErrorAt pred "invalid 'calc' step, left-hand-side is{indentD m!"{lhs} : {← Meta.inferType lhs}"}\nprevious right-hand-side is{indentD m!"{prevRhs} : {← Meta.inferType prevRhs}"}"
-- Creates a mvar to represent the proof that the calc tactic solves the -- Creates a mvar to represent the proof that the calc tactic solves the
-- current branch -- current branch

View File

@ -138,17 +138,36 @@ def goalSerialize (state: GoalState) (options: @&Protocol.Options): CoreM (Array
runMetaM <| state.serializeGoals (parent := .none) options runMetaM <| state.serializeGoals (parent := .none) options
@[export pantograph_goal_print_m] @[export pantograph_goal_print_m]
def goalPrint (state: GoalState) (options: @&Protocol.Options): CoreM Protocol.GoalPrintResult := def goalPrint (state: GoalState) (rootExpr: Bool) (parentExpr: Bool) (goals: Bool) (extraMVars : Array String) (options: @&Protocol.Options)
runMetaM do : CoreM Protocol.GoalPrintResult := runMetaM do
state.restoreMetaM state.restoreMetaM
return {
root? := ← state.rootExpr?.mapM (λ expr => let root? ← if rootExpr then
state.withRootContext do state.rootExpr?.mapM λ expr => state.withRootContext do
serializeExpression options (← instantiateAll expr)), serializeExpression options (← instantiateAll expr)
parent? := ← state.parentExpr?.mapM (λ expr => else
state.withParentContext do pure .none
serializeExpression options (← instantiateAll expr)), let parent? ← if parentExpr then
} state.parentExpr?.mapM λ expr => state.withParentContext do
serializeExpression options (← instantiateAll expr)
else
pure .none
let goals ← if goals then
goalSerialize state options
else
pure #[]
let extraMVars ← extraMVars.mapM λ mvarId => do
let mvarId: MVarId := { name := mvarId.toName }
let .some _ ← mvarId.findDecl? | return {}
state.withContext mvarId do
let .some expr ← getExprMVarAssignment? mvarId | return {}
serializeExpression options (← instantiateAll expr)
return {
root?,
parent?,
goals,
extraMVars,
}
@[export pantograph_goal_tactic_m] @[export pantograph_goal_tactic_m]
def goalTactic (state: GoalState) (goal: MVarId) (tactic: String): CoreM TacticResult := def goalTactic (state: GoalState) (goal: MVarId) (tactic: String): CoreM TacticResult :=

View File

@ -271,12 +271,23 @@ structure GoalDeleteResult where
structure GoalPrint where structure GoalPrint where
stateId: Nat stateId: Nat
-- Print root?
rootExpr?: Option Bool := .some False
-- Print the parent expr?
parentExpr?: Option Bool := .some False
-- Print goals?
goals?: Option Bool := .some False
-- Print values of extra mvars?
extraMVars?: Option (Array String) := .none
deriving Lean.FromJson deriving Lean.FromJson
structure GoalPrintResult where 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
parent?: Option Expression parent?: Option Expression := .none
goals: Array Goal := #[]
extraMVars: Array Expression := #[]
deriving Lean.ToJson deriving Lean.ToJson
-- Diagnostic Options, not available in REPL -- Diagnostic Options, not available in REPL
@ -312,6 +323,8 @@ structure FrontendProcess where
invocations: Bool := false invocations: Bool := false
-- If set to true, collect `sorry`s -- If set to true, collect `sorry`s
sorrys: Bool := false sorrys: Bool := false
-- If set to true, extract new constants
newConstants: Bool := false
deriving Lean.FromJson deriving Lean.FromJson
structure InvokedTactic where structure InvokedTactic where
goalBefore: String goalBefore: String
@ -325,11 +338,16 @@ structure InvokedTactic where
structure CompilationUnit where structure CompilationUnit where
-- String boundaries of compilation units -- String boundaries of compilation units
boundary: (Nat × Nat) boundary: (Nat × Nat)
messages: Array String := #[]
-- Tactic invocations -- Tactic invocations
invocations?: Option (List InvokedTactic) := .none invocations?: Option (List InvokedTactic) := .none
goalStateId?: Option Nat := .none goalStateId?: Option Nat := .none
goals: Array Goal := #[] goals?: Option (Array Goal) := .none
messages: Array String := #[] -- Code segments which generated the goals
goalSrcBoundaries?: Option (Array (Nat × Nat)) := .none
-- New constants defined in compilation unit
newConstants?: Option (Array String) := .none
deriving Lean.ToJson deriving Lean.ToJson
structure FrontendProcessResult where structure FrontendProcessResult where
units: List CompilationUnit units: List CompilationUnit

View File

@ -1,6 +1,6 @@
namespace Pantograph namespace Pantograph
@[export pantograph_version] @[export pantograph_version]
def version := "0.2.19" def version := "0.2.23"
end Pantograph end Pantograph

View File

@ -7,7 +7,7 @@ A Machine-to-Machine interaction system for Lean 4.
Pantograph provides interfaces to execute proofs, construct expressions, and Pantograph provides interfaces to execute proofs, construct expressions, and
examine the symbol list of a Lean project for machine learning. examine the symbol list of a Lean project for machine learning.
See [documentations](doc/) for design rationale and references. See [documentations](doc/rationale.md) for design rationale and references.
## Installation ## Installation

View File

@ -79,6 +79,7 @@ def execute (command: Protocol.Command): MainM Lean.Json := do
let state ← get let state ← get
let nGoals := state.goalStates.size let nGoals := state.goalStates.size
set { state with nextId := 0, goalStates := .empty } set { state with nextId := 0, goalStates := .empty }
Lean.Core.resetMessageLog
return .ok { nGoals } return .ok { nGoals }
stat (_: Protocol.Stat): MainM (CR Protocol.StatResult) := do stat (_: Protocol.Stat): MainM (CR Protocol.StatResult) := do
let state ← get let state ← get
@ -222,7 +223,13 @@ def execute (command: Protocol.Command): MainM Lean.Json := do
let state ← get let state ← get
let .some goalState := state.goalStates[args.stateId]? | let .some goalState := state.goalStates[args.stateId]? |
return .error $ errorIndex s!"Invalid state index {args.stateId}" return .error $ errorIndex s!"Invalid state index {args.stateId}"
let result ← runMetaInMainM <| goalPrint goalState state.options let result ← runMetaInMainM <| goalPrint
goalState
(rootExpr := args.rootExpr?.getD False)
(parentExpr := args.parentExpr?.getD False)
(goals := args.goals?.getD False)
(extraMVars := args.extraMVars?.getD #[])
(options := state.options)
return .ok result return .ok result
goal_save (args: Protocol.GoalSave): MainM (CR Protocol.GoalSaveResult) := do goal_save (args: Protocol.GoalSave): MainM (CR Protocol.GoalSaveResult) := do
let state ← get let state ← get
@ -257,27 +264,38 @@ def execute (command: Protocol.Command): MainM Lean.Json := do
pure $ .some invocations pure $ .some invocations
else else
pure .none pure .none
let sorrys := if args.sorrys then let sorrys if args.sorrys then
Frontend.collectSorrys step Frontend.collectSorrys step
else else
[] pure []
let messages ← step.messageStrings let messages ← step.messageStrings
return (step.before, boundary, invocations?, sorrys, messages) let newConstants ← if args.newConstants then
Frontend.collectNewDefinedConstants step
else
pure []
return (step.before, boundary, invocations?, sorrys, messages, newConstants)
let li ← frontendM.run context |>.run' state let li ← frontendM.run context |>.run' state
let units ← li.mapM λ (env, boundary, invocations?, sorrys, messages) => Lean.withEnv env do let units ← li.mapM λ (env, boundary, invocations?, sorrys, messages, newConstants) => Lean.withEnv env do
let (goalStateId?, goals) ← if sorrys.isEmpty then do let newConstants? := if args.newConstants then
pure (.none, #[]) .some $ newConstants.toArray.map λ name => name.toString
else
.none
let (goalStateId?, goals?, goalSrcBoundaries?) ← if sorrys.isEmpty then do
pure (.none, .none, .none)
else do else do
let goalState ← runMetaInMainM $ Frontend.sorrysToGoalState sorrys let { state, srcBoundaries } ← runMetaInMainM $ Frontend.sorrysToGoalState sorrys
let stateId ← newGoalState goalState let stateId ← newGoalState state
let goals ← goalSerialize goalState options let goals ← goalSerialize state options
pure (.some stateId, goals) let srcBoundaries := srcBoundaries.toArray.map (λ (b, e) => (b.byteIdx, e.byteIdx))
pure (.some stateId, .some goals, .some srcBoundaries)
return { return {
boundary, boundary,
messages,
invocations?, invocations?,
goalStateId?, goalStateId?,
goals, goals?,
messages, goalSrcBoundaries?,
newConstants?,
} }
return .ok { units } return .ok { units }
catch e => catch e =>

View File

@ -95,19 +95,19 @@ def runTermElabMSeq (env: Environment) (termElabM: Elab.TermElabM LSpec.TestSeq)
def exprToStr (e: Expr): Lean.MetaM String := toString <$> Meta.ppExpr e def exprToStr (e: Expr): Lean.MetaM String := toString <$> Meta.ppExpr e
def strToTermSyntax [Monad m] [MonadEnv m] (s: String): m Syntax := do def strToTermSyntax (s: String): CoreM Syntax := do
let .ok stx := Parser.runParserCategory let .ok stx := Parser.runParserCategory
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
(catName := `term) (catName := `term)
(input := s) (input := s)
(fileName := filename) | panic! s!"Failed to parse {s}" (fileName := ← getFileName) | panic! s!"Failed to parse {s}"
return stx return stx
def parseSentence (s: String): Elab.TermElabM Expr := do def parseSentence (s: String): Elab.TermElabM Expr := do
let stx ← match Parser.runParserCategory let stx ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
(catName := `term) (catName := `term)
(input := s) (input := s)
(fileName := filename) with (fileName := ← getFileName) with
| .ok syn => pure syn | .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}" | .error error => throwError "Failed to parse: {error}"
Elab.Term.elabTerm (stx := stx) .none Elab.Term.elabTerm (stx := stx) .none
@ -123,23 +123,29 @@ def mvarUserNameAndType (mvarId: MVarId): MetaM (Name × String) := do
-- Monadic testing -- Monadic testing
abbrev TestT := StateT LSpec.TestSeq abbrev TestT := StateRefT' IO.RealWorld LSpec.TestSeq
def addTest [Monad m] (test: LSpec.TestSeq) : TestT m Unit := do section Monadic
variable [Monad m] [MonadLiftT (ST IO.RealWorld) m]
def addTest (test: LSpec.TestSeq) : TestT m Unit := do
set $ (← get) ++ test set $ (← get) ++ test
def checkEq [Monad m] [DecidableEq α] (desc : String) (lhs rhs : α) : TestT m Unit := do def checkEq [DecidableEq α] [Repr α] (desc : String) (lhs rhs : α) : TestT m Unit := do
addTest $ LSpec.check desc (lhs == rhs) addTest $ LSpec.check desc (lhs = rhs)
def checkTrue [Monad m] (desc : String) (flag : Bool) : TestT m Unit := do def checkTrue (desc : String) (flag : Bool) : TestT m Unit := do
addTest $ LSpec.check desc flag addTest $ LSpec.check desc flag
def fail [Monad m] (desc : String) : TestT m Unit := do def fail (desc : String) : TestT m Unit := do
addTest $ LSpec.check desc false addTest $ LSpec.check desc false
def runTest [Monad m] (t: TestT m Unit): m LSpec.TestSeq := def runTest (t: TestT m Unit): m LSpec.TestSeq :=
Prod.snd <$> t.run LSpec.TestSeq.done Prod.snd <$> t.run LSpec.TestSeq.done
def runTestWithResult { α } [Monad m] (t: TestT m α): m (α × LSpec.TestSeq) := def runTestWithResult { α } (t: TestT m α): m (α × LSpec.TestSeq) :=
t.run LSpec.TestSeq.done t.run LSpec.TestSeq.done
end Monadic
def runTestTermElabM (env: Environment) (t: TestT Elab.TermElabM Unit): def runTestTermElabM (env: Environment) (t: TestT Elab.TermElabM Unit):
IO LSpec.TestSeq := IO LSpec.TestSeq :=
runTermElabMSeq env $ runTest t runTermElabMSeq env $ runTest t

View File

@ -10,13 +10,13 @@ def collectSorrysFromSource (source: String) : MetaM (List GoalState) := do
let filename := "<anonymous>" let filename := "<anonymous>"
let (context, state) ← do Frontend.createContextStateFromFile source filename (← getEnv) {} let (context, state) ← do Frontend.createContextStateFromFile source filename (← getEnv) {}
let m := Frontend.mapCompilationSteps λ step => do let m := Frontend.mapCompilationSteps λ step => do
return (step.before, Frontend.collectSorrys step) return (step.before, Frontend.collectSorrys step)
let li ← m.run context |>.run' state let li ← m.run context |>.run' state
let goalStates ← li.filterMapM λ (env, sorrys) => withEnv env do let goalStates ← li.filterMapM λ (env, sorrys) => withEnv env do
if sorrys.isEmpty then if sorrys.isEmpty then
return .none return .none
let goalState ← Frontend.sorrysToGoalState sorrys let { state, .. } ← Frontend.sorrysToGoalState sorrys
return .some goalState return .some state
return goalStates return goalStates
def test_multiple_sorrys_in_proof : TestT MetaM Unit := do def test_multiple_sorrys_in_proof : TestT MetaM Unit := do
@ -177,6 +177,47 @@ example (n: Nat) : mystery n + 1 = n + 2 := sorry
} }
]) ])
def test_capture_type_mismatch : TestT MetaM Unit := do
let input := "
def mystery (k: Nat) : Nat := true
"
let goalStates ← (collectSorrysFromSource input).run' {}
let [goalState] := goalStates | panic! s!"Incorrect number of states: {goalStates.length}"
checkEq "goals" ((← goalState.serializeGoals (options := {})).map (·.devolatilize)) #[
{
target := { pp? := "Nat" },
vars := #[{
userName := "k",
type? := .some { pp? := "Nat" },
}],
}
]
def collectNewConstants (source: String) : MetaM (List (List Name)) := do
let filename := "<anonymous>"
let (context, state) ← do Frontend.createContextStateFromFile source filename (← getEnv) {}
let m := Frontend.mapCompilationSteps λ step => do
Frontend.collectNewDefinedConstants step
m.run context |>.run' state
def test_collect_one_constant : TestT MetaM Unit := do
let input := "
def mystery : Nat := 123
"
let names ← collectNewConstants input
checkEq "constants" names [[`mystery]]
def test_collect_one_theorem : TestT MetaM Unit := do
let input := "
theorem mystery [SizeOf α] (as : List α) (i : Fin as.length) : sizeOf (as.get i) < sizeOf as := by
match as, i with
| a::as, ⟨0, _⟩ => simp_arith [get]
| a::as, ⟨i+1, h⟩ =>
have ih := sizeOf_get as ⟨i, Nat.le_of_succ_le_succ h⟩
apply Nat.lt_trans ih
simp_arith
"
let names ← collectNewConstants input
checkEq "constants" names [[`mystery]]
def suite (env : Environment): List (String × IO LSpec.TestSeq) := def suite (env : Environment): List (String × IO LSpec.TestSeq) :=
let tests := [ let tests := [
@ -185,6 +226,9 @@ def suite (env : Environment): List (String × IO LSpec.TestSeq) :=
("sorry_in_induction", test_sorry_in_induction), ("sorry_in_induction", test_sorry_in_induction),
("sorry_in_coupled", test_sorry_in_coupled), ("sorry_in_coupled", test_sorry_in_coupled),
("environment_capture", test_environment_capture), ("environment_capture", test_environment_capture),
("capture_type_mismatch", test_capture_type_mismatch),
("collect_one_constant", test_collect_one_constant),
("collect_one_theorem", test_collect_one_theorem),
] ]
tests.map (fun (name, test) => (name, runMetaMSeq env $ runTest test)) tests.map (fun (name, test) => (name, runMetaMSeq env $ runTest test))

View File

@ -72,8 +72,8 @@ def test_tactic : Test :=
({ stateId := 0, root := "_uniq.9" }: Protocol.GoalStartResult), ({ stateId := 0, root := "_uniq.9" }: Protocol.GoalStartResult),
step "goal.tactic" [("stateId", .num 0), ("goalId", .num 0), ("tactic", .str "intro x")] step "goal.tactic" [("stateId", .num 0), ("goalId", .num 0), ("tactic", .str "intro x")]
({ nextStateId? := .some 1, goals? := #[goal1], }: Protocol.GoalTacticResult), ({ nextStateId? := .some 1, goals? := #[goal1], }: Protocol.GoalTacticResult),
step "goal.print" [("stateId", .num 1)] step "goal.print" [("stateId", .num 1), ("parentExpr", .bool true), ("rootExpr", .bool true)]
({ parent? := .some { pp? := .some "fun x => ?m.11" }, }: Protocol.GoalPrintResult), ({ parent? := .some { pp? := .some "fun x => ?m.12 x" }, }: Protocol.GoalPrintResult),
step "goal.tactic" [("stateId", .num 1), ("goalId", .num 0), ("tactic", .str "intro y")] step "goal.tactic" [("stateId", .num 1), ("goalId", .num 0), ("tactic", .str "intro y")]
({ nextStateId? := .some 2, goals? := #[goal2], }: Protocol.GoalTacticResult), ({ nextStateId? := .some 2, goals? := #[goal2], }: Protocol.GoalTacticResult),
] ]
@ -174,6 +174,7 @@ def test_frontend_process : Test :=
("file", .str file), ("file", .str file),
("invocations", .bool true), ("invocations", .bool true),
("sorrys", .bool false), ("sorrys", .bool false),
("newConstants", .bool false),
] ]
({ ({
units := [{ units := [{
@ -214,6 +215,7 @@ def test_frontend_process_sorry : Test :=
("file", .str file), ("file", .str file),
("invocations", .bool false), ("invocations", .bool false),
("sorrys", .bool true), ("sorrys", .bool true),
("newConstants", .bool false),
] ]
({ ({
units := [{ units := [{
@ -221,7 +223,8 @@ def test_frontend_process_sorry : Test :=
}, { }, {
boundary := (solved.utf8ByteSize, solved.utf8ByteSize + withSorry.utf8ByteSize), boundary := (solved.utf8ByteSize, solved.utf8ByteSize + withSorry.utf8ByteSize),
goalStateId? := .some 0, goalStateId? := .some 0,
goals := #[goal1], goals? := .some #[goal1],
goalSrcBoundaries? := .some #[(57, 62)],
messages := #["<anonymous>:2:0: warning: declaration uses 'sorry'\n"], messages := #["<anonymous>:2:0: warning: declaration uses 'sorry'\n"],
}], }],
}: Protocol.FrontendProcessResult), }: Protocol.FrontendProcessResult),

View File

@ -8,10 +8,7 @@ namespace Pantograph.Test.Metavar
open Pantograph open Pantograph
open Lean open Lean
abbrev TestM := StateRefT LSpec.TestSeq (ReaderT Protocol.Options Elab.TermElabM) abbrev TestM := TestT $ ReaderT Protocol.Options Elab.TermElabM
def addTest (test: LSpec.TestSeq): TestM Unit := do
set $ (← get) ++ test
-- Tests that all delay assigned mvars are instantiated -- Tests that all delay assigned mvars are instantiated
def test_instantiate_mvar: TestM Unit := do def test_instantiate_mvar: TestM Unit := do
@ -32,8 +29,6 @@ def test_instantiate_mvar: TestM Unit := do
"((:c LE.le) (:c Nat) (:c instLENat) ((:c OfNat.ofNat) (:mv _uniq.2) (:lit 2) (:mv _uniq.3)) ((:c OfNat.ofNat) (:mv _uniq.14) (:lit 5) (:mv _uniq.15)))") "((:c LE.le) (:c Nat) (:c instLENat) ((:c OfNat.ofNat) (:mv _uniq.2) (:lit 2) (:mv _uniq.3)) ((:c OfNat.ofNat) (:mv _uniq.14) (:lit 5) (:mv _uniq.15)))")
return () return ()
def startProof (expr: String): TestM (Option GoalState) := do def startProof (expr: String): TestM (Option GoalState) := do
let env ← Lean.MonadEnv.getEnv let env ← Lean.MonadEnv.getEnv
let syn? := parseTerm env expr let syn? := parseTerm env expr

View File

@ -14,10 +14,7 @@ inductive Start where
| copy (name: String) -- Start from some name in the environment | copy (name: String) -- Start from some name in the environment
| expr (expr: String) -- Start from some expression | expr (expr: String) -- Start from some expression
abbrev TestM := StateRefT LSpec.TestSeq (ReaderT Protocol.Options Elab.TermElabM) abbrev TestM := TestT $ ReaderT Protocol.Options $ Elab.TermElabM
def addTest (test: LSpec.TestSeq): TestM Unit := do
set $ (← get) ++ test
def startProof (start: Start): TestM (Option GoalState) := do def startProof (start: Start): TestM (Option GoalState) := do
let env ← Lean.MonadEnv.getEnv let env ← Lean.MonadEnv.getEnv
@ -704,6 +701,56 @@ def test_nat_zero_add_alt: 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 state0 ← match state? with
| .some state => pure state
| .none => do
addTest $ assertUnreachable "Goal could not parse"
return ()
let tactic := "intro p"
let state1 ← match ← state0.tacticOn 0 tactic with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
let tactic := "exact ⟨0, by simp⟩"
let .failure messages ← state1.tacticOn 0 tactic | addTest $ assertUnreachable s!"{tactic} should fail"
checkEq s!"{tactic} fails" messages #[s!"{← getFileName}:0:12: error: unsolved goals\np : Nat → Prop\n⊢ p 0\n"]
def test_tactic_failure_synthesize_placeholder : TestM Unit := do
let state? ← startProof (.expr "∀ (p q r : Prop) (h : p → q), q ∧ r")
let state0 ← match state? with
| .some state => pure state
| .none => do
addTest $ assertUnreachable "Goal could not parse"
return ()
let tactic := "intro p q r h"
let state1 ← match ← state0.tacticOn 0 tactic with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
let tactic := "simpa [h] using And.imp_left h _"
let state2 ← match ← state1.tacticOn 0 tactic with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
checkEq tactic ((← state2.serializeGoals).map (·.devolatilize)) #[
buildGoal [("p", "Prop"), ("q", "Prop"), ("r", "Prop"), ("h", "p → q")] "p ∧ r"
]
--let .failure messages ← state1.tacticOn 0 tactic | addTest $ assertUnreachable s!"{tactic} should fail"
--let message := s!"<Pantograph>:0:31: error: don't know how to synthesize placeholder\ncontext:\np q r : Prop\nh : p → q\n⊢ p ∧ r\n"
--checkEq s!"{tactic} fails" messages #[message]
def suite (env: Environment): List (String × IO LSpec.TestSeq) := def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
let tests := [ let tests := [
("identity", test_identity), ("identity", test_identity),
@ -716,6 +763,8 @@ def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
("calc", test_calc), ("calc", test_calc),
("Nat.zero_add", test_nat_zero_add), ("Nat.zero_add", test_nat_zero_add),
("Nat.zero_add alt", test_nat_zero_add_alt), ("Nat.zero_add alt", test_nat_zero_add_alt),
("tactic failure with unresolved goals", test_tactic_failure_unresolved_goals),
("tactic failure with synthesize placeholder", test_tactic_failure_synthesize_placeholder),
] ]
tests.map (fun (name, test) => (name, proofRunner env test)) tests.map (fun (name, test) => (name, proofRunner env test))

View File

@ -28,7 +28,7 @@ def test_nat_brec_on : TestT Elab.TermElabM Unit := do
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
(catName := `term) (catName := `term)
(input := "@Nat.brecOn") (input := "@Nat.brecOn")
(fileName := filename) with (fileName := ← getFileName) with
| .ok syn => pure syn | .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}" | .error error => throwError "Failed to parse: {error}"
-- Apply the tactic -- Apply the tactic
@ -52,7 +52,7 @@ def test_list_brec_on : TestT Elab.TermElabM Unit := do
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
(catName := `term) (catName := `term)
(input := "@List.brecOn") (input := "@List.brecOn")
(fileName := filename) with (fileName := ← getFileName) with
| .ok syn => pure syn | .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}" | .error error => throwError "Failed to parse: {error}"
-- Apply the tactic -- Apply the tactic
@ -74,7 +74,7 @@ def test_partial_motive_instantiation : TestT Elab.TermElabM Unit := do
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
(catName := `term) (catName := `term)
(input := "@Nat.brecOn") (input := "@Nat.brecOn")
(fileName := filename) with (fileName := ← getFileName) with
| .ok syn => pure syn | .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}" | .error error => throwError "Failed to parse: {error}"
let expr ← parseSentence expr let expr ← parseSentence expr

View File

@ -15,7 +15,7 @@ def test_nat : TestT Elab.TermElabM Unit := do
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
(catName := `term) (catName := `term)
(input := "h") (input := "h")
(fileName := filename) with (fileName := ← getFileName) with
| .ok syn => pure syn | .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}" | .error error => throwError "Failed to parse: {error}"
-- Apply the tactic -- Apply the tactic
@ -32,7 +32,7 @@ def test_nat_fail : TestT Elab.TermElabM Unit := do
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
(catName := `term) (catName := `term)
(input := "h") (input := "h")
(fileName := filename) with (fileName := ← getFileName) with
| .ok syn => pure syn | .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}" | .error error => throwError "Failed to parse: {error}"
-- Apply the tactic -- Apply the tactic
@ -52,7 +52,7 @@ def test_list : TestT Elab.TermElabM Unit := do
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
(catName := `term) (catName := `term)
(input := "h") (input := "h")
(fileName := filename) with (fileName := ← getFileName) with
| .ok syn => pure syn | .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}" | .error error => throwError "Failed to parse: {error}"
-- Apply the tactic -- Apply the tactic

View File

@ -15,7 +15,7 @@ def test_define : TestT Elab.TermElabM Unit := do
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
(catName := `term) (catName := `term)
(input := "Or.inl h") (input := "Or.inl h")
(fileName := filename) with (fileName := ← getFileName) with
| .ok syn => pure syn | .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}" | .error error => throwError "Failed to parse: {error}"
-- Apply the tactic -- Apply the tactic

View File

@ -24,6 +24,35 @@ The name Pantograph is a pun. It means two things
a locomotive. In comparison the (relatively) simple Pantograph software powers a locomotive. In comparison the (relatively) simple Pantograph software powers
theorem proving projects. theorem proving projects.
## Caveats and Limitations
Pantograph does not exactly mimic Lean LSP's behaviour. That would not grant the
flexibility it offers. To support tree search means Pantograph has to act
differently from Lean in some times, but never at the sacrifice of soundness.
- When Lean LSP says "don't know how to synthesize placeholder", this indicates
the human operator needs to manually move the cursor to the placeholder and
type in the correct expression. This error therefore should not halt the proof
process, and the placeholder should be turned into a goal.
- When Lean LSP says "unresolved goals", that means a proof cannot finish where
it is supposed to finish at the end of a `by` block. Pantograph will raise the
error in this case, since it indicates the termination of a proof search branch.
- `pick_goal` or `swap` will not work since they run contrary to tree search
paradigms. However, if there are tactics which perform non-trivial operations
to multiple goals at the same time, this constrain could potentially be
relaxed at a cost of great bookkeeping overhead to the user.
Pantograph cannot perform things that are inherently constrained by Lean. These
include:
- If a tactic loses track of metavariables, it will not be caught until the end
of the proof search. This is a bug in the tactic itself.
- Timeouts for executing tactics is not available. Maybe this will change in the
future.
- Interceptions of parsing errors generally cannot be turned into goals (e.g.
`def mystery : Nat := :=`) due to Lean's parsing system.
## References ## References
* [Pantograph Paper](https://arxiv.org/abs/2410.16429) * [Pantograph Paper](https://arxiv.org/abs/2410.16429)

View File

@ -44,9 +44,11 @@ See `Pantograph/Protocol.lean` for a description of the parameters and return va
state. The user is responsible to ensure the sender/receiver instances share state. The user is responsible to ensure the sender/receiver instances share
the same environment. the same environment.
* `frontend.process { ["fileName": <fileName>,] ["file": <str>], invocations: * `frontend.process { ["fileName": <fileName>,] ["file": <str>], invocations:
<bool>, sorrys: <bool> }`: Executes the Lean frontend on a file, collecting <bool>, sorrys: <bool>, newConstants: <bool> }`: Executes the Lean frontend on
either the tactic invocations (`"invocations": true`) or the sorrys into goal a file, collecting the tactic invocations (`"invocations": true`), the
states (`"sorrys": true`) sorrys and type errors into goal states (`"sorrys": true`), and new constants
(`"newConstants": true`). In the case of `sorrys`, this command additionally
outputs the position of each captured `sorry`.
## Errors ## Errors

View File

@ -18,9 +18,19 @@
lean4-nix, lean4-nix,
lspec, lspec,
... ...
}: } : flake-parts.lib.mkFlake { inherit inputs; } {
flake-parts.lib.mkFlake {inherit inputs;} { flake = {
flake = { };
systems = [
"aarch64-linux"
"aarch64-darwin"
"x86_64-linux"
"x86_64-darwin"
];
perSystem = { system, pkgs, ... }: let
pkgs = import nixpkgs {
inherit system;
overlays = [ (lean4-nix.readToolchainFile ./lean-toolchain) ];
}; };
systems = [ systems = [
"aarch64-linux" "aarch64-linux"