doc: Add limitations #145

Merged
aniva merged 2 commits from doc/rationale into dev 2024-12-11 16:51:24 -08:00
27 changed files with 791 additions and 251 deletions
Showing only changes of commit fecab387dc - Show all commits

6
.gitignore vendored
View File

@ -1,6 +1,4 @@
.*
!.gitignore
*.olean
/build
/lake-packages
*.[io]lean
/result

View File

@ -327,11 +327,11 @@ partial def serializeExpressionSexp (expr: Expr) (sanitize: Bool := true): MetaM
-- Lean these are handled using a `#` prefix.
pure s!"{deBruijnIndex}"
| .fvar fvarId =>
let name := ofName fvarId.name
let name := fvarId.name
pure s!"(:fv {name})"
| .mvar mvarId => do
let pref := if ← mvarId.isDelayedAssigned then "mvd" else "mv"
let name := ofName mvarId.name
let name := mvarId.name
pure s!"(:{pref} {name})"
| .sort level =>
let level := serializeSortLevel level sanitize
@ -346,20 +346,20 @@ partial def serializeExpressionSexp (expr: Expr) (sanitize: Bool := true): MetaM
let args := " ".intercalate args
pure s!"({fn'} {args})"
| .lam binderName binderType body binderInfo => do
let binderName' := ofName binderName
let binderName' := binderName.eraseMacroScopes
let binderType' ← self binderType
let body' ← self body
let binderInfo' := binderInfoSexp binderInfo
pure s!"(:lambda {binderName'} {binderType'} {body'}{binderInfo'})"
| .forallE binderName binderType body binderInfo => do
let binderName' := ofName binderName
let binderName' := binderName.eraseMacroScopes
let binderType' ← self binderType
let body' ← self body
let binderInfo' := binderInfoSexp binderInfo
pure s!"(:forall {binderName'} {binderType'} {body'}{binderInfo'})"
| .letE name type value body _ => do
-- Dependent boolean flag diacarded
let name' := serializeName name
let name' := name.eraseMacroScopes
let type' ← self type
let value' ← self value
let body' ← self body
@ -387,7 +387,6 @@ partial def serializeExpressionSexp (expr: Expr) (sanitize: Bool := true): MetaM
| .implicit => " :implicit"
| .strictImplicit => " :strictImplicit"
| .instImplicit => " :instImplicit"
ofName (name: Name) := serializeName name sanitize
def serializeExpression (options: @&Protocol.Options) (e: Expr): MetaM Protocol.Expression := do
let pp?: Option String ← match options.printExprPretty with
@ -420,13 +419,13 @@ def serializeGoal (options: @&Protocol.Options) (goal: MVarId) (mvarDecl: Metava
match localDecl with
| .cdecl _ fvarId userName _ _ _ =>
return {
name := ofName fvarId.name,
name := fvarId.name.toString,
userName:= ofName userName.simpMacroScopes,
isInaccessible := userName.isInaccessibleUserName
}
| .ldecl _ fvarId userName _ _ _ _ => do
return {
name := ofName fvarId.name,
name := fvarId.name.toString,
userName := toString userName.simpMacroScopes,
isInaccessible := userName.isInaccessibleUserName
}
@ -436,7 +435,7 @@ def serializeGoal (options: @&Protocol.Options) (goal: MVarId) (mvarDecl: Metava
let userName := userName.simpMacroScopes
let type ← instantiate type
return {
name := ofName fvarId.name,
name := fvarId.name.toString,
userName:= ofName userName,
isInaccessible := userName.isInaccessibleUserName
type? := .some (← serializeExpression options type)
@ -450,7 +449,7 @@ def serializeGoal (options: @&Protocol.Options) (goal: MVarId) (mvarDecl: Metava
else
pure $ .none
return {
name := ofName fvarId.name,
name := fvarId.name.toString,
userName:= ofName userName,
isInaccessible := userName.isInaccessibleUserName
type? := .some (← serializeExpression options type)
@ -469,7 +468,7 @@ def serializeGoal (options: @&Protocol.Options) (goal: MVarId) (mvarDecl: Metava
| false => ppVar localDecl
return var::acc
return {
name := ofName goal.name,
name := goal.name.toString,
userName? := if mvarDecl.userName == .anonymous then .none else .some (ofName mvarDecl.userName),
isConversion := isLHSGoal? mvarDecl.type |>.isSome,
target := (← serializeExpression options (← instantiate mvarDecl.type)),

View File

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

View File

@ -30,6 +30,13 @@ end Lean.PersistentArray
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
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.Command
import Lean.Elab.InfoTree
import Lean.DeclarationRange
import Pantograph.Frontend.Basic
import Pantograph.Frontend.MetaTranslate
import Pantograph.Goal
import Pantograph.Protocol
import Pantograph.Frontend.InfoTree
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)).join.toPArray']
else
(children.toList.map (filter p m)).join
| .hole mvar => if m mvar then [.hole mvar] else []
end Lean.Elab.InfoTree
namespace Pantograph.Frontend
-- Info tree filtering functions
/- Adapted from lean-training-data -/
structure TacticInvocation where
info : Elab.TacticInfo
ctx : Elab.ContextInfo
@ -131,19 +65,10 @@ protected def usedConstants (t: TacticInvocation) : NameSet :=
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.bind (fun t => findAllInfo t context? pred)
| _ => []
/-- Return all `TacticInfo` nodes in an `InfoTree` corresponding to tactics,
each equipped with its relevant `ContextInfo`, and any children info trees. -/
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
| _ => false
infos.filterMap fun p => match p with
@ -162,9 +87,11 @@ def collectTacticsFromCompilationStep (step : CompilationStep) : IO (List Protoc
tactics.mapM λ invocation => do
let goalBefore := (Format.joinSep (← invocation.goalState) "\n").pretty
let goalAfter := (Format.joinSep (← invocation.goalStateAfter) "\n").pretty
let tactic ← invocation.ctx.runMetaM {} do
let t ← PrettyPrinter.ppTactic ⟨invocation.info.stx⟩
return t.pretty
let tactic ← invocation.ctx.runMetaM {} <| Meta.withMCtx invocation.info.mctxBefore do
return (← invocation.ctx.ppSyntax {} invocation.info.stx).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
return {
goalBefore,
@ -177,47 +104,79 @@ structure InfoWithContext where
info: Elab.Info
context?: Option Elab.ContextInfo := .none
private def collectSorrysInTree (t : Elab.InfoTree) : List InfoWithContext :=
let infos := findAllInfo t none fun i => match i with
| .ofTermInfo { expectedType?, expr, stx, .. } =>
expr.isSorry ∧ expectedType?.isSome ∧ stx.isOfKind `Lean.Parser.Term.sorry
private def collectSorrysInTree (t : Elab.InfoTree) : IO (List InfoWithContext) := do
let infos ← t.findAllInfoM none fun i ctx? => match i with
| .ofTermInfo { expectedType?, expr, stx, lctx, .. } => do
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, .. } =>
-- The `sorry` term is distinct from the `sorry` tactic
let isSorry := stx.isOfKind `Lean.Parser.Tactic.tacticSorry
isSorry ∧ !goalsBefore.isEmpty
| _ => false
infos.map fun (info, context?, _) => { info, context? }
return (isSorry ∧ !goalsBefore.isEmpty, ¬ isSorry)
| _ => return (false, true)
return infos.map fun (info, context?, _) => { info, context? }
-- NOTE: Plural deliberately not spelled "sorries"
@[export pantograph_frontend_collect_sorrys_m]
def collectSorrys (step: CompilationStep) : List InfoWithContext :=
step.trees.bind collectSorrysInTree
def collectSorrys (step: CompilationStep) : IO (List InfoWithContext) := do
return (← step.trees.mapM collectSorrysInTree).join
structure AnnotatedGoalState where
state : GoalState
srcBoundaries : List (String.Pos × String.Pos)
/--
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
the current `MetavarContext`.
-/
@[export pantograph_frontend_sorrys_to_goal_state]
def sorrysToGoalState (sorrys : List InfoWithContext) : MetaM GoalState := do
@[export pantograph_frontend_sorrys_to_goal_state_m]
def sorrysToGoalState (sorrys : List InfoWithContext) : MetaM AnnotatedGoalState := do
assert! !sorrys.isEmpty
let goalsM := sorrys.mapM λ i => do
match i.info with
| .ofTermInfo termInfo => do
let mvarId ← MetaTranslate.translateMVarFromTermInfo termInfo i.context?
return [mvarId]
return [(mvarId, stxByteRange termInfo.stx)]
| .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"
let goals := List.join (← goalsM.run {} |>.run' {})
let annotatedGoals := List.join (← goalsM.run {} |>.run' {})
let goals := annotatedGoals.map Prod.fst
let srcBoundaries := annotatedGoals.map Prod.snd
let root := match goals with
| [] => panic! "No MVars generated"
| [g] => g
| _ => { 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

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)).join.toPArray']
else
(children.toList.map (filter p m)).join
| .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.bind (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).join
| _ => 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
open Lean
def filename: String := "<pantograph>"
/--
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 }
protected def GoalState.metaState (state: GoalState): Meta.State :=
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
mvarId.withContext m |>.run' (← read) state.metaState
@ -177,16 +177,51 @@ protected def GoalState.getMVarEAssignment (goalState: GoalState) (mvarId: MVarI
--- 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
unless (← getMCtx).decls.contains goal do
throwError s!"Goal is not in context: {goal.name}"
goal.checkNotAssigned `GoalState.step
let (_, newGoals) ← tacticM { elaborator := .anonymous } |>.run { goals := [goal] }
let (_, { goals }) ← tacticM { elaborator := .anonymous } |>.run { goals := [goal] }
let nextElabState ← MonadBacktrack.saveState
Elab.Term.synthesizeSyntheticMVarsNoPostponing
let goals ← if guardMVarErrors then
pure $ mergeMVarLists goals (← collectAllErroredMVars goal)
else
pure goals
return {
state with
savedState := { term := nextElabState, tactic := newGoals },
savedState := { term := nextElabState, tactic := { goals }, },
parentMVar? := .some goal,
calcPrevRhs? := .none,
}
@ -202,16 +237,28 @@ inductive TacticResult where
-- The given action cannot be executed in the state
| invalidAction (message: String)
/-- Executes a `TacticM` monads on this `GoalState`, collecting the errors as necessary -/
protected def GoalState.tryTacticM (state: GoalState) (goal: MVarId) (tacticM: Elab.Tactic.TacticM Unit):
/-- 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) (guardMVarErrors : Bool := false):
Elab.TermElabM TacticResult := do
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
catch exception =>
return .failure #[← exception.toMessageData.toString]
/-- 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):
Elab.TermElabM TacticResult := do
state.restoreElabM
@ -219,10 +266,10 @@ protected def GoalState.tryTactic (state: GoalState) (goal: MVarId) (tactic: Str
(env := ← MonadEnv.getEnv)
(catName := if state.isConv then `conv else `tactic)
(input := tactic)
(fileName := filename) with
(fileName := ← getFileName) with
| .ok stx => pure $ stx
| .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):
Elab.TermElabM TacticResult := do
@ -231,7 +278,7 @@ protected def GoalState.tryAssign (state: GoalState) (goal: MVarId) (expr: Strin
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := expr)
(fileName := filename) with
(fileName := ← getFileName) with
| .ok syn => pure syn
| .error error => return .parseError error
state.tryTacticM goal $ Tactic.evalAssign expr
@ -245,7 +292,7 @@ protected def GoalState.tryLet (state: GoalState) (goal: MVarId) (binderName: St
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := type)
(fileName := filename) with
(fileName := ← getFileName) with
| .ok syn => pure syn
| .error error => return .parseError error
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)
(catName := `term)
(input := pred)
(fileName := filename) with
(fileName := ← getFileName) with
| .ok syn => pure syn
| .error error => return .parseError error
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}"
if let some prevRhs := calcPrevRhs? then
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
-- current branch

View File

@ -289,6 +289,19 @@ structure GoalDiag where
instantiate: Bool := true
printSexp: Bool := false
structure GoalSave where
id: Nat
path: System.FilePath
deriving Lean.FromJson
structure GoalSaveResult where
deriving Lean.ToJson
structure GoalLoad where
path: System.FilePath
deriving Lean.FromJson
structure GoalLoadResult where
id: Nat
deriving Lean.ToJson
/-- Executes the Lean compiler on a single file -/
structure FrontendProcess where
@ -299,6 +312,8 @@ structure FrontendProcess where
invocations: Bool := false
-- If set to true, collect `sorry`s
sorrys: Bool := false
-- If set to true, extract new constants
newConstants: Bool := false
deriving Lean.FromJson
structure InvokedTactic where
goalBefore: String
@ -312,11 +327,16 @@ structure InvokedTactic where
structure CompilationUnit where
-- String boundaries of compilation units
boundary: (Nat × Nat)
messages: Array String := #[]
-- Tactic invocations
invocations?: Option (List InvokedTactic) := .none
goalStateId?: Option Nat := .none
goals: Array Goal := #[]
messages: Array String := #[]
goals?: Option (Array Goal) := .none
-- 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
structure FrontendProcessResult where
units: List CompilationUnit

View File

@ -2,6 +2,7 @@ import Lean.Environment
import Lean.Replay
import Init.System.IOError
import Std.Data.HashMap
import Pantograph.Goal
/-!
Input/Output functions
@ -55,7 +56,7 @@ and when unpickling, we build a fresh `Environment` from the imports,
and then add the new constants.
-/
@[export pantograph_env_pickle_m]
def env_pickle (env : Environment) (path : System.FilePath) : IO Unit :=
def environmentPickle (env : Environment) (path : System.FilePath) : IO Unit :=
Pantograph.pickle path (env.header.imports, env.constants.map₂)
/--
@ -65,9 +66,97 @@ We construct a fresh `Environment` with the relevant imports,
and then replace the new constants.
-/
@[export pantograph_env_unpickle_m]
def env_unpickle (path : System.FilePath) : IO (Environment × CompactedRegion) := unsafe do
def environmentUnpickle (path : System.FilePath) : IO (Environment × CompactedRegion) := unsafe do
let ((imports, map₂), region) ← Pantograph.unpickle (Array Import × PHashMap Name ConstantInfo) path
let env ← importModules imports {} 0
return (← env.replay (Std.HashMap.ofList map₂.toList), region)
open Lean.Core in
structure CompactCoreState where
-- env : Environment
nextMacroScope : MacroScope := firstFrontendMacroScope + 1
ngen : NameGenerator := {}
-- traceState : TraceState := {}
-- cache : Cache := {}
-- messages : MessageLog := {}
-- infoState : Elab.InfoState := {}
@[export pantograph_goal_state_pickle_m]
def goalStatePickle (goalState : GoalState) (path : System.FilePath) : IO Unit :=
let {
savedState := {
term := {
meta := {
core,
meta,
}
«elab»,
},
tactic
}
root,
parentMVar?,
convMVar?,
calcPrevRhs?,
} := goalState
--let env := core.env
Pantograph.pickle path (
({ core with } : CompactCoreState),
meta,
«elab»,
tactic,
root,
parentMVar?,
convMVar?,
calcPrevRhs?,
)
@[export pantograph_goal_state_unpickle_m]
def goalStateUnpickle (path : System.FilePath) (env : Environment)
: IO (GoalState × CompactedRegion) := unsafe do
let ((
compactCore,
meta,
«elab»,
tactic,
root,
parentMVar?,
convMVar?,
calcPrevRhs?,
), region) ← Pantograph.unpickle (
CompactCoreState ×
Meta.State ×
Elab.Term.State ×
Elab.Tactic.State ×
MVarId ×
Option MVarId ×
Option (MVarId × MVarId × List MVarId) ×
Option (MVarId × Expr)
) path
let goalState := {
savedState := {
term := {
meta := {
core := {
compactCore with
passedHeartbeats := 0,
env,
},
meta,
},
«elab»,
},
tactic,
},
root,
parentMVar?,
convMVar?,
calcPrevRhs?,
}
return (goalState, region)
end Pantograph

View File

@ -1,6 +1,6 @@
namespace Pantograph
@[export pantograph_version]
def version := "0.2.19"
def version := "0.2.22"
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
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
@ -69,62 +69,7 @@ stat
```
where the application of `assumption` should lead to a failure.
### Commands
See `Pantograph/Protocol.lean` for a description of the parameters and return values in JSON.
* `reset`: Delete all cached expressions and proof trees
* `stat`: Display resource usage
* `expr.echo {"expr": <expr>, "type": <optional expected type>, ["levels": [<levels>]]}`: Determine the
type of an expression and format it.
* `env.catalog`: Display a list of all safe Lean symbols in the current environment
* `env.inspect {"name": <name>, "value": <bool>}`: Show the type and package of a
given symbol; If value flag is set, the value is printed or hidden. By default
only the values of definitions are printed.
* `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`
One particular option for interest for machine learning researchers is the
automatic mode (flag: `"automaticMode"`). By default it is turned on, with
all goals automatically resuming. This makes Pantograph act like a gym,
with no resumption necessary to manage your goals.
* `options.print`: Display the current set of options
* `goal.start {["name": <name>], ["expr": <expr>], ["levels": [<levels>]], ["copyFrom": <symbol>]}`:
Start a new proof from a given expression or symbol
* `goal.tactic {"stateId": <id>, "goalId": <id>, ...}`: Execute a tactic string on a
given goal. The tactic is supplied as additional key-value pairs in one of the following formats:
- `{ "tactic": <tactic> }`: Execute an ordinary tactic
- `{ "expr": <expr> }`: Assign the given proof term to the current goal
- `{ "have": <expr>, "binderName": <name> }`: Execute `have` and creates a branch goal
- `{ "calc": <expr> }`: Execute one step of a `calc` tactic. Each step must
be of the form `lhs op rhs`. An `lhs` of `_` indicates that it should be set
to the previous `rhs`.
- `{ "conv": <bool> }`: Enter or exit conversion tactic mode. In the case of
exit, the goal id is ignored.
* `goal.continue {"stateId": <id>, ["branch": <id>], ["goals": <names>]}`:
Execute continuation/resumption
- `{ "branch": <id> }`: Continue on branch state. The current state must have no goals.
- `{ "goals": <names> }`: Resume the given goals
* `goal.remove {"stateIds": [<id>]}"`: Drop the goal states specified in the list
* `goal.print {"stateId": <id>}"`: Print a goal state
* `frontend.process { ["fileName": <fileName>",] ["file": <str>], invocations:
<bool>, sorrys: <bool> }`: Executes the Lean frontend on a file, collecting
either the tactic invocations (`"invocations": true`) or the sorrys into goal
states (`"sorrys": true`)
### Errors
When an error pertaining to the execution of a command happens, the returning JSON structure is
``` json
{ "error": "type", "desc": "description" }
```
Common error forms:
* `command`: Indicates malformed command structure which results from either
invalid command or a malformed JSON structure that cannot be fed to an
individual command.
* `index`: Indicates an invariant maintained by the output of one command and
input of another is broken. For example, attempting to query a symbol not
existing in the library or indexing into a non-existent proof state.
For a list of commands, see [REPL Documentation](doc/repl.md).
### Project Environment
@ -135,7 +80,7 @@ the environment might be setup like this:
``` sh
LIB="../lib"
LIB_MATHLIB="$LIB/mathlib4/lake-packages"
LIB_MATHLIB="$LIB/mathlib4/.lake"
export LEAN_PATH="$LIB/mathlib4/build/lib:$LIB_MATHLIB/aesop/build/lib:$LIB_MATHLIB/Qq/build/lib:$LIB_MATHLIB/std/build/lib"
LEAN_PATH=$LEAN_PATH build/bin/pantograph $@

View File

@ -15,6 +15,16 @@ structure State where
/-- Main state monad for executing commands -/
abbrev MainM := ReaderT Context (StateT State Lean.CoreM)
def newGoalState (goalState: GoalState) : MainM Nat := do
let state ← get
let stateId := state.nextId
set { state with
goalStates := state.goalStates.insert stateId goalState,
nextId := state.nextId + 1
}
return stateId
-- HACK: For some reason writing `CommandM α := MainM (Except ... α)` disables
-- certain monadic features in `MainM`
abbrev CR α := Except Protocol.InteractionError α
@ -50,6 +60,8 @@ def execute (command: Protocol.Command): MainM Lean.Json := do
| "goal.continue" => run goal_continue
| "goal.delete" => run goal_delete
| "goal.print" => run goal_print
| "goal.save" => run goal_save
| "goal.load" => run goal_load
| "frontend.process" => run frontend_process
| cmd =>
let error: Protocol.InteractionError :=
@ -62,19 +74,12 @@ def execute (command: Protocol.Command): MainM Lean.Json := do
errorCommand := errorI "command"
errorIndex := errorI "index"
errorIO := errorI "io"
newGoalState (goalState: GoalState) : MainM Nat := do
let state ← get
let stateId := state.nextId
set { state with
goalStates := state.goalStates.insert stateId goalState,
nextId := state.nextId + 1
}
return stateId
-- Command Functions
reset (_: Protocol.Reset): MainM (CR Protocol.StatResult) := do
let state ← get
let nGoals := state.goalStates.size
set { state with nextId := 0, goalStates := .empty }
Lean.Core.resetMessageLog
return .ok { nGoals }
stat (_: Protocol.Stat): MainM (CR Protocol.StatResult) := do
let state ← get
@ -90,10 +95,10 @@ def execute (command: Protocol.Command): MainM Lean.Json := do
Environment.addDecl args
env_save (args: Protocol.EnvSaveLoad): MainM (CR Protocol.EnvSaveLoadResult) := do
let env ← Lean.MonadEnv.getEnv
env_pickle env args.path
environmentPickle env args.path
return .ok {}
env_load (args: Protocol.EnvSaveLoad): MainM (CR Protocol.EnvSaveLoadResult) := do
let (env, _) ← env_unpickle args.path
let (env, _) ← environmentUnpickle args.path
Lean.setEnv env
return .ok {}
expr_echo (args: Protocol.ExprEcho): MainM (CR Protocol.ExprEchoResult) := do
@ -203,11 +208,7 @@ def execute (command: Protocol.Command): MainM Lean.Json := do
match nextState? with
| .error error => return .error <| errorI "structure" error
| .ok nextGoalState =>
let nextStateId := state.nextId
set { state with
goalStates := state.goalStates.insert nextStateId nextGoalState,
nextId := state.nextId + 1
}
let nextStateId ← newGoalState nextGoalState
let goals ← goalSerialize nextGoalState (options := state.options)
return .ok {
nextStateId,
@ -224,6 +225,16 @@ def execute (command: Protocol.Command): MainM Lean.Json := do
return .error $ errorIndex s!"Invalid state index {args.stateId}"
let result ← runMetaInMainM <| goalPrint goalState state.options
return .ok result
goal_save (args: Protocol.GoalSave): MainM (CR Protocol.GoalSaveResult) := do
let state ← get
let .some goalState := state.goalStates[args.id]? |
return .error $ errorIndex s!"Invalid state index {args.id}"
goalStatePickle goalState args.path
return .ok {}
goal_load (args: Protocol.GoalLoad): MainM (CR Protocol.GoalLoadResult) := do
let (goalState, _) ← goalStateUnpickle args.path (← Lean.MonadEnv.getEnv)
let id ← newGoalState goalState
return .ok { id }
frontend_process (args: Protocol.FrontendProcess): MainM (CR Protocol.FrontendProcessResult) := do
let options := (← get).options
try
@ -247,27 +258,38 @@ def execute (command: Protocol.Command): MainM Lean.Json := do
pure $ .some invocations
else
pure .none
let sorrys := if args.sorrys then
let sorrys if args.sorrys then
Frontend.collectSorrys step
else
[]
pure []
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 units ← li.mapM λ (env, boundary, invocations?, sorrys, messages) => Lean.withEnv env do
let (goalStateId?, goals) ← if sorrys.isEmpty then do
pure (.none, #[])
let units ← li.mapM λ (env, boundary, invocations?, sorrys, messages, newConstants) => Lean.withEnv env do
let newConstants? := if args.newConstants then
.some $ newConstants.toArray.map λ name => name.toString
else
.none
let (goalStateId?, goals?, goalSrcBoundaries?) ← if sorrys.isEmpty then do
pure (.none, .none, .none)
else do
let goalState ← runMetaInMainM $ Frontend.sorrysToGoalState sorrys
let stateId ← newGoalState goalState
let goals ← goalSerialize goalState options
pure (.some stateId, goals)
let { state, srcBoundaries } ← runMetaInMainM $ Frontend.sorrysToGoalState sorrys
let stateId ← newGoalState state
let goals ← goalSerialize state options
let srcBoundaries := srcBoundaries.toArray.map (λ (b, e) => (b.byteIdx, e.byteIdx))
pure (.some stateId, .some goals, .some srcBoundaries)
return {
boundary,
messages,
invocations?,
goalStateId?,
goals,
messages,
goals?,
goalSrcBoundaries?,
newConstants?,
}
return .ok { units }
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 strToTermSyntax [Monad m] [MonadEnv m] (s: String): m Syntax := do
def strToTermSyntax (s: String): CoreM Syntax := do
let .ok stx := Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := s)
(fileName := filename) | panic! s!"Failed to parse {s}"
(fileName := ← getFileName) | panic! s!"Failed to parse {s}"
return stx
def parseSentence (s: String): Elab.TermElabM Expr := do
let stx ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := s)
(fileName := filename) with
(fileName := ← getFileName) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
Elab.Term.elabTerm (stx := stx) .none
@ -123,13 +123,28 @@ def mvarUserNameAndType (mvarId: MVarId): MetaM (Name × String) := do
-- 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
def runTest [Monad m] (t: TestT m Unit): m LSpec.TestSeq :=
def checkEq [DecidableEq α] [Repr α] (desc : String) (lhs rhs : α) : TestT m Unit := do
addTest $ LSpec.check desc (lhs = rhs)
def checkTrue (desc : String) (flag : Bool) : TestT m Unit := do
addTest $ LSpec.check desc flag
def fail (desc : String) : TestT m Unit := do
addTest $ LSpec.check desc false
def runTest (t: TestT m Unit): m LSpec.TestSeq :=
Prod.snd <$> t.run LSpec.TestSeq.done
def runTestWithResult { α } (t: TestT m α): m (α × LSpec.TestSeq) :=
t.run LSpec.TestSeq.done
end Monadic
def runTestTermElabM (env: Environment) (t: TestT Elab.TermElabM Unit):
IO LSpec.TestSeq :=

View File

@ -32,7 +32,7 @@ def test_expr_to_binder (env: Environment): IO LSpec.TestSeq := do
def test_sexp_of_symbol (env: Environment): IO LSpec.TestSeq := do
let entries: List (String × String) := [
-- This one contains unhygienic variable names which must be suppressed
("Nat.add", "(:forall _ (:c Nat) (:forall _ (:c Nat) (:c Nat)))"),
("Nat.add", "(:forall a (:c Nat) (:forall a (:c Nat) (:c Nat)))"),
-- These ones are normal and easy
("Nat.add_one", "(:forall n (:c Nat) ((:c Eq) (:c Nat) ((:c HAdd.hAdd) (:c Nat) (:c Nat) (:c Nat) ((:c instHAdd) (:c Nat) (:c instAddNat)) 0 ((:c OfNat.ofNat) (:c Nat) (:lit 1) ((:c instOfNatNat) (:lit 1)))) ((:c Nat.succ) 0)))"),
("Nat.le_of_succ_le", "(:forall n (:c Nat) (:forall m (:c Nat) (:forall h ((:c LE.le) (:c Nat) (:c instLENat) ((:c Nat.succ) 1) 0) ((:c LE.le) (:c Nat) (:c instLENat) 2 1)) :implicit) :implicit)"),

View File

@ -10,13 +10,13 @@ def collectSorrysFromSource (source: String) : MetaM (List GoalState) := do
let filename := "<anonymous>"
let (context, state) ← do Frontend.createContextStateFromFile source filename (← getEnv) {}
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 goalStates ← li.filterMapM λ (env, sorrys) => withEnv env do
if sorrys.isEmpty then
return .none
let goalState ← Frontend.sorrysToGoalState sorrys
return .some goalState
let { state, .. } ← Frontend.sorrysToGoalState sorrys
return .some state
return goalStates
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) :=
let tests := [
@ -185,6 +226,9 @@ def suite (env : Environment): List (String × IO LSpec.TestSeq) :=
("sorry_in_induction", test_sorry_in_induction),
("sorry_in_coupled", test_sorry_in_coupled),
("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))

View File

@ -174,6 +174,7 @@ def test_frontend_process : Test :=
("file", .str file),
("invocations", .bool true),
("sorrys", .bool false),
("newConstants", .bool false),
]
({
units := [{
@ -214,6 +215,7 @@ def test_frontend_process_sorry : Test :=
("file", .str file),
("invocations", .bool false),
("sorrys", .bool true),
("newConstants", .bool false),
]
({
units := [{
@ -221,7 +223,8 @@ def test_frontend_process_sorry : Test :=
}, {
boundary := (solved.utf8ByteSize, solved.utf8ByteSize + withSorry.utf8ByteSize),
goalStateId? := .some 0,
goals := #[goal1],
goals? := .some #[goal1],
goalSrcBoundaries? := .some #[(57, 62)],
messages := #["<anonymous>:2:0: warning: declaration uses 'sorry'\n"],
}],
}: Protocol.FrontendProcessResult),

View File

@ -24,7 +24,7 @@ def test_expr_echo (env: Environment): IO LSpec.TestSeq := do
},
expr := {
pp? := "⟨∀ (x : Prop), x → x, fun x h => h⟩",
sexp? := "((:c PSigma.mk) (:sort 0) (:lambda p (:sort 0) 0) (:forall x (:sort 0) (:forall _ 0 1)) (:lambda x (:sort 0) (:lambda h 0 0)))",
sexp? := "((:c PSigma.mk) (:sort 0) (:lambda p (:sort 0) 0) (:forall x (:sort 0) (:forall a 0 1)) (:lambda x (:sort 0) (:lambda h 0 0)))",
}
}))
return tests

View File

@ -1,11 +1,12 @@
import LSpec
import Test.Delate
import Test.Environment
import Test.Frontend
import Test.Integration
import Test.Library
import Test.Metavar
import Test.Proofs
import Test.Delate
import Test.Serial
import Test.Tactic
-- Test running infrastructure
@ -51,6 +52,7 @@ def main (args: List String) := do
("Metavar", Metavar.suite env_default),
("Proofs", Proofs.suite env_default),
("Delate", Delate.suite env_default),
("Serial", Serial.suite env_default),
("Tactic/Congruence", Tactic.Congruence.suite env_default),
("Tactic/Motivated Apply", Tactic.MotivatedApply.suite env_default),
("Tactic/No Confuse", Tactic.NoConfuse.suite env_default),

View File

@ -8,10 +8,7 @@ namespace Pantograph.Test.Metavar
open Pantograph
open Lean
abbrev TestM := StateRefT LSpec.TestSeq (ReaderT Protocol.Options Elab.TermElabM)
def addTest (test: LSpec.TestSeq): TestM Unit := do
set $ (← get) ++ test
abbrev TestM := TestT $ ReaderT Protocol.Options Elab.TermElabM
-- Tests that all delay assigned mvars are instantiated
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)))")
return ()
def startProof (expr: String): TestM (Option GoalState) := do
let env ← Lean.MonadEnv.getEnv
let syn? := parseTerm env expr

View File

@ -14,10 +14,7 @@ inductive Start where
| copy (name: String) -- Start from some name in the environment
| expr (expr: String) -- Start from some expression
abbrev TestM := StateRefT LSpec.TestSeq (ReaderT Protocol.Options Elab.TermElabM)
def addTest (test: LSpec.TestSeq): TestM Unit := do
set $ (← get) ++ test
abbrev TestM := TestT $ ReaderT Protocol.Options $ Elab.TermElabM
def startProof (start: Start): TestM (Option GoalState) := do
let env ← Lean.MonadEnv.getEnv
@ -282,9 +279,9 @@ def test_or_comm: TestM Unit := do
serializeExpressionSexp (← instantiateAll state2.parentExpr?.get!) (sanitize := false)
let orPQ := s!"((:c Or) (:fv {fvP}) (:fv {fvQ}))"
let orQP := s!"((:c Or) (:fv {fvQ}) (:fv {fvP}))"
let motive := s!"(:lambda t._@._hyg.26 {orPQ} (:forall h ((:c Eq) ((:c Or) (:fv {fvP}) (:fv {fvQ})) (:fv {fvH}) 0) {orQP}))"
let caseL := s!"(:lambda h._@._hyg.27 (:fv {fvP}) (:lambda h._@._hyg.28 ((:c Eq) {orPQ} (:fv {fvH}) ((:c Or.inl) (:fv {fvP}) (:fv {fvQ}) 0)) (:subst (:mv {caseL}) (:fv {fvP}) (:fv {fvQ}) 1)))"
let caseR := s!"(:lambda h._@._hyg.29 (:fv {fvQ}) (:lambda h._@._hyg.30 ((:c Eq) {orPQ} (:fv {fvH}) ((:c Or.inr) (:fv {fvP}) (:fv {fvQ}) 0)) (:subst (:mv {caseR}) (:fv {fvP}) (:fv {fvQ}) 1)))"
let motive := s!"(:lambda t {orPQ} (:forall h ((:c Eq) ((:c Or) (:fv {fvP}) (:fv {fvQ})) (:fv {fvH}) 0) {orQP}))"
let caseL := s!"(:lambda h (:fv {fvP}) (:lambda h ((:c Eq) {orPQ} (:fv {fvH}) ((:c Or.inl) (:fv {fvP}) (:fv {fvQ}) 0)) (:subst (:mv {caseL}) (:fv {fvP}) (:fv {fvQ}) 1)))"
let caseR := s!"(:lambda h (:fv {fvQ}) (:lambda h ((:c Eq) {orPQ} (:fv {fvH}) ((:c Or.inr) (:fv {fvP}) (:fv {fvQ}) 0)) (:subst (:mv {caseR}) (:fv {fvP}) (:fv {fvQ}) 1)))"
let conduit := s!"((:c Eq.refl) {orPQ} (:fv {fvH}))"
addTest $ LSpec.test "(2 parent)" (state2parent ==
s!"((:c Or.casesOn) (:fv {fvP}) (:fv {fvQ}) {motive} (:fv {fvH}) {caseL} {caseR} {conduit})")
@ -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) :=
let tests := [
("identity", test_identity),
@ -716,6 +763,8 @@ def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
("calc", test_calc),
("Nat.zero_add", test_nat_zero_add),
("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))

109
Test/Serial.lean Normal file
View File

@ -0,0 +1,109 @@
import LSpec
import Test.Common
import Lean
import Pantograph.Library
open Lean
namespace Pantograph.Test.Serial
def tempPath : IO System.FilePath := do
Prod.snd <$> IO.FS.createTempFile
structure MultiState where
coreContext : Core.Context
env: Environment
abbrev TestM := TestT $ StateRefT MultiState $ IO
instance : MonadEnv TestM where
getEnv := return (← getThe MultiState).env
modifyEnv f := do modifyThe MultiState fun s => { s with env := f s.env }
def runCoreM { α } (state : Core.State) (testCoreM : TestT CoreM α) : TestM (α × Core.State) := do
let multiState ← getThe MultiState
let coreM := runTestWithResult testCoreM
match ← (coreM.run multiState.coreContext state).toBaseIO with
| .error e => do
throw $ .userError $ ← e.toMessageData.toString
| .ok ((a, tests), state') => do
set $ (← getThe LSpec.TestSeq) ++ tests
return (a, state')
def test_environment_pickling : TestM Unit := do
let coreSrc : Core.State := { env := ← getEnv }
let coreDst : Core.State := { env := ← getEnv }
let name := `mystery
let envPicklePath ← tempPath
let ((), _) ← runCoreM coreSrc do
let type: Expr := .forallE `p (.sort 0) (.forallE `h (.bvar 0) (.bvar 1) .default) .default
let value: Expr := .lam `p (.sort 0) (.lam `h (.bvar 0) (.bvar 0) .default) .default
let c := Lean.Declaration.defnDecl <| Lean.mkDefinitionValEx
(name := name)
(levelParams := [])
(type := type)
(value := value)
(hints := Lean.mkReducibilityHintsRegularEx 1)
(safety := Lean.DefinitionSafety.safe)
(all := [])
let env' ← match (← getEnv).addDecl (← getOptions) c with
| .error e => do
let error ← (e.toMessageData (← getOptions)).toString
throwError error
| .ok env' => pure env'
environmentPickle env' envPicklePath
let _ ← runCoreM coreDst do
let (env', _) ← environmentUnpickle envPicklePath
checkTrue s!"Has symbol {name}" (env'.find? name).isSome
let anotherName := `mystery2
checkTrue s!"Doesn't have symbol {anotherName}" (env'.find? anotherName).isNone
IO.FS.removeFile envPicklePath
def test_goal_state_pickling_simple : TestM Unit := do
let coreSrc : Core.State := { env := ← getEnv }
let coreDst : Core.State := { env := ← getEnv }
let statePath ← tempPath
let type: Expr := .forallE `p (.sort 0) (.forallE `h (.bvar 0) (.bvar 1) .default) .default
let stateGenerate : MetaM GoalState := runTermElabMInMeta do
GoalState.create type
let ((), _) ← runCoreM coreSrc do
let state ← stateGenerate.run'
goalStatePickle state statePath
let ((), _) ← runCoreM coreDst do
let (goalState, _) ← goalStateUnpickle statePath (← getEnv)
let metaM : MetaM (List Expr) := do
goalState.goals.mapM λ goal => goalState.withContext goal goal.getType
let types ← metaM.run'
checkTrue "Goals" $ types[0]!.equal type
IO.FS.removeFile statePath
structure Test where
name : String
routine: TestM Unit
protected def Test.run (test: Test) (env: Lean.Environment) : IO LSpec.TestSeq := do
-- Create the state
let state : MultiState := {
coreContext := ← createCoreContext #[],
env,
}
match ← ((runTest $ test.routine).run' state).toBaseIO with
| .ok e => return e
| .error e =>
return LSpec.check s!"Emitted exception: {e.toString}" (e.toString == "")
def suite (env : Lean.Environment): List (String × IO LSpec.TestSeq) :=
let tests: List Test := [
{ name := "environment_pickling", routine := test_environment_pickling, },
{ name := "goal_state_pickling_simple", routine := test_goal_state_pickling_simple, },
]
tests.map (fun test => (test.name, test.run env))
end Pantograph.Test.Serial

View File

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

View File

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

View File

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

View File

@ -24,6 +24,22 @@ The name Pantograph is a pun. It means two things
a locomotive. In comparison the (relatively) simple Pantograph software powers
theorem proving projects.
## Caveats
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.
## References
* [Pantograph Paper](https://arxiv.org/abs/2410.16429)

66
doc/repl.md Normal file
View File

@ -0,0 +1,66 @@
# REPL
## Commands
See `Pantograph/Protocol.lean` for a description of the parameters and return values in JSON.
* `reset`: Delete all cached expressions and proof trees
* `stat`: Display resource usage
* `expr.echo {"expr": <expr>, "type": <optional expected type>, ["levels": [<levels>]]}`: Determine the
type of an expression and format it.
* `env.catalog`: Display a list of all safe Lean symbols in the current environment
* `env.inspect {"name": <name>, "value": <bool>}`: Show the type and package of a
given symbol; If value flag is set, the value is printed or hidden. By default
only the values of definitions are printed.
* `env.save { "path": <fileName> }`, `env.load { "path": <fileName> }`: Save/Load the
current environment to/from a file
* `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`
One particular option for interest for machine learning researchers is the
automatic mode (flag: `"automaticMode"`). By default it is turned on, with
all goals automatically resuming. This makes Pantograph act like a gym,
with no resumption necessary to manage your goals.
* `options.print`: Display the current set of options
* `goal.start {["name": <name>], ["expr": <expr>], ["levels": [<levels>]], ["copyFrom": <symbol>]}`:
Start a new proof from a given expression or symbol
* `goal.tactic {"stateId": <id>, "goalId": <id>, ...}`: Execute a tactic string on a
given goal. The tactic is supplied as additional key-value pairs in one of the following formats:
- `{ "tactic": <tactic> }`: Execute an ordinary tactic
- `{ "expr": <expr> }`: Assign the given proof term to the current goal
- `{ "have": <expr>, "binderName": <name> }`: Execute `have` and creates a branch goal
- `{ "calc": <expr> }`: Execute one step of a `calc` tactic. Each step must
be of the form `lhs op rhs`. An `lhs` of `_` indicates that it should be set
to the previous `rhs`.
- `{ "conv": <bool> }`: Enter or exit conversion tactic mode. In the case of
exit, the goal id is ignored.
* `goal.continue {"stateId": <id>, ["branch": <id>], ["goals": <names>]}`:
Execute continuation/resumption
- `{ "branch": <id> }`: Continue on branch state. The current state must have no goals.
- `{ "goals": <names> }`: Resume the given goals
* `goal.remove {"stateIds": [<id>]}"`: Drop the goal states specified in the list
* `goal.print {"stateId": <id>}"`: Print a goal state
* `goal.save { "id": <id>, "path": <fileName> }`, `goal.load { "path": <fileName> }`:
Save/Load a goal state to/from a file. The environment is not carried with the
state. The user is responsible to ensure the sender/receiver instances share
the same environment.
* `frontend.process { ["fileName": <fileName>,] ["file": <str>], invocations:
<bool>, sorrys: <bool>, newConstants: <bool> }`: Executes the Lean frontend on
a file, collecting the tactic invocations (`"invocations": true`), the
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
When an error pertaining to the execution of a command happens, the returning JSON structure is
``` json
{ "error": "type", "desc": "description" }
```
Common error forms:
* `command`: Indicates malformed command structure which results from either
invalid command or a malformed JSON structure that cannot be fed to an
individual command.
* `index`: Indicates an invariant maintained by the output of one command and
input of another is broken. For example, attempting to query a symbol not
existing in the library or indexing into a non-existent proof state.

View File

@ -22,6 +22,8 @@
flake = {
};
systems = [
"aarch64-linux"
"aarch64-darwin"
"x86_64-linux"
"x86_64-darwin"
];