Compare commits
No commits in common. "dev" and "v0.3.0" have entirely different histories.
20
Main.lean
20
Main.lean
|
@ -8,12 +8,6 @@ import Repl
|
||||||
open Pantograph.Repl
|
open Pantograph.Repl
|
||||||
open Pantograph.Protocol
|
open Pantograph.Protocol
|
||||||
|
|
||||||
/-- Print a string to stdout without buffering -/
|
|
||||||
def printImmediate (s : String) : IO Unit := do
|
|
||||||
let stdout ← IO.getStdout
|
|
||||||
stdout.putStr (s ++ "\n")
|
|
||||||
stdout.flush
|
|
||||||
|
|
||||||
/-- Parse a command either in `{ "cmd": ..., "payload": ... }` form or `cmd { ... }` form. -/
|
/-- Parse a command either in `{ "cmd": ..., "payload": ... }` form or `cmd { ... }` form. -/
|
||||||
def parseCommand (s: String): Except String Command := do
|
def parseCommand (s: String): Except String Command := do
|
||||||
match s.trim.get? 0 with
|
match s.trim.get? 0 with
|
||||||
|
@ -40,27 +34,27 @@ partial def loop : MainM Unit := do repeat do
|
||||||
| .error error =>
|
| .error error =>
|
||||||
let error := Lean.toJson ({ error := "command", desc := error }: InteractionError)
|
let error := Lean.toJson ({ error := "command", desc := error }: InteractionError)
|
||||||
-- Using `Lean.Json.compress` here to prevent newline
|
-- Using `Lean.Json.compress` here to prevent newline
|
||||||
printImmediate error.compress
|
IO.println error.compress
|
||||||
| .ok command =>
|
| .ok command =>
|
||||||
try
|
try
|
||||||
let ret ← execute command
|
let ret ← execute command
|
||||||
let str := match state.options.printJsonPretty with
|
let str := match state.options.printJsonPretty with
|
||||||
| true => ret.pretty
|
| true => ret.pretty
|
||||||
| false => ret.compress
|
| false => ret.compress
|
||||||
printImmediate str
|
IO.println str
|
||||||
catch e =>
|
catch e =>
|
||||||
let message := e.toString
|
let message := e.toString
|
||||||
let error := Lean.toJson ({ error := "main", desc := message }: InteractionError)
|
let error := Lean.toJson ({ error := "main", desc := message }: InteractionError)
|
||||||
printImmediate error.compress
|
IO.println error.compress
|
||||||
|
|
||||||
def main (args: List String): IO Unit := do
|
|
||||||
|
unsafe def main (args: List String): IO Unit := do
|
||||||
-- NOTE: A more sophisticated scheme of command line argument handling is needed.
|
-- NOTE: A more sophisticated scheme of command line argument handling is needed.
|
||||||
if args == ["--version"] then do
|
if args == ["--version"] then do
|
||||||
IO.println s!"{Pantograph.version}"
|
IO.println s!"{Pantograph.version}"
|
||||||
return
|
return
|
||||||
|
|
||||||
unsafe do
|
Pantograph.initSearch ""
|
||||||
Pantograph.initSearch ""
|
|
||||||
|
|
||||||
-- Separate imports and options
|
-- Separate imports and options
|
||||||
let (options, imports) := args.partition (·.startsWith "--")
|
let (options, imports) := args.partition (·.startsWith "--")
|
||||||
|
@ -68,7 +62,7 @@ def main (args: List String): IO Unit := do
|
||||||
let coreState ← Pantograph.createCoreState imports.toArray
|
let coreState ← Pantograph.createCoreState imports.toArray
|
||||||
try
|
try
|
||||||
let mainM := loop.run { coreContext } |>.run' { env := coreState.env }
|
let mainM := loop.run { coreContext } |>.run' { env := coreState.env }
|
||||||
printImmediate "ready."
|
IO.println "ready."
|
||||||
mainM
|
mainM
|
||||||
catch ex =>
|
catch ex =>
|
||||||
let message := ex.toString
|
let message := ex.toString
|
||||||
|
|
|
@ -63,16 +63,12 @@ def exprProjToApp (env : Environment) (e : Expr) : Expr :=
|
||||||
(List.range numFields)
|
(List.range numFields)
|
||||||
mkAppN callee (typeArgs ++ [motive, major, induct]).toArray
|
mkAppN callee (typeArgs ++ [motive, major, induct]).toArray
|
||||||
|
|
||||||
def isAuxLemma (n : Name) : Bool :=
|
def _root_.Lean.Name.isAuxLemma (n : Lean.Name) : Bool := n matches .num (.str _ "_auxLemma") _
|
||||||
match n with
|
|
||||||
-- `mkAuxLemma` generally allows for arbitrary prefixes but these are the ones produced by core.
|
|
||||||
| .str _ s => "_proof_".isPrefixOf s || "_simp_".isPrefixOf s
|
|
||||||
| _ => false
|
|
||||||
|
|
||||||
/-- Unfold all lemmas created by `Lean.Meta.mkAuxLemma`. These end in `_auxLemma.nn` where `nn` is a number. -/
|
/-- Unfold all lemmas created by `Lean.Meta.mkAuxLemma`. These end in `_auxLemma.nn` where `nn` is a number. -/
|
||||||
@[export pantograph_unfold_aux_lemmas_m]
|
@[export pantograph_unfold_aux_lemmas_m]
|
||||||
def unfoldAuxLemmas (e : Expr) : CoreM Expr := do
|
def unfoldAuxLemmas : Expr → CoreM Expr :=
|
||||||
Meta.deltaExpand e isAuxLemma
|
(Meta.deltaExpand · Lean.Name.isAuxLemma)
|
||||||
/-- Unfold all matcher applications -/
|
/-- Unfold all matcher applications -/
|
||||||
@[export pantograph_unfold_matchers_m]
|
@[export pantograph_unfold_matchers_m]
|
||||||
def unfoldMatchers (expr : Expr) : CoreM Expr :=
|
def unfoldMatchers (expr : Expr) : CoreM Expr :=
|
||||||
|
@ -168,7 +164,7 @@ partial def instantiateDelayedMVars (expr : Expr) : MetaM Expr :=
|
||||||
self e := instantiateDelayedMVars e
|
self e := instantiateDelayedMVars e
|
||||||
|
|
||||||
/--
|
/--
|
||||||
Convert an expression to an equivalent form with
|
Convert an expression to an equiavlent form with
|
||||||
1. No nested delayed assigned mvars
|
1. No nested delayed assigned mvars
|
||||||
2. No aux lemmas or matchers
|
2. No aux lemmas or matchers
|
||||||
3. No assigned mvars
|
3. No assigned mvars
|
||||||
|
@ -455,6 +451,7 @@ def serializeExpression (options: @&Protocol.Options) (e: Expr): MetaM Protocol.
|
||||||
dependentMVars?,
|
dependentMVars?,
|
||||||
}
|
}
|
||||||
|
|
||||||
|
|
||||||
/-- Adapted from ppGoal -/
|
/-- Adapted from ppGoal -/
|
||||||
def serializeGoal (options: @&Protocol.Options) (goal: MVarId) (mvarDecl: MetavarDecl) (parentDecl?: Option MetavarDecl := .none)
|
def serializeGoal (options: @&Protocol.Options) (goal: MVarId) (mvarDecl: MetavarDecl) (parentDecl?: Option MetavarDecl := .none)
|
||||||
: MetaM Protocol.Goal := do
|
: MetaM Protocol.Goal := do
|
||||||
|
@ -520,6 +517,7 @@ def serializeGoal (options: @&Protocol.Options) (goal: MVarId) (mvarDecl: Metava
|
||||||
return {
|
return {
|
||||||
name := goal.name.toString,
|
name := goal.name.toString,
|
||||||
userName? := if mvarDecl.userName == .anonymous then .none else .some (ofName mvarDecl.userName),
|
userName? := if mvarDecl.userName == .anonymous then .none else .some (ofName mvarDecl.userName),
|
||||||
|
isConversion := isLHSGoal? mvarDecl.type |>.isSome,
|
||||||
target := (← serializeExpression options (← instantiate mvarDecl.type)),
|
target := (← serializeExpression options (← instantiate mvarDecl.type)),
|
||||||
vars := vars.reverse.toArray
|
vars := vars.reverse.toArray
|
||||||
}
|
}
|
||||||
|
@ -529,20 +527,17 @@ def serializeGoal (options: @&Protocol.Options) (goal: MVarId) (mvarDecl: Metava
|
||||||
|
|
||||||
protected def GoalState.serializeGoals
|
protected def GoalState.serializeGoals
|
||||||
(state: GoalState)
|
(state: GoalState)
|
||||||
|
(parent: Option GoalState := .none)
|
||||||
(options: @&Protocol.Options := {}):
|
(options: @&Protocol.Options := {}):
|
||||||
MetaM (Array Protocol.Goal):= do
|
MetaM (Array Protocol.Goal):= do
|
||||||
state.restoreMetaM
|
state.restoreMetaM
|
||||||
let goals := state.goals.toArray
|
let goals := state.goals.toArray
|
||||||
|
let parentDecl? := parent.bind (λ parentState => parentState.mctx.findDecl? state.parentMVar?.get!)
|
||||||
goals.mapM fun goal => do
|
goals.mapM fun goal => do
|
||||||
let fragment := match state.fragments[goal]? with
|
|
||||||
| .none => .tactic
|
|
||||||
| .some $ .calc .. => .calc
|
|
||||||
| .some $ .conv .. => .conv
|
|
||||||
| .some $ .convSentinel .. => .conv
|
|
||||||
match state.mctx.findDecl? goal with
|
match state.mctx.findDecl? goal with
|
||||||
| .some mvarDecl =>
|
| .some mvarDecl =>
|
||||||
let serializedGoal ← serializeGoal options goal mvarDecl (parentDecl? := .none)
|
let serializedGoal ← serializeGoal options goal mvarDecl (parentDecl? := parentDecl?)
|
||||||
pure { serializedGoal with fragment }
|
pure serializedGoal
|
||||||
| .none => throwError s!"Metavariable does not exist in context {goal.name}"
|
| .none => throwError s!"Metavariable does not exist in context {goal.name}"
|
||||||
|
|
||||||
/-- Print the metavariables in a readable format -/
|
/-- Print the metavariables in a readable format -/
|
||||||
|
@ -608,9 +603,7 @@ protected def GoalState.diag (goalState: GoalState) (parent?: Option GoalState :
|
||||||
userNameToString : Name → String
|
userNameToString : Name → String
|
||||||
| .anonymous => ""
|
| .anonymous => ""
|
||||||
| other => s!"[{other}]"
|
| other => s!"[{other}]"
|
||||||
parentHasMVar (mvarId: MVarId): Bool := match parent? with
|
parentHasMVar (mvarId: MVarId): Bool := parent?.map (λ state => state.mctx.decls.contains mvarId) |>.getD true
|
||||||
| .some state => state.mctx.decls.contains mvarId
|
|
||||||
| .none => true
|
|
||||||
|
|
||||||
initialize
|
initialize
|
||||||
registerTraceClass `Pantograph.Delate
|
registerTraceClass `Pantograph.Delate
|
||||||
|
|
|
@ -4,7 +4,6 @@ import Pantograph.Protocol
|
||||||
import Pantograph.Serial
|
import Pantograph.Serial
|
||||||
import Lean.Environment
|
import Lean.Environment
|
||||||
import Lean.Replay
|
import Lean.Replay
|
||||||
import Lean.Util.Path
|
|
||||||
|
|
||||||
open Lean
|
open Lean
|
||||||
open Pantograph
|
open Pantograph
|
||||||
|
@ -14,7 +13,7 @@ namespace Pantograph.Environment
|
||||||
@[export pantograph_is_name_internal]
|
@[export pantograph_is_name_internal]
|
||||||
def isNameInternal (n: Name): Bool :=
|
def isNameInternal (n: Name): Bool :=
|
||||||
-- Returns true if the name is an implementation detail which should not be shown to the user.
|
-- Returns true if the name is an implementation detail which should not be shown to the user.
|
||||||
isAuxLemma n ∨ n.hasMacroScopes
|
n.isAuxLemma ∨ n.hasMacroScopes
|
||||||
|
|
||||||
/-- Catalog all the non-internal and safe names -/
|
/-- Catalog all the non-internal and safe names -/
|
||||||
@[export pantograph_environment_catalog]
|
@[export pantograph_environment_catalog]
|
||||||
|
@ -131,19 +130,17 @@ def inspect (args: Protocol.EnvInspect) (options: @&Protocol.Options): Protocol.
|
||||||
} }
|
} }
|
||||||
| _ => pure core
|
| _ => pure core
|
||||||
let result ← if args.source?.getD false then
|
let result ← if args.source?.getD false then
|
||||||
try
|
let srcSearchPath ← initSrcSearchPath
|
||||||
let sourceUri? ← module?.bindM (Server.documentUriFromModule? ·)
|
let sourceUri? ← module?.bindM (Server.documentUriFromModule srcSearchPath ·)
|
||||||
let declRange? ← findDeclarationRanges? name
|
let declRange? ← findDeclarationRanges? name
|
||||||
let sourceStart? := declRange?.map (·.range.pos)
|
let sourceStart? := declRange?.map (·.range.pos)
|
||||||
let sourceEnd? := declRange?.map (·.range.endPos)
|
let sourceEnd? := declRange?.map (·.range.endPos)
|
||||||
.pure {
|
.pure {
|
||||||
result with
|
result with
|
||||||
sourceUri? := sourceUri?.map (toString ·),
|
sourceUri?,
|
||||||
sourceStart?,
|
sourceStart?,
|
||||||
sourceEnd?,
|
sourceEnd?,
|
||||||
}
|
}
|
||||||
catch _e =>
|
|
||||||
.pure result
|
|
||||||
else
|
else
|
||||||
.pure result
|
.pure result
|
||||||
return result
|
return result
|
||||||
|
|
|
@ -52,6 +52,10 @@ structure CompilationStep where
|
||||||
|
|
||||||
namespace CompilationStep
|
namespace CompilationStep
|
||||||
|
|
||||||
|
@[export pantograph_frontend_compilation_step_message_strings_m]
|
||||||
|
def messageStrings (step: CompilationStep) : IO (Array String) := do
|
||||||
|
List.toArray <$> step.msgs.mapM (·.toString)
|
||||||
|
|
||||||
end CompilationStep
|
end CompilationStep
|
||||||
|
|
||||||
|
|
||||||
|
|
|
@ -154,31 +154,33 @@ the draft tactic instead.
|
||||||
-/
|
-/
|
||||||
@[export pantograph_frontend_sorrys_to_goal_state_m]
|
@[export pantograph_frontend_sorrys_to_goal_state_m]
|
||||||
def sorrysToGoalState (sorrys : List InfoWithContext) : MetaM AnnotatedGoalState := do
|
def sorrysToGoalState (sorrys : List InfoWithContext) : MetaM AnnotatedGoalState := do
|
||||||
|
let env := sorrys.head? >>= (·.context?) |>.map (·.env) |>.getD (← getEnv)
|
||||||
assert! !sorrys.isEmpty
|
assert! !sorrys.isEmpty
|
||||||
let goalsM := sorrys.mapM λ i => do
|
withEnv env do
|
||||||
match i.info with
|
let goalsM := sorrys.mapM λ i => do
|
||||||
| .ofTermInfo termInfo => do
|
match i.info with
|
||||||
let mvarId ← MetaTranslate.translateMVarFromTermInfo termInfo i.context?
|
| .ofTermInfo termInfo => do
|
||||||
if (← mvarId.getType).hasSorry then
|
let mvarId ← MetaTranslate.translateMVarFromTermInfo termInfo i.context?
|
||||||
throwError s!"Coupling is not allowed in drafting"
|
|
||||||
return [(mvarId, stxByteRange termInfo.stx)]
|
|
||||||
| .ofTacticInfo tacticInfo => do
|
|
||||||
let mvarIds ← MetaTranslate.translateMVarFromTacticInfoBefore tacticInfo i.context?
|
|
||||||
for mvarId in mvarIds do
|
|
||||||
if (← mvarId.getType).hasSorry then
|
if (← mvarId.getType).hasSorry then
|
||||||
throwError s!"Coupling is not allowed in drafting"
|
throwError s!"Coupling is not allowed in drafting"
|
||||||
let range := stxByteRange tacticInfo.stx
|
return [(mvarId, stxByteRange termInfo.stx)]
|
||||||
return mvarIds.map (·, range)
|
| .ofTacticInfo tacticInfo => do
|
||||||
| _ => panic! "Invalid info"
|
let mvarIds ← MetaTranslate.translateMVarFromTacticInfoBefore tacticInfo i.context?
|
||||||
let annotatedGoals := List.flatten (← goalsM.run {} |>.run' {})
|
for mvarId in mvarIds do
|
||||||
let goals := annotatedGoals.map Prod.fst
|
if (← mvarId.getType).hasSorry then
|
||||||
let srcBoundaries := annotatedGoals.map Prod.snd
|
throwError s!"Coupling is not allowed in drafting"
|
||||||
let root := match goals with
|
let range := stxByteRange tacticInfo.stx
|
||||||
| [] => panic! "No MVars generated"
|
return mvarIds.map (·, range)
|
||||||
| [g] => g
|
| _ => panic! "Invalid info"
|
||||||
| _ => { name := .anonymous }
|
let annotatedGoals := List.flatten (← goalsM.run {} |>.run' {})
|
||||||
let state ← GoalState.createFromMVars goals root
|
let goals := annotatedGoals.map Prod.fst
|
||||||
return { state, srcBoundaries }
|
let srcBoundaries := annotatedGoals.map Prod.snd
|
||||||
|
let root := match goals with
|
||||||
|
| [] => panic! "No MVars generated"
|
||||||
|
| [g] => g
|
||||||
|
| _ => { name := .anonymous }
|
||||||
|
let state ← GoalState.createFromMVars goals root
|
||||||
|
return { state, srcBoundaries }
|
||||||
|
|
||||||
|
|
||||||
@[export pantograph_frontend_collect_new_defined_constants_m]
|
@[export pantograph_frontend_collect_new_defined_constants_m]
|
||||||
|
|
|
@ -10,7 +10,7 @@ namespace Lean.Elab
|
||||||
private def elaboratorToString : Name → String
|
private def elaboratorToString : Name → String
|
||||||
| .anonymous => ""
|
| .anonymous => ""
|
||||||
| n => s!"⟨{n}⟩ "
|
| n => s!"⟨{n}⟩ "
|
||||||
private def indent (s : String) : String := "\n".intercalate $ s.splitOn "\n" |>.map (" " ++ .)
|
private def indent (s : String) : String := "\n".intercalate $ s.splitOn "\n" |>.map ("\t" ++ .)
|
||||||
|
|
||||||
/-- The `Syntax` for a `Lean.Elab.Info`, if there is one. -/
|
/-- The `Syntax` for a `Lean.Elab.Info`, if there is one. -/
|
||||||
protected def Info.stx? : Info → Option Syntax
|
protected def Info.stx? : Info → Option Syntax
|
||||||
|
@ -131,16 +131,16 @@ partial def InfoTree.toString (t : InfoTree) (ctx?: Option Elab.ContextInfo := .
|
||||||
| .node info children =>
|
| .node info children =>
|
||||||
if let some ctx := ctx? then
|
if let some ctx := ctx? then
|
||||||
let node : String ← match info with
|
let node : String ← match info with
|
||||||
| .ofTermInfo info => pure s!"[term] {info.stx}"
|
| .ofTermInfo info => pure s!"[term] {(← info.toString ctx)}"
|
||||||
| .ofCommandInfo info => pure s!"[command] {info.stx}"
|
| .ofCommandInfo info => pure s!"[command] {(← info.toString ctx)}"
|
||||||
| .ofTacticInfo info => pure s!"[tactic] {info.stx}"
|
| .ofTacticInfo info => pure s!"[tactic] {(← info.toString ctx)}"
|
||||||
| .ofMacroExpansionInfo _ => pure "[macro_exp]"
|
| .ofMacroExpansionInfo _ => pure "[macro_exp]"
|
||||||
| .ofOptionInfo info => pure s!"[option] {info.stx}"
|
| .ofOptionInfo _ => pure "[option]"
|
||||||
| .ofFieldInfo _ => pure "[field]"
|
| .ofFieldInfo _ => pure "[field]"
|
||||||
| .ofCompletionInfo info => pure s!"[completion] {info.stx}"
|
| .ofCompletionInfo _ => pure "[completion]"
|
||||||
| .ofUserWidgetInfo _ => pure "[user_widget]"
|
| .ofUserWidgetInfo _ => pure "[user_widget]"
|
||||||
| .ofCustomInfo _ => pure "[custom]"
|
| .ofCustomInfo _ => pure "[custom]"
|
||||||
| .ofFVarAliasInfo _ => pure "[fvar_alias]"
|
| .ofFVarAliasInfo _ => pure "[fvar]"
|
||||||
| .ofFieldRedeclInfo _ => pure "[field_redecl]"
|
| .ofFieldRedeclInfo _ => pure "[field_redecl]"
|
||||||
| .ofChoiceInfo _ => pure "[choice]"
|
| .ofChoiceInfo _ => pure "[choice]"
|
||||||
| .ofPartialTermInfo _ => pure "[partial_term]"
|
| .ofPartialTermInfo _ => pure "[partial_term]"
|
||||||
|
|
|
@ -118,7 +118,6 @@ partial def translateLCtx : MetaTranslateM LocalContext := do
|
||||||
|
|
||||||
partial def translateMVarId (srcMVarId: MVarId) : MetaTranslateM MVarId := do
|
partial def translateMVarId (srcMVarId: MVarId) : MetaTranslateM MVarId := do
|
||||||
if let .some mvarId' := (← get).mvarMap[srcMVarId]? then
|
if let .some mvarId' := (← get).mvarMap[srcMVarId]? then
|
||||||
trace[Pantograph.Frontend.MetaTranslate] "Existing mvar id {srcMVarId.name} → {mvarId'.name}"
|
|
||||||
return mvarId'
|
return mvarId'
|
||||||
let mvarId' ← Meta.withLCtx .empty #[] do
|
let mvarId' ← Meta.withLCtx .empty #[] do
|
||||||
let srcDecl := (← getSourceMCtx).findDecl? srcMVarId |>.get!
|
let srcDecl := (← getSourceMCtx).findDecl? srcMVarId |>.get!
|
||||||
|
@ -135,7 +134,6 @@ partial def translateMVarId (srcMVarId: MVarId) : MetaTranslateM MVarId := do
|
||||||
let fvars' ← mvarIdPending'.withContext $ fvars.mapM translateExpr
|
let fvars' ← mvarIdPending'.withContext $ fvars.mapM translateExpr
|
||||||
assignDelayedMVar mvarId' fvars' mvarIdPending'
|
assignDelayedMVar mvarId' fvars' mvarIdPending'
|
||||||
pure mvarId'
|
pure mvarId'
|
||||||
trace[Pantograph.Frontend.MetaTranslate] "Translated {srcMVarId.name} → {mvarId'.name}"
|
|
||||||
addTranslatedMVar srcMVarId mvarId'
|
addTranslatedMVar srcMVarId mvarId'
|
||||||
return mvarId'
|
return mvarId'
|
||||||
end
|
end
|
||||||
|
@ -150,7 +148,6 @@ def translateMVarFromTermInfo (termInfo : Elab.TermInfo) (context? : Option Elab
|
||||||
let lctx' ← translateLCtx
|
let lctx' ← translateLCtx
|
||||||
let mvar ← Meta.withLCtx lctx' #[] do
|
let mvar ← Meta.withLCtx lctx' #[] do
|
||||||
let type' ← translateExpr type
|
let type' ← translateExpr type
|
||||||
trace[Pantograph.Frontend.MetaTranslate] "Translating from term info {← Meta.ppExpr type'}"
|
|
||||||
Meta.mkFreshExprSyntheticOpaqueMVar type'
|
Meta.mkFreshExprSyntheticOpaqueMVar type'
|
||||||
return mvar.mvarId!
|
return mvar.mvarId!
|
||||||
|
|
||||||
|
|
|
@ -10,72 +10,24 @@ import Lean
|
||||||
namespace Pantograph
|
namespace Pantograph
|
||||||
open Lean
|
open Lean
|
||||||
|
|
||||||
/-- The acting area of a tactic -/
|
|
||||||
inductive Site where
|
|
||||||
-- Dormant all other goals
|
|
||||||
| focus (goal : MVarId)
|
|
||||||
-- Move the goal to the first in the list
|
|
||||||
| prefer (goal : MVarId)
|
|
||||||
-- Execute as-is, no goals go dormant
|
|
||||||
| unfocus
|
|
||||||
deriving BEq, Inhabited
|
|
||||||
|
|
||||||
instance : Coe MVarId Site where
|
|
||||||
coe := .focus
|
|
||||||
instance : ToString Site where
|
|
||||||
toString
|
|
||||||
| .focus { name } => s!"[{name}]"
|
|
||||||
| .prefer { name } => s!"[{name},...]"
|
|
||||||
| .unfocus => "[*]"
|
|
||||||
|
|
||||||
/-- Executes a `TacticM` on a site and return affected goals -/
|
|
||||||
protected def Site.runTacticM (site : Site)
|
|
||||||
{ m } [Monad m] [MonadLiftT Elab.Tactic.TacticM m] [MonadControlT Elab.Tactic.TacticM m] [MonadMCtx m] [MonadError m]
|
|
||||||
(f : m α) : m (α × List MVarId) :=
|
|
||||||
match site with
|
|
||||||
| .focus goal => do
|
|
||||||
Elab.Tactic.setGoals [goal]
|
|
||||||
let a ← f
|
|
||||||
return (a, [goal])
|
|
||||||
| .prefer goal => do
|
|
||||||
let before ← Elab.Tactic.getUnsolvedGoals
|
|
||||||
let otherGoals := before.filter (· != goal)
|
|
||||||
Elab.Tactic.setGoals (goal :: otherGoals)
|
|
||||||
let a ← f
|
|
||||||
let after ← Elab.Tactic.getUnsolvedGoals
|
|
||||||
let parents := before.filter (¬ after.contains ·)
|
|
||||||
Elab.Tactic.pruneSolvedGoals
|
|
||||||
return (a, parents)
|
|
||||||
| .unfocus => do
|
|
||||||
let before ← Elab.Tactic.getUnsolvedGoals
|
|
||||||
let a ← f
|
|
||||||
let after ← Elab.Tactic.getUnsolvedGoals
|
|
||||||
let parents := before.filter (¬ after.contains ·)
|
|
||||||
Elab.Tactic.pruneSolvedGoals
|
|
||||||
return (a, parents)
|
|
||||||
|
|
||||||
/--
|
/--
|
||||||
Kernel view of the state of a proof
|
Represents an interconnected set of metavariables, or a state in proof search
|
||||||
-/
|
-/
|
||||||
structure GoalState where
|
structure GoalState where
|
||||||
-- Captured `TacticM` state
|
|
||||||
savedState : Elab.Tactic.SavedState
|
savedState : Elab.Tactic.SavedState
|
||||||
|
|
||||||
-- The root goal which is the search target
|
-- The root hole which is the search target
|
||||||
root: MVarId
|
root: MVarId
|
||||||
|
|
||||||
/--
|
-- Parent state metavariable source
|
||||||
Parent goals which became assigned or fragmented to produce this state.
|
parentMVar?: Option MVarId
|
||||||
Note that due to the existence of tactic fragments, parent goals do not
|
|
||||||
necessarily have an expression assignment.
|
|
||||||
-/
|
|
||||||
parentMVars : List MVarId := []
|
|
||||||
|
|
||||||
-- Any goal associated with a fragment has a partial tactic which has not
|
-- Existence of this field shows that we are currently in `conv` mode.
|
||||||
-- finished executing.
|
-- (convRhs, goal, dormant)
|
||||||
fragments : FragmentMap := .empty
|
convMVar?: Option (MVarId × MVarId × List MVarId) := .none
|
||||||
|
-- Previous RHS for calc, so we don't have to repeat it every time
|
||||||
def throwNoGoals { m α } [Monad m] [MonadError m] : m α := throwError "no goals to be solved"
|
-- WARNING: If using `state with` outside of `calc`, this must be set to `.none`
|
||||||
|
calcPrevRhs?: Option (MVarId × Expr) := .none
|
||||||
|
|
||||||
@[export pantograph_goal_state_create_m]
|
@[export pantograph_goal_state_create_m]
|
||||||
protected def GoalState.create (expr: Expr): Elab.TermElabM GoalState := do
|
protected def GoalState.create (expr: Expr): Elab.TermElabM GoalState := do
|
||||||
|
@ -90,6 +42,7 @@ protected def GoalState.create (expr: Expr): Elab.TermElabM GoalState := do
|
||||||
return {
|
return {
|
||||||
root := root.mvarId!,
|
root := root.mvarId!,
|
||||||
savedState,
|
savedState,
|
||||||
|
parentMVar? := .none,
|
||||||
}
|
}
|
||||||
@[export pantograph_goal_state_create_from_mvars_m]
|
@[export pantograph_goal_state_create_from_mvars_m]
|
||||||
protected def GoalState.createFromMVars (goals: List MVarId) (root: MVarId): MetaM GoalState := do
|
protected def GoalState.createFromMVars (goals: List MVarId) (root: MVarId): MetaM GoalState := do
|
||||||
|
@ -98,19 +51,13 @@ protected def GoalState.createFromMVars (goals: List MVarId) (root: MVarId): Met
|
||||||
return {
|
return {
|
||||||
root,
|
root,
|
||||||
savedState,
|
savedState,
|
||||||
|
parentMVar? := .none,
|
||||||
}
|
}
|
||||||
@[always_inline]
|
@[export pantograph_goal_state_is_conv]
|
||||||
|
protected def GoalState.isConv (state: GoalState): Bool :=
|
||||||
|
state.convMVar?.isSome
|
||||||
protected def GoalState.goals (state: GoalState): List MVarId :=
|
protected def GoalState.goals (state: GoalState): List MVarId :=
|
||||||
state.savedState.tactic.goals
|
state.savedState.tactic.goals
|
||||||
@[always_inline]
|
|
||||||
protected def GoalState.mainGoal? (state : GoalState) : Option MVarId :=
|
|
||||||
state.goals.head?
|
|
||||||
@[always_inline]
|
|
||||||
protected def GoalState.actingGoal? (state : GoalState) (site : Site) : Option MVarId := do
|
|
||||||
match site with
|
|
||||||
| .focus goal | .prefer goal => return goal
|
|
||||||
| .unfocus => state.mainGoal?
|
|
||||||
|
|
||||||
@[export pantograph_goal_state_goals]
|
@[export pantograph_goal_state_goals]
|
||||||
protected def GoalState.goalsArray (state: GoalState): Array MVarId := state.goals.toArray
|
protected def GoalState.goalsArray (state: GoalState): Array MVarId := state.goals.toArray
|
||||||
protected def GoalState.mctx (state: GoalState): MetavarContext :=
|
protected def GoalState.mctx (state: GoalState): MetavarContext :=
|
||||||
|
@ -122,10 +69,8 @@ protected def GoalState.env (state: GoalState): Environment :=
|
||||||
protected def GoalState.metaContextOfGoal (state: GoalState) (mvarId: MVarId): Option Meta.Context := do
|
protected def GoalState.metaContextOfGoal (state: GoalState) (mvarId: MVarId): Option Meta.Context := do
|
||||||
let mvarDecl ← state.mctx.findDecl? mvarId
|
let mvarDecl ← state.mctx.findDecl? mvarId
|
||||||
return { lctx := mvarDecl.lctx, localInstances := mvarDecl.localInstances }
|
return { lctx := mvarDecl.lctx, localInstances := mvarDecl.localInstances }
|
||||||
@[always_inline]
|
|
||||||
protected def GoalState.metaState (state: GoalState): Meta.State :=
|
protected def GoalState.metaState (state: GoalState): Meta.State :=
|
||||||
state.savedState.term.meta.meta
|
state.savedState.term.meta.meta
|
||||||
@[always_inline]
|
|
||||||
protected def GoalState.coreState (state: GoalState): Core.SavedState :=
|
protected def GoalState.coreState (state: GoalState): Core.SavedState :=
|
||||||
state.savedState.term.meta.core
|
state.savedState.term.meta.core
|
||||||
|
|
||||||
|
@ -133,19 +78,18 @@ protected def GoalState.withContext' (state: GoalState) (mvarId: MVarId) (m: Met
|
||||||
mvarId.withContext m |>.run' (← read) state.metaState
|
mvarId.withContext m |>.run' (← read) state.metaState
|
||||||
protected def GoalState.withContext { m } [MonadControlT MetaM m] [Monad m] (state: GoalState) (mvarId: MVarId) : m α → m α :=
|
protected def GoalState.withContext { m } [MonadControlT MetaM m] [Monad m] (state: GoalState) (mvarId: MVarId) : m α → m α :=
|
||||||
Meta.mapMetaM <| state.withContext' mvarId
|
Meta.mapMetaM <| state.withContext' mvarId
|
||||||
/-- Uses context of the first parent -/
|
|
||||||
protected def GoalState.withParentContext { n } [MonadControlT MetaM n] [Monad n] (state: GoalState): n α → n α :=
|
protected def GoalState.withParentContext { n } [MonadControlT MetaM n] [Monad n] (state: GoalState): n α → n α :=
|
||||||
Meta.mapMetaM <| state.withContext' state.parentMVars[0]!
|
Meta.mapMetaM <| state.withContext' state.parentMVar?.get!
|
||||||
protected def GoalState.withRootContext { n } [MonadControlT MetaM n] [Monad n] (state: GoalState): n α → n α :=
|
protected def GoalState.withRootContext { n } [MonadControlT MetaM n] [Monad n] (state: GoalState): n α → n α :=
|
||||||
Meta.mapMetaM <| state.withContext' state.root
|
Meta.mapMetaM <| state.withContext' state.root
|
||||||
|
|
||||||
private def restoreCoreMExtra (state : Core.SavedState) : CoreM Unit :=
|
private def GoalState.mvars (state: GoalState): SSet MVarId :=
|
||||||
let { nextMacroScope, ngen, .. } := state
|
state.mctx.decls.foldl (init := .empty) fun acc k _ => acc.insert k
|
||||||
modifyGetThe Core.State (fun st => ((),
|
|
||||||
{ st with nextMacroScope, ngen }))
|
|
||||||
-- Restore the name generator and macro scopes of the core state
|
-- Restore the name generator and macro scopes of the core state
|
||||||
protected def GoalState.restoreCoreMExtra (state: GoalState): CoreM Unit :=
|
protected def GoalState.restoreCoreMExtra (state: GoalState): CoreM Unit := do
|
||||||
restoreCoreMExtra state.coreState
|
let savedCore := state.coreState
|
||||||
|
modifyGetThe Core.State (fun st => ((),
|
||||||
|
{ st with nextMacroScope := savedCore.nextMacroScope, ngen := savedCore.ngen }))
|
||||||
protected def GoalState.restoreMetaM (state: GoalState): MetaM Unit := do
|
protected def GoalState.restoreMetaM (state: GoalState): MetaM Unit := do
|
||||||
state.restoreCoreMExtra
|
state.restoreCoreMExtra
|
||||||
state.savedState.term.meta.restore
|
state.savedState.term.meta.restore
|
||||||
|
@ -156,30 +100,57 @@ private def GoalState.restoreTacticM (state: GoalState) (goal: MVarId): Elab.Tac
|
||||||
state.restoreElabM
|
state.restoreElabM
|
||||||
Elab.Tactic.setGoals [goal]
|
Elab.Tactic.setGoals [goal]
|
||||||
|
|
||||||
/--
|
@[export pantograph_goal_state_focus]
|
||||||
Brings into scope a list of goals. User must ensure `goals` are distinct.
|
protected def GoalState.focus (state: GoalState) (goalId: Nat): Option GoalState := do
|
||||||
-/
|
let goal ← state.savedState.tactic.goals[goalId]?
|
||||||
@[export pantograph_goal_state_resume]
|
|
||||||
protected def GoalState.resume (state : GoalState) (goals : List MVarId) : Except String GoalState := do
|
|
||||||
if ¬ (goals.all (state.mctx.decls.contains ·)) then
|
|
||||||
let invalid_goals := goals.filter (λ goal => ¬ state.mctx.decls.contains goal) |>.map (·.name.toString)
|
|
||||||
.error s!"Goals {invalid_goals} are not in scope"
|
|
||||||
-- Set goals to the goals that have not been assigned yet, similar to the `focus` tactic.
|
|
||||||
let unassigned := goals.filter λ goal =>
|
|
||||||
let isSolved := state.mctx.eAssignment.contains goal || state.mctx.dAssignment.contains goal
|
|
||||||
¬ isSolved
|
|
||||||
return {
|
return {
|
||||||
state with
|
state with
|
||||||
savedState := {
|
savedState := {
|
||||||
term := state.savedState.term,
|
state.savedState with
|
||||||
tactic := { goals := unassigned },
|
tactic := { goals := [goal] },
|
||||||
|
},
|
||||||
|
calcPrevRhs? := .none,
|
||||||
|
}
|
||||||
|
|
||||||
|
/-- Immediately bring all parent goals back into scope. Used in automatic mode -/
|
||||||
|
@[export pantograph_goal_state_immediate_resume_parent]
|
||||||
|
protected def GoalState.immediateResume (state: GoalState) (parent: GoalState): GoalState :=
|
||||||
|
-- Prune parents solved goals
|
||||||
|
let mctx := state.mctx
|
||||||
|
let parentGoals := parent.goals.filter $ λ goal => mctx.eAssignment.contains goal
|
||||||
|
{
|
||||||
|
state with
|
||||||
|
savedState := {
|
||||||
|
state.savedState with
|
||||||
|
tactic := { goals := state.goals ++ parentGoals },
|
||||||
},
|
},
|
||||||
}
|
}
|
||||||
|
|
||||||
|
/--
|
||||||
|
Brings into scope a list of goals
|
||||||
|
-/
|
||||||
|
@[export pantograph_goal_state_resume]
|
||||||
|
protected def GoalState.resume (state: GoalState) (goals: List MVarId): Except String GoalState :=
|
||||||
|
if ¬ (goals.all (λ goal => state.mvars.contains goal)) then
|
||||||
|
let invalid_goals := goals.filter (λ goal => ¬ state.mvars.contains goal) |>.map (·.name.toString)
|
||||||
|
.error s!"Goals {invalid_goals} are not in scope"
|
||||||
|
else
|
||||||
|
-- Set goals to the goals that have not been assigned yet, similar to the `focus` tactic.
|
||||||
|
let unassigned := goals.filter (λ goal =>
|
||||||
|
let mctx := state.mctx
|
||||||
|
¬(mctx.eAssignment.contains goal || mctx.dAssignment.contains goal))
|
||||||
|
.ok {
|
||||||
|
state with
|
||||||
|
savedState := {
|
||||||
|
term := state.savedState.term,
|
||||||
|
tactic := { goals := unassigned },
|
||||||
|
},
|
||||||
|
}
|
||||||
/--
|
/--
|
||||||
Brings into scope all goals from `branch`
|
Brings into scope all goals from `branch`
|
||||||
-/
|
-/
|
||||||
@[export pantograph_goal_state_continue]
|
@[export pantograph_goal_state_continue]
|
||||||
protected def GoalState.continue (target : GoalState) (branch : GoalState) : Except String GoalState :=
|
protected def GoalState.continue (target: GoalState) (branch: GoalState): Except String GoalState :=
|
||||||
if !target.goals.isEmpty then
|
if !target.goals.isEmpty then
|
||||||
.error s!"Target state has unresolved goals"
|
.error s!"Target state has unresolved goals"
|
||||||
else if target.root != branch.root then
|
else if target.root != branch.root then
|
||||||
|
@ -188,239 +159,39 @@ protected def GoalState.continue (target : GoalState) (branch : GoalState) : Exc
|
||||||
target.resume (goals := branch.goals)
|
target.resume (goals := branch.goals)
|
||||||
|
|
||||||
@[export pantograph_goal_state_root_expr]
|
@[export pantograph_goal_state_root_expr]
|
||||||
protected def GoalState.rootExpr? (goalState : GoalState) : Option Expr := do
|
protected def GoalState.rootExpr? (goalState: GoalState): Option Expr := do
|
||||||
if goalState.root.name == .anonymous then
|
if goalState.root.name == .anonymous then
|
||||||
.none
|
.none
|
||||||
let expr ← goalState.mctx.eAssignment.find? goalState.root
|
let expr ← goalState.mctx.eAssignment.find? goalState.root
|
||||||
let (expr, _) := instantiateMVarsCore (mctx := goalState.mctx) (e := expr)
|
let (expr, _) := instantiateMVarsCore (mctx := goalState.mctx) (e := expr)
|
||||||
|
if expr.hasExprMVar then
|
||||||
|
-- Must not assert that the goal state is empty here. We could be in a branch goal.
|
||||||
|
--assert! ¬goalState.goals.isEmpty
|
||||||
|
.none
|
||||||
|
else
|
||||||
|
assert! goalState.goals.isEmpty
|
||||||
|
return expr
|
||||||
|
@[export pantograph_goal_state_parent_expr]
|
||||||
|
protected def GoalState.parentExpr? (goalState: GoalState): Option Expr := do
|
||||||
|
let parent ← goalState.parentMVar?
|
||||||
|
let expr := goalState.mctx.eAssignment.find! parent
|
||||||
|
let (expr, _) := instantiateMVarsCore (mctx := goalState.mctx) (e := expr)
|
||||||
return expr
|
return expr
|
||||||
@[export pantograph_goal_state_is_solved]
|
|
||||||
protected def GoalState.isSolved (goalState : GoalState) : Bool :=
|
|
||||||
let solvedRoot := match goalState.rootExpr? with
|
|
||||||
| .some e => ¬ e.hasExprMVar
|
|
||||||
| .none => true
|
|
||||||
goalState.goals.isEmpty && solvedRoot
|
|
||||||
@[export pantograph_goal_state_get_mvar_e_assignment]
|
@[export pantograph_goal_state_get_mvar_e_assignment]
|
||||||
protected def GoalState.getMVarEAssignment (goalState: GoalState) (mvarId: MVarId): Option Expr := do
|
protected def GoalState.getMVarEAssignment (goalState: GoalState) (mvarId: MVarId): Option Expr := do
|
||||||
let expr ← goalState.mctx.eAssignment.find? mvarId
|
let expr ← goalState.mctx.eAssignment.find? mvarId
|
||||||
let (expr, _) := instantiateMVarsCore (mctx := goalState.mctx) (e := expr)
|
let (expr, _) := instantiateMVarsCore (mctx := goalState.mctx) (e := expr)
|
||||||
return expr
|
return expr
|
||||||
@[export pantograph_goal_state_parent_exprs]
|
|
||||||
protected def GoalState.parentExprs (state : GoalState) : List (Except Fragment Expr) :=
|
|
||||||
state.parentMVars.map λ goal => match state.getMVarEAssignment goal with
|
|
||||||
| .some e => .ok e
|
|
||||||
-- A parent goal which is not assigned must have a fragment
|
|
||||||
| .none => .error state.fragments[goal]!
|
|
||||||
@[always_inline]
|
|
||||||
protected def GoalState.hasUniqueParent (state : GoalState) : Bool :=
|
|
||||||
state.parentMVars.length == 1
|
|
||||||
@[always_inline]
|
|
||||||
protected def GoalState.parentExpr! (state : GoalState) : Expr :=
|
|
||||||
assert! state.parentMVars.length == 1
|
|
||||||
(state.getMVarEAssignment state.parentMVars[0]!).get!
|
|
||||||
|
|
||||||
deriving instance BEq for DelayedMetavarAssignment
|
|
||||||
|
|
||||||
/-- Given states `dst`, `src`, and `src'`, where `dst` and `src'` are
|
|
||||||
descendants of `src`, replay the differential `src' - src` in `dst`. Colliding
|
|
||||||
metavariable and lemma names will be automatically renamed to ensure there is no
|
|
||||||
collision. This implements branch unification. Unification might be impossible
|
|
||||||
if conflicting assignments exist. We also assume the monotonicity property: In a
|
|
||||||
chain of descending goal states, a mvar cannot be unassigned, and once assigned
|
|
||||||
its assignment cannot change. -/
|
|
||||||
@[export pantograph_goal_state_replay_m]
|
|
||||||
protected def GoalState.replay (dst : GoalState) (src src' : GoalState) : CoreM GoalState :=
|
|
||||||
withTraceNode `Pantograph.GoalState.replay (fun _ => return m!"replay") do
|
|
||||||
let srcNGen := src.coreState.ngen
|
|
||||||
let srcNGen' := src'.coreState.ngen
|
|
||||||
let dstNGen := dst.coreState.ngen
|
|
||||||
assert! srcNGen.namePrefix == srcNGen'.namePrefix
|
|
||||||
assert! srcNGen.namePrefix == dstNGen.namePrefix
|
|
||||||
assert! src.mctx.depth == src'.mctx.depth
|
|
||||||
assert! src.mctx.depth == dst.mctx.depth
|
|
||||||
|
|
||||||
let diffNGenIdx := dst.coreState.ngen.idx - srcNGen.idx
|
|
||||||
|
|
||||||
let env ← dst.coreState.env.replayConsts src.env src'.env (skipExisting := true)
|
|
||||||
|
|
||||||
trace[Pantograph.GoalState.replay] "Merging ngen {srcNGen.idx} -> ({srcNGen'.idx}, {dstNGen.idx})"
|
|
||||||
-- True if the name is generated after `src`
|
|
||||||
let isNewName : Name → Bool
|
|
||||||
| .num pref n =>
|
|
||||||
pref == srcNGen.namePrefix ∧ n ≥ srcNGen.idx
|
|
||||||
| _ => false
|
|
||||||
let mapId : Name → Name
|
|
||||||
| id@(.num pref n) =>
|
|
||||||
if isNewName id then
|
|
||||||
.num pref (n + diffNGenIdx)
|
|
||||||
else
|
|
||||||
id
|
|
||||||
| id => id
|
|
||||||
let mapMVar : MVarId → MVarId
|
|
||||||
| { name } => ⟨mapId name⟩
|
|
||||||
let rec mapLevel : Level → Level
|
|
||||||
| .succ x => .succ (mapLevel x)
|
|
||||||
| .max l1 l2 => .max (mapLevel l1) (mapLevel l2)
|
|
||||||
| .imax l1 l2 => .imax (mapLevel l1) (mapLevel l2)
|
|
||||||
| .mvar { name } => .mvar ⟨mapId name⟩
|
|
||||||
| l => l
|
|
||||||
let mapExpr (e : Expr) : CoreM Expr := Core.transform e λ
|
|
||||||
| .sort level => pure $ .done $ .sort (mapLevel level)
|
|
||||||
| .mvar { name } => pure $ .done $ .mvar ⟨mapId name⟩
|
|
||||||
| _ => pure .continue
|
|
||||||
let mapDelayedAssignment (d : DelayedMetavarAssignment) : CoreM DelayedMetavarAssignment := do
|
|
||||||
let { mvarIdPending, fvars } := d
|
|
||||||
return {
|
|
||||||
mvarIdPending := mapMVar mvarIdPending,
|
|
||||||
fvars := ← fvars.mapM mapExpr,
|
|
||||||
}
|
|
||||||
let mapLocalDecl (ldecl : LocalDecl) : CoreM LocalDecl := do
|
|
||||||
let ldecl := ldecl.setType (← mapExpr ldecl.type)
|
|
||||||
if let .some value := ldecl.value? then
|
|
||||||
return ldecl.setValue (← mapExpr value)
|
|
||||||
else
|
|
||||||
return ldecl
|
|
||||||
|
|
||||||
let { term := savedTerm@{ meta := savedMeta@{ core, meta := meta@{ mctx, .. } }, .. }, .. } := dst.savedState
|
|
||||||
trace[Pantograph.GoalState.replay] "Merging mvars {src.mctx.mvarCounter} -> ({src'.mctx.mvarCounter}, {dst.mctx.mvarCounter})"
|
|
||||||
let mctx := {
|
|
||||||
mctx with
|
|
||||||
mvarCounter := mctx.mvarCounter + (src'.mctx.mvarCounter - src.mctx.mvarCounter),
|
|
||||||
lDepth := src'.mctx.lDepth.foldl (init := mctx.lDepth) λ acc lmvarId@{ name } depth =>
|
|
||||||
if src.mctx.lDepth.contains lmvarId then
|
|
||||||
acc
|
|
||||||
else
|
|
||||||
acc.insert ⟨mapId name⟩ depth
|
|
||||||
decls := ← src'.mctx.decls.foldlM (init := mctx.decls) λ acc _mvarId@{ name } decl => do
|
|
||||||
if decl.index < src.mctx.mvarCounter then
|
|
||||||
return acc
|
|
||||||
let mvarId := ⟨mapId name⟩
|
|
||||||
let decl := {
|
|
||||||
decl with
|
|
||||||
lctx := ← decl.lctx.foldlM (init := .empty) λ acc decl => do
|
|
||||||
let decl ← mapLocalDecl decl
|
|
||||||
return acc.addDecl decl,
|
|
||||||
type := ← mapExpr decl.type,
|
|
||||||
}
|
|
||||||
return acc.insert mvarId decl
|
|
||||||
|
|
||||||
-- Merge mvar assignments
|
|
||||||
userNames := src'.mctx.userNames.foldl (init := mctx.userNames) λ acc userName mvarId =>
|
|
||||||
if acc.contains userName then
|
|
||||||
acc
|
|
||||||
else
|
|
||||||
acc.insert userName mvarId,
|
|
||||||
lAssignment := src'.mctx.lAssignment.foldl (init := mctx.lAssignment) λ acc lmvarId' l =>
|
|
||||||
let lmvarId := ⟨mapId lmvarId'.name⟩
|
|
||||||
if mctx.lAssignment.contains lmvarId then
|
|
||||||
-- Skip the intersecting assignments for now
|
|
||||||
acc
|
|
||||||
else
|
|
||||||
let l := mapLevel l
|
|
||||||
acc.insert lmvarId l,
|
|
||||||
eAssignment := ← src'.mctx.eAssignment.foldlM (init := mctx.eAssignment) λ acc mvarId' e => do
|
|
||||||
let mvarId := ⟨mapId mvarId'.name⟩
|
|
||||||
if mctx.eAssignment.contains mvarId then
|
|
||||||
-- Skip the intersecting assignments for now
|
|
||||||
return acc
|
|
||||||
else
|
|
||||||
let e ← mapExpr e
|
|
||||||
return acc.insert mvarId e,
|
|
||||||
dAssignment := ← src'.mctx.dAssignment.foldlM (init := mctx.dAssignment) λ acc mvarId' d => do
|
|
||||||
let mvarId := ⟨mapId mvarId'.name⟩
|
|
||||||
if mctx.dAssignment.contains mvarId then
|
|
||||||
return acc
|
|
||||||
else
|
|
||||||
let d ← mapDelayedAssignment d
|
|
||||||
return acc.insert mvarId d
|
|
||||||
}
|
|
||||||
let ngen := {
|
|
||||||
core.ngen with
|
|
||||||
idx := core.ngen.idx + (srcNGen'.idx - srcNGen.idx)
|
|
||||||
}
|
|
||||||
-- Merge conflicting lmvar and mvar assignments using `isDefEq`
|
|
||||||
|
|
||||||
let savedMeta := {
|
|
||||||
savedMeta with
|
|
||||||
core := {
|
|
||||||
core with
|
|
||||||
ngen,
|
|
||||||
env,
|
|
||||||
-- Reset the message log when declaration uses `sorry`
|
|
||||||
messages := {}
|
|
||||||
}
|
|
||||||
meta := {
|
|
||||||
meta with
|
|
||||||
mctx,
|
|
||||||
}
|
|
||||||
}
|
|
||||||
let m : MetaM Meta.SavedState := Meta.withMCtx mctx do
|
|
||||||
restoreCoreMExtra savedMeta.core
|
|
||||||
savedMeta.restore
|
|
||||||
|
|
||||||
for (lmvarId, l') in src'.mctx.lAssignment do
|
|
||||||
if isNewName lmvarId.name then
|
|
||||||
continue
|
|
||||||
let .some l ← getLevelMVarAssignment? lmvarId | continue
|
|
||||||
let l' := mapLevel l'
|
|
||||||
trace[Pantograph.GoalState.replay] "Merging level assignments on {lmvarId.name}"
|
|
||||||
unless ← Meta.isLevelDefEq l l' do
|
|
||||||
throwError "Conflicting assignment of level metavariable {lmvarId.name}"
|
|
||||||
for (mvarId, e') in src'.mctx.eAssignment do
|
|
||||||
if isNewName mvarId.name then
|
|
||||||
continue
|
|
||||||
if ← mvarId.isDelayedAssigned then
|
|
||||||
throwError "Conflicting assignment of expr metavariable (e != d) {mvarId.name}"
|
|
||||||
let .some e ← getExprMVarAssignment? mvarId | continue
|
|
||||||
let e' ← mapExpr e'
|
|
||||||
trace[Pantograph.GoalState.replay] "Merging expr assignments on {mvarId.name}"
|
|
||||||
unless ← Meta.isDefEq e e' do
|
|
||||||
throwError "Conflicting assignment of expr metavariable (e != e) {mvarId.name}"
|
|
||||||
for (mvarId, d') in src'.mctx.dAssignment do
|
|
||||||
if isNewName mvarId.name then
|
|
||||||
continue
|
|
||||||
if ← mvarId.isAssigned then
|
|
||||||
throwError "Conflicting assignment of expr metavariable (d != e) {mvarId.name}"
|
|
||||||
let .some d ← getDelayedMVarAssignment? mvarId | continue
|
|
||||||
trace[Pantograph.GoalState.replay] "Merging expr (delayed) assignments on {mvarId.name}"
|
|
||||||
unless d == d' do
|
|
||||||
throwError "Conflicting assignment of expr metavariable (d != d) {mvarId.name}"
|
|
||||||
|
|
||||||
Meta.saveState
|
|
||||||
let goals := dst.savedState.tactic.goals ++
|
|
||||||
src'.savedState.tactic.goals.map (⟨mapId ·.name⟩)
|
|
||||||
let fragments ← src'.fragments.foldM (init := dst.fragments) λ acc mvarId' fragment' => do
|
|
||||||
let mvarId := ⟨mapId mvarId'.name⟩
|
|
||||||
let fragment ← fragment'.map mapExpr
|
|
||||||
if let .some _fragment0 := acc[mvarId]? then
|
|
||||||
throwError "Conflicting fragments on {mvarId.name}"
|
|
||||||
return acc.insert mvarId fragment
|
|
||||||
return {
|
|
||||||
dst with
|
|
||||||
savedState := {
|
|
||||||
tactic := {
|
|
||||||
goals
|
|
||||||
},
|
|
||||||
term := {
|
|
||||||
savedTerm with
|
|
||||||
meta := ← m.run',
|
|
||||||
},
|
|
||||||
},
|
|
||||||
parentMVars := dst.parentMVars ++ src.parentMVars.map mapMVar,
|
|
||||||
fragments,
|
|
||||||
}
|
|
||||||
|
|
||||||
--- Tactic execution functions ---
|
--- Tactic execution functions ---
|
||||||
|
|
||||||
/--
|
-- Mimics `Elab.Term.logUnassignedUsingErrorInfos`
|
||||||
These descendants serve as "seed" mvars. If a MVarError's mvar is related to one
|
|
||||||
of these seed mvars, it means an error has occurred when a tactic was executing
|
|
||||||
on `src`. `evalTactic`, will not capture these mvars, so we need to manually
|
|
||||||
find them and save them into the goal list. See the rationales document for the
|
|
||||||
inspiration of this function.
|
|
||||||
-/
|
|
||||||
private def collectAllErroredMVars (src : MVarId) : Elab.TermElabM (List MVarId) := do
|
private def collectAllErroredMVars (src : MVarId) : Elab.TermElabM (List MVarId) := do
|
||||||
-- Mimics `Elab.Term.logUnassignedUsingErrorInfos`
|
-- These descendants serve as "seed" mvars. If a MVarError's mvar is related
|
||||||
|
-- to one of these seed mvars, it means an error has occurred when a tactic
|
||||||
|
-- was executing on `src`. `evalTactic`, will not capture these mvars, so we
|
||||||
|
-- need to manually find them and save them into the goal list.
|
||||||
|
|
||||||
let descendants ← Meta.getMVars (.mvar src)
|
let descendants ← Meta.getMVars (.mvar src)
|
||||||
--let _ ← Elab.Term.logUnassignedUsingErrorInfos descendants
|
--let _ ← Elab.Term.logUnassignedUsingErrorInfos descendants
|
||||||
let mut alreadyVisited : MVarIdSet := {}
|
let mut alreadyVisited : MVarIdSet := {}
|
||||||
|
@ -435,7 +206,6 @@ private def collectAllErroredMVars (src : MVarId) : Elab.TermElabM (List MVarId)
|
||||||
result := mvarDeps.foldl (·.insert ·) result
|
result := mvarDeps.foldl (·.insert ·) result
|
||||||
return result.toList
|
return result.toList
|
||||||
|
|
||||||
/-- Merger of two unique lists -/
|
|
||||||
private def mergeMVarLists (li1 li2 : List MVarId) : List MVarId :=
|
private def mergeMVarLists (li1 li2 : List MVarId) : List MVarId :=
|
||||||
let li2' := li2.filter (¬ li1.contains ·)
|
let li2' := li2.filter (¬ li1.contains ·)
|
||||||
li1 ++ li2'
|
li1 ++ li2'
|
||||||
|
@ -445,125 +215,97 @@ Set `guardMVarErrors` to true to capture mvar errors. Lean will not
|
||||||
automatically collect mvars from text tactics (vide
|
automatically collect mvars from text tactics (vide
|
||||||
`test_tactic_failure_synthesize_placeholder`)
|
`test_tactic_failure_synthesize_placeholder`)
|
||||||
-/
|
-/
|
||||||
protected def GoalState.step' { α } (state : GoalState) (site : Site) (tacticM : Elab.Tactic.TacticM α) (guardMVarErrors : Bool := false)
|
protected def GoalState.step (state: GoalState) (goal: MVarId) (tacticM: Elab.Tactic.TacticM Unit) (guardMVarErrors : Bool := false)
|
||||||
: Elab.TermElabM (α × GoalState) := do
|
: Elab.TermElabM GoalState := do
|
||||||
let ((a, parentMVars), { goals }) ← site.runTacticM tacticM
|
unless (← getMCtx).decls.contains goal do
|
||||||
|>.run { elaborator := .anonymous }
|
throwError s!"Goal is not in context: {goal.name}"
|
||||||
|>.run state.savedState.tactic
|
goal.checkNotAssigned `GoalState.step
|
||||||
|
let (_, { goals }) ← tacticM { elaborator := .anonymous } |>.run { goals := [goal] }
|
||||||
let nextElabState ← MonadBacktrack.saveState
|
let nextElabState ← MonadBacktrack.saveState
|
||||||
--Elab.Term.synthesizeSyntheticMVarsNoPostponing
|
--Elab.Term.synthesizeSyntheticMVarsNoPostponing
|
||||||
|
|
||||||
let goals ← if guardMVarErrors then
|
let goals ← if guardMVarErrors then
|
||||||
parentMVars.foldlM (init := goals) λ goals parent => do
|
pure $ mergeMVarLists goals (← collectAllErroredMVars goal)
|
||||||
let errors ← collectAllErroredMVars parent
|
|
||||||
return mergeMVarLists goals errors
|
|
||||||
else
|
else
|
||||||
pure goals
|
pure goals
|
||||||
let state' := {
|
return {
|
||||||
state with
|
state with
|
||||||
savedState := { term := nextElabState, tactic := { goals }, },
|
savedState := { term := nextElabState, tactic := { goals }, },
|
||||||
parentMVars,
|
parentMVar? := .some goal,
|
||||||
|
calcPrevRhs? := .none,
|
||||||
}
|
}
|
||||||
return (a, state')
|
|
||||||
protected def GoalState.step (state : GoalState) (site : Site) (tacticM : Elab.Tactic.TacticM Unit) (guardMVarErrors : Bool := false)
|
|
||||||
: Elab.TermElabM GoalState :=
|
|
||||||
Prod.snd <$> GoalState.step' state site tacticM guardMVarErrors
|
|
||||||
|
|
||||||
/-- Response for executing a tactic -/
|
/-- Response for executing a tactic -/
|
||||||
inductive TacticResult where
|
inductive TacticResult where
|
||||||
-- Goes to next state
|
-- Goes to next state
|
||||||
| success (state : GoalState) (messages : Array Message)
|
| success (state: GoalState)
|
||||||
-- Tactic failed with messages
|
-- Tactic failed with messages
|
||||||
| failure (messages : Array Message)
|
| failure (messages: Array String)
|
||||||
-- Could not parse tactic
|
-- Could not parse tactic
|
||||||
| parseError (message : String)
|
| parseError (message: String)
|
||||||
-- The given action cannot be executed in the state
|
-- The given action cannot be executed in the state
|
||||||
| invalidAction (message : String)
|
| invalidAction (message: String)
|
||||||
|
|
||||||
private def dumpMessageLog (prevMessageLength : Nat := 0) : CoreM (Bool × List Message) := do
|
private def dumpMessageLog (prevMessageLength : Nat) : CoreM (Array String) := do
|
||||||
let newMessages := (← Core.getMessageLog).toList.drop prevMessageLength
|
let newMessages ← (← Core.getMessageLog).toList.drop prevMessageLength
|
||||||
let hasErrors := newMessages.any (·.severity == .error)
|
|>.filterMapM λ m => do
|
||||||
|
if m.severity == .error then
|
||||||
|
return .some $ ← m.toString
|
||||||
|
else
|
||||||
|
return .none
|
||||||
Core.resetMessageLog
|
Core.resetMessageLog
|
||||||
return (hasErrors, newMessages)
|
return newMessages.toArray
|
||||||
|
|
||||||
/-- Execute a `TermElabM` producing a goal state, capturing the error and turn it into a `TacticResult` -/
|
|
||||||
def withCapturingError (elabM : Elab.Term.TermElabM GoalState) : Elab.TermElabM TacticResult := do
|
|
||||||
let messageLog ← Core.getMessageLog
|
|
||||||
unless messageLog.toList.isEmpty do
|
|
||||||
IO.eprintln s!"{← messageLog.toList.mapM (·.toString)}"
|
|
||||||
assert! messageLog.toList.isEmpty
|
|
||||||
try
|
|
||||||
let state ← elabM
|
|
||||||
|
|
||||||
-- Check if error messages have been generated in the core.
|
|
||||||
let (hasError, newMessages) ← dumpMessageLog
|
|
||||||
if hasError then
|
|
||||||
return .failure newMessages.toArray
|
|
||||||
else
|
|
||||||
return .success state newMessages.toArray
|
|
||||||
catch exception =>
|
|
||||||
match exception with
|
|
||||||
| .internal _ =>
|
|
||||||
let (_, messages) ← dumpMessageLog
|
|
||||||
return .failure messages.toArray
|
|
||||||
| _ =>
|
|
||||||
let (_, messages) ← dumpMessageLog
|
|
||||||
let message := {
|
|
||||||
fileName := ← getFileName,
|
|
||||||
pos := ← getRefPosition,
|
|
||||||
data := exception.toMessageData,
|
|
||||||
}
|
|
||||||
return .failure (message :: messages).toArray
|
|
||||||
|
|
||||||
/-- Executes a `TacticM` monad on this `GoalState`, collecting the errors as necessary -/
|
/-- Executes a `TacticM` monad on this `GoalState`, collecting the errors as necessary -/
|
||||||
protected def GoalState.tryTacticM
|
protected def GoalState.tryTacticM
|
||||||
(state: GoalState) (site : Site)
|
(state: GoalState) (goal: MVarId) (tacticM: Elab.Tactic.TacticM Unit)
|
||||||
(tacticM: Elab.Tactic.TacticM Unit)
|
|
||||||
(guardMVarErrors : Bool := false)
|
(guardMVarErrors : Bool := false)
|
||||||
: Elab.TermElabM TacticResult := do
|
: Elab.TermElabM TacticResult := do
|
||||||
state.restoreElabM
|
let prevMessageLength := state.coreState.messages.toList.length
|
||||||
withCapturingError do
|
try
|
||||||
state.step site tacticM guardMVarErrors
|
let nextState ← state.step goal tacticM guardMVarErrors
|
||||||
|
|
||||||
|
-- Check if error messages have been generated in the core.
|
||||||
|
let newMessages ← dumpMessageLog prevMessageLength
|
||||||
|
if ¬ newMessages.isEmpty then
|
||||||
|
return .failure newMessages
|
||||||
|
return .success nextState
|
||||||
|
catch exception =>
|
||||||
|
match exception with
|
||||||
|
| .internal _ => return .failure $ ← dumpMessageLog prevMessageLength
|
||||||
|
| _ => 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]
|
@[export pantograph_goal_state_try_tactic_m]
|
||||||
protected def GoalState.tryTactic (state: GoalState) (site : Site) (tactic: String):
|
protected def GoalState.tryTactic (state: GoalState) (goal: MVarId) (tactic: String):
|
||||||
Elab.TermElabM TacticResult := do
|
Elab.TermElabM TacticResult := do
|
||||||
state.restoreElabM
|
state.restoreElabM
|
||||||
let .some goal := state.actingGoal? site | throwNoGoals
|
|
||||||
if let .some fragment := state.fragments[goal]? then
|
|
||||||
return ← withCapturingError do
|
|
||||||
let (fragments, state') ← state.step' site do
|
|
||||||
fragment.step goal tactic $ state.fragments.erase goal
|
|
||||||
return { state' with fragments }
|
|
||||||
-- Normal tactic without fragment
|
|
||||||
let tactic ← match Parser.runParserCategory
|
let tactic ← match Parser.runParserCategory
|
||||||
(env := ← getEnv)
|
(env := ← MonadEnv.getEnv)
|
||||||
(catName := `tactic)
|
(catName := if state.isConv then `conv else `tactic)
|
||||||
(input := tactic)
|
(input := tactic)
|
||||||
(fileName := ← getFileName) with
|
(fileName := ← getFileName) with
|
||||||
| .ok stx => pure $ stx
|
| .ok stx => pure $ stx
|
||||||
| .error error => return .parseError error
|
| .error error => return .parseError error
|
||||||
let tacticM := Elab.Tactic.evalTactic tactic
|
assert! ¬ (← goal.isAssigned)
|
||||||
withCapturingError do
|
state.tryTacticM goal (Elab.Tactic.evalTactic tactic) true
|
||||||
state.step site tacticM (guardMVarErrors := true)
|
|
||||||
|
|
||||||
-- Specialized Tactics
|
protected def GoalState.tryAssign (state: GoalState) (goal: MVarId) (expr: String):
|
||||||
|
Elab.TermElabM TacticResult := do
|
||||||
protected def GoalState.tryAssign (state : GoalState) (site : Site) (expr : String)
|
|
||||||
: Elab.TermElabM TacticResult := do
|
|
||||||
state.restoreElabM
|
state.restoreElabM
|
||||||
let expr ← match Parser.runParserCategory
|
let expr ← match Parser.runParserCategory
|
||||||
(env := ← getEnv)
|
(env := ← MonadEnv.getEnv)
|
||||||
(catName := `term)
|
(catName := `term)
|
||||||
(input := expr)
|
(input := expr)
|
||||||
(fileName := ← getFileName) with
|
(fileName := ← getFileName) with
|
||||||
| .ok syn => pure syn
|
| .ok syn => pure syn
|
||||||
| .error error => return .parseError error
|
| .error error => return .parseError error
|
||||||
state.tryTacticM site $ Tactic.evalAssign expr
|
state.tryTacticM goal $ Tactic.evalAssign expr
|
||||||
|
|
||||||
protected def GoalState.tryLet (state : GoalState) (site : Site) (binderName : String) (type : String)
|
-- Specialized Tactics
|
||||||
: Elab.TermElabM TacticResult := do
|
|
||||||
|
protected def GoalState.tryLet (state: GoalState) (goal: MVarId) (binderName: String) (type: String):
|
||||||
|
Elab.TermElabM TacticResult := do
|
||||||
state.restoreElabM
|
state.restoreElabM
|
||||||
let type ← match Parser.runParserCategory
|
let type ← match Parser.runParserCategory
|
||||||
(env := ← MonadEnv.getEnv)
|
(env := ← MonadEnv.getEnv)
|
||||||
|
@ -572,59 +314,150 @@ protected def GoalState.tryLet (state : GoalState) (site : Site) (binderName : S
|
||||||
(fileName := ← getFileName) with
|
(fileName := ← getFileName) with
|
||||||
| .ok syn => pure syn
|
| .ok syn => pure syn
|
||||||
| .error error => return .parseError error
|
| .error error => return .parseError error
|
||||||
state.tryTacticM site $ Tactic.evalLet binderName.toName type
|
state.tryTacticM goal $ Tactic.evalLet binderName.toName type
|
||||||
|
|
||||||
/-- Enter conv tactic mode -/
|
/-- Enter conv tactic mode -/
|
||||||
@[export pantograph_goal_state_conv_enter_m]
|
protected def GoalState.conv (state: GoalState) (goal: MVarId):
|
||||||
protected def GoalState.convEnter (state : GoalState) (site : Site) :
|
|
||||||
Elab.TermElabM TacticResult := do
|
Elab.TermElabM TacticResult := do
|
||||||
let .some goal := state.actingGoal? site | throwNoGoals
|
if state.convMVar?.isSome then
|
||||||
if let .some (.conv ..) := state.fragments[goal]? then
|
|
||||||
return .invalidAction "Already in conv state"
|
return .invalidAction "Already in conv state"
|
||||||
state.restoreElabM
|
goal.checkNotAssigned `GoalState.conv
|
||||||
withCapturingError do
|
let tacticM : Elab.Tactic.TacticM (Elab.Tactic.SavedState × MVarId) := do
|
||||||
let (fragments, state') ← state.step' site Fragment.enterConv
|
state.restoreTacticM goal
|
||||||
return {
|
|
||||||
state' with
|
|
||||||
fragments := fragments.fold (init := state'.fragments) λ acc goal fragment =>
|
|
||||||
acc.insert goal fragment
|
|
||||||
}
|
|
||||||
|
|
||||||
/-- Exit from a tactic fragment. -/
|
-- See Lean.Elab.Tactic.Conv.convTarget
|
||||||
@[export pantograph_goal_state_fragment_exit_m]
|
let convMVar ← Elab.Tactic.withMainContext do
|
||||||
protected def GoalState.fragmentExit (state : GoalState) (site : Site):
|
let (rhs, newGoal) ← Elab.Tactic.Conv.mkConvGoalFor (← Elab.Tactic.getMainTarget)
|
||||||
|
Elab.Tactic.replaceMainGoal [newGoal.mvarId!]
|
||||||
|
pure rhs.mvarId!
|
||||||
|
return (← MonadBacktrack.saveState, convMVar)
|
||||||
|
try
|
||||||
|
let (nextSavedState, convRhs) ← tacticM { elaborator := .anonymous } |>.run' state.savedState.tactic
|
||||||
|
-- Other goals are now dormant
|
||||||
|
let otherGoals := state.goals.filter $ λ g => g != goal
|
||||||
|
return .success {
|
||||||
|
root := state.root,
|
||||||
|
savedState := nextSavedState
|
||||||
|
parentMVar? := .some goal,
|
||||||
|
convMVar? := .some (convRhs, goal, otherGoals),
|
||||||
|
calcPrevRhs? := .none
|
||||||
|
}
|
||||||
|
catch exception =>
|
||||||
|
return .failure #[← exception.toMessageData.toString]
|
||||||
|
|
||||||
|
/-- Exit from `conv` mode. Resumes all goals before the mode starts and applys the conv -/
|
||||||
|
@[export pantograph_goal_state_conv_exit_m]
|
||||||
|
protected def GoalState.convExit (state: GoalState):
|
||||||
Elab.TermElabM TacticResult := do
|
Elab.TermElabM TacticResult := do
|
||||||
let .some goal := state.actingGoal? site | throwNoGoals
|
let (convRhs, convGoal, _) ← match state.convMVar? with
|
||||||
let .some fragment := state.fragments[goal]? |
|
| .some mvar => pure mvar
|
||||||
return .invalidAction "Goal does not have a fragment"
|
| .none => return .invalidAction "Not in conv state"
|
||||||
state.restoreElabM
|
let tacticM : Elab.Tactic.TacticM Elab.Tactic.SavedState:= do
|
||||||
withCapturingError do
|
-- Vide `Lean.Elab.Tactic.Conv.convert`
|
||||||
let (fragments, state') ← state.step' goal (fragment.exit goal state.fragments)
|
state.savedState.restore
|
||||||
return {
|
|
||||||
state' with
|
-- Close all existing goals with `refl`
|
||||||
fragments,
|
for mvarId in (← Elab.Tactic.getGoals) do
|
||||||
|
liftM <| mvarId.refl <|> mvarId.inferInstance <|> pure ()
|
||||||
|
Elab.Tactic.pruneSolvedGoals
|
||||||
|
unless (← Elab.Tactic.getGoals).isEmpty do
|
||||||
|
throwError "convert tactic failed, there are unsolved goals\n{Elab.goalsToMessageData (← Elab.Tactic.getGoals)}"
|
||||||
|
|
||||||
|
Elab.Tactic.setGoals [convGoal]
|
||||||
|
|
||||||
|
let targetNew ← instantiateMVars (.mvar convRhs)
|
||||||
|
let proof ← instantiateMVars (.mvar convGoal)
|
||||||
|
|
||||||
|
Elab.Tactic.liftMetaTactic1 fun mvarId => mvarId.replaceTargetEq targetNew proof
|
||||||
|
MonadBacktrack.saveState
|
||||||
|
try
|
||||||
|
let nextSavedState ← tacticM { elaborator := .anonymous } |>.run' state.savedState.tactic
|
||||||
|
return .success {
|
||||||
|
root := state.root,
|
||||||
|
savedState := nextSavedState
|
||||||
|
parentMVar? := .some convGoal,
|
||||||
|
convMVar? := .none
|
||||||
|
calcPrevRhs? := .none
|
||||||
}
|
}
|
||||||
|
catch exception =>
|
||||||
|
return .failure #[← exception.toMessageData.toString]
|
||||||
|
|
||||||
protected def GoalState.calcPrevRhsOf? (state : GoalState) (goal : MVarId) : Option Expr := do
|
protected def GoalState.calcPrevRhsOf? (state: GoalState) (goal: MVarId): Option Expr := do
|
||||||
let .some (.calc prevRhs?) := state.fragments[goal]? | .none
|
let (mvarId, rhs) ← state.calcPrevRhs?
|
||||||
prevRhs?
|
if mvarId == goal then
|
||||||
|
.some rhs
|
||||||
@[export pantograph_goal_state_calc_enter_m]
|
else
|
||||||
protected def GoalState.calcEnter (state : GoalState) (site : Site)
|
.none
|
||||||
: Elab.TermElabM TacticResult := do
|
@[export pantograph_goal_state_try_calc_m]
|
||||||
let .some goal := state.actingGoal? site | throwNoGoals
|
protected def GoalState.tryCalc (state: GoalState) (goal: MVarId) (pred: String):
|
||||||
if let .some _ := state.fragments[goal]? then
|
Elab.TermElabM TacticResult := do
|
||||||
return .invalidAction "Goal already has a fragment"
|
|
||||||
state.restoreElabM
|
state.restoreElabM
|
||||||
withCapturingError do
|
if state.convMVar?.isSome then
|
||||||
let fragment := Fragment.enterCalc
|
return .invalidAction "Cannot initiate `calc` while in `conv` state"
|
||||||
let fragments := state.fragments.insert goal fragment
|
let `(term|$pred) ← match Parser.runParserCategory
|
||||||
return {
|
(env := state.env)
|
||||||
state with
|
(catName := `term)
|
||||||
fragments,
|
(input := pred)
|
||||||
}
|
(fileName := ← getFileName) with
|
||||||
|
| .ok syn => pure syn
|
||||||
|
| .error error => return .parseError error
|
||||||
|
goal.checkNotAssigned `GoalState.tryCalc
|
||||||
|
let calcPrevRhs? := state.calcPrevRhsOf? goal
|
||||||
|
let decl ← goal.getDecl
|
||||||
|
let target ← instantiateMVars decl.type
|
||||||
|
let tag := decl.userName
|
||||||
|
try
|
||||||
|
goal.withContext do
|
||||||
|
|
||||||
initialize
|
let mut step ← Elab.Term.elabType <| ← do
|
||||||
registerTraceClass `Pantograph.GoalState.replay
|
if let some prevRhs := calcPrevRhs? then
|
||||||
|
Elab.Term.annotateFirstHoleWithType pred (← Meta.inferType prevRhs)
|
||||||
|
else
|
||||||
|
pure pred
|
||||||
|
|
||||||
|
let some (_, lhs, rhs) ← Elab.Term.getCalcRelation? step |
|
||||||
|
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}"}"
|
||||||
|
|
||||||
|
-- Creates a mvar to represent the proof that the calc tactic solves the
|
||||||
|
-- current branch
|
||||||
|
-- In the Lean `calc` tactic this is gobbled up by
|
||||||
|
-- `withCollectingNewGoalsFrom`
|
||||||
|
let mut proof ← Meta.mkFreshExprMVarAt (← getLCtx) (← Meta.getLocalInstances) step
|
||||||
|
(userName := tag ++ `calc)
|
||||||
|
let mvarBranch := proof.mvarId!
|
||||||
|
|
||||||
|
let mut proofType ← Meta.inferType proof
|
||||||
|
let mut remainder? := Option.none
|
||||||
|
|
||||||
|
-- The calc tactic either solves the main goal or leaves another relation.
|
||||||
|
-- Replace the main goal, and save the new goal if necessary
|
||||||
|
unless ← Meta.isDefEq proofType target do
|
||||||
|
let rec throwFailed :=
|
||||||
|
throwError "'calc' tactic failed, has type{indentExpr proofType}\nbut it is expected to have type{indentExpr target}"
|
||||||
|
let some (_, _, rhs) ← Elab.Term.getCalcRelation? proofType | throwFailed
|
||||||
|
let some (r, _, rhs') ← Elab.Term.getCalcRelation? target | throwFailed
|
||||||
|
let lastStep := mkApp2 r rhs rhs'
|
||||||
|
let lastStepGoal ← Meta.mkFreshExprSyntheticOpaqueMVar lastStep tag
|
||||||
|
(proof, proofType) ← Elab.Term.mkCalcTrans proof proofType lastStepGoal lastStep
|
||||||
|
unless ← Meta.isDefEq proofType target do throwFailed
|
||||||
|
remainder? := .some lastStepGoal.mvarId!
|
||||||
|
goal.assign proof
|
||||||
|
|
||||||
|
let goals := [ mvarBranch ] ++ remainder?.toList
|
||||||
|
let calcPrevRhs? := remainder?.map $ λ g => (g, rhs)
|
||||||
|
return .success {
|
||||||
|
root := state.root,
|
||||||
|
savedState := {
|
||||||
|
term := ← MonadBacktrack.saveState,
|
||||||
|
tactic := { goals },
|
||||||
|
},
|
||||||
|
parentMVar? := .some goal,
|
||||||
|
calcPrevRhs?
|
||||||
|
}
|
||||||
|
catch exception =>
|
||||||
|
return .failure #[← exception.toMessageData.toString]
|
||||||
|
|
||||||
end Pantograph
|
end Pantograph
|
||||||
|
|
|
@ -69,10 +69,9 @@ def createCoreContext (options: Array String): IO Core.Context := do
|
||||||
@[export pantograph_create_core_state]
|
@[export pantograph_create_core_state]
|
||||||
def createCoreState (imports: Array String): IO Core.State := do
|
def createCoreState (imports: Array String): IO Core.State := do
|
||||||
let env ← Lean.importModules
|
let env ← Lean.importModules
|
||||||
(imports := imports.map (λ str => { module := str.toName }))
|
(imports := imports.map (λ str => { module := str.toName, runtimeOnly := false }))
|
||||||
(opts := {})
|
(opts := {})
|
||||||
(trustLevel := 1)
|
(trustLevel := 1)
|
||||||
(loadExts := true)
|
|
||||||
return { env := env }
|
return { env := env }
|
||||||
|
|
||||||
@[export pantograph_parse_elab_type_m]
|
@[export pantograph_parse_elab_type_m]
|
||||||
|
@ -117,24 +116,21 @@ def goalStartExpr (expr: String) : Protocol.FallibleT Elab.TermElabM GoalState :
|
||||||
|
|
||||||
@[export pantograph_goal_serialize_m]
|
@[export pantograph_goal_serialize_m]
|
||||||
def goalSerialize (state: GoalState) (options: @&Protocol.Options): CoreM (Array Protocol.Goal) :=
|
def goalSerialize (state: GoalState) (options: @&Protocol.Options): CoreM (Array Protocol.Goal) :=
|
||||||
runMetaM <| state.serializeGoals options
|
runMetaM <| state.serializeGoals (parent := .none) options
|
||||||
|
|
||||||
@[export pantograph_goal_print_m]
|
@[export pantograph_goal_print_m]
|
||||||
def goalPrint (state: GoalState) (rootExpr: Bool) (parentExprs: Bool) (goals: Bool) (extraMVars : Array String) (options: @&Protocol.Options)
|
def goalPrint (state: GoalState) (rootExpr: Bool) (parentExpr: Bool) (goals: Bool) (extraMVars : Array String) (options: @&Protocol.Options)
|
||||||
: CoreM Protocol.GoalPrintResult := runMetaM do
|
: CoreM Protocol.GoalPrintResult := runMetaM do
|
||||||
state.restoreMetaM
|
state.restoreMetaM
|
||||||
|
|
||||||
let rootExpr? := state.rootExpr?
|
|
||||||
let root? ← if rootExpr then
|
let root? ← if rootExpr then
|
||||||
rootExpr?.mapM λ expr => state.withRootContext do
|
state.rootExpr?.mapM λ expr => state.withRootContext do
|
||||||
serializeExpression options (← instantiateAll expr)
|
serializeExpression options (← instantiateAll expr)
|
||||||
else
|
else
|
||||||
pure .none
|
pure .none
|
||||||
let parentExprs? ← if parentExprs then
|
let parent? ← if parentExpr then
|
||||||
.some <$> state.parentMVars.mapM λ parent => parent.withContext do
|
state.parentExpr?.mapM λ expr => state.withParentContext do
|
||||||
let val? := state.getMVarEAssignment parent
|
serializeExpression options (← instantiateAll expr)
|
||||||
val?.mapM λ val => do
|
|
||||||
serializeExpression options (← instantiateAll val)
|
|
||||||
else
|
else
|
||||||
pure .none
|
pure .none
|
||||||
let goals ← if goals then
|
let goals ← if goals then
|
||||||
|
@ -147,38 +143,34 @@ def goalPrint (state: GoalState) (rootExpr: Bool) (parentExprs: Bool) (goals: Bo
|
||||||
state.withContext mvarId do
|
state.withContext mvarId do
|
||||||
let .some expr ← getExprMVarAssignment? mvarId | return {}
|
let .some expr ← getExprMVarAssignment? mvarId | return {}
|
||||||
serializeExpression options (← instantiateAll expr)
|
serializeExpression options (← instantiateAll expr)
|
||||||
let env ← getEnv
|
|
||||||
return {
|
return {
|
||||||
root?,
|
root?,
|
||||||
parentExprs?,
|
parent?,
|
||||||
goals,
|
goals,
|
||||||
extraMVars,
|
extraMVars,
|
||||||
rootHasSorry := rootExpr?.map (·.hasSorry) |>.getD false,
|
|
||||||
rootHasUnsafe := rootExpr?.map (env.hasUnsafe ·) |>.getD false,
|
|
||||||
rootHasMVar := rootExpr?.map (·.hasExprMVar) |>.getD false,
|
|
||||||
}
|
}
|
||||||
|
|
||||||
@[export pantograph_goal_have_m]
|
@[export pantograph_goal_have_m]
|
||||||
protected def GoalState.tryHave (state: GoalState) (site : Site) (binderName: String) (type: String): Elab.TermElabM TacticResult := do
|
protected def GoalState.tryHave (state: GoalState) (goal: MVarId) (binderName: String) (type: String): Elab.TermElabM TacticResult := do
|
||||||
let type ← match (← parseTermM type) with
|
let type ← match (← parseTermM type) with
|
||||||
| .ok syn => pure syn
|
| .ok syn => pure syn
|
||||||
| .error error => return .parseError error
|
| .error error => return .parseError error
|
||||||
state.restoreElabM
|
state.restoreElabM
|
||||||
state.tryTacticM site $ Tactic.evalHave binderName.toName type
|
state.tryTacticM goal $ Tactic.evalHave binderName.toName type
|
||||||
@[export pantograph_goal_try_define_m]
|
@[export pantograph_goal_try_define_m]
|
||||||
protected def GoalState.tryDefine (state: GoalState) (site : Site) (binderName: String) (expr: String): Elab.TermElabM TacticResult := do
|
protected def GoalState.tryDefine (state: GoalState) (goal: MVarId) (binderName: String) (expr: String): Elab.TermElabM TacticResult := do
|
||||||
let expr ← match (← parseTermM expr) with
|
let expr ← match (← parseTermM expr) with
|
||||||
| .ok syn => pure syn
|
| .ok syn => pure syn
|
||||||
| .error error => return .parseError error
|
| .error error => return .parseError error
|
||||||
state.restoreElabM
|
state.restoreElabM
|
||||||
state.tryTacticM site $ Tactic.evalDefine binderName.toName expr
|
state.tryTacticM goal (Tactic.evalDefine binderName.toName expr)
|
||||||
@[export pantograph_goal_try_draft_m]
|
@[export pantograph_goal_try_draft_m]
|
||||||
protected def GoalState.tryDraft (state: GoalState) (site : Site) (expr: String): Elab.TermElabM TacticResult := do
|
protected def GoalState.tryDraft (state: GoalState) (goal: MVarId) (expr: String): Elab.TermElabM TacticResult := do
|
||||||
let expr ← match (← parseTermM expr) with
|
let expr ← match (← parseTermM expr) with
|
||||||
| .ok syn => pure syn
|
| .ok syn => pure syn
|
||||||
| .error error => return .parseError error
|
| .error error => return .parseError error
|
||||||
state.restoreElabM
|
state.restoreElabM
|
||||||
state.tryTacticM site $ Tactic.evalDraft expr
|
state.tryTacticM goal (Tactic.evalDraft expr)
|
||||||
|
|
||||||
-- Cancel the token after a timeout.
|
-- Cancel the token after a timeout.
|
||||||
@[export pantograph_run_cancel_token_with_timeout_m]
|
@[export pantograph_run_cancel_token_with_timeout_m]
|
||||||
|
|
|
@ -6,7 +6,6 @@ its field names to avoid confusion with error messages generated by the REPL.
|
||||||
-/
|
-/
|
||||||
import Lean.Data.Json
|
import Lean.Data.Json
|
||||||
import Lean.Data.Position
|
import Lean.Data.Position
|
||||||
import Lean.Message
|
|
||||||
|
|
||||||
namespace Pantograph.Protocol
|
namespace Pantograph.Protocol
|
||||||
|
|
||||||
|
@ -61,23 +60,20 @@ structure Variable where
|
||||||
type?: Option Expression := .none
|
type?: Option Expression := .none
|
||||||
value?: Option Expression := .none
|
value?: Option Expression := .none
|
||||||
deriving Lean.ToJson
|
deriving Lean.ToJson
|
||||||
inductive Fragment where
|
|
||||||
| tactic
|
|
||||||
| conv
|
|
||||||
| calc
|
|
||||||
deriving BEq, DecidableEq, Repr, Lean.ToJson
|
|
||||||
structure Goal where
|
structure Goal where
|
||||||
|
name: String := ""
|
||||||
/-- Name of the metavariable -/
|
/-- Name of the metavariable -/
|
||||||
name : String := ""
|
userName?: Option String := .none
|
||||||
/-- User-facing name -/
|
/-- Is the goal in conversion mode -/
|
||||||
userName? : Option String := .none
|
isConversion: Bool := false
|
||||||
fragment : Fragment := .tactic
|
|
||||||
/-- target expression type -/
|
/-- target expression type -/
|
||||||
target : Expression
|
target: Expression
|
||||||
/-- Variables -/
|
/-- Variables -/
|
||||||
vars : Array Variable := #[]
|
vars: Array Variable := #[]
|
||||||
deriving Lean.ToJson
|
deriving Lean.ToJson
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
--- Individual Commands and return types ---
|
--- Individual Commands and return types ---
|
||||||
|
|
||||||
structure Command where
|
structure Command where
|
||||||
|
@ -91,10 +87,12 @@ structure InteractionError where
|
||||||
deriving Lean.ToJson
|
deriving Lean.ToJson
|
||||||
|
|
||||||
def errorIndex (desc: String): InteractionError := { error := "index", desc }
|
def errorIndex (desc: String): InteractionError := { error := "index", desc }
|
||||||
def errorOperation (desc: String): InteractionError := { error := "operation", desc }
|
|
||||||
def errorExpr (desc: String): InteractionError := { error := "expr", desc }
|
def errorExpr (desc: String): InteractionError := { error := "expr", desc }
|
||||||
|
|
||||||
|
|
||||||
|
--- Individual command and return types ---
|
||||||
|
|
||||||
|
|
||||||
structure Reset where
|
structure Reset where
|
||||||
deriving Lean.FromJson
|
deriving Lean.FromJson
|
||||||
structure Stat where
|
structure Stat where
|
||||||
|
@ -250,17 +248,17 @@ structure GoalStartResult where
|
||||||
root: String
|
root: String
|
||||||
deriving Lean.ToJson
|
deriving Lean.ToJson
|
||||||
structure GoalTactic where
|
structure GoalTactic where
|
||||||
|
-- Identifiers for tree, state, and goal
|
||||||
stateId: Nat
|
stateId: Nat
|
||||||
-- If omitted, act on the first goal
|
goalId: Nat := 0
|
||||||
goalId?: Option Nat := .none
|
|
||||||
-- If set to true, goal will not go dormant. Defaults to `automaticMode`
|
|
||||||
autoResume?: Option Bool := .none
|
|
||||||
-- One of the fields here must be filled
|
-- One of the fields here must be filled
|
||||||
tactic?: Option String := .none
|
tactic?: Option String := .none
|
||||||
mode?: Option String := .none -- Changes the current category to {"tactic", "calc", "conv"}
|
|
||||||
expr?: Option String := .none
|
expr?: Option String := .none
|
||||||
have?: Option String := .none
|
have?: Option String := .none
|
||||||
let?: Option String := .none
|
let?: Option String := .none
|
||||||
|
calc?: Option String := .none
|
||||||
|
-- true to enter `conv`, `false` to exit. In case of exit the `goalId` is ignored.
|
||||||
|
conv?: Option Bool := .none
|
||||||
draft?: Option String := .none
|
draft?: Option String := .none
|
||||||
|
|
||||||
-- In case of the `have` tactic, the new free variable name is provided here
|
-- In case of the `have` tactic, the new free variable name is provided here
|
||||||
|
@ -270,17 +268,14 @@ structure GoalTactic where
|
||||||
structure GoalTacticResult where
|
structure GoalTacticResult where
|
||||||
-- The next goal state id. Existence of this field shows success
|
-- The next goal state id. Existence of this field shows success
|
||||||
nextStateId?: Option Nat := .none
|
nextStateId?: Option Nat := .none
|
||||||
-- If the array is empty, it shows the goals have been fully resolved. If this
|
-- If the array is empty, it shows the goals have been fully resolved.
|
||||||
-- is .none, there has been a tactic error.
|
goals?: Option (Array Goal) := .none
|
||||||
goals?: Option (Array Goal) := .none
|
|
||||||
|
|
||||||
messages? : Option (Array Lean.SerialMessage) := .some #[]
|
-- Existence of this field shows tactic execution failure
|
||||||
|
tacticErrors?: Option (Array String) := .none
|
||||||
|
|
||||||
-- Existence of this field shows the tactic parsing has failed
|
-- Existence of this field shows the tactic parsing has failed
|
||||||
parseError? : Option String := .none
|
parseError?: Option String := .none
|
||||||
|
|
||||||
hasSorry : Bool := false
|
|
||||||
hasUnsafe : Bool := false
|
|
||||||
deriving Lean.ToJson
|
deriving Lean.ToJson
|
||||||
structure GoalContinue where
|
structure GoalContinue where
|
||||||
-- State from which the continuation acquires the context
|
-- State from which the continuation acquires the context
|
||||||
|
@ -310,8 +305,8 @@ structure GoalPrint where
|
||||||
|
|
||||||
-- Print root?
|
-- Print root?
|
||||||
rootExpr?: Option Bool := .some False
|
rootExpr?: Option Bool := .some False
|
||||||
-- Print the parent expressions
|
-- Print the parent expr?
|
||||||
parentExprs?: Option Bool := .some False
|
parentExpr?: Option Bool := .some False
|
||||||
-- Print goals?
|
-- Print goals?
|
||||||
goals?: Option Bool := .some False
|
goals?: Option Bool := .some False
|
||||||
-- Print values of extra mvars?
|
-- Print values of extra mvars?
|
||||||
|
@ -321,13 +316,9 @@ structure GoalPrintResult where
|
||||||
-- The root expression
|
-- The root expression
|
||||||
root?: Option Expression := .none
|
root?: Option Expression := .none
|
||||||
-- The filling expression of the parent goal
|
-- The filling expression of the parent goal
|
||||||
parentExprs?: Option (List (Option Expression)) := .none
|
parent?: Option Expression := .none
|
||||||
goals: Array Goal := #[]
|
goals: Array Goal := #[]
|
||||||
extraMVars: Array Expression := #[]
|
extraMVars: Array Expression := #[]
|
||||||
|
|
||||||
rootHasSorry : Bool := false
|
|
||||||
rootHasUnsafe : Bool := false
|
|
||||||
rootHasMVar : Bool := true
|
|
||||||
deriving Lean.ToJson
|
deriving Lean.ToJson
|
||||||
|
|
||||||
-- Diagnostic Options, not available in REPL
|
-- Diagnostic Options, not available in REPL
|
||||||
|
@ -363,8 +354,8 @@ structure FrontendProcess where
|
||||||
readHeader : Bool := false
|
readHeader : Bool := false
|
||||||
-- Alter the REPL environment after the compilation units.
|
-- Alter the REPL environment after the compilation units.
|
||||||
inheritEnv : Bool := false
|
inheritEnv : Bool := false
|
||||||
-- collect tactic invocations and output to a given file
|
-- collect tactic invocations
|
||||||
invocations?: Option String := .none
|
invocations: Bool := false
|
||||||
-- collect `sorry`s
|
-- collect `sorry`s
|
||||||
sorrys: Bool := false
|
sorrys: Bool := false
|
||||||
-- collect type errors
|
-- collect type errors
|
||||||
|
@ -384,9 +375,9 @@ 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 Lean.SerialMessage := #[]
|
messages: Array String := #[]
|
||||||
-- Number of tactic invocations
|
-- Tactic invocations
|
||||||
nInvocations?: Option Nat := .none
|
invocations?: Option (List InvokedTactic) := .none
|
||||||
goalStateId?: Option Nat := .none
|
goalStateId?: Option Nat := .none
|
||||||
goals?: Option (Array Goal) := .none
|
goals?: Option (Array Goal) := .none
|
||||||
-- Code segments which generated the goals
|
-- Code segments which generated the goals
|
||||||
|
@ -398,12 +389,6 @@ structure CompilationUnit where
|
||||||
structure FrontendProcessResult where
|
structure FrontendProcessResult where
|
||||||
units: List CompilationUnit
|
units: List CompilationUnit
|
||||||
deriving Lean.ToJson
|
deriving Lean.ToJson
|
||||||
structure FrontendDataUnit where
|
|
||||||
invocations? : Option (List Protocol.InvokedTactic) := .none
|
|
||||||
deriving Lean.ToJson
|
|
||||||
structure FrontendData where
|
|
||||||
units : List FrontendDataUnit
|
|
||||||
deriving Lean.ToJson
|
|
||||||
|
|
||||||
abbrev FallibleT := ExceptT InteractionError
|
abbrev FallibleT := ExceptT InteractionError
|
||||||
|
|
||||||
|
|
|
@ -1,19 +1,19 @@
|
||||||
import Pantograph.Goal
|
|
||||||
|
|
||||||
import Lean.Environment
|
import Lean.Environment
|
||||||
import Lean.Replay
|
import Lean.Replay
|
||||||
|
import Init.System.IOError
|
||||||
import Std.Data.HashMap
|
import Std.Data.HashMap
|
||||||
|
import Pantograph.Goal
|
||||||
open Lean
|
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
Input/Output functions borrowed from REPL
|
Input/Output functions
|
||||||
|
|
||||||
# Pickling and unpickling objects
|
# Pickling and unpickling objects
|
||||||
|
|
||||||
By abusing `saveModuleData` and `readModuleData` we can pickle and unpickle objects to disk.
|
By abusing `saveModuleData` and `readModuleData` we can pickle and unpickle objects to disk.
|
||||||
-/
|
-/
|
||||||
|
|
||||||
|
open Lean
|
||||||
|
|
||||||
namespace Pantograph
|
namespace Pantograph
|
||||||
|
|
||||||
/--
|
/--
|
||||||
|
@ -46,21 +46,6 @@ unsafe def withUnpickle [Monad m] [MonadLiftT IO m] {α β : Type}
|
||||||
region.free
|
region.free
|
||||||
pure r
|
pure r
|
||||||
|
|
||||||
abbrev ConstArray := Array (Name × ConstantInfo)
|
|
||||||
abbrev DistilledEnvironment := Array Import × ConstArray
|
|
||||||
|
|
||||||
/-- Boil an environment down to minimal components -/
|
|
||||||
def distillEnvironment (env : Environment) (background? : Option Environment := .none)
|
|
||||||
: DistilledEnvironment :=
|
|
||||||
let filter : Name → Bool := match background? with
|
|
||||||
| .some env => (¬ env.contains ·)
|
|
||||||
| .none => λ _ => true
|
|
||||||
let constants : ConstArray := env.constants.map₂.foldl (init := #[]) λ acc name info =>
|
|
||||||
if filter name then
|
|
||||||
acc.push (name, info)
|
|
||||||
else
|
|
||||||
acc
|
|
||||||
(env.header.imports, constants)
|
|
||||||
/--
|
/--
|
||||||
Pickle an `Environment` to disk.
|
Pickle an `Environment` to disk.
|
||||||
|
|
||||||
|
@ -71,23 +56,15 @@ and when unpickling, we build a fresh `Environment` from the imports,
|
||||||
and then add the new constants.
|
and then add the new constants.
|
||||||
-/
|
-/
|
||||||
@[export pantograph_env_pickle_m]
|
@[export pantograph_env_pickle_m]
|
||||||
def environmentPickle (env : Environment) (path : System.FilePath) (background? : Option Environment := .none)
|
def environmentPickle (env : Environment) (path : System.FilePath) : IO Unit :=
|
||||||
: IO Unit :=
|
Pantograph.pickle path (env.header.imports, env.constants.map₂)
|
||||||
pickle path $ distillEnvironment env background?
|
|
||||||
|
|
||||||
deriving instance BEq for Import
|
|
||||||
|
|
||||||
def resurrectEnvironment
|
def resurrectEnvironment
|
||||||
(distilled : DistilledEnvironment)
|
(imports : Array Import)
|
||||||
(background? : Option Environment := .none)
|
(map₂ : PHashMap Name ConstantInfo)
|
||||||
: IO Environment := do
|
: IO Environment := do
|
||||||
let (imports, constArray) := distilled
|
let env ← importModules imports {} 0
|
||||||
let env ← match background? with
|
env.replay (Std.HashMap.ofList map₂.toList)
|
||||||
| .none => importModules imports {} 0 (loadExts := true)
|
|
||||||
| .some env =>
|
|
||||||
assert! env.imports == imports
|
|
||||||
pure env
|
|
||||||
env.replay (Std.HashMap.ofList constArray.toList)
|
|
||||||
/--
|
/--
|
||||||
Unpickle an `Environment` from disk.
|
Unpickle an `Environment` from disk.
|
||||||
|
|
||||||
|
@ -95,10 +72,9 @@ We construct a fresh `Environment` with the relevant imports,
|
||||||
and then replace the new constants.
|
and then replace the new constants.
|
||||||
-/
|
-/
|
||||||
@[export pantograph_env_unpickle_m]
|
@[export pantograph_env_unpickle_m]
|
||||||
def environmentUnpickle (path : System.FilePath) (background? : Option Environment := .none)
|
def environmentUnpickle (path : System.FilePath) : IO (Environment × CompactedRegion) := unsafe do
|
||||||
: IO (Environment × CompactedRegion) := unsafe do
|
let ((imports, map₂), region) ← Pantograph.unpickle (Array Import × PHashMap Name ConstantInfo) path
|
||||||
let (distilled, region) ← unpickle (Array Import × ConstArray) path
|
return (← resurrectEnvironment imports map₂, region)
|
||||||
return (← resurrectEnvironment distilled background?, region)
|
|
||||||
|
|
||||||
|
|
||||||
open Lean.Core in
|
open Lean.Core in
|
||||||
|
@ -112,8 +88,7 @@ structure CompactCoreState where
|
||||||
-- infoState : Elab.InfoState := {}
|
-- infoState : Elab.InfoState := {}
|
||||||
|
|
||||||
@[export pantograph_goal_state_pickle_m]
|
@[export pantograph_goal_state_pickle_m]
|
||||||
def goalStatePickle (goalState : GoalState) (path : System.FilePath) (background? : Option Environment := .none)
|
def goalStatePickle (goalState : GoalState) (path : System.FilePath) : IO Unit :=
|
||||||
: IO Unit :=
|
|
||||||
let {
|
let {
|
||||||
savedState := {
|
savedState := {
|
||||||
term := {
|
term := {
|
||||||
|
@ -128,11 +103,12 @@ def goalStatePickle (goalState : GoalState) (path : System.FilePath) (background
|
||||||
tactic
|
tactic
|
||||||
}
|
}
|
||||||
root,
|
root,
|
||||||
parentMVars,
|
parentMVar?,
|
||||||
fragments,
|
convMVar?,
|
||||||
|
calcPrevRhs?,
|
||||||
} := goalState
|
} := goalState
|
||||||
pickle path (
|
Pantograph.pickle path (
|
||||||
distillEnvironment env background?,
|
env.constants.map₂,
|
||||||
|
|
||||||
({ nextMacroScope, ngen } : CompactCoreState),
|
({ nextMacroScope, ngen } : CompactCoreState),
|
||||||
meta,
|
meta,
|
||||||
|
@ -140,15 +116,16 @@ def goalStatePickle (goalState : GoalState) (path : System.FilePath) (background
|
||||||
tactic,
|
tactic,
|
||||||
|
|
||||||
root,
|
root,
|
||||||
parentMVars,
|
parentMVar?,
|
||||||
fragments,
|
convMVar?,
|
||||||
|
calcPrevRhs?,
|
||||||
)
|
)
|
||||||
|
|
||||||
@[export pantograph_goal_state_unpickle_m]
|
@[export pantograph_goal_state_unpickle_m]
|
||||||
def goalStateUnpickle (path : System.FilePath) (background? : Option Environment := .none)
|
def goalStateUnpickle (path : System.FilePath) (env : Environment)
|
||||||
: IO (GoalState × CompactedRegion) := unsafe do
|
: IO (GoalState × CompactedRegion) := unsafe do
|
||||||
let ((
|
let ((
|
||||||
distilledEnv,
|
map₂,
|
||||||
|
|
||||||
compactCore,
|
compactCore,
|
||||||
meta,
|
meta,
|
||||||
|
@ -156,10 +133,11 @@ def goalStateUnpickle (path : System.FilePath) (background? : Option Environment
|
||||||
tactic,
|
tactic,
|
||||||
|
|
||||||
root,
|
root,
|
||||||
parentMVars,
|
parentMVar?,
|
||||||
fragments,
|
convMVar?,
|
||||||
|
calcPrevRhs?,
|
||||||
), region) ← Pantograph.unpickle (
|
), region) ← Pantograph.unpickle (
|
||||||
DistilledEnvironment ×
|
PHashMap Name ConstantInfo ×
|
||||||
|
|
||||||
CompactCoreState ×
|
CompactCoreState ×
|
||||||
Meta.State ×
|
Meta.State ×
|
||||||
|
@ -167,10 +145,11 @@ def goalStateUnpickle (path : System.FilePath) (background? : Option Environment
|
||||||
Elab.Tactic.State ×
|
Elab.Tactic.State ×
|
||||||
|
|
||||||
MVarId ×
|
MVarId ×
|
||||||
List MVarId ×
|
Option MVarId ×
|
||||||
FragmentMap
|
Option (MVarId × MVarId × List MVarId) ×
|
||||||
|
Option (MVarId × Expr)
|
||||||
) path
|
) path
|
||||||
let env ← resurrectEnvironment distilledEnv background?
|
let env ← env.replay (Std.HashMap.ofList map₂.toList)
|
||||||
let goalState := {
|
let goalState := {
|
||||||
savedState := {
|
savedState := {
|
||||||
term := {
|
term := {
|
||||||
|
@ -187,8 +166,9 @@ def goalStateUnpickle (path : System.FilePath) (background? : Option Environment
|
||||||
tactic,
|
tactic,
|
||||||
},
|
},
|
||||||
root,
|
root,
|
||||||
parentMVars,
|
parentMVar?,
|
||||||
fragments,
|
convMVar?,
|
||||||
|
calcPrevRhs?,
|
||||||
}
|
}
|
||||||
return (goalState, region)
|
return (goalState, region)
|
||||||
|
|
||||||
|
|
|
@ -1,3 +1,2 @@
|
||||||
import Pantograph.Tactic.Assign
|
import Pantograph.Tactic.Assign
|
||||||
import Pantograph.Tactic.Fragment
|
|
||||||
import Pantograph.Tactic.Prograde
|
import Pantograph.Tactic.Prograde
|
||||||
|
|
|
@ -1,189 +0,0 @@
|
||||||
/- Fragmented tactics are the tactics which can give incremental feedback and
|
|
||||||
whose integrity as a block is crucial to its operation. e.g. `calc` or `conv`.
|
|
||||||
Here, a unified system handles all fragments.
|
|
||||||
|
|
||||||
Inside a tactic fragment, the parser category may be different. An incomplete
|
|
||||||
fragmented tactic may not be elaboratable..
|
|
||||||
|
|
||||||
In line with continuation/resumption paradigms, the exit function of a fragment
|
|
||||||
tactic is responsible for resuming incomplete goals with fragments. For example,
|
|
||||||
when a conversion tactic finishes, the sentinels should resume the root of the
|
|
||||||
conversion tactic goal. The user cannot be expected to execute this resumption,
|
|
||||||
since the root is automatically dormanted at the entry of the conversion tactic
|
|
||||||
mode.
|
|
||||||
-/
|
|
||||||
import Lean.Meta
|
|
||||||
import Lean.Elab
|
|
||||||
|
|
||||||
open Lean
|
|
||||||
|
|
||||||
namespace Pantograph
|
|
||||||
|
|
||||||
inductive Fragment where
|
|
||||||
| calc (prevRhs? : Option Expr)
|
|
||||||
| conv (rhs : MVarId)
|
|
||||||
-- This goal is spawned from a `conv`
|
|
||||||
| convSentinel (parent : MVarId)
|
|
||||||
deriving BEq, Inhabited
|
|
||||||
|
|
||||||
abbrev FragmentMap := Std.HashMap MVarId Fragment
|
|
||||||
def FragmentMap.empty : FragmentMap := Std.HashMap.emptyWithCapacity 2
|
|
||||||
protected def FragmentMap.filter (map : FragmentMap) (pred : MVarId → Fragment → Bool) : FragmentMap :=
|
|
||||||
map.fold (init := FragmentMap.empty) λ acc mvarId fragment =>
|
|
||||||
if pred mvarId fragment then
|
|
||||||
acc.insert mvarId fragment
|
|
||||||
else
|
|
||||||
acc
|
|
||||||
|
|
||||||
protected def Fragment.map (fragment : Fragment) (mapExpr : Expr → CoreM Expr) : CoreM Fragment :=
|
|
||||||
let mapMVar (mvarId : MVarId) : CoreM MVarId :=
|
|
||||||
return (← mapExpr (.mvar mvarId)) |>.mvarId!
|
|
||||||
match fragment with
|
|
||||||
| .calc prevRhs? => return .calc (← prevRhs?.mapM mapExpr)
|
|
||||||
| .conv rhs => do
|
|
||||||
let rhs' ← mapMVar rhs
|
|
||||||
return .conv rhs'
|
|
||||||
| .convSentinel parent => do
|
|
||||||
return .convSentinel (← mapMVar parent)
|
|
||||||
|
|
||||||
protected def Fragment.enterCalc : Fragment := .calc .none
|
|
||||||
protected def Fragment.enterConv : Elab.Tactic.TacticM FragmentMap := do
|
|
||||||
let goal ← Elab.Tactic.getMainGoal
|
|
||||||
goal.checkNotAssigned `GoalState.conv
|
|
||||||
let (rhs, newGoal) ← goal.withContext do
|
|
||||||
let target ← instantiateMVars (← goal.getType)
|
|
||||||
let (rhs, newGoal) ← Elab.Tactic.Conv.mkConvGoalFor target
|
|
||||||
pure (rhs.mvarId!, newGoal.mvarId!)
|
|
||||||
Elab.Tactic.replaceMainGoal [newGoal]
|
|
||||||
return FragmentMap.empty
|
|
||||||
|>.insert goal (.conv rhs)
|
|
||||||
|>.insert newGoal (.convSentinel goal)
|
|
||||||
|
|
||||||
protected partial def Fragment.exit (fragment : Fragment) (goal : MVarId) (fragments : FragmentMap)
|
|
||||||
: Elab.Tactic.TacticM FragmentMap :=
|
|
||||||
match fragment with
|
|
||||||
| .calc .. => do
|
|
||||||
Elab.Tactic.setGoals [goal]
|
|
||||||
return fragments.erase goal
|
|
||||||
| .conv rhs => do
|
|
||||||
let goals := (← Elab.Tactic.getGoals).filter λ descendant =>
|
|
||||||
match fragments[descendant]? with
|
|
||||||
| .some s => (.convSentinel goal) == s
|
|
||||||
| _ => false -- Not a conv goal from this
|
|
||||||
-- Close all existing goals with `refl`
|
|
||||||
for mvarId in goals do
|
|
||||||
liftM <| mvarId.refl <|> mvarId.inferInstance <|> pure ()
|
|
||||||
unless (← goals.filterM (·.isAssignedOrDelayedAssigned)).isEmpty do
|
|
||||||
throwError "convert tactic failed, there are unsolved goals\n{Elab.goalsToMessageData (goals)}"
|
|
||||||
|
|
||||||
-- Ensure the meta tactic runs on `goal` even if its dormant by forcing resumption
|
|
||||||
Elab.Tactic.setGoals $ goal :: (← Elab.Tactic.getGoals)
|
|
||||||
let targetNew ← instantiateMVars (.mvar rhs)
|
|
||||||
let proof ← instantiateMVars (.mvar goal)
|
|
||||||
|
|
||||||
Elab.Tactic.liftMetaTactic1 (·.replaceTargetEq targetNew proof)
|
|
||||||
|
|
||||||
-- Try to solve maiinline by rfl
|
|
||||||
let mvarId ← Elab.Tactic.getMainGoal
|
|
||||||
liftM <| mvarId.refl <|> mvarId.inferInstance <|> pure ()
|
|
||||||
Elab.Tactic.pruneSolvedGoals
|
|
||||||
return fragments.filter λ mvarId fragment =>
|
|
||||||
!(mvarId == goal || fragment == .convSentinel goal)
|
|
||||||
| .convSentinel parent =>
|
|
||||||
let parentFragment := fragments[parent]!
|
|
||||||
parentFragment.exit parent (fragments.erase goal)
|
|
||||||
|
|
||||||
protected def Fragment.step (fragment : Fragment) (goal : MVarId) (s : String) (map : FragmentMap)
|
|
||||||
: Elab.Tactic.TacticM FragmentMap := goal.withContext do
|
|
||||||
assert! ¬ (← goal.isAssigned)
|
|
||||||
match fragment with
|
|
||||||
| .calc prevRhs? => do
|
|
||||||
let .ok stx := Parser.runParserCategory
|
|
||||||
(env := ← getEnv)
|
|
||||||
(catName := `term)
|
|
||||||
(input := s)
|
|
||||||
(fileName := ← getFileName) | throwError s!"Failed to parse calc element {s}"
|
|
||||||
let `(term|$pred) := stx
|
|
||||||
let decl ← goal.getDecl
|
|
||||||
let target ← instantiateMVars decl.type
|
|
||||||
let tag := decl.userName
|
|
||||||
|
|
||||||
let mut step ← Elab.Term.elabType <| ← do
|
|
||||||
if let some prevRhs := prevRhs? then
|
|
||||||
Elab.Term.annotateFirstHoleWithType pred (← Meta.inferType prevRhs)
|
|
||||||
else
|
|
||||||
pure pred
|
|
||||||
|
|
||||||
let some (_, lhs, rhs) ← Elab.Term.getCalcRelation? step |
|
|
||||||
throwErrorAt pred "invalid 'calc' step, relation expected{indentExpr step}"
|
|
||||||
if let some prevRhs := prevRhs? 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}"}"
|
|
||||||
|
|
||||||
-- Creates a mvar to represent the proof that the calc tactic solves the
|
|
||||||
-- current branch
|
|
||||||
-- In the Lean `calc` tactic this is gobbled up by
|
|
||||||
-- `withCollectingNewGoalsFrom`
|
|
||||||
let mut proof ← Meta.mkFreshExprMVarAt (← getLCtx) (← Meta.getLocalInstances) step
|
|
||||||
(userName := tag ++ `calc)
|
|
||||||
let mvarBranch := proof.mvarId!
|
|
||||||
|
|
||||||
let mut proofType ← Meta.inferType proof
|
|
||||||
let mut remainder? := Option.none
|
|
||||||
|
|
||||||
-- The calc tactic either solves the main goal or leaves another relation.
|
|
||||||
-- Replace the main goal, and save the new goal if necessary
|
|
||||||
unless ← Meta.isDefEq proofType target do
|
|
||||||
let rec throwFailed :=
|
|
||||||
throwError "'calc' tactic failed, has type{indentExpr proofType}\nbut it is expected to have type{indentExpr target}"
|
|
||||||
let some (_, _, rhs) ← Elab.Term.getCalcRelation? proofType | throwFailed
|
|
||||||
let some (r, _, rhs') ← Elab.Term.getCalcRelation? target | throwFailed
|
|
||||||
let lastStep := mkApp2 r rhs rhs'
|
|
||||||
let lastStepGoal ← Meta.mkFreshExprSyntheticOpaqueMVar lastStep tag
|
|
||||||
(proof, proofType) ← Elab.Term.mkCalcTrans proof proofType lastStepGoal lastStep
|
|
||||||
unless ← Meta.isDefEq proofType target do throwFailed
|
|
||||||
remainder? := .some lastStepGoal.mvarId!
|
|
||||||
goal.assign proof
|
|
||||||
let map := map.erase goal
|
|
||||||
|
|
||||||
let goals := [ mvarBranch ] ++ remainder?.toList
|
|
||||||
Elab.Tactic.setGoals goals
|
|
||||||
match remainder? with
|
|
||||||
| .some goal => return map.insert goal $ .calc (.some rhs)
|
|
||||||
| .none => return map
|
|
||||||
| .conv .. => do
|
|
||||||
throwError "Direct operation on conversion tactic parent goal is not allowed"
|
|
||||||
| fragment@(.convSentinel parent) => do
|
|
||||||
assert! isLHSGoal? (← goal.getType) |>.isSome
|
|
||||||
let tactic ← match Parser.runParserCategory
|
|
||||||
(env := ← MonadEnv.getEnv)
|
|
||||||
(catName := `conv)
|
|
||||||
(input := s)
|
|
||||||
(fileName := ← getFileName) with
|
|
||||||
| .ok stx => pure $ stx
|
|
||||||
| .error error => throwError error
|
|
||||||
let oldGoals ← Elab.Tactic.getGoals
|
|
||||||
-- Label newly generated goals as conv sentinels
|
|
||||||
Elab.Tactic.evalTactic tactic
|
|
||||||
let newConvGoals ← (← Elab.Tactic.getUnsolvedGoals).filterM λ g => do
|
|
||||||
-- conv tactic might generate non-conv goals
|
|
||||||
if oldGoals.contains g then
|
|
||||||
return false
|
|
||||||
return isLHSGoal? (← g.getType) |>.isSome
|
|
||||||
-- Conclude the conv by exiting the parent fragment if new goals is empty
|
|
||||||
if newConvGoals.isEmpty then
|
|
||||||
let hasSiblingFragment := map.fold (init := false) λ flag _ fragment =>
|
|
||||||
if flag then
|
|
||||||
true
|
|
||||||
else match fragment with
|
|
||||||
| .convSentinel parent' => parent == parent'
|
|
||||||
| _ => false
|
|
||||||
if ¬ hasSiblingFragment then
|
|
||||||
-- This fragment must exist since we have conv goals
|
|
||||||
let parentFragment := map[parent]!
|
|
||||||
-- All descendants exhausted. Exit from the parent conv.
|
|
||||||
return ← parentFragment.exit parent map
|
|
||||||
return newConvGoals.foldl (init := map) λ acc g =>
|
|
||||||
acc.insert g fragment
|
|
||||||
|
|
||||||
end Pantograph
|
|
|
@ -1,6 +1,6 @@
|
||||||
namespace Pantograph
|
namespace Pantograph
|
||||||
|
|
||||||
@[export pantograph_version]
|
@[export pantograph_version]
|
||||||
def version := "0.3.4"
|
def version := "0.3.0"
|
||||||
|
|
||||||
end Pantograph
|
end Pantograph
|
||||||
|
|
|
@ -60,11 +60,11 @@ Example proving a theorem: (alternatively use `goal.start {"copyFrom": "Nat.add_
|
||||||
```
|
```
|
||||||
$ pantograph Init
|
$ pantograph Init
|
||||||
goal.start {"expr": "∀ (n m : Nat), n + m = m + n"}
|
goal.start {"expr": "∀ (n m : Nat), n + m = m + n"}
|
||||||
goal.tactic {"stateId": 0, "tactic": "intro n m"}
|
goal.tactic {"stateId": 0, "goalId": 0, "tactic": "intro n m"}
|
||||||
goal.tactic {"stateId": 1, "tactic": "assumption"}
|
goal.tactic {"stateId": 1, "goalId": 0, "tactic": "assumption"}
|
||||||
goal.delete {"stateIds": [0]}
|
goal.delete {"stateIds": [0]}
|
||||||
stat {}
|
stat {}
|
||||||
goal.tactic {"stateId": 1, "tactic": "rw [Nat.add_comm]"}
|
goal.tactic {"stateId": 1, "goalId": 0, "tactic": "rw [Nat.add_comm]"}
|
||||||
stat
|
stat
|
||||||
```
|
```
|
||||||
where the application of `assumption` should lead to a failure.
|
where the application of `assumption` should lead to a failure.
|
||||||
|
@ -81,7 +81,7 @@ the environment might be setup like this:
|
||||||
``` sh
|
``` sh
|
||||||
LIB="../lib"
|
LIB="../lib"
|
||||||
LIB_MATHLIB="$LIB/mathlib4/.lake"
|
LIB_MATHLIB="$LIB/mathlib4/.lake"
|
||||||
export LEAN_PATH="$LIB_MATHLIB:$LIB_MATHLIB/aesop/build/lib:$LIB_MATHLIB/Qq/build/lib:$LIB_MATHLIB/std/build/lib"
|
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 $@
|
LEAN_PATH=$LEAN_PATH build/bin/pantograph $@
|
||||||
```
|
```
|
||||||
|
|
209
Repl.lean
209
Repl.lean
|
@ -7,14 +7,12 @@ open Lean
|
||||||
|
|
||||||
structure Context where
|
structure Context where
|
||||||
coreContext : Core.Context
|
coreContext : Core.Context
|
||||||
-- If true, the environment will change after running `CoreM`
|
|
||||||
inheritEnv : Bool := false
|
|
||||||
|
|
||||||
/-- Stores state of the REPL -/
|
/-- Stores state of the REPL -/
|
||||||
structure State where
|
structure State where
|
||||||
options : Protocol.Options := {}
|
options : Protocol.Options := {}
|
||||||
nextId : Nat := 0
|
nextId : Nat := 0
|
||||||
goalStates : Std.HashMap Nat GoalState := Std.HashMap.emptyWithCapacity
|
goalStates : Std.HashMap Nat GoalState := Std.HashMap.empty
|
||||||
|
|
||||||
env : Environment
|
env : Environment
|
||||||
-- Parser state
|
-- Parser state
|
||||||
|
@ -30,10 +28,7 @@ instance : MonadEnv MainM where
|
||||||
getEnv := return (← get).env
|
getEnv := return (← get).env
|
||||||
modifyEnv f := modify fun s => { s with env := f s.env }
|
modifyEnv f := modify fun s => { s with env := f s.env }
|
||||||
|
|
||||||
def withInheritEnv [Monad m] [MonadWithReaderOf Context m] [MonadLift MainM m] { α } (z : m α) : m α := do
|
def newGoalState (goalState: GoalState) : MainM Nat := do
|
||||||
withTheReader Context ({ · with inheritEnv := true }) z
|
|
||||||
|
|
||||||
def newGoalState (goalState : GoalState) : MainM Nat := do
|
|
||||||
let state ← get
|
let state ← get
|
||||||
let stateId := state.nextId
|
let stateId := state.nextId
|
||||||
set { state with
|
set { state with
|
||||||
|
@ -44,7 +39,7 @@ def newGoalState (goalState : GoalState) : MainM Nat := do
|
||||||
|
|
||||||
def runCoreM { α } (coreM : CoreM α) : EMainM α := do
|
def runCoreM { α } (coreM : CoreM α) : EMainM α := do
|
||||||
let scope := (← get).scope
|
let scope := (← get).scope
|
||||||
let options := (← get).options
|
let options :=(← get).options
|
||||||
let cancelTk? ← match options.timeout with
|
let cancelTk? ← match options.timeout with
|
||||||
| 0 => pure .none
|
| 0 => pure .none
|
||||||
| _ => .some <$> IO.CancelToken.new
|
| _ => .some <$> IO.CancelToken.new
|
||||||
|
@ -65,31 +60,27 @@ def runCoreM { α } (coreM : CoreM α) : EMainM α := do
|
||||||
Except.ok <$> coreM
|
Except.ok <$> coreM
|
||||||
catch ex =>
|
catch ex =>
|
||||||
let desc ← ex.toMessageData.toString
|
let desc ← ex.toMessageData.toString
|
||||||
return Except.error ({ error := "exception", desc } : Protocol.InteractionError)
|
return Except.error $ ({ error := "exception", desc } : Protocol.InteractionError)
|
||||||
finally
|
|
||||||
for {msg, ..} in (← getTraceState).traces do
|
|
||||||
IO.eprintln (← msg.format.toIO)
|
|
||||||
resetTraceState
|
|
||||||
if let .some token := cancelTk? then
|
if let .some token := cancelTk? then
|
||||||
runCancelTokenWithTimeout token (timeout := .ofBitVec options.timeout)
|
runCancelTokenWithTimeout token (timeout := .ofBitVec options.timeout)
|
||||||
let (result, state') ← match ← (coreM'.run coreCtx coreState).toIO' with
|
let (result, state') ← match ← (coreM'.run coreCtx coreState).toIO' with
|
||||||
| Except.error (Exception.error _ msg) => Protocol.throw $ { error := "core", desc := ← msg.toString }
|
| Except.error (Exception.error _ msg) => Protocol.throw $ { error := "core", desc := ← msg.toString }
|
||||||
| Except.error (Exception.internal id _) => Protocol.throw $ { error := "internal", desc := (← id.getName).toString }
|
| Except.error (Exception.internal id _) => Protocol.throw $ { error := "internal", desc := (← id.getName).toString }
|
||||||
| Except.ok a => pure a
|
| Except.ok a => pure a
|
||||||
if (← read).inheritEnv && result matches .ok _ then
|
if result matches .ok _ then
|
||||||
setEnv state'.env
|
modifyEnv λ _ => state'.env
|
||||||
liftExcept result
|
liftExcept result
|
||||||
|
|
||||||
def runCoreM' { α } (coreM : Protocol.FallibleT CoreM α) : EMainM α := do
|
def runCoreM' { α } (coreM : Protocol.FallibleT CoreM α) : EMainM α := do
|
||||||
liftExcept $ ← runCoreM coreM.run
|
liftExcept $ ← runCoreM coreM.run
|
||||||
|
|
||||||
|
|
||||||
def liftMetaM { α } (metaM : MetaM α): EMainM α :=
|
def liftMetaM { α } (metaM : MetaM α): EMainM α :=
|
||||||
runCoreM metaM.run'
|
runCoreM metaM.run'
|
||||||
def liftTermElabM { α } (termElabM : Elab.TermElabM α) (levelNames : List Name := [])
|
def liftTermElabM { α } (termElabM: Elab.TermElabM α) (levelNames : List Name := [])
|
||||||
: EMainM α := do
|
: EMainM α := do
|
||||||
let scope := (← get).scope
|
let scope := (← get).scope
|
||||||
let context := {
|
let context := {
|
||||||
errToSorry := false,
|
|
||||||
isNoncomputableSection := scope.isNoncomputable,
|
isNoncomputableSection := scope.isNoncomputable,
|
||||||
}
|
}
|
||||||
let state := {
|
let state := {
|
||||||
|
@ -97,96 +88,18 @@ def liftTermElabM { α } (termElabM : Elab.TermElabM α) (levelNames : List Name
|
||||||
}
|
}
|
||||||
runCoreM $ termElabM.run' context state |>.run'
|
runCoreM $ termElabM.run' context state |>.run'
|
||||||
|
|
||||||
section Goal
|
|
||||||
|
|
||||||
def goal_tactic (args: Protocol.GoalTactic): EMainM Protocol.GoalTacticResult := do
|
|
||||||
let state ← getMainState
|
|
||||||
let .some goalState := state.goalStates[args.stateId]? |
|
|
||||||
Protocol.throw $ Protocol.errorIndex s!"Invalid state index {args.stateId}"
|
|
||||||
let unshielded := args.autoResume?.getD state.options.automaticMode
|
|
||||||
let site ← match args.goalId?, unshielded with
|
|
||||||
| .some goalId, true => do
|
|
||||||
let .some goal := goalState.goals[goalId]? |
|
|
||||||
Protocol.throw $ Protocol.errorIndex s!"Invalid goal index {goalId}"
|
|
||||||
pure (.prefer goal)
|
|
||||||
| .some goalId, false => do
|
|
||||||
let .some goal := goalState.goals[goalId]? |
|
|
||||||
Protocol.throw $ Protocol.errorIndex s!"Invalid goal index {goalId}"
|
|
||||||
pure (.focus goal)
|
|
||||||
| .none, true => pure .unfocus
|
|
||||||
| .none, false => do
|
|
||||||
let .some goal := goalState.mainGoal? |
|
|
||||||
Protocol.throw $ Protocol.errorIndex s!"No goals to be solved"
|
|
||||||
pure (.focus goal)
|
|
||||||
let nextGoalState?: Except _ TacticResult ← liftTermElabM do
|
|
||||||
-- NOTE: Should probably use a macro to handle this...
|
|
||||||
match args.tactic?, args.mode?, args.expr?, args.have?, args.let?, args.draft? with
|
|
||||||
| .some tactic, .none, .none, .none, .none, .none => do
|
|
||||||
pure $ Except.ok $ ← goalState.tryTactic site tactic
|
|
||||||
| .none, .some mode, .none, .none, .none, .none => match mode with
|
|
||||||
| "tactic" => do -- Exit from the current fragment
|
|
||||||
pure $ Except.ok $ ← goalState.fragmentExit site
|
|
||||||
| "conv" => do
|
|
||||||
pure $ Except.ok $ ← goalState.convEnter site
|
|
||||||
| "calc" => do
|
|
||||||
pure $ Except.ok $ ← goalState.calcEnter site
|
|
||||||
| _ => pure $ .error $ Protocol.errorOperation s!"Invalid mode {mode}"
|
|
||||||
| .none, .none, .some expr, .none, .none, .none => do
|
|
||||||
pure $ Except.ok $ ← goalState.tryAssign site expr
|
|
||||||
| .none, .none, .none, .some type, .none, .none => do
|
|
||||||
let binderName := args.binderName?.getD ""
|
|
||||||
pure $ Except.ok $ ← goalState.tryHave site binderName type
|
|
||||||
| .none, .none, .none, .none, .some type, .none => do
|
|
||||||
let binderName := args.binderName?.getD ""
|
|
||||||
pure $ Except.ok $ ← goalState.tryLet site binderName type
|
|
||||||
| .none, .none, .none, .none, .none, .some draft => do
|
|
||||||
pure $ Except.ok $ ← goalState.tryDraft site draft
|
|
||||||
| _, _, _, _, _, _ =>
|
|
||||||
pure $ .error $ Protocol.errorOperation
|
|
||||||
"Exactly one of {tactic, mode, expr, have, let, draft} must be supplied"
|
|
||||||
match nextGoalState? with
|
|
||||||
| .error error => Protocol.throw error
|
|
||||||
| .ok (.success nextGoalState messages) => do
|
|
||||||
let env ← getEnv
|
|
||||||
let nextStateId ← newGoalState nextGoalState
|
|
||||||
let parentExprs := nextGoalState.parentExprs
|
|
||||||
let hasSorry := parentExprs.any λ
|
|
||||||
| .ok e => e.hasSorry
|
|
||||||
| .error _ => false
|
|
||||||
let hasUnsafe := parentExprs.any λ
|
|
||||||
| .ok e => env.hasUnsafe e
|
|
||||||
| .error _ => false
|
|
||||||
let goals ← runCoreM $ nextGoalState.serializeGoals (options := state.options) |>.run'
|
|
||||||
let messages ← messages.mapM (·.serialize)
|
|
||||||
return {
|
|
||||||
nextStateId? := .some nextStateId,
|
|
||||||
goals? := .some goals,
|
|
||||||
messages? := .some messages,
|
|
||||||
hasSorry,
|
|
||||||
hasUnsafe,
|
|
||||||
}
|
|
||||||
| .ok (.parseError message) =>
|
|
||||||
return { messages? := .none, parseError? := .some message }
|
|
||||||
| .ok (.invalidAction message) =>
|
|
||||||
Protocol.throw $ errorI "invalid" message
|
|
||||||
| .ok (.failure messages) =>
|
|
||||||
let messages ← messages.mapM (·.serialize)
|
|
||||||
return { messages? := .some messages }
|
|
||||||
|
|
||||||
end Goal
|
|
||||||
|
|
||||||
section Frontend
|
section Frontend
|
||||||
|
|
||||||
structure CompilationUnit where
|
structure CompilationUnit where
|
||||||
-- Environment immediately before the unit
|
-- Should be the penultimate environment, but this is ok
|
||||||
env : Environment
|
env : Environment
|
||||||
boundary : Nat × Nat
|
boundary : Nat × Nat
|
||||||
invocations : List Protocol.InvokedTactic
|
invocations : List Protocol.InvokedTactic
|
||||||
sorrys : List Frontend.InfoWithContext
|
sorrys : List Frontend.InfoWithContext
|
||||||
messages : Array SerialMessage
|
messages : Array String
|
||||||
newConstants : List Name
|
newConstants : List Name
|
||||||
|
|
||||||
def frontend_process (args: Protocol.FrontendProcess): EMainM Protocol.FrontendProcessResult := do
|
def frontend_process_inner (args: Protocol.FrontendProcess): EMainM Protocol.FrontendProcessResult := do
|
||||||
let options := (← getMainState).options
|
let options := (← getMainState).options
|
||||||
let (fileName, file) ← match args.fileName?, args.file? with
|
let (fileName, file) ← match args.fileName?, args.file? with
|
||||||
| .some fileName, .none => do
|
| .some fileName, .none => do
|
||||||
|
@ -203,7 +116,7 @@ def frontend_process (args: Protocol.FrontendProcess): EMainM Protocol.FrontendP
|
||||||
let frontendM: Elab.Frontend.FrontendM (List CompilationUnit) :=
|
let frontendM: Elab.Frontend.FrontendM (List CompilationUnit) :=
|
||||||
Frontend.mapCompilationSteps λ step => do
|
Frontend.mapCompilationSteps λ step => do
|
||||||
let boundary := (step.src.startPos.byteIdx, step.src.stopPos.byteIdx)
|
let boundary := (step.src.startPos.byteIdx, step.src.stopPos.byteIdx)
|
||||||
let invocations: Option (List Protocol.InvokedTactic) ← if args.invocations?.isSome then
|
let invocations: Option (List Protocol.InvokedTactic) ← if args.invocations then
|
||||||
Frontend.collectTacticsFromCompilationStep step
|
Frontend.collectTacticsFromCompilationStep step
|
||||||
else
|
else
|
||||||
pure []
|
pure []
|
||||||
|
@ -211,7 +124,7 @@ def frontend_process (args: Protocol.FrontendProcess): EMainM Protocol.FrontendP
|
||||||
Frontend.collectSorrys step (options := { collectTypeErrors := args.typeErrorsAsGoals })
|
Frontend.collectSorrys step (options := { collectTypeErrors := args.typeErrorsAsGoals })
|
||||||
else
|
else
|
||||||
pure []
|
pure []
|
||||||
let messages ← step.msgs.toArray.mapM (·.serialize)
|
let messages ← step.messageStrings
|
||||||
let newConstants ← if args.newConstants then
|
let newConstants ← if args.newConstants then
|
||||||
Frontend.collectNewDefinedConstants step
|
Frontend.collectNewDefinedConstants step
|
||||||
else
|
else
|
||||||
|
@ -230,16 +143,12 @@ def frontend_process (args: Protocol.FrontendProcess): EMainM Protocol.FrontendP
|
||||||
if let .some scope := state'.commandState.scopes.head? then
|
if let .some scope := state'.commandState.scopes.head? then
|
||||||
-- modify the scope
|
-- modify the scope
|
||||||
set { ← getMainState with scope }
|
set { ← getMainState with scope }
|
||||||
if let .some fileName := args.invocations? then
|
|
||||||
let units := li.map λ unit => { invocations? := .some unit.invocations }
|
|
||||||
let data : Protocol.FrontendData := { units }
|
|
||||||
IO.FS.writeFile fileName (toJson data |>.compress)
|
|
||||||
let units ← li.mapM λ step => withEnv step.env do
|
let units ← li.mapM λ step => withEnv step.env do
|
||||||
let newConstants? := if args.newConstants then
|
let newConstants? := if args.newConstants then
|
||||||
.some $ step.newConstants.toArray.map λ name => name.toString
|
.some $ step.newConstants.toArray.map λ name => name.toString
|
||||||
else
|
else
|
||||||
.none
|
.none
|
||||||
let (goalStateId?, goals?, goalSrcBoundaries?) ← if step.sorrys.isEmpty then
|
let (goalStateId?, goals?, goalSrcBoundaries?) ← if step.sorrys.isEmpty then do
|
||||||
pure (.none, .none, .none)
|
pure (.none, .none, .none)
|
||||||
else do
|
else do
|
||||||
let ({ state, srcBoundaries }, goals) ← liftMetaM do
|
let ({ state, srcBoundaries }, goals) ← liftMetaM do
|
||||||
|
@ -249,11 +158,11 @@ def frontend_process (args: Protocol.FrontendProcess): EMainM Protocol.FrontendP
|
||||||
let stateId ← newGoalState state
|
let stateId ← newGoalState state
|
||||||
let srcBoundaries := srcBoundaries.toArray.map (λ (b, e) => (b.byteIdx, e.byteIdx))
|
let srcBoundaries := srcBoundaries.toArray.map (λ (b, e) => (b.byteIdx, e.byteIdx))
|
||||||
pure (.some stateId, .some goals, .some srcBoundaries)
|
pure (.some stateId, .some goals, .some srcBoundaries)
|
||||||
let nInvocations? := if args.invocations?.isSome then .some step.invocations.length else .none
|
let invocations? := if args.invocations then .some step.invocations else .none
|
||||||
return {
|
return {
|
||||||
boundary := step.boundary,
|
boundary := step.boundary,
|
||||||
messages := step.messages,
|
messages := step.messages,
|
||||||
nInvocations?,
|
invocations?,
|
||||||
goalStateId?,
|
goalStateId?,
|
||||||
goals?,
|
goals?,
|
||||||
goalSrcBoundaries?,
|
goalSrcBoundaries?,
|
||||||
|
@ -269,10 +178,7 @@ def execute (command: Protocol.Command): MainM Json := do
|
||||||
try
|
try
|
||||||
match fromJson? command.payload with
|
match fromJson? command.payload with
|
||||||
| .ok args => do
|
| .ok args => do
|
||||||
let (msg, result) ← IO.FS.withIsolatedStreams (isolateStderr := false) $ comm args
|
match (← comm args |>.run) with
|
||||||
if !msg.isEmpty then
|
|
||||||
IO.eprint s!"stdout: {msg}"
|
|
||||||
match result with
|
|
||||||
| .ok result => return toJson result
|
| .ok result => return toJson result
|
||||||
| .error ierror => return toJson ierror
|
| .error ierror => return toJson ierror
|
||||||
| .error error => return toJson $ errorCommand s!"Unable to parse json: {error}"
|
| .error error => return toJson $ errorCommand s!"Unable to parse json: {error}"
|
||||||
|
@ -306,12 +212,13 @@ def execute (command: Protocol.Command): MainM Json := do
|
||||||
return toJson error
|
return toJson error
|
||||||
where
|
where
|
||||||
errorCommand := errorI "command"
|
errorCommand := errorI "command"
|
||||||
|
errorIndex := errorI "index"
|
||||||
errorIO := errorI "io"
|
errorIO := errorI "io"
|
||||||
-- Command Functions
|
-- Command Functions
|
||||||
reset (_: Protocol.Reset): EMainM Protocol.StatResult := do
|
reset (_: Protocol.Reset): EMainM Protocol.StatResult := do
|
||||||
let state ← getMainState
|
let state ← getMainState
|
||||||
let nGoals := state.goalStates.size
|
let nGoals := state.goalStates.size
|
||||||
set { state with nextId := 0, goalStates := .emptyWithCapacity }
|
set { state with nextId := 0, goalStates := .empty }
|
||||||
return { nGoals }
|
return { nGoals }
|
||||||
stat (_: Protocol.Stat): EMainM Protocol.StatResult := do
|
stat (_: Protocol.Stat): EMainM Protocol.StatResult := do
|
||||||
let state ← getMainState
|
let state ← getMainState
|
||||||
|
@ -328,7 +235,7 @@ def execute (command: Protocol.Command): MainM Json := do
|
||||||
env_inspect (args: Protocol.EnvInspect): EMainM Protocol.EnvInspectResult := do
|
env_inspect (args: Protocol.EnvInspect): EMainM Protocol.EnvInspectResult := do
|
||||||
let state ← getMainState
|
let state ← getMainState
|
||||||
runCoreM' $ Environment.inspect args state.options
|
runCoreM' $ Environment.inspect args state.options
|
||||||
env_add (args: Protocol.EnvAdd): EMainM Protocol.EnvAddResult := withInheritEnv do
|
env_add (args: Protocol.EnvAdd): EMainM Protocol.EnvAddResult := do
|
||||||
runCoreM' $ Environment.addDecl args.name (args.levels?.getD #[]) args.type? args.value args.isTheorem
|
runCoreM' $ Environment.addDecl args.name (args.levels?.getD #[]) args.type? args.value args.isTheorem
|
||||||
env_save (args: Protocol.EnvSaveLoad): EMainM Protocol.EnvSaveLoadResult := do
|
env_save (args: Protocol.EnvSaveLoad): EMainM Protocol.EnvSaveLoadResult := do
|
||||||
let env ← MonadEnv.getEnv
|
let env ← MonadEnv.getEnv
|
||||||
|
@ -370,7 +277,7 @@ def execute (command: Protocol.Command): MainM Json := do
|
||||||
| .some expr, .none => goalStartExpr expr |>.run
|
| .some expr, .none => goalStartExpr expr |>.run
|
||||||
| .none, .some copyFrom => do
|
| .none, .some copyFrom => do
|
||||||
(match (← getEnv).find? <| copyFrom.toName with
|
(match (← getEnv).find? <| copyFrom.toName with
|
||||||
| .none => return .error <| Protocol.errorIndex s!"Symbol not found: {copyFrom}"
|
| .none => return .error <| errorIndex s!"Symbol not found: {copyFrom}"
|
||||||
| .some cInfo => return .ok (← GoalState.create cInfo.type))
|
| .some cInfo => return .ok (← GoalState.create cInfo.type))
|
||||||
| _, _ =>
|
| _, _ =>
|
||||||
return .error <| errorI "arguments" "Exactly one of {expr, copyFrom} must be supplied"
|
return .error <| errorI "arguments" "Exactly one of {expr, copyFrom} must be supplied"
|
||||||
|
@ -379,14 +286,72 @@ def execute (command: Protocol.Command): MainM Json := do
|
||||||
| .ok goalState =>
|
| .ok goalState =>
|
||||||
let stateId ← newGoalState goalState
|
let stateId ← newGoalState goalState
|
||||||
return { stateId, root := goalState.root.name.toString }
|
return { stateId, root := goalState.root.name.toString }
|
||||||
|
goal_tactic (args: Protocol.GoalTactic): EMainM Protocol.GoalTacticResult := do
|
||||||
|
let state ← getMainState
|
||||||
|
let .some goalState := state.goalStates[args.stateId]? |
|
||||||
|
Protocol.throw $ errorIndex s!"Invalid state index {args.stateId}"
|
||||||
|
let .some goal := goalState.goals[args.goalId]? |
|
||||||
|
Protocol.throw $ errorIndex s!"Invalid goal index {args.goalId}"
|
||||||
|
let nextGoalState?: Except _ TacticResult ← liftTermElabM do
|
||||||
|
-- NOTE: Should probably use a macro to handle this...
|
||||||
|
match args.tactic?, args.expr?, args.have?, args.let?, args.calc?, args.conv?, args.draft? with
|
||||||
|
| .some tactic, .none, .none, .none, .none, .none, .none => do
|
||||||
|
pure <| Except.ok <| ← goalState.tryTactic goal tactic
|
||||||
|
| .none, .some expr, .none, .none, .none, .none, .none => do
|
||||||
|
pure <| Except.ok <| ← goalState.tryAssign goal expr
|
||||||
|
| .none, .none, .some type, .none, .none, .none, .none => do
|
||||||
|
let binderName := args.binderName?.getD ""
|
||||||
|
pure <| Except.ok <| ← goalState.tryHave goal binderName type
|
||||||
|
| .none, .none, .none, .some type, .none, .none, .none => do
|
||||||
|
let binderName := args.binderName?.getD ""
|
||||||
|
pure <| Except.ok <| ← goalState.tryLet goal binderName type
|
||||||
|
| .none, .none, .none, .none, .some pred, .none, .none => do
|
||||||
|
pure <| Except.ok <| ← goalState.tryCalc goal pred
|
||||||
|
| .none, .none, .none, .none, .none, .some true, .none => do
|
||||||
|
pure <| Except.ok <| ← goalState.conv goal
|
||||||
|
| .none, .none, .none, .none, .none, .some false, .none => do
|
||||||
|
pure <| Except.ok <| ← goalState.convExit
|
||||||
|
| .none, .none, .none, .none, .none, .none, .some draft => do
|
||||||
|
pure <| Except.ok <| ← goalState.tryDraft goal draft
|
||||||
|
| _, _, _, _, _, _, _ =>
|
||||||
|
let error := errorI "arguments" "Exactly one of {tactic, expr, have, let, calc, conv, draft} must be supplied"
|
||||||
|
pure $ .error error
|
||||||
|
match nextGoalState? with
|
||||||
|
| .error error => Protocol.throw error
|
||||||
|
| .ok (.success nextGoalState) => do
|
||||||
|
let nextGoalState ← match state.options.automaticMode, args.conv? with
|
||||||
|
| true, .none => do
|
||||||
|
let .ok result := nextGoalState.resume (nextGoalState.goals ++ goalState.goals) |
|
||||||
|
Protocol.throw $ errorIO "Resuming known goals"
|
||||||
|
pure result
|
||||||
|
| true, .some true => pure nextGoalState
|
||||||
|
| true, .some false => do
|
||||||
|
let .some (_, _, dormantGoals) := goalState.convMVar? |
|
||||||
|
Protocol.throw $ errorIO "If conv exit succeeded this should not fail"
|
||||||
|
let .ok result := nextGoalState.resume (nextGoalState.goals ++ dormantGoals) |
|
||||||
|
Protocol.throw $ errorIO "Resuming known goals"
|
||||||
|
pure result
|
||||||
|
| false, _ => pure nextGoalState
|
||||||
|
let nextStateId ← newGoalState nextGoalState
|
||||||
|
let goals ← runCoreM $ nextGoalState.serializeGoals (parent := .some goalState) (options := state.options) |>.run'
|
||||||
|
return {
|
||||||
|
nextStateId? := .some nextStateId,
|
||||||
|
goals? := .some goals,
|
||||||
|
}
|
||||||
|
| .ok (.parseError message) =>
|
||||||
|
return { parseError? := .some message }
|
||||||
|
| .ok (.invalidAction message) =>
|
||||||
|
Protocol.throw $ errorI "invalid" message
|
||||||
|
| .ok (.failure messages) =>
|
||||||
|
return { tacticErrors? := .some messages }
|
||||||
goal_continue (args: Protocol.GoalContinue): EMainM Protocol.GoalContinueResult := do
|
goal_continue (args: Protocol.GoalContinue): EMainM Protocol.GoalContinueResult := do
|
||||||
let state ← getMainState
|
let state ← getMainState
|
||||||
let .some target := state.goalStates[args.target]? |
|
let .some target := state.goalStates[args.target]? |
|
||||||
Protocol.throw $ Protocol.errorIndex s!"Invalid state index {args.target}"
|
Protocol.throw $ errorIndex s!"Invalid state index {args.target}"
|
||||||
let nextGoalState? : GoalState ← match args.branch?, args.goals? with
|
let nextGoalState? : GoalState ← match args.branch?, args.goals? with
|
||||||
| .some branchId, .none => do
|
| .some branchId, .none => do
|
||||||
match state.goalStates[branchId]? with
|
match state.goalStates[branchId]? with
|
||||||
| .none => Protocol.throw $ Protocol.errorIndex s!"Invalid state index {branchId}"
|
| .none => Protocol.throw $ errorIndex s!"Invalid state index {branchId}"
|
||||||
| .some branch => pure $ target.continue branch
|
| .some branch => pure $ target.continue branch
|
||||||
| .none, .some goals =>
|
| .none, .some goals =>
|
||||||
let goals := goals.toList.map (λ n => { name := n.toName })
|
let goals := goals.toList.map (λ n => { name := n.toName })
|
||||||
|
@ -409,11 +374,11 @@ def execute (command: Protocol.Command): MainM Json := do
|
||||||
goal_print (args: Protocol.GoalPrint): EMainM Protocol.GoalPrintResult := do
|
goal_print (args: Protocol.GoalPrint): EMainM Protocol.GoalPrintResult := do
|
||||||
let state ← getMainState
|
let state ← getMainState
|
||||||
let .some goalState := state.goalStates[args.stateId]? |
|
let .some goalState := state.goalStates[args.stateId]? |
|
||||||
Protocol.throw $ Protocol.errorIndex s!"Invalid state index {args.stateId}"
|
Protocol.throw $ errorIndex s!"Invalid state index {args.stateId}"
|
||||||
let result ← liftMetaM <| goalPrint
|
let result ← liftMetaM <| goalPrint
|
||||||
goalState
|
goalState
|
||||||
(rootExpr := args.rootExpr?.getD False)
|
(rootExpr := args.rootExpr?.getD False)
|
||||||
(parentExprs := args.parentExprs?.getD False)
|
(parentExpr := args.parentExpr?.getD False)
|
||||||
(goals := args.goals?.getD False)
|
(goals := args.goals?.getD False)
|
||||||
(extraMVars := args.extraMVars?.getD #[])
|
(extraMVars := args.extraMVars?.getD #[])
|
||||||
(options := state.options)
|
(options := state.options)
|
||||||
|
@ -421,12 +386,14 @@ def execute (command: Protocol.Command): MainM Json := do
|
||||||
goal_save (args: Protocol.GoalSave): EMainM Protocol.GoalSaveResult := do
|
goal_save (args: Protocol.GoalSave): EMainM Protocol.GoalSaveResult := do
|
||||||
let state ← getMainState
|
let state ← getMainState
|
||||||
let .some goalState := state.goalStates[args.id]? |
|
let .some goalState := state.goalStates[args.id]? |
|
||||||
Protocol.throw $ Protocol.errorIndex s!"Invalid state index {args.id}"
|
Protocol.throw $ errorIndex s!"Invalid state index {args.id}"
|
||||||
goalStatePickle goalState args.path (background? := .some $ ← getEnv)
|
goalStatePickle goalState args.path
|
||||||
return {}
|
return {}
|
||||||
goal_load (args: Protocol.GoalLoad): EMainM Protocol.GoalLoadResult := do
|
goal_load (args: Protocol.GoalLoad): EMainM Protocol.GoalLoadResult := do
|
||||||
let (goalState, _) ← goalStateUnpickle args.path (background? := .some $ ← getEnv)
|
let (goalState, _) ← goalStateUnpickle args.path (← MonadEnv.getEnv)
|
||||||
let id ← newGoalState goalState
|
let id ← newGoalState goalState
|
||||||
return { id }
|
return { id }
|
||||||
|
frontend_process (args: Protocol.FrontendProcess): EMainM Protocol.FrontendProcessResult := do
|
||||||
|
frontend_process_inner args
|
||||||
|
|
||||||
end Pantograph.Repl
|
end Pantograph.Repl
|
||||||
|
|
|
@ -69,13 +69,12 @@ end Condensed
|
||||||
|
|
||||||
def GoalState.get! (state: GoalState) (i: Nat): MVarId := state.goals[i]!
|
def GoalState.get! (state: GoalState) (i: Nat): MVarId := state.goals[i]!
|
||||||
def GoalState.tacticOn (state: GoalState) (goalId: Nat) (tactic: String) := state.tryTactic (state.get! goalId) tactic
|
def GoalState.tacticOn (state: GoalState) (goalId: Nat) (tactic: String) := state.tryTactic (state.get! goalId) tactic
|
||||||
def GoalState.tacticOn' (state: GoalState) (goalId: Nat) (tactic: TSyntax `tactic) :=
|
|
||||||
state.tryTacticM (state.get! goalId) (Elab.Tactic.evalTactic tactic) true
|
|
||||||
|
|
||||||
def TacticResult.toString : TacticResult → String
|
def TacticResult.toString : TacticResult → String
|
||||||
| .success state _messages => s!".success ({state.goals.length} goals)"
|
| .success state => s!".success ({state.goals.length} goals)"
|
||||||
| .failure messages =>
|
| .failure messages =>
|
||||||
s!".failure ({messages.size} messages)"
|
let messages := "\n".intercalate messages.toList
|
||||||
|
s!".failure {messages}"
|
||||||
| .parseError error => s!".parseError {error}"
|
| .parseError error => s!".parseError {error}"
|
||||||
| .invalidAction error => s!".invalidAction {error}"
|
| .invalidAction error => s!".invalidAction {error}"
|
||||||
|
|
||||||
|
@ -95,14 +94,12 @@ def runCoreMSeq (env: Environment) (coreM: CoreM LSpec.TestSeq) (options: Array
|
||||||
| .ok a => return a
|
| .ok a => return a
|
||||||
def runMetaMSeq (env: Environment) (metaM: MetaM LSpec.TestSeq): IO LSpec.TestSeq :=
|
def runMetaMSeq (env: Environment) (metaM: MetaM LSpec.TestSeq): IO LSpec.TestSeq :=
|
||||||
runCoreMSeq env metaM.run'
|
runCoreMSeq env metaM.run'
|
||||||
def runTermElabMInMeta { α } (termElabM: Elab.TermElabM α): MetaM α :=
|
def runTermElabMInMeta { α } (termElabM: Lean.Elab.TermElabM α): Lean.MetaM α :=
|
||||||
termElabM.run' (ctx := defaultElabContext)
|
termElabM.run' (ctx := defaultElabContext)
|
||||||
def runTermElabMInCore { α } (termElabM: Elab.TermElabM α): CoreM α :=
|
|
||||||
(runTermElabMInMeta termElabM).run'
|
|
||||||
def runTermElabMSeq (env: Environment) (termElabM: Elab.TermElabM LSpec.TestSeq): IO LSpec.TestSeq :=
|
def runTermElabMSeq (env: Environment) (termElabM: Elab.TermElabM LSpec.TestSeq): IO LSpec.TestSeq :=
|
||||||
runMetaMSeq env $ termElabM.run' (ctx := defaultElabContext)
|
runMetaMSeq env $ termElabM.run' (ctx := defaultElabContext)
|
||||||
|
|
||||||
def exprToStr (e: Expr): MetaM String := toString <$> Meta.ppExpr e
|
def exprToStr (e: Expr): Lean.MetaM String := toString <$> Meta.ppExpr e
|
||||||
|
|
||||||
def strToTermSyntax (s: String): CoreM Syntax := do
|
def strToTermSyntax (s: String): CoreM Syntax := do
|
||||||
let .ok stx := Parser.runParserCategory
|
let .ok stx := Parser.runParserCategory
|
||||||
|
@ -160,13 +157,6 @@ 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
|
||||||
def transformTestT { α } { μ μ' : Type → Type }
|
|
||||||
[Monad μ] [Monad μ'] [MonadLiftT (ST IO.RealWorld) μ] [MonadLiftT (ST IO.RealWorld) μ']
|
|
||||||
(tr : {β : Type} → μ β → μ' β) (m : TestT μ α) : TestT μ' α := do
|
|
||||||
let tests ← get
|
|
||||||
let (a, tests) ← tr (β := α × LSpec.TestSeq) (m.run tests)
|
|
||||||
set tests
|
|
||||||
return a
|
|
||||||
|
|
||||||
def cdeclOf (userName: Name) (type: Expr): Condensed.LocalDecl :=
|
def cdeclOf (userName: Name) (type: Expr): Condensed.LocalDecl :=
|
||||||
{ userName, type }
|
{ userName, type }
|
||||||
|
@ -182,29 +172,6 @@ def buildGoal (nameType: List (String × String)) (target: String) (userName?: O
|
||||||
})).toArray
|
})).toArray
|
||||||
}
|
}
|
||||||
|
|
||||||
namespace Tactic
|
|
||||||
|
|
||||||
/-- Create an aux lemma and assigns it to `mvarId`, which is circuitous, but
|
|
||||||
exercises the aux lemma generator. -/
|
|
||||||
def assignWithAuxLemma (type : Expr) (value? : Option Expr := .none) : Elab.Tactic.TacticM Unit := do
|
|
||||||
let type ← instantiateMVars type
|
|
||||||
let value ← match value? with
|
|
||||||
| .some value => instantiateMVars value
|
|
||||||
| .none => Meta.mkSorry type (synthetic := false)
|
|
||||||
if type.hasExprMVar then
|
|
||||||
throwError "Type has expression mvar"
|
|
||||||
if value.hasExprMVar then
|
|
||||||
throwError "value has expression mvar"
|
|
||||||
let goal ← Elab.Tactic.getMainGoal
|
|
||||||
goal.withContext do
|
|
||||||
let name ← Meta.mkAuxLemma [] type value
|
|
||||||
unless ← Meta.isDefEq type (← goal.getType) do
|
|
||||||
throwError "Type provided is incorrect"
|
|
||||||
goal.assign (.const name [])
|
|
||||||
Elab.Tactic.pruneSolvedGoals
|
|
||||||
|
|
||||||
end Tactic
|
|
||||||
|
|
||||||
end Test
|
end Test
|
||||||
|
|
||||||
end Pantograph
|
end Pantograph
|
||||||
|
|
|
@ -48,9 +48,10 @@ def test_sexp_of_symbol (env: Environment): IO LSpec.TestSeq := do
|
||||||
def test_sexp_of_elab (env: Environment): IO LSpec.TestSeq := do
|
def test_sexp_of_elab (env: Environment): IO LSpec.TestSeq := do
|
||||||
let entries: List (String × (List Name) × String) := [
|
let entries: List (String × (List Name) × String) := [
|
||||||
("λ x: Nat × Bool => x.1", [], "(:lambda x ((:c Prod) (:c Nat) (:c Bool)) ((:c Prod.fst) (:c Nat) (:c Bool) 0))"),
|
("λ x: Nat × Bool => x.1", [], "(:lambda x ((:c Prod) (:c Nat) (:c Bool)) ((:c Prod.fst) (:c Nat) (:c Bool) 0))"),
|
||||||
|
("λ x: Array Nat => x.data", [], "(:lambda x ((:c Array) (:c Nat)) ((:c Array.data) (:c Nat) 0))"),
|
||||||
("λ {α: Sort (u + 1)} => List α", [`u], "(:lambda α (:sort (+ u 1)) ((:c List) 0) :i)"),
|
("λ {α: Sort (u + 1)} => List α", [`u], "(:lambda α (:sort (+ u 1)) ((:c List) 0) :i)"),
|
||||||
("λ {α} => List α", [], "(:lambda α (:sort (+ (:mv _uniq.4) 1)) ((:c List) 0) :i)"),
|
("λ {α} => List α", [], "(:lambda α (:sort (+ (:mv _uniq.4) 1)) ((:c List) 0) :i)"),
|
||||||
("(2: Nat) <= (5: Nat)", [], "((:c LE.le) (:mv _uniq.20) (:mv _uniq.21) ((:c OfNat.ofNat) (:mv _uniq.4) (:lit 2) (:mv _uniq.5)) ((:c OfNat.ofNat) (:mv _uniq.15) (:lit 5) (:mv _uniq.16)))"),
|
("(2: Nat) <= (5: Nat)", [], "((:c LE.le) (:mv _uniq.18) (:mv _uniq.19) ((:c OfNat.ofNat) (:mv _uniq.4) (:lit 2) (:mv _uniq.5)) ((:c OfNat.ofNat) (:mv _uniq.14) (:lit 5) (:mv _uniq.15)))"),
|
||||||
]
|
]
|
||||||
entries.foldlM (λ suites (source, levels, target) =>
|
entries.foldlM (λ suites (source, levels, target) =>
|
||||||
let termElabM := do
|
let termElabM := do
|
||||||
|
|
|
@ -15,7 +15,11 @@ deriving instance DecidableEq, Repr for Protocol.RecursorRule
|
||||||
deriving instance DecidableEq, Repr for Protocol.RecursorInfo
|
deriving instance DecidableEq, Repr for Protocol.RecursorInfo
|
||||||
deriving instance DecidableEq, Repr for Protocol.EnvInspectResult
|
deriving instance DecidableEq, Repr for Protocol.EnvInspectResult
|
||||||
|
|
||||||
def test_catalog (env : Environment) : IO LSpec.TestSeq := do
|
def test_catalog: IO LSpec.TestSeq := do
|
||||||
|
let env: Environment ← importModules
|
||||||
|
(imports := #[`Init])
|
||||||
|
(opts := {})
|
||||||
|
(trustLevel := 1)
|
||||||
let inner: CoreM LSpec.TestSeq := do
|
let inner: CoreM LSpec.TestSeq := do
|
||||||
let catalogResult ← Environment.catalog {}
|
let catalogResult ← Environment.catalog {}
|
||||||
let symbolsWithNum := env.constants.fold (init := #[]) (λ acc name info =>
|
let symbolsWithNum := env.constants.fold (init := #[]) (λ acc name info =>
|
||||||
|
@ -30,6 +34,7 @@ def test_symbol_visibility: IO LSpec.TestSeq := do
|
||||||
let entries: List (Name × Bool) := [
|
let entries: List (Name × Bool) := [
|
||||||
("Nat.add_comm".toName, false),
|
("Nat.add_comm".toName, false),
|
||||||
("foo.bla.Init.Data.List.Basic.2.1.Init.Lean.Expr._hyg.4".toName, true),
|
("foo.bla.Init.Data.List.Basic.2.1.Init.Lean.Expr._hyg.4".toName, true),
|
||||||
|
("Init.Data.Nat.Basic._auxLemma.4".toName, true),
|
||||||
]
|
]
|
||||||
let suite := entries.foldl (λ suites (symbol, target) =>
|
let suite := entries.foldl (λ suites (symbol, target) =>
|
||||||
let test := LSpec.check symbol.toString ((Environment.isNameInternal symbol) == target)
|
let test := LSpec.check symbol.toString ((Environment.isNameInternal symbol) == target)
|
||||||
|
@ -41,7 +46,11 @@ inductive ConstantCat where
|
||||||
| ctor (info: Protocol.ConstructorInfo)
|
| ctor (info: Protocol.ConstructorInfo)
|
||||||
| recursor (info: Protocol.RecursorInfo)
|
| recursor (info: Protocol.RecursorInfo)
|
||||||
|
|
||||||
def test_inspect (env : Environment) : IO LSpec.TestSeq := do
|
def test_inspect: IO LSpec.TestSeq := do
|
||||||
|
let env: Environment ← importModules
|
||||||
|
(imports := #[`Init])
|
||||||
|
(opts := {})
|
||||||
|
(trustLevel := 1)
|
||||||
let testCases: List (String × ConstantCat) := [
|
let testCases: List (String × ConstantCat) := [
|
||||||
("Or", ConstantCat.induct {
|
("Or", ConstantCat.induct {
|
||||||
numParams := 2,
|
numParams := 2,
|
||||||
|
@ -88,29 +97,37 @@ def test_inspect (env : Environment) : IO LSpec.TestSeq := do
|
||||||
) LSpec.TestSeq.done
|
) LSpec.TestSeq.done
|
||||||
runCoreMSeq env inner
|
runCoreMSeq env inner
|
||||||
|
|
||||||
def test_symbol_location (env : Environment) : TestT IO Unit := do
|
def test_symbol_location : TestT IO Unit := do
|
||||||
|
let env: Environment ← importModules
|
||||||
|
(imports := #[`Init])
|
||||||
|
(opts := {})
|
||||||
|
(trustLevel := 1)
|
||||||
addTest $ ← runTestCoreM env do
|
addTest $ ← runTestCoreM env do
|
||||||
let .ok result ← (Environment.inspect { name := "Nat.le_of_succ_le", source? := .some true } (options := {})).run | fail "Inspect failed"
|
let .ok result ← (Environment.inspect { name := "Nat.le_of_succ_le", source? := .some true } (options := {})).run | fail "Inspect failed"
|
||||||
checkEq "module" result.module? <| .some "Init.Data.Nat.Basic"
|
checkEq "module" result.module? <| .some "Init.Data.Nat.Basic"
|
||||||
|
|
||||||
-- Extraction of source doesn't work for symbols in `Init` for some reason
|
-- Extraction of source doesn't work for symbols in `Init` for some reason
|
||||||
checkTrue "file" result.sourceUri?.isNone
|
checkTrue "file" result.sourceUri?.isNone
|
||||||
checkEq "sourceStart" (result.sourceStart?.map (·.column)) <| .some 0
|
checkEq "pos" (result.sourceStart?.map (·.column)) <| .some 0
|
||||||
checkEq "sourceEnd" (result.sourceEnd?.map (·.column)) <| .some 88
|
checkEq "pos" (result.sourceEnd?.map (·.column)) <| .some 88
|
||||||
let { imports, constNames, .. } ← Environment.moduleRead ⟨"Init.Data.Nat.Basic"⟩
|
let { imports, constNames, .. } ← Environment.moduleRead ⟨"Init.Data.Nat.Basic"⟩
|
||||||
checkEq "imports" imports #["Init.SimpLemmas", "Init.Data.NeZero", "Init.Grind.Tactics"]
|
checkEq "imports" imports #["Init.SimpLemmas", "Init.Data.NeZero"]
|
||||||
checkTrue "constNames" $ constNames.contains "Nat.succ_add"
|
checkTrue "constNames" $ constNames.contains "Nat.succ_add"
|
||||||
|
|
||||||
def test_matcher (env : Environment) : TestT IO Unit := do
|
def test_matcher : TestT IO Unit := do
|
||||||
|
let env: Environment ← importModules
|
||||||
|
(imports := #[`Init])
|
||||||
|
(opts := {})
|
||||||
|
(trustLevel := 1)
|
||||||
checkTrue "not matcher" $ ¬ Meta.isMatcherCore env `Nat.strongRecOn
|
checkTrue "not matcher" $ ¬ Meta.isMatcherCore env `Nat.strongRecOn
|
||||||
|
|
||||||
def suite (env : Environment) : List (String × IO LSpec.TestSeq) :=
|
def suite: List (String × IO LSpec.TestSeq) :=
|
||||||
[
|
[
|
||||||
("Catalog", test_catalog env),
|
("Catalog", test_catalog),
|
||||||
("Symbol Visibility", test_symbol_visibility),
|
("Symbol Visibility", test_symbol_visibility),
|
||||||
("Inspect", test_inspect env),
|
("Inspect", test_inspect),
|
||||||
("Symbol Location", runTest $ test_symbol_location env),
|
("Symbol Location", runTest test_symbol_location),
|
||||||
("Matcher", runTest $ test_matcher env),
|
("Matcher", runTest test_matcher),
|
||||||
]
|
]
|
||||||
|
|
||||||
end Pantograph.Test.Environment
|
end Pantograph.Test.Environment
|
||||||
|
|
|
@ -17,62 +17,53 @@ deriving instance Lean.ToJson for Protocol.GoalStart
|
||||||
deriving instance Lean.ToJson for Protocol.GoalPrint
|
deriving instance Lean.ToJson for Protocol.GoalPrint
|
||||||
deriving instance Lean.ToJson for Protocol.GoalTactic
|
deriving instance Lean.ToJson for Protocol.GoalTactic
|
||||||
deriving instance Lean.ToJson for Protocol.FrontendProcess
|
deriving instance Lean.ToJson for Protocol.FrontendProcess
|
||||||
deriving instance Lean.ToJson for Protocol.FrontendDataUnit
|
|
||||||
deriving instance Lean.ToJson for Protocol.FrontendData
|
|
||||||
|
|
||||||
abbrev TestM α := TestT MainM α
|
|
||||||
abbrev Test := TestM Unit
|
|
||||||
|
|
||||||
def getFileName : TestM String := do
|
|
||||||
return (← read).coreContext.fileName
|
|
||||||
|
|
||||||
def step { α β } [Lean.ToJson α] [Lean.ToJson β] (cmd: String) (payload: α)
|
def step { α β } [Lean.ToJson α] [Lean.ToJson β] (cmd: String) (payload: α)
|
||||||
(expected: β) (name? : Option String := .none) : TestM Unit := do
|
(expected: β) (name? : Option String := .none): MainM LSpec.TestSeq := do
|
||||||
let payload := Lean.toJson payload
|
let payload := Lean.toJson payload
|
||||||
let name := name?.getD s!"{cmd} {payload.compress}"
|
let name := name?.getD s!"{cmd} {payload.compress}"
|
||||||
let result ← Repl.execute { cmd, payload }
|
let result ← Repl.execute { cmd, payload }
|
||||||
checkEq name result.compress (Lean.toJson expected).compress
|
return LSpec.test name (result.pretty = (Lean.toJson expected).pretty)
|
||||||
def stepFile { α } [Lean.ToJson α] (name : String) (path : String)
|
|
||||||
(expected : α) : TestM Unit := do
|
abbrev Test := List (MainM LSpec.TestSeq)
|
||||||
let content ← IO.FS.readFile path
|
|
||||||
let payload? : Except String Lean.Json := Lean.Json.parse content
|
|
||||||
match payload? with
|
|
||||||
| .ok payload =>
|
|
||||||
let expected := Lean.toJson expected
|
|
||||||
checkEq name payload.compress expected.compress
|
|
||||||
| .error e => fail s!"{name} {e}"
|
|
||||||
|
|
||||||
def test_expr_echo : Test :=
|
def test_expr_echo : Test :=
|
||||||
step "expr.echo"
|
[
|
||||||
({ expr := "λ {α : Sort (u + 1)} => List α", levels? := .some #["u"]}: Protocol.ExprEcho)
|
step "expr.echo"
|
||||||
({
|
({ expr := "λ {α : Sort (u + 1)} => List α", levels? := .some #["u"]}: Protocol.ExprEcho)
|
||||||
type := { pp? := .some "{α : Type u} → Type u" },
|
({
|
||||||
expr := { pp? := .some "fun {α} => List α" }
|
type := { pp? := .some "{α : Type u} → Type u" },
|
||||||
}: Protocol.ExprEchoResult)
|
expr := { pp? := .some "fun {α} => List α" }
|
||||||
|
}: Protocol.ExprEchoResult),
|
||||||
|
]
|
||||||
|
|
||||||
def test_option_modify : Test := do
|
def test_option_modify : Test :=
|
||||||
let pp? := Option.some "∀ (n : Nat), n + 1 = n.succ"
|
let pp? := Option.some "∀ (n : Nat), n + 1 = n.succ"
|
||||||
let sexp? := Option.some "(: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)))"
|
let sexp? := Option.some "(: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)))"
|
||||||
let module? := Option.some "Init.Data.Nat.Basic"
|
let module? := Option.some "Init.Data.Nat.Basic"
|
||||||
let options: Protocol.Options := {}
|
let options: Protocol.Options := {}
|
||||||
step "env.inspect" ({ name := "Nat.add_one" } : Protocol.EnvInspect)
|
[
|
||||||
({ type := { pp? }, module? }: Protocol.EnvInspectResult)
|
step "env.inspect" ({ name := "Nat.add_one" } : Protocol.EnvInspect)
|
||||||
step "options.set" ({ printExprAST? := .some true } : Protocol.OptionsSet)
|
({ type := { pp? }, module? }: Protocol.EnvInspectResult),
|
||||||
({ }: Protocol.OptionsSetResult)
|
step "options.set" ({ printExprAST? := .some true } : Protocol.OptionsSet)
|
||||||
step "env.inspect" ({ name := "Nat.add_one" } : Protocol.EnvInspect)
|
({ }: Protocol.OptionsSetResult),
|
||||||
({ type := { pp?, sexp? }, module? }: Protocol.EnvInspectResult)
|
step "env.inspect" ({ name := "Nat.add_one" } : Protocol.EnvInspect)
|
||||||
step "options.print" ({} : Protocol.OptionsPrint)
|
({ type := { pp?, sexp? }, module? }: Protocol.EnvInspectResult),
|
||||||
({ options with printExprAST := true }: Protocol.Options)
|
step "options.print" ({} : Protocol.OptionsPrint)
|
||||||
|
({ options with printExprAST := true }: Protocol.Options),
|
||||||
def test_malformed_command : Test := do
|
]
|
||||||
|
def test_malformed_command : Test :=
|
||||||
let invalid := "invalid"
|
let invalid := "invalid"
|
||||||
step invalid ({ name := "Nat.add_one" }: Protocol.EnvInspect)
|
[
|
||||||
({ error := "command", desc := s!"Unknown command {invalid}" }: Protocol.InteractionError)
|
step invalid ({ name := "Nat.add_one" }: Protocol.EnvInspect)
|
||||||
(name? := .some "Invalid Command")
|
({ error := "command", desc := s!"Unknown command {invalid}" }: Protocol.InteractionError)
|
||||||
step "expr.echo" (Lean.Json.mkObj [(invalid, .str "Random garbage data")])
|
(name? := .some "Invalid Command"),
|
||||||
({ error := "command", desc := s!"Unable to parse json: Pantograph.Protocol.ExprEcho.expr: String expected" }:
|
step "expr.echo" (Lean.Json.mkObj [(invalid, .str "Random garbage data")])
|
||||||
Protocol.InteractionError) (name? := .some "JSON Deserialization")
|
({ error := "command", desc := s!"Unable to parse json: Pantograph.Protocol.ExprEcho.expr: String expected" }:
|
||||||
def test_tactic : Test := do
|
Protocol.InteractionError)
|
||||||
|
(name? := .some "JSON Deserialization")
|
||||||
|
]
|
||||||
|
def test_tactic : Test :=
|
||||||
let varX := { name := "_uniq.10", userName := "x", type? := .some { pp? := .some "Prop" }}
|
let varX := { name := "_uniq.10", userName := "x", type? := .some { pp? := .some "Prop" }}
|
||||||
let goal1: Protocol.Goal := {
|
let goal1: Protocol.Goal := {
|
||||||
name := "_uniq.11",
|
name := "_uniq.11",
|
||||||
|
@ -87,45 +78,38 @@ def test_tactic : Test := do
|
||||||
{ name := "_uniq.13", userName := "y", type? := .some { pp? := .some "Prop" }}
|
{ name := "_uniq.13", userName := "y", type? := .some { pp? := .some "Prop" }}
|
||||||
],
|
],
|
||||||
}
|
}
|
||||||
step "goal.start" ({ expr := "∀ (p q: Prop), p ∨ q → q ∨ p" }: Protocol.GoalStart)
|
[
|
||||||
({ stateId := 0, root := "_uniq.9" }: Protocol.GoalStartResult)
|
step "goal.start" ({ expr := "∀ (p q: Prop), p ∨ q → q ∨ p" }: Protocol.GoalStart)
|
||||||
step "goal.tactic" ({ stateId := 0, tactic? := .some "intro x" }: Protocol.GoalTactic)
|
({ stateId := 0, root := "_uniq.9" }: Protocol.GoalStartResult),
|
||||||
({ nextStateId? := .some 1, goals? := #[goal1], }: Protocol.GoalTacticResult)
|
step "goal.tactic" ({ stateId := 0, tactic? := .some "intro x" }: Protocol.GoalTactic)
|
||||||
step "goal.print" ({ stateId := 1, parentExprs? := .some true, rootExpr? := .some true }: Protocol.GoalPrint)
|
({ nextStateId? := .some 1, goals? := #[goal1], }: Protocol.GoalTacticResult),
|
||||||
({
|
step "goal.print" ({ stateId := 1, parentExpr? := .some true, rootExpr? := .some true }: Protocol.GoalPrint)
|
||||||
root? := .some { pp? := "fun x => ?m.11"},
|
({ parent? := .some { pp? := .some "fun x => ?m.11" }, }: Protocol.GoalPrintResult),
|
||||||
parentExprs? := .some [.some { pp? := .some "fun x => ?m.11" }],
|
step "goal.tactic" ({ stateId := 1, tactic? := .some "intro y" }: Protocol.GoalTactic)
|
||||||
}: Protocol.GoalPrintResult)
|
({ nextStateId? := .some 2, goals? := #[goal2], }: Protocol.GoalTacticResult),
|
||||||
step "goal.tactic" ({ stateId := 1, tactic? := .some "intro y" }: Protocol.GoalTactic)
|
step "goal.tactic" ({ stateId := 1, tactic? := .some "apply Nat.le_of_succ_le" }: Protocol.GoalTactic)
|
||||||
({ nextStateId? := .some 2, goals? := #[goal2], }: Protocol.GoalTacticResult)
|
({ tacticErrors? := .some #["tactic 'apply' failed, failed to unify\n ∀ {m : Nat}, Nat.succ ?n ≤ m → ?n ≤ m\nwith\n ∀ (q : Prop), x ∨ q → q ∨ x\nx : Prop\n⊢ ∀ (q : Prop), x ∨ q → q ∨ x"] }:
|
||||||
step "goal.tactic" ({ stateId := 1, tactic? := .some "apply Nat.le_of_succ_le" }: Protocol.GoalTactic)
|
Protocol.GoalTacticResult)
|
||||||
({
|
]
|
||||||
messages? := .some #[{
|
|
||||||
fileName := ← getFileName,
|
|
||||||
kind := .anonymous,
|
|
||||||
pos := ⟨0, 0⟩,
|
|
||||||
data := "tactic 'apply' failed, could not unify the conclusion of `@Nat.le_of_succ_le`\n ∀ {m : Nat}, Nat.succ ?n ≤ m → ?n ≤ m\nwith the goal\n ∀ (q : Prop), x ∨ q → q ∨ x\n\nNote: The full type of `@Nat.le_of_succ_le` is\n ∀ {n m : Nat}, n.succ ≤ m → n ≤ m\nx : Prop\n⊢ ∀ (q : Prop), x ∨ q → q ∨ x",
|
|
||||||
}]
|
|
||||||
}: Protocol.GoalTacticResult)
|
|
||||||
step "goal.tactic" ({ stateId := 0, tactic? := .some "sorry" }: Protocol.GoalTactic)
|
|
||||||
({ nextStateId? := .some 3, goals? := .some #[], hasSorry := true }: Protocol.GoalTacticResult)
|
|
||||||
example : (1 : Nat) + (2 * 3) = 1 + (4 - 3) + (6 - 4) + 3 := by
|
example : (1 : Nat) + (2 * 3) = 1 + (4 - 3) + (6 - 4) + 3 := by
|
||||||
simp
|
simp
|
||||||
def test_tactic_timeout : Test := do
|
def test_tactic_timeout : Test :=
|
||||||
step "goal.start" ({ expr := "(1 : Nat) + (2 * 3) = 1 + (4 - 3) + (6 - 4) + 3" }: Protocol.GoalStart)
|
[
|
||||||
({ stateId := 0, root := "_uniq.370" }: Protocol.GoalStartResult)
|
step "goal.start" ({ expr := "(1 : Nat) + (2 * 3) = 1 + (4 - 3) + (6 - 4) + 3" }: Protocol.GoalStart)
|
||||||
-- timeout of 10 milliseconds
|
({ stateId := 0, root := "_uniq.319" }: Protocol.GoalStartResult),
|
||||||
step "options.set" ({ timeout? := .some 10 } : Protocol.OptionsSet)
|
-- timeout of 10 milliseconds
|
||||||
({ }: Protocol.OptionsSetResult)
|
step "options.set" ({ timeout? := .some 10 } : Protocol.OptionsSet)
|
||||||
step "goal.tactic" ({ stateId := 0, expr? := .some "by\nsleep 1000; simp" }: Protocol.GoalTactic)
|
({ }: Protocol.OptionsSetResult),
|
||||||
({ error := "internal", desc := "interrupt" }: Protocol.InteractionError)
|
step "goal.tactic" ({ stateId := 0, expr? := .some "by\nsleep 1000; simp" }: Protocol.GoalTactic)
|
||||||
-- ensure graceful recovery
|
({ error := "internal", desc := "interrupt" }: Protocol.InteractionError),
|
||||||
step "options.set" ({ timeout? := .some 0 } : Protocol.OptionsSet)
|
-- ensure graceful recovery
|
||||||
({ }: Protocol.OptionsSetResult)
|
step "options.set" ({ timeout? := .some 0 } : Protocol.OptionsSet)
|
||||||
step "goal.tactic" ({ stateId := 0, tactic? := .some "simp" }: Protocol.GoalTactic)
|
({ }: Protocol.OptionsSetResult),
|
||||||
({ nextStateId? := .some 1, goals? := .some #[], }: Protocol.GoalTacticResult)
|
step "goal.tactic" ({ stateId := 0, tactic? := .some "simp" }: Protocol.GoalTactic)
|
||||||
|
({ nextStateId? := .some 1, goals? := .some #[], }: Protocol.GoalTacticResult),
|
||||||
|
]
|
||||||
|
|
||||||
def test_automatic_mode (automatic: Bool): Test := do
|
def test_automatic_mode (automatic: Bool): Test :=
|
||||||
let varsPQ := #[
|
let varsPQ := #[
|
||||||
{ name := "_uniq.10", userName := "p", type? := .some { pp? := .some "Prop" }},
|
{ name := "_uniq.10", userName := "p", type? := .some { pp? := .some "Prop" }},
|
||||||
{ name := "_uniq.13", userName := "q", type? := .some { pp? := .some "Prop" }}
|
{ name := "_uniq.13", userName := "q", type? := .some { pp? := .some "Prop" }}
|
||||||
|
@ -161,248 +145,168 @@ def test_automatic_mode (automatic: Bool): Test := do
|
||||||
{ name := "_uniq.49", userName := "h✝", type? := .some { pp? := .some "p" }, isInaccessible := true}
|
{ name := "_uniq.49", userName := "h✝", type? := .some { pp? := .some "p" }, isInaccessible := true}
|
||||||
],
|
],
|
||||||
}
|
}
|
||||||
step "options.set" ({automaticMode? := .some automatic}: Protocol.OptionsSet)
|
[
|
||||||
({}: Protocol.OptionsSetResult)
|
step "options.set" ({automaticMode? := .some automatic}: Protocol.OptionsSet)
|
||||||
step "goal.start" ({ expr := "∀ (p q: Prop), p ∨ q → q ∨ p"} : Protocol.GoalStart)
|
({}: Protocol.OptionsSetResult),
|
||||||
({ stateId := 0, root := "_uniq.9" }: Protocol.GoalStartResult)
|
step "goal.start" ({ expr := "∀ (p q: Prop), p ∨ q → q ∨ p"} : Protocol.GoalStart)
|
||||||
step "goal.tactic" ({ stateId := 0, tactic? := .some "intro p q h" }: Protocol.GoalTactic)
|
({ stateId := 0, root := "_uniq.9" }: Protocol.GoalStartResult),
|
||||||
({ nextStateId? := .some 1, goals? := #[goal1], }: Protocol.GoalTacticResult)
|
step "goal.tactic" ({ stateId := 0, tactic? := .some "intro p q h" }: Protocol.GoalTactic)
|
||||||
step "goal.tactic" ({ stateId := 1, tactic? := .some "cases h" }: Protocol.GoalTactic)
|
({ nextStateId? := .some 1, goals? := #[goal1], }: Protocol.GoalTacticResult),
|
||||||
({ nextStateId? := .some 2, goals? := #[goal2l, goal2r], }: Protocol.GoalTacticResult)
|
step "goal.tactic" ({ stateId := 1, tactic? := .some "cases h" }: Protocol.GoalTactic)
|
||||||
let goals? := if automatic then #[goal3l, goal2r] else #[goal3l]
|
({ nextStateId? := .some 2, goals? := #[goal2l, goal2r], }: Protocol.GoalTacticResult),
|
||||||
step "goal.tactic" ({ stateId := 2, tactic? := .some "apply Or.inr" }: Protocol.GoalTactic)
|
let goals? := if automatic then #[goal3l, goal2r] else #[goal3l]
|
||||||
({ nextStateId? := .some 3, goals?, }: Protocol.GoalTacticResult)
|
step "goal.tactic" ({ stateId := 2, tactic? := .some "apply Or.inr" }: Protocol.GoalTactic)
|
||||||
|
({ nextStateId? := .some 3, goals?, }: Protocol.GoalTacticResult),
|
||||||
def test_conv_calc : Test := do
|
|
||||||
step "options.set" ({automaticMode? := .some false}: Protocol.OptionsSet)
|
|
||||||
({}: Protocol.OptionsSetResult)
|
|
||||||
step "goal.start" ({ expr := "∀ (a b: Nat), (b = 2) -> 1 + a + 1 = a + b"} : Protocol.GoalStart)
|
|
||||||
({ stateId := 0, root := "_uniq.171" }: Protocol.GoalStartResult)
|
|
||||||
let vars := #[
|
|
||||||
{ name := "_uniq.172", userName := "a", type? := .some { pp? := .some "Nat" }},
|
|
||||||
{ name := "_uniq.175", userName := "b", type? := .some { pp? := .some "Nat" }},
|
|
||||||
{ name := "_uniq.178", userName := "h", type? := .some { pp? := .some "b = 2" }},
|
|
||||||
]
|
]
|
||||||
let goal : Protocol.Goal := {
|
|
||||||
vars,
|
|
||||||
name := "_uniq.179",
|
|
||||||
target := { pp? := "1 + a + 1 = a + b" },
|
|
||||||
}
|
|
||||||
step "goal.tactic" ({ stateId := 0, tactic? := .some "intro a b h" }: Protocol.GoalTactic)
|
|
||||||
({ nextStateId? := .some 1, goals? := #[goal], }: Protocol.GoalTacticResult)
|
|
||||||
step "goal.tactic" ({ stateId := 1, mode? := .some "calc" }: Protocol.GoalTactic)
|
|
||||||
({ nextStateId? := .some 2, goals? := #[{ goal with fragment := .calc }], }: Protocol.GoalTacticResult)
|
|
||||||
let goalCalc : Protocol.Goal := {
|
|
||||||
vars,
|
|
||||||
name := "_uniq.381",
|
|
||||||
userName? := .some "calc",
|
|
||||||
target := { pp? := "1 + a + 1 = a + 1 + 1" },
|
|
||||||
}
|
|
||||||
let goalMain : Protocol.Goal := {
|
|
||||||
vars,
|
|
||||||
name := "_uniq.400",
|
|
||||||
fragment := .calc,
|
|
||||||
target := { pp? := "a + 1 + 1 = a + b" },
|
|
||||||
}
|
|
||||||
step "goal.tactic" ({ stateId := 2, tactic? := .some "1 + a + 1 = a + 1 + 1" }: Protocol.GoalTactic)
|
|
||||||
({ nextStateId? := .some 3, goals? := #[goalCalc, goalMain], }: Protocol.GoalTacticResult)
|
|
||||||
let goalConv : Protocol.Goal := {
|
|
||||||
goalCalc with
|
|
||||||
fragment := .conv,
|
|
||||||
userName? := .none,
|
|
||||||
name := "_uniq.468",
|
|
||||||
}
|
|
||||||
step "goal.tactic" ({ stateId := 3, mode? := .some "conv" }: Protocol.GoalTactic)
|
|
||||||
({ nextStateId? := .some 4, goals? := #[goalConv], }: Protocol.GoalTacticResult)
|
|
||||||
|
|
||||||
def test_env_add_inspect : Test := do
|
def test_env_add_inspect : Test :=
|
||||||
let name1 := "Pantograph.mystery"
|
let name1 := "Pantograph.mystery"
|
||||||
let name2 := "Pantograph.mystery2"
|
let name2 := "Pantograph.mystery2"
|
||||||
let name3 := "Pantograph.mystery3"
|
let name3 := "Pantograph.mystery3"
|
||||||
step "env.add"
|
[
|
||||||
({
|
step "env.add"
|
||||||
name := name1,
|
({
|
||||||
value := "λ (a b: Prop) => Or a b",
|
name := name1,
|
||||||
isTheorem := false
|
value := "λ (a b: Prop) => Or a b",
|
||||||
}: Protocol.EnvAdd)
|
isTheorem := false
|
||||||
({}: Protocol.EnvAddResult)
|
}: Protocol.EnvAdd)
|
||||||
step "env.inspect" ({name := name1, value? := .some true} : Protocol.EnvInspect)
|
({}: Protocol.EnvAddResult),
|
||||||
({
|
step "env.inspect" ({name := name1, value? := .some true} : Protocol.EnvInspect)
|
||||||
value? := .some { pp? := .some "fun a b => a ∨ b" },
|
({
|
||||||
type := { pp? := .some "Prop → Prop → Prop" },
|
value? := .some { pp? := .some "fun a b => a ∨ b" },
|
||||||
}: Protocol.EnvInspectResult)
|
type := { pp? := .some "Prop → Prop → Prop" },
|
||||||
step "env.add"
|
}:
|
||||||
({
|
Protocol.EnvInspectResult),
|
||||||
name := name2,
|
step "env.add"
|
||||||
type? := "Nat → Int",
|
({
|
||||||
value := "λ (a: Nat) => a + 1",
|
name := name2,
|
||||||
isTheorem := false
|
type? := "Nat → Int",
|
||||||
}: Protocol.EnvAdd)
|
value := "λ (a: Nat) => a + 1",
|
||||||
({}: Protocol.EnvAddResult)
|
isTheorem := false
|
||||||
step "env.inspect" ({name := name2, value? := .some true} : Protocol.EnvInspect)
|
}: Protocol.EnvAdd)
|
||||||
({
|
({}: Protocol.EnvAddResult),
|
||||||
value? := .some { pp? := .some "fun a => ↑a + 1" },
|
step "env.inspect" ({name := name2, value? := .some true} : Protocol.EnvInspect)
|
||||||
type := { pp? := .some "Nat → Int" },
|
({
|
||||||
}: Protocol.EnvInspectResult)
|
value? := .some { pp? := .some "fun a => ↑a + 1" },
|
||||||
step "env.add"
|
type := { pp? := .some "Nat → Int" },
|
||||||
({
|
}:
|
||||||
name := name3,
|
Protocol.EnvInspectResult),
|
||||||
levels? := .some #["u"]
|
step "env.add"
|
||||||
type? := "(α : Type u) → α → (α × α)",
|
({
|
||||||
value := "λ (α : Type u) (x : α) => (x, x)",
|
name := name3,
|
||||||
isTheorem := false
|
levels? := .some #["u"]
|
||||||
}: Protocol.EnvAdd)
|
type? := "(α : Type u) → α → (α × α)",
|
||||||
({}: Protocol.EnvAddResult)
|
value := "λ (α : Type u) (x : α) => (x, x)",
|
||||||
step "env.inspect" ({name := name3} : Protocol.EnvInspect)
|
isTheorem := false
|
||||||
({
|
}: Protocol.EnvAdd)
|
||||||
type := { pp? := .some "(α : Type u) → α → α × α" },
|
({}: Protocol.EnvAddResult),
|
||||||
}: Protocol.EnvInspectResult)
|
step "env.inspect" ({name := name3} : Protocol.EnvInspect)
|
||||||
|
({
|
||||||
|
type := { pp? := .some "(α : Type u) → α → α × α" },
|
||||||
|
}:
|
||||||
|
Protocol.EnvInspectResult),
|
||||||
|
]
|
||||||
|
|
||||||
example : ∀ (p: Prop), p → p := by
|
example : ∀ (p: Prop), p → p := by
|
||||||
intro p h
|
intro p h
|
||||||
exact h
|
exact h
|
||||||
|
|
||||||
def test_frontend_process : Test := do
|
def test_frontend_process : Test :=
|
||||||
let file := "example : ∀ (p q: Prop), p → p ∨ q := by\n intro p q h\n exact Or.inl h"
|
let file := "example : ∀ (p q: Prop), p → p ∨ q := by\n intro p q h\n exact Or.inl h"
|
||||||
let goal1 := "p q : Prop\nh : p\n⊢ p ∨ q"
|
let goal1 := "p q : Prop\nh : p\n⊢ p ∨ q"
|
||||||
IO.FS.withTempDir λ tempdir => do
|
[
|
||||||
let filename := s!"{tempdir}/invocations.jsonl"
|
step "frontend.process"
|
||||||
step "frontend.process"
|
({
|
||||||
({
|
file? := .some file,
|
||||||
file? := .some file,
|
invocations := true,
|
||||||
invocations? := .some filename,
|
}: Protocol.FrontendProcess)
|
||||||
}: Protocol.FrontendProcess)
|
({
|
||||||
({
|
units := [{
|
||||||
units := [{
|
boundary := (0, file.utf8ByteSize),
|
||||||
boundary := (0, file.utf8ByteSize),
|
invocations? := .some [
|
||||||
nInvocations? := .some 2,
|
{
|
||||||
}],
|
goalBefore := "⊢ ∀ (p q : Prop), p → p ∨ q",
|
||||||
}: Protocol.FrontendProcessResult)
|
goalAfter := goal1,
|
||||||
stepFile (α := Protocol.FrontendData) "invocations" filename
|
tactic := "intro p q h",
|
||||||
{ units := [{
|
usedConstants := #[],
|
||||||
invocations? := .some [
|
},
|
||||||
{
|
{
|
||||||
goalBefore := "⊢ ∀ (p q : Prop), p → p ∨ q",
|
goalBefore := goal1 ,
|
||||||
goalAfter := goal1,
|
goalAfter := "",
|
||||||
tactic := "intro p q h",
|
tactic := "exact Or.inl h",
|
||||||
usedConstants := #[],
|
usedConstants := #["Or.inl"],
|
||||||
},
|
},
|
||||||
{
|
]
|
||||||
goalBefore := goal1 ,
|
}],
|
||||||
goalAfter := "",
|
}: Protocol.FrontendProcessResult),
|
||||||
tactic := "exact Or.inl h",
|
]
|
||||||
usedConstants := #["Or.inl"],
|
|
||||||
},
|
|
||||||
]
|
|
||||||
} ] }
|
|
||||||
|
|
||||||
example : 1 + 2 = 3 := rfl
|
example : 1 + 2 = 3 := rfl
|
||||||
example (p: Prop): p → p := by simp
|
example (p: Prop): p → p := by simp
|
||||||
|
|
||||||
def test_frontend_process_sorry : Test := do
|
def test_frontend_process_sorry : Test :=
|
||||||
let solved := "example : 1 + 2 = 3 := rfl\n"
|
let solved := "example : 1 + 2 = 3 := rfl\n"
|
||||||
let withSorry := "example (p: Prop): p → p := sorry"
|
let withSorry := "example (p: Prop): p → p := sorry"
|
||||||
let file := s!"{solved}{withSorry}"
|
[
|
||||||
let goal1: Protocol.Goal := {
|
let file := s!"{solved}{withSorry}"
|
||||||
name := "_uniq.6",
|
let goal1: Protocol.Goal := {
|
||||||
target := { pp? := .some "p → p" },
|
name := "_uniq.6",
|
||||||
vars := #[{ name := "_uniq.4", userName := "p", type? := .some { pp? := .some "Prop" }}],
|
target := { pp? := .some "p → p" },
|
||||||
}
|
vars := #[{ name := "_uniq.4", userName := "p", type? := .some { pp? := .some "Prop" }}],
|
||||||
step "frontend.process"
|
}
|
||||||
({
|
step "frontend.process"
|
||||||
file? := .some file,
|
({
|
||||||
sorrys := true,
|
file? := .some file,
|
||||||
}: Protocol.FrontendProcess)
|
sorrys := true,
|
||||||
({
|
}: Protocol.FrontendProcess)
|
||||||
units := [{
|
({
|
||||||
boundary := (0, solved.utf8ByteSize),
|
units := [{
|
||||||
}, {
|
boundary := (0, solved.utf8ByteSize),
|
||||||
boundary := (solved.utf8ByteSize, solved.utf8ByteSize + withSorry.utf8ByteSize),
|
}, {
|
||||||
goalStateId? := .some 0,
|
boundary := (solved.utf8ByteSize, solved.utf8ByteSize + withSorry.utf8ByteSize),
|
||||||
goals? := .some #[goal1],
|
goalStateId? := .some 0,
|
||||||
goalSrcBoundaries? := .some #[(57, 62)],
|
goals? := .some #[goal1],
|
||||||
messages := #[{
|
goalSrcBoundaries? := .some #[(57, 62)],
|
||||||
fileName := "<anonymous>",
|
messages := #["<anonymous>:2:0: warning: declaration uses 'sorry'\n"],
|
||||||
kind := `hasSorry,
|
|
||||||
pos := ⟨2, 0⟩,
|
|
||||||
endPos := .some ⟨2, 7⟩,
|
|
||||||
severity := .warning,
|
|
||||||
data := "declaration uses 'sorry'",
|
|
||||||
}],
|
}],
|
||||||
}],
|
}: Protocol.FrontendProcessResult),
|
||||||
}: Protocol.FrontendProcessResult)
|
]
|
||||||
|
|
||||||
def test_import_open : Test := do
|
def test_import_open : Test :=
|
||||||
let header := "import Init\nopen Nat\nuniverse u"
|
let header := "import Init\nopen Nat\nuniverse u"
|
||||||
let goal1: Protocol.Goal := {
|
let goal1: Protocol.Goal := {
|
||||||
name := "_uniq.81",
|
name := "_uniq.67",
|
||||||
target := { pp? := .some "n + 1 = n.succ" },
|
target := { pp? := .some "n + 1 = n.succ" },
|
||||||
vars := #[{ name := "_uniq.80", userName := "n", type? := .some { pp? := .some "Nat" }}],
|
vars := #[{ name := "_uniq.66", userName := "n", type? := .some { pp? := .some "Nat" }}],
|
||||||
}
|
}
|
||||||
step "frontend.process"
|
[
|
||||||
({
|
step "frontend.process"
|
||||||
file? := .some header,
|
({
|
||||||
readHeader := true,
|
file? := .some header,
|
||||||
inheritEnv := true,
|
readHeader := true,
|
||||||
}: Protocol.FrontendProcess)
|
inheritEnv := true,
|
||||||
({
|
}: Protocol.FrontendProcess)
|
||||||
units := [
|
({
|
||||||
{ boundary := (12, 21) },
|
units := [
|
||||||
{ boundary := (21, header.utf8ByteSize) },
|
{ boundary := (12, 21) },
|
||||||
],
|
{ boundary := (21, header.utf8ByteSize) },
|
||||||
}: Protocol.FrontendProcessResult)
|
],
|
||||||
step "goal.start" ({ expr := "∀ (n : Nat), n + 1 = Nat.succ n"} : Protocol.GoalStart)
|
}: Protocol.FrontendProcessResult),
|
||||||
({ stateId := 0, root := "_uniq.79" }: Protocol.GoalStartResult)
|
step "goal.start" ({ expr := "∀ (n : Nat), n + 1 = Nat.succ n"} : Protocol.GoalStart)
|
||||||
step "goal.tactic" ({ stateId := 0, tactic? := .some "intro n" }: Protocol.GoalTactic)
|
({ stateId := 0, root := "_uniq.65" }: Protocol.GoalStartResult),
|
||||||
({ nextStateId? := .some 1, goals? := #[goal1], }: Protocol.GoalTacticResult)
|
step "goal.tactic" ({ stateId := 0, tactic? := .some "intro n" }: Protocol.GoalTactic)
|
||||||
step "goal.tactic" ({ stateId := 1, tactic? := .some "apply add_one" }: Protocol.GoalTactic)
|
({ nextStateId? := .some 1, goals? := #[goal1], }: Protocol.GoalTacticResult),
|
||||||
({ nextStateId? := .some 2, goals? := .some #[], }: Protocol.GoalTacticResult)
|
step "goal.tactic" ({ stateId := 1, tactic? := .some "apply add_one" }: Protocol.GoalTactic)
|
||||||
step "goal.start" ({ expr := "∀ (x : Sort u), Sort (u + 1)"} : Protocol.GoalStart)
|
({ nextStateId? := .some 2, goals? := .some #[], }: Protocol.GoalTacticResult),
|
||||||
({ stateId := 3, root := "_uniq.5" }: Protocol.GoalStartResult)
|
step "goal.start" ({ expr := "∀ (x : Sort u), Sort (u + 1)"} : Protocol.GoalStart)
|
||||||
|
({ stateId := 3, root := "_uniq.5" }: Protocol.GoalStartResult),
|
||||||
|
]
|
||||||
|
|
||||||
/-- Ensure there cannot be circular references -/
|
def runTest (env: Lean.Environment) (steps: Test): IO LSpec.TestSeq := do
|
||||||
def test_frontend_process_circular : Test := do
|
|
||||||
let withSorry := "theorem mystery : 1 + 2 = 2 + 3 := sorry"
|
|
||||||
let goal1: Protocol.Goal := {
|
|
||||||
name := "_uniq.2",
|
|
||||||
target := { pp? := .some "1 + 2 = 2 + 3" },
|
|
||||||
vars := #[],
|
|
||||||
}
|
|
||||||
step "frontend.process"
|
|
||||||
({
|
|
||||||
file? := .some withSorry,
|
|
||||||
sorrys := true,
|
|
||||||
}: Protocol.FrontendProcess)
|
|
||||||
({
|
|
||||||
units := [{
|
|
||||||
boundary := (0, withSorry.utf8ByteSize),
|
|
||||||
goalStateId? := .some 0,
|
|
||||||
goals? := .some #[goal1],
|
|
||||||
goalSrcBoundaries? := .some #[(35, 40)],
|
|
||||||
messages := #[{
|
|
||||||
fileName := "<anonymous>",
|
|
||||||
kind := `hasSorry,
|
|
||||||
pos := ⟨1, 8⟩,
|
|
||||||
endPos := .some ⟨1, 15⟩,
|
|
||||||
severity := .warning,
|
|
||||||
data := "declaration uses 'sorry'"
|
|
||||||
}],
|
|
||||||
}],
|
|
||||||
} : Protocol.FrontendProcessResult)
|
|
||||||
step "goal.tactic" ({ stateId := 0, tactic? := .some "exact?" }: Protocol.GoalTactic)
|
|
||||||
({
|
|
||||||
messages? := .some #[{
|
|
||||||
fileName := ← getFileName,
|
|
||||||
kind := .anonymous,
|
|
||||||
pos := ⟨0, 0⟩,
|
|
||||||
data := "`exact?` could not close the goal. Try `apply?` to see partial suggestions."
|
|
||||||
}]
|
|
||||||
} : Protocol.GoalTacticResult)
|
|
||||||
|
|
||||||
def runTestSuite (env : Lean.Environment) (steps : Test): IO LSpec.TestSeq := do
|
|
||||||
-- Setup the environment for execution
|
-- Setup the environment for execution
|
||||||
let coreContext ← createCoreContext #[]
|
let coreContext ← createCoreContext #[]
|
||||||
let mainM : MainM LSpec.TestSeq := runTest steps
|
let mainM : MainM LSpec.TestSeq :=
|
||||||
|
steps.foldlM (λ suite step => do return suite ++ (← step)) LSpec.TestSeq.done
|
||||||
mainM.run { coreContext } |>.run' { env }
|
mainM.run { coreContext } |>.run' { env }
|
||||||
|
|
||||||
def suite (env : Lean.Environment): List (String × IO LSpec.TestSeq) :=
|
def suite (env : Lean.Environment): List (String × IO LSpec.TestSeq) :=
|
||||||
|
@ -414,14 +318,12 @@ def suite (env : Lean.Environment): List (String × IO LSpec.TestSeq) :=
|
||||||
("Tactic Timeout", test_tactic_timeout),
|
("Tactic Timeout", test_tactic_timeout),
|
||||||
("Manual Mode", test_automatic_mode false),
|
("Manual Mode", test_automatic_mode false),
|
||||||
("Automatic Mode", test_automatic_mode true),
|
("Automatic Mode", test_automatic_mode true),
|
||||||
("Conv-Calc", test_conv_calc),
|
|
||||||
("env.add env.inspect", test_env_add_inspect),
|
("env.add env.inspect", test_env_add_inspect),
|
||||||
("frontend.process invocation", test_frontend_process),
|
("frontend.process invocation", test_frontend_process),
|
||||||
("frontend.process sorry", test_frontend_process_sorry),
|
("frontend.process sorry", test_frontend_process_sorry),
|
||||||
("frontend.process import", test_import_open),
|
("frontend.process import", test_import_open),
|
||||||
("frontend.process circular", test_frontend_process_circular),
|
|
||||||
]
|
]
|
||||||
tests.map (fun (name, test) => (name, runTestSuite env test))
|
tests.map (fun (name, test) => (name, runTest env test))
|
||||||
|
|
||||||
|
|
||||||
end Pantograph.Test.Integration
|
end Pantograph.Test.Integration
|
||||||
|
|
|
@ -39,28 +39,22 @@ open Pantograph.Test
|
||||||
def main (args: List String) := do
|
def main (args: List String) := do
|
||||||
let nameFilter? := args.head?
|
let nameFilter? := args.head?
|
||||||
Lean.initSearchPath (← Lean.findSysroot)
|
Lean.initSearchPath (← Lean.findSysroot)
|
||||||
let env_default : Lean.Environment ← Lean.importModules
|
let env_default: Lean.Environment ← Lean.importModules
|
||||||
(imports := #[`Init])
|
(imports := #[`Init])
|
||||||
(opts := {})
|
(opts := {})
|
||||||
(trustLevel := 1)
|
(trustLevel := 1)
|
||||||
(loadExts := true)
|
|
||||||
|
|
||||||
let suites: List (String × (Lean.Environment → List (String × IO LSpec.TestSeq))) := [
|
let suites: List (String × List (String × IO LSpec.TestSeq)) := [
|
||||||
("Environment", Environment.suite),
|
("Environment", Environment.suite),
|
||||||
("Frontend", Frontend.suite),
|
("Frontend", Frontend.suite env_default),
|
||||||
("Integration", Integration.suite),
|
("Integration", Integration.suite env_default),
|
||||||
("Library", Library.suite),
|
("Library", Library.suite env_default),
|
||||||
("Metavar", Metavar.suite),
|
("Metavar", Metavar.suite env_default),
|
||||||
("Proofs", Proofs.suite),
|
("Proofs", Proofs.suite env_default),
|
||||||
("Delate", Delate.suite),
|
("Delate", Delate.suite env_default),
|
||||||
("Serial", Serial.suite),
|
("Serial", Serial.suite env_default),
|
||||||
("Tactic/Assign", Tactic.Assign.suite),
|
("Tactic/Assign", Tactic.Assign.suite env_default),
|
||||||
("Tactic/Fragment", Tactic.Fragment.suite),
|
("Tactic/Prograde", Tactic.Prograde.suite env_default),
|
||||||
("Tactic/Prograde", Tactic.Prograde.suite),
|
|
||||||
("Tactic/Special", Tactic.Special.suite),
|
|
||||||
]
|
]
|
||||||
let suiterunner (f : Lean.Environment → List (String × IO LSpec.TestSeq)) :=
|
let tests: List (String × IO LSpec.TestSeq) := suites.foldl (λ acc (name, suite) => acc ++ (addPrefix name suite)) []
|
||||||
f env_default
|
|
||||||
let tests : List (String × IO LSpec.TestSeq) := suites.foldl (init := []) λ acc (name, suite) =>
|
|
||||||
acc ++ (addPrefix name $ suiterunner suite)
|
|
||||||
LSpec.lspecEachIO [()] (λ () => runTestGroup nameFilter? tests)
|
LSpec.lspecEachIO [()] (λ () => runTestGroup nameFilter? tests)
|
||||||
|
|
|
@ -25,8 +25,8 @@ def test_instantiate_mvar: TestM Unit := do
|
||||||
addTest $ assertUnreachable e
|
addTest $ assertUnreachable e
|
||||||
return ()
|
return ()
|
||||||
let t ← Lean.Meta.inferType expr
|
let t ← Lean.Meta.inferType expr
|
||||||
checkEq "typing" (toString (← serializeExpressionSexp t))
|
addTest $ LSpec.check "typing" ((toString (← serializeExpressionSexp t)) =
|
||||||
"((:c LE.le) (:c Nat) (:c instLENat) ((:c OfNat.ofNat) (:mv _uniq.2) (:lit 2) (:mv _uniq.3)) ((:c OfNat.ofNat) (:mv _uniq.15) (:lit 5) (:mv _uniq.16)))"
|
"((: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
|
||||||
|
@ -79,20 +79,20 @@ def test_m_couple: TestM Unit := do
|
||||||
return ()
|
return ()
|
||||||
|
|
||||||
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := "apply Nat.le_trans") with
|
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := "apply Nat.le_trans") with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
addTest $ LSpec.check "apply Nat.le_trans" ((← state1.serializeGoals (options := ← read)).map (·.target.pp?) =
|
addTest $ LSpec.check "apply Nat.le_trans" ((← state1.serializeGoals (options := ← read)).map (·.target.pp?) =
|
||||||
#[.some "2 ≤ ?m", .some "?m ≤ 5", .some "Nat"])
|
#[.some "2 ≤ ?m", .some "?m ≤ 5", .some "Nat"])
|
||||||
checkTrue "(1 root)" $ ¬ state1.isSolved
|
addTest $ LSpec.test "(1 root)" state1.rootExpr?.isNone
|
||||||
-- Set m to 3
|
-- Set m to 3
|
||||||
let state2 ← match ← state1.tacticOn (goalId := 2) (tactic := "exact 3") with
|
let state2 ← match ← state1.tacticOn (goalId := 2) (tactic := "exact 3") with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
checkTrue "(1b root)" $ ¬ state2.isSolved
|
addTest $ LSpec.test "(1b root)" state2.rootExpr?.isNone
|
||||||
let state1b ← match state2.continue state1 with
|
let state1b ← match state2.continue state1 with
|
||||||
| .error msg => do
|
| .error msg => do
|
||||||
addTest $ assertUnreachable $ msg
|
addTest $ assertUnreachable $ msg
|
||||||
|
@ -100,7 +100,7 @@ def test_m_couple: TestM Unit := do
|
||||||
| .ok state => pure state
|
| .ok state => pure state
|
||||||
addTest $ LSpec.check "exact 3" ((← state1b.serializeGoals (options := ← read)).map (·.target.pp?) =
|
addTest $ LSpec.check "exact 3" ((← state1b.serializeGoals (options := ← read)).map (·.target.pp?) =
|
||||||
#[.some "2 ≤ 3", .some "3 ≤ 5"])
|
#[.some "2 ≤ 3", .some "3 ≤ 5"])
|
||||||
checkTrue "(2 root)" $ ¬ state1b.isSolved
|
addTest $ LSpec.test "(2 root)" state1b.rootExpr?.isNone
|
||||||
|
|
||||||
def test_m_couple_simp: TestM Unit := do
|
def test_m_couple_simp: TestM Unit := do
|
||||||
let state? ← startProof "(2: Nat) ≤ 5"
|
let state? ← startProof "(2: Nat) ≤ 5"
|
||||||
|
@ -111,23 +111,22 @@ def test_m_couple_simp: TestM Unit := do
|
||||||
return ()
|
return ()
|
||||||
|
|
||||||
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := "apply Nat.le_trans") with
|
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := "apply Nat.le_trans") with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
let serializedState1 ← state1.serializeGoals (options := { ← read with printDependentMVars := true })
|
let serializedState1 ← state1.serializeGoals (options := { ← read with printDependentMVars := true })
|
||||||
addTest $ LSpec.check "apply Nat.le_trans" (serializedState1.map (·.target.pp?) =
|
addTest $ LSpec.check "apply Nat.le_trans" (serializedState1.map (·.target.pp?) =
|
||||||
#[.some "2 ≤ ?m", .some "?m ≤ 5", .some "Nat"])
|
#[.some "2 ≤ ?m", .some "?m ≤ 5", .some "Nat"])
|
||||||
let n := state1.goals[2]!
|
|
||||||
addTest $ LSpec.check "(metavariables)" (serializedState1.map (·.target.dependentMVars?.get!) =
|
addTest $ LSpec.check "(metavariables)" (serializedState1.map (·.target.dependentMVars?.get!) =
|
||||||
#[#[toString n.name], #[toString n.name], #[]])
|
#[#["_uniq.38"], #["_uniq.38"], #[]])
|
||||||
|
|
||||||
let state2 ← match ← state1.tacticOn (goalId := 2) (tactic := "exact 2") with
|
let state2 ← match ← state1.tacticOn (goalId := 2) (tactic := "exact 2") with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
checkTrue "(1b root)" $ ¬ state2.isSolved
|
addTest $ LSpec.test "(1b root)" state2.rootExpr?.isNone
|
||||||
let state1b ← match state2.continue state1 with
|
let state1b ← match state2.continue state1 with
|
||||||
| .error msg => do
|
| .error msg => do
|
||||||
addTest $ assertUnreachable $ msg
|
addTest $ assertUnreachable $ msg
|
||||||
|
@ -135,9 +134,9 @@ def test_m_couple_simp: TestM Unit := do
|
||||||
| .ok state => pure state
|
| .ok state => pure state
|
||||||
addTest $ LSpec.check "exact 2" ((← state1b.serializeGoals (options := ← read)).map (·.target.pp?) =
|
addTest $ LSpec.check "exact 2" ((← state1b.serializeGoals (options := ← read)).map (·.target.pp?) =
|
||||||
#[.some "2 ≤ 2", .some "2 ≤ 5"])
|
#[.some "2 ≤ 2", .some "2 ≤ 5"])
|
||||||
checkTrue "(2 root)" $ ¬ state1b.isSolved
|
addTest $ LSpec.test "(2 root)" state1b.rootExpr?.isNone
|
||||||
let state3 ← match ← state1b.tacticOn (goalId := 0) (tactic := "simp") with
|
let state3 ← match ← state1b.tacticOn (goalId := 0) (tactic := "simp") with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -147,7 +146,7 @@ def test_m_couple_simp: TestM Unit := do
|
||||||
return ()
|
return ()
|
||||||
| .ok state => pure state
|
| .ok state => pure state
|
||||||
let state5 ← match ← state4.tacticOn (goalId := 0) (tactic := "simp") with
|
let state5 ← match ← state4.tacticOn (goalId := 0) (tactic := "simp") with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -159,10 +158,10 @@ def test_m_couple_simp: TestM Unit := do
|
||||||
addTest $ assertUnreachable "(5 root)"
|
addTest $ assertUnreachable "(5 root)"
|
||||||
return ()
|
return ()
|
||||||
let rootStr: String := toString (← Lean.Meta.ppExpr root)
|
let rootStr: String := toString (← Lean.Meta.ppExpr root)
|
||||||
--checkEq "(5 root)" rootStr "Nat.le_trans (of_eq_true (_proof_4✝ 2)) (of_eq_true (eq_true_of_decide (Eq.refl true)))"
|
addTest $ LSpec.check "(5 root)" (rootStr = "Nat.le_trans (of_eq_true (Init.Data.Nat.Basic._auxLemma.4 2)) (of_eq_true (eq_true_of_decide (Eq.refl true)))")
|
||||||
let unfoldedRoot ← unfoldAuxLemmas root
|
let unfoldedRoot ← unfoldAuxLemmas root
|
||||||
checkEq "(5 root)" (toString (← Lean.Meta.ppExpr unfoldedRoot))
|
addTest $ LSpec.check "(5 root)" ((toString (← Lean.Meta.ppExpr unfoldedRoot)) =
|
||||||
"Nat.le_trans (of_eq_true (eq_true (Nat.le_refl 2))) (of_eq_true (eq_true_of_decide (Eq.refl true)))"
|
"Nat.le_trans (of_eq_true (eq_true (Nat.le_refl 2))) (of_eq_true (eq_true_of_decide (Eq.refl true)))")
|
||||||
return ()
|
return ()
|
||||||
|
|
||||||
def test_proposition_generation: TestM Unit := do
|
def test_proposition_generation: TestM Unit := do
|
||||||
|
@ -174,7 +173,7 @@ def test_proposition_generation: TestM Unit := do
|
||||||
return ()
|
return ()
|
||||||
|
|
||||||
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := "apply PSigma.mk") with
|
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := "apply PSigma.mk") with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -185,27 +184,27 @@ def test_proposition_generation: TestM Unit := do
|
||||||
])
|
])
|
||||||
if let #[goal1, goal2] := ← state1.serializeGoals (options := { (← read) with printExprAST := true }) then
|
if let #[goal1, goal2] := ← state1.serializeGoals (options := { (← read) with printExprAST := true }) then
|
||||||
addTest $ LSpec.test "(1 reference)" (goal1.target.sexp? = .some s!"(:mv {goal2.name})")
|
addTest $ LSpec.test "(1 reference)" (goal1.target.sexp? = .some s!"(:mv {goal2.name})")
|
||||||
checkTrue "(1 root)" $ ¬ state1.isSolved
|
addTest $ LSpec.test "(1 root)" state1.rootExpr?.isNone
|
||||||
|
|
||||||
let state2 ← match ← state1.tryAssign (state1.get! 0) (expr := "λ (x: Nat) => _") with
|
let state2 ← match ← state1.tryAssign (state1.get! 0) (expr := "λ (x: Nat) => _") with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
addTest $ LSpec.check ":= λ (x: Nat), _" ((← state2.serializeGoals (options := ← read)).map (·.target.pp?) =
|
addTest $ LSpec.check ":= λ (x: Nat), _" ((← state2.serializeGoals (options := ← read)).map (·.target.pp?) =
|
||||||
#[.some "?m.30 x"])
|
#[.some "?m.30 x"])
|
||||||
checkTrue "(2 root)" $ ¬ state2.isSolved
|
addTest $ LSpec.test "(2 root)" state2.rootExpr?.isNone
|
||||||
|
|
||||||
let assign := "Eq.refl x"
|
let assign := "Eq.refl x"
|
||||||
let state3 ← match ← state2.tryAssign (state2.get! 0) (expr := assign) with
|
let state3 ← match ← state2.tryAssign (state2.get! 0) (expr := assign) with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
addTest $ LSpec.check s!":= {assign}" ((← state3.serializeGoals (options := ← read)).map (·.target.pp?) =
|
addTest $ LSpec.check s!":= {assign}" ((← state3.serializeGoals (options := ← read)).map (·.target.pp?) =
|
||||||
#[])
|
#[])
|
||||||
|
|
||||||
checkTrue "(3 root)" state3.isSolved
|
addTest $ LSpec.test "(3 root)" state3.rootExpr?.isSome
|
||||||
return ()
|
return ()
|
||||||
|
|
||||||
def test_partial_continuation: TestM Unit := do
|
def test_partial_continuation: TestM Unit := do
|
||||||
|
@ -217,7 +216,7 @@ def test_partial_continuation: TestM Unit := do
|
||||||
return ()
|
return ()
|
||||||
|
|
||||||
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := "apply Nat.le_trans") with
|
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := "apply Nat.le_trans") with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -225,7 +224,7 @@ def test_partial_continuation: TestM Unit := do
|
||||||
#[.some "2 ≤ ?m", .some "?m ≤ 5", .some "Nat"])
|
#[.some "2 ≤ ?m", .some "?m ≤ 5", .some "Nat"])
|
||||||
|
|
||||||
let state2 ← match ← state1.tacticOn (goalId := 2) (tactic := "apply Nat.succ") with
|
let state2 ← match ← state1.tacticOn (goalId := 2) (tactic := "apply Nat.succ") with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -239,74 +238,34 @@ def test_partial_continuation: TestM Unit := do
|
||||||
addTest $ assertUnreachable $ msg
|
addTest $ assertUnreachable $ msg
|
||||||
return ()
|
return ()
|
||||||
| .ok state => pure state
|
| .ok state => pure state
|
||||||
addTest $ LSpec.check "(continue 1)" ((← state1b.serializeGoals (options := ← read)).map (·.target.pp?) =
|
addTest $ LSpec.check "(continue)" ((← state1b.serializeGoals (options := ← read)).map (·.target.pp?) =
|
||||||
#[.some "2 ≤ Nat.succ ?m", .some "Nat.succ ?m ≤ 5", .some "Nat"])
|
#[.some "2 ≤ Nat.succ ?m", .some "Nat.succ ?m ≤ 5", .some "Nat"])
|
||||||
checkTrue "(2 root)" state1b.rootExpr?.get!.hasExprMVar
|
addTest $ LSpec.test "(2 root)" state1b.rootExpr?.isNone
|
||||||
|
|
||||||
-- Roundtrip
|
-- Roundtrip
|
||||||
--let coupled_goals := coupled_goals.map (λ g =>
|
--let coupled_goals := coupled_goals.map (λ g =>
|
||||||
-- { name := str_to_name $ serializeName g.name (sanitize := false)})
|
-- { name := str_to_name $ serializeName g.name (sanitize := false)})
|
||||||
let coupled_goals := coupled_goals.map (·.name.toString)
|
let coupled_goals := coupled_goals.map (λ g => serializeName g.name (sanitize := false))
|
||||||
let coupled_goals := coupled_goals.map ({ name := ·.toName })
|
let coupled_goals := coupled_goals.map (λ g => { name := g.toName })
|
||||||
let state1b ← match state2.resume (goals := coupled_goals) with
|
let state1b ← match state2.resume (goals := coupled_goals) with
|
||||||
| .error msg => do
|
| .error msg => do
|
||||||
addTest $ assertUnreachable $ msg
|
addTest $ assertUnreachable $ msg
|
||||||
return ()
|
return ()
|
||||||
| .ok state => pure state
|
| .ok state => pure state
|
||||||
checkEq "(continue 2)" ((← state1b.serializeGoals (options := ← read)).map (·.target.pp?))
|
addTest $ LSpec.check "(continue)" ((← state1b.serializeGoals (options := ← read)).map (·.target.pp?) =
|
||||||
#[.some "2 ≤ Nat.succ ?m", .some "Nat.succ ?m ≤ 5", .some "Nat"]
|
#[.some "2 ≤ Nat.succ ?m", .some "Nat.succ ?m ≤ 5", .some "Nat"])
|
||||||
checkTrue "(2 root)" state1b.rootExpr?.get!.hasExprMVar
|
addTest $ LSpec.test "(2 root)" state1b.rootExpr?.isNone
|
||||||
|
|
||||||
-- Continuation should fail if the state does not exist:
|
-- Continuation should fail if the state does not exist:
|
||||||
match state0.resume coupled_goals with
|
match state0.resume coupled_goals with
|
||||||
| .error error => checkEq "(continuation failure message)" error "Goals [_uniq.45, _uniq.46, _uniq.43, _uniq.52] are not in scope"
|
| .error error => addTest $ LSpec.check "(continuation failure message)" (error = "Goals [_uniq.40, _uniq.41, _uniq.38, _uniq.47] are not in scope")
|
||||||
| .ok _ => fail "(continuation should fail)"
|
| .ok _ => addTest $ assertUnreachable "(continuation failure)"
|
||||||
-- Continuation should fail if some goals have not been solved
|
-- Continuation should fail if some goals have not been solved
|
||||||
match state2.continue state1 with
|
match state2.continue state1 with
|
||||||
| .error error => checkEq "(continuation failure message)" error "Target state has unresolved goals"
|
| .error error => addTest $ LSpec.check "(continuation failure message)" (error = "Target state has unresolved goals")
|
||||||
| .ok _ => fail "(continuation should fail)"
|
| .ok _ => addTest $ assertUnreachable "(continuation failure)"
|
||||||
return ()
|
return ()
|
||||||
|
|
||||||
def test_branch_unification : TestM Unit := do
|
|
||||||
let .ok rootTarget ← elabTerm (← `(term|∀ (p q : Prop), p → p ∧ (p ∨ q))) .none | unreachable!
|
|
||||||
let state ← GoalState.create rootTarget
|
|
||||||
let .success state _ ← state.tacticOn' 0 (← `(tactic|intro p q h)) | fail "intro failed to run"
|
|
||||||
let .success state _ ← state.tacticOn' 0 (← `(tactic|apply And.intro)) | fail "apply And.intro failed to run"
|
|
||||||
let .success state1 _ ← state.tacticOn' 0 (← `(tactic|exact h)) | fail "exact h failed to run"
|
|
||||||
let .success state2 _ ← state.tacticOn' 1 (← `(tactic|apply Or.inl)) | fail "apply Or.inl failed to run"
|
|
||||||
checkEq "(state2 goals)" state2.goals.length 1
|
|
||||||
let state' ← state2.replay state state1
|
|
||||||
checkEq "(state' goals)" state'.goals.length 1
|
|
||||||
let .success stateT _ ← state'.tacticOn' 0 (← `(tactic|exact h)) | fail "exact h failed to run"
|
|
||||||
let .some root := stateT.rootExpr? | fail "Root expression must exist"
|
|
||||||
checkEq "(root)" (toString $ ← Meta.ppExpr root) "fun p q h => ⟨h, Or.inl h⟩"
|
|
||||||
|
|
||||||
/-- Test merger when both branches have new aux lemmas -/
|
|
||||||
def test_replay_environment : TestM Unit := do
|
|
||||||
let .ok rootTarget ← elabTerm (← `(term|(2: Nat) ≤ 3 ∧ (3: Nat) ≤ 5)) .none | unreachable!
|
|
||||||
let state ← GoalState.create rootTarget
|
|
||||||
let .success state _ ← state.tacticOn' 0 (← `(tactic|apply And.intro)) | fail "apply And.intro failed to run"
|
|
||||||
let goal := state.goals[0]!
|
|
||||||
let type ← goal.withContext do
|
|
||||||
let .ok type ← elabTerm (← `(term|(2: Nat) ≤ 3)) (.some $ .sort 0) | unreachable!
|
|
||||||
pure type
|
|
||||||
let .success state1 _ ← state.tryTacticM goal (Tactic.assignWithAuxLemma type) | fail "left"
|
|
||||||
|
|
||||||
state.restoreMetaM
|
|
||||||
let goal := state.goals[1]!
|
|
||||||
let type ← goal.withContext do
|
|
||||||
let .ok type ← elabTerm (← `(term|(3: Nat) ≤ 5)) (.some $ .sort 0) | unreachable!
|
|
||||||
pure type
|
|
||||||
let .success state2 _ ← state.tryTacticM goal (Tactic.assignWithAuxLemma type) | fail "right"
|
|
||||||
checkEq "(state1 goals)" state1.goals.length 0
|
|
||||||
checkEq "(state2 goals)" state2.goals.length 0
|
|
||||||
let stateT ← state2.replay state state1
|
|
||||||
checkEq "(stateT goals)" stateT.goals.length 0
|
|
||||||
let .some root := stateT.rootExpr? | fail "Root expression must exist"
|
|
||||||
checkTrue "root has aux lemma" $ root.getUsedConstants.any isAuxLemma
|
|
||||||
checkEq "(root)" (toString $ ← Meta.ppExpr root) "⟨_proof_1, _proof_2⟩"
|
|
||||||
let root ← unfoldAuxLemmas root
|
|
||||||
checkEq "(root unfold)" (toString $ ← Meta.ppExpr root) "⟨sorry, sorry⟩"
|
|
||||||
|
|
||||||
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
|
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
|
||||||
let tests := [
|
let tests := [
|
||||||
|
@ -314,9 +273,7 @@ def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
|
||||||
("2 < 5", test_m_couple),
|
("2 < 5", test_m_couple),
|
||||||
("2 < 5", test_m_couple_simp),
|
("2 < 5", test_m_couple_simp),
|
||||||
("Proposition Generation", test_proposition_generation),
|
("Proposition Generation", test_proposition_generation),
|
||||||
("Partial Continuation", test_partial_continuation),
|
("Partial Continuation", test_partial_continuation)
|
||||||
("Branch Unification", test_branch_unification),
|
|
||||||
--("Replay Environment", test_replay_environment),
|
|
||||||
]
|
]
|
||||||
tests.map (fun (name, test) => (name, proofRunner env test))
|
tests.map (fun (name, test) => (name, proofRunner env test))
|
||||||
|
|
||||||
|
|
362
Test/Proofs.lean
362
Test/Proofs.lean
|
@ -11,7 +11,7 @@ open Pantograph
|
||||||
open Lean
|
open Lean
|
||||||
|
|
||||||
inductive Start where
|
inductive Start where
|
||||||
| copy (name: Name) -- 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 := TestT $ ReaderT Protocol.Options $ Elab.TermElabM
|
abbrev TestM := TestT $ ReaderT Protocol.Options $ Elab.TermElabM
|
||||||
|
@ -20,7 +20,7 @@ def startProof (start: Start): TestM (Option GoalState) := do
|
||||||
let env ← Lean.MonadEnv.getEnv
|
let env ← Lean.MonadEnv.getEnv
|
||||||
match start with
|
match start with
|
||||||
| .copy name =>
|
| .copy name =>
|
||||||
let cInfo? := name |> env.find?
|
let cInfo? := name.toName |> env.find?
|
||||||
addTest $ LSpec.check s!"Symbol exists {name}" cInfo?.isSome
|
addTest $ LSpec.check s!"Symbol exists {name}" cInfo?.isSome
|
||||||
match cInfo? with
|
match cInfo? with
|
||||||
| .some cInfo =>
|
| .some cInfo =>
|
||||||
|
@ -29,10 +29,24 @@ def startProof (start: Start): TestM (Option GoalState) := do
|
||||||
| .none =>
|
| .none =>
|
||||||
return Option.none
|
return Option.none
|
||||||
| .expr expr =>
|
| .expr expr =>
|
||||||
let expr ← parseSentence expr
|
let syn? := parseTerm env expr
|
||||||
return .some $ ← GoalState.create (expr := expr)
|
addTest $ LSpec.check s!"Parsing {expr}" (syn?.isOk)
|
||||||
|
match syn? with
|
||||||
|
| .error error =>
|
||||||
|
IO.println error
|
||||||
|
return Option.none
|
||||||
|
| .ok syn =>
|
||||||
|
let expr? ← elabType syn
|
||||||
|
addTest $ LSpec.check s!"Elaborating" expr?.isOk
|
||||||
|
match expr? with
|
||||||
|
| .error error =>
|
||||||
|
IO.println error
|
||||||
|
return Option.none
|
||||||
|
| .ok expr =>
|
||||||
|
let goal ← GoalState.create (expr := expr)
|
||||||
|
return Option.some goal
|
||||||
|
|
||||||
private def buildNamedGoal (name: String) (nameType: List (String × String)) (target: String)
|
def buildNamedGoal (name: String) (nameType: List (String × String)) (target: String)
|
||||||
(userName?: Option String := .none): Protocol.Goal :=
|
(userName?: Option String := .none): Protocol.Goal :=
|
||||||
{
|
{
|
||||||
name,
|
name,
|
||||||
|
@ -43,7 +57,7 @@ private def buildNamedGoal (name: String) (nameType: List (String × String)) (t
|
||||||
type? := .some { pp? := .some x.snd },
|
type? := .some { pp? := .some x.snd },
|
||||||
})).toArray
|
})).toArray
|
||||||
}
|
}
|
||||||
private def buildGoal (nameType: List (String × String)) (target: String) (userName?: Option String := .none):
|
def buildGoal (nameType: List (String × String)) (target: String) (userName?: Option String := .none):
|
||||||
Protocol.Goal :=
|
Protocol.Goal :=
|
||||||
{
|
{
|
||||||
userName?,
|
userName?,
|
||||||
|
@ -53,7 +67,7 @@ private def buildGoal (nameType: List (String × String)) (target: String) (user
|
||||||
type? := .some { pp? := .some x.snd },
|
type? := .some { pp? := .some x.snd },
|
||||||
})).toArray
|
})).toArray
|
||||||
}
|
}
|
||||||
private def proofRunner (env: Lean.Environment) (tests: TestM Unit): IO LSpec.TestSeq := do
|
def proofRunner (env: Lean.Environment) (tests: TestM Unit): IO LSpec.TestSeq := do
|
||||||
let termElabM := tests.run LSpec.TestSeq.done |>.run {} -- with default options
|
let termElabM := tests.run LSpec.TestSeq.done |>.run {} -- with default options
|
||||||
|
|
||||||
let coreContext: Lean.Core.Context ← createCoreContext #[]
|
let coreContext: Lean.Core.Context ← createCoreContext #[]
|
||||||
|
@ -66,18 +80,24 @@ private def proofRunner (env: Lean.Environment) (tests: TestM Unit): IO LSpec.Te
|
||||||
return a
|
return a
|
||||||
|
|
||||||
def test_identity: TestM Unit := do
|
def test_identity: TestM Unit := do
|
||||||
let rootTarget ← Elab.Term.elabTerm (← `(term|∀ (p: Prop), p → p)) .none
|
let state? ← startProof (.expr "∀ (p: Prop), p → p")
|
||||||
let state0 ← GoalState.create (expr := rootTarget)
|
let state0 ← match state? with
|
||||||
let state1 ← match ← state0.tacticOn' 0 (← `(tactic|intro p h)) with
|
| .some state => pure state
|
||||||
| .success state _ => pure state
|
| .none => do
|
||||||
| other => do
|
addTest $ assertUnreachable "Goal could not parse"
|
||||||
fail other.toString
|
|
||||||
return ()
|
return ()
|
||||||
let inner := "_uniq.11"
|
|
||||||
addTest $ LSpec.check "intro" ((← state1.serializeGoals (options := ← read)).map (·.name) =
|
let tactic := "intro p h"
|
||||||
|
let state1 ← match ← state0.tacticOn 0 tactic with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
let inner := "_uniq.12"
|
||||||
|
addTest $ LSpec.check tactic ((← state1.serializeGoals (options := ← read)).map (·.name) =
|
||||||
#[inner])
|
#[inner])
|
||||||
let state1parent ← state1.withParentContext do
|
let state1parent ← state1.withParentContext do
|
||||||
serializeExpressionSexp (← instantiateAll state1.parentExpr!)
|
serializeExpressionSexp (← instantiateAll state1.parentExpr?.get!)
|
||||||
addTest $ LSpec.test "(1 parent)" (state1parent == s!"(:lambda p (:sort 0) (:lambda h 0 (:subst (:mv {inner}) 1 0)))")
|
addTest $ LSpec.test "(1 parent)" (state1parent == s!"(:lambda p (:sort 0) (:lambda h 0 (:subst (:mv {inner}) 1 0)))")
|
||||||
|
|
||||||
-- Individual test cases
|
-- Individual test cases
|
||||||
|
@ -86,7 +106,7 @@ example: ∀ (a b: Nat), a + b = b + a := by
|
||||||
rw [Nat.add_comm]
|
rw [Nat.add_comm]
|
||||||
def test_nat_add_comm (manual: Bool): TestM Unit := do
|
def test_nat_add_comm (manual: Bool): TestM Unit := do
|
||||||
let state? ← startProof <| match manual with
|
let state? ← startProof <| match manual with
|
||||||
| false => .copy `Nat.add_comm
|
| false => .copy "Nat.add_comm"
|
||||||
| true => .expr "∀ (a b: Nat), a + b = b + a"
|
| true => .expr "∀ (a b: Nat), a + b = b + a"
|
||||||
addTest $ LSpec.check "Start goal" state?.isSome
|
addTest $ LSpec.check "Start goal" state?.isSome
|
||||||
let state0 ← match state? with
|
let state0 ← match state? with
|
||||||
|
@ -96,7 +116,7 @@ def test_nat_add_comm (manual: Bool): TestM Unit := do
|
||||||
return ()
|
return ()
|
||||||
|
|
||||||
let state1 ← match ← state0.tacticOn 0 "intro n m" with
|
let state1 ← match ← state0.tacticOn 0 "intro n m" with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -105,20 +125,53 @@ def test_nat_add_comm (manual: Bool): TestM Unit := do
|
||||||
|
|
||||||
match ← state1.tacticOn 0 "assumption" with
|
match ← state1.tacticOn 0 "assumption" with
|
||||||
| .failure #[message] =>
|
| .failure #[message] =>
|
||||||
checkEq "assumption"
|
addTest $ LSpec.check "assumption" (message = "tactic 'assumption' failed\nn m : Nat\n⊢ n + m = m + n")
|
||||||
(← message.toString)
|
|
||||||
s!"{← getFileName}:0:0: error: tactic 'assumption' failed\nn m : Nat\n⊢ n + m = m + n\n"
|
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
|
||||||
let state2 ← match ← state1.tacticOn 0 "rw [Nat.add_comm]" with
|
let state2 ← match ← state1.tacticOn 0 "rw [Nat.add_comm]" with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
addTest $ LSpec.test "rw [Nat.add_comm]" state2.goals.isEmpty
|
addTest $ LSpec.test "rw [Nat.add_comm]" state2.goals.isEmpty
|
||||||
|
|
||||||
return ()
|
return ()
|
||||||
|
def test_delta_variable: TestM Unit := do
|
||||||
|
let options: Protocol.Options := { noRepeat := true }
|
||||||
|
let state? ← startProof <| .expr "∀ (a b: Nat), a + b = b + a"
|
||||||
|
addTest $ LSpec.check "Start goal" state?.isSome
|
||||||
|
let state0 ← match state? with
|
||||||
|
| .some state => pure state
|
||||||
|
| .none => do
|
||||||
|
addTest $ assertUnreachable "Goal could not parse"
|
||||||
|
return ()
|
||||||
|
|
||||||
|
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := "intro n") with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check "intro n" ((← state1.serializeGoals (parent := state0) options).map (·.devolatilize) =
|
||||||
|
#[buildGoalSelective [("n", .some "Nat")] "∀ (b : Nat), n + b = b + n"])
|
||||||
|
let state2 ← match ← state1.tacticOn (goalId := 0) (tactic := "intro m") with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check "intro m" ((← state2.serializeGoals (parent := state1) options).map (·.devolatilize) =
|
||||||
|
#[buildGoalSelective [("n", .none), ("m", .some "Nat")] "n + m = m + n"])
|
||||||
|
return ()
|
||||||
|
where
|
||||||
|
-- Like `buildGoal` but allow certain variables to be elided.
|
||||||
|
buildGoalSelective (nameType: List (String × Option String)) (target: String): Protocol.Goal :=
|
||||||
|
{
|
||||||
|
target := { pp? := .some target},
|
||||||
|
vars := (nameType.map fun x => ({
|
||||||
|
userName := x.fst,
|
||||||
|
type? := x.snd.map (λ type => { pp? := type }),
|
||||||
|
})).toArray
|
||||||
|
}
|
||||||
|
|
||||||
example (w x y z : Nat) (p : Nat → Prop)
|
example (w x y z : Nat) (p : Nat → Prop)
|
||||||
(h : p (x * y + z * w * x)) : p (x * w * z + y * x) := by
|
(h : p (x * y + z * w * x)) : p (x * w * z + y * x) := by
|
||||||
|
@ -134,27 +187,27 @@ def test_arith: TestM Unit := do
|
||||||
|
|
||||||
let tactic := "intros"
|
let tactic := "intros"
|
||||||
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
addTest $ LSpec.check tactic (state1.goals.length = 1)
|
addTest $ LSpec.check tactic (state1.goals.length = 1)
|
||||||
checkTrue "(1 root)" state1.rootExpr?.get!.hasExprMVar
|
addTest $ LSpec.test "(1 root)" state1.rootExpr?.isNone
|
||||||
let state2 ← match ← state1.tacticOn (goalId := 0) (tactic := "simp [Nat.add_assoc, Nat.add_comm, Nat.add_left_comm, Nat.mul_comm, Nat.mul_assoc, Nat.mul_left_comm] at *") with
|
let state2 ← match ← state1.tacticOn (goalId := 0) (tactic := "simp [Nat.add_assoc, Nat.add_comm, Nat.add_left_comm, Nat.mul_comm, Nat.mul_assoc, Nat.mul_left_comm] at *") with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
addTest $ LSpec.check "simp ..." (state2.goals.length = 1)
|
addTest $ LSpec.check "simp ..." (state2.goals.length = 1)
|
||||||
checkTrue "(2 root)" state2.rootExpr?.get!.hasExprMVar
|
addTest $ LSpec.check "(2 root)" state2.rootExpr?.isNone
|
||||||
let tactic := "assumption"
|
let tactic := "assumption"
|
||||||
let state3 ← match ← state2.tacticOn (goalId := 0) (tactic := tactic) with
|
let state3 ← match ← state2.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
addTest $ LSpec.test tactic state3.goals.isEmpty
|
addTest $ LSpec.test tactic state3.goals.isEmpty
|
||||||
checkTrue "(3 root)" $ ¬ state3.rootExpr?.get!.hasExprMVar
|
addTest $ LSpec.check "(3 root)" state3.rootExpr?.isSome
|
||||||
return ()
|
return ()
|
||||||
|
|
||||||
-- Two ways to write the same theorem
|
-- Two ways to write the same theorem
|
||||||
|
@ -179,12 +232,12 @@ def test_or_comm: TestM Unit := do
|
||||||
| .none => do
|
| .none => do
|
||||||
addTest $ assertUnreachable "Goal could not parse"
|
addTest $ assertUnreachable "Goal could not parse"
|
||||||
return ()
|
return ()
|
||||||
checkTrue "(0 parent)" state0.parentMVars.isEmpty
|
addTest $ LSpec.check "(0 parent)" state0.parentExpr?.isNone
|
||||||
checkTrue "(0 root)" state0.rootExpr?.isNone
|
addTest $ LSpec.check "(0 root)" state0.rootExpr?.isNone
|
||||||
|
|
||||||
let tactic := "intro p q h"
|
let tactic := "intro p q h"
|
||||||
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -204,15 +257,15 @@ def test_or_comm: TestM Unit := do
|
||||||
{ name := fvH, userName := "h", type? := .some { pp? := .some "p ∨ q" } }
|
{ name := fvH, userName := "h", type? := .some { pp? := .some "p ∨ q" } }
|
||||||
]
|
]
|
||||||
}])
|
}])
|
||||||
checkTrue "(1 parent)" state1.hasUniqueParent
|
addTest $ LSpec.check "(1 parent)" state1.parentExpr?.isSome
|
||||||
checkTrue "(1 root)" $ ¬ state1.isSolved
|
addTest $ LSpec.check "(1 root)" state1.rootExpr?.isNone
|
||||||
|
|
||||||
let state1parent ← state1.withParentContext do
|
let state1parent ← state1.withParentContext do
|
||||||
serializeExpressionSexp (← instantiateAll state1.parentExpr!)
|
serializeExpressionSexp (← instantiateAll state1.parentExpr?.get!)
|
||||||
addTest $ LSpec.test "(1 parent)" (state1parent == s!"(:lambda p (:sort 0) (:lambda q (:sort 0) (:lambda h ((:c Or) 1 0) (:subst (:mv {state1g0}) 2 1 0))))")
|
addTest $ LSpec.test "(1 parent)" (state1parent == s!"(:lambda p (:sort 0) (:lambda q (:sort 0) (:lambda h ((:c Or) 1 0) (:subst (:mv {state1g0}) 2 1 0))))")
|
||||||
let tactic := "cases h"
|
let tactic := "cases h"
|
||||||
let state2 ← match ← state1.tacticOn (goalId := 0) (tactic := tactic) with
|
let state2 ← match ← state1.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -223,11 +276,11 @@ def test_or_comm: TestM Unit := do
|
||||||
let (caseL, caseR) := (state2g0.name.toString, state2g1.name.toString)
|
let (caseL, caseR) := (state2g0.name.toString, state2g1.name.toString)
|
||||||
addTest $ LSpec.check tactic ((← state2.serializeGoals (options := ← read)).map (·.name) =
|
addTest $ LSpec.check tactic ((← state2.serializeGoals (options := ← read)).map (·.name) =
|
||||||
#[caseL, caseR])
|
#[caseL, caseR])
|
||||||
checkTrue "(2 parent exists)" state2.hasUniqueParent
|
addTest $ LSpec.check "(2 parent exists)" state2.parentExpr?.isSome
|
||||||
checkTrue "(2 root)" $ ¬ state2.isSolved
|
addTest $ LSpec.check "(2 root)" state2.rootExpr?.isNone
|
||||||
|
|
||||||
let state2parent ← state2.withParentContext do
|
let state2parent ← state2.withParentContext do
|
||||||
serializeExpressionSexp (← instantiateAll state2.parentExpr!)
|
serializeExpressionSexp (← instantiateAll state2.parentExpr?.get!)
|
||||||
let orPQ := s!"((:c Or) (:fv {fvP}) (:fv {fvQ}))"
|
let orPQ := s!"((:c Or) (:fv {fvP}) (:fv {fvQ}))"
|
||||||
let orQP := s!"((:c Or) (:fv {fvQ}) (:fv {fvP}))"
|
let orQP := s!"((:c Or) (:fv {fvQ}) (:fv {fvP}))"
|
||||||
let motive := s!"(:lambda t {orPQ} (:forall h ((:c Eq) ((:c Or) (:fv {fvP}) (:fv {fvQ})) (:fv {fvH}) 0) {orQP}))"
|
let motive := s!"(:lambda t {orPQ} (:forall h ((:c Eq) ((:c Or) (:fv {fvP}) (:fv {fvQ})) (:fv {fvH}) 0) {orQP}))"
|
||||||
|
@ -238,37 +291,37 @@ def test_or_comm: TestM Unit := do
|
||||||
s!"((:c Or.casesOn) (:fv {fvP}) (:fv {fvQ}) {motive} (:fv {fvH}) {caseL} {caseR} {conduit})")
|
s!"((:c Or.casesOn) (:fv {fvP}) (:fv {fvQ}) {motive} (:fv {fvH}) {caseL} {caseR} {conduit})")
|
||||||
|
|
||||||
let state3_1 ← match ← state2.tacticOn (goalId := 0) (tactic := "apply Or.inr") with
|
let state3_1 ← match ← state2.tacticOn (goalId := 0) (tactic := "apply Or.inr") with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
let state3_1parent ← state3_1.withParentContext do
|
let state3_1parent ← state3_1.withParentContext do
|
||||||
serializeExpressionSexp (← instantiateAll state3_1.parentExpr!)
|
serializeExpressionSexp (← instantiateAll state3_1.parentExpr?.get!)
|
||||||
let [state3_1goal0] := state3_1.goals | fail "Should have 1 goal"
|
let [state3_1goal0] := state3_1.goals | fail "Should have 1 goal"
|
||||||
addTest $ LSpec.test "(3_1 parent)" (state3_1parent == s!"((:c Or.inr) (:fv {fvQ}) (:fv {fvP}) (:mv {state3_1goal0}))")
|
addTest $ LSpec.test "(3_1 parent)" (state3_1parent == s!"((:c Or.inr) (:fv {fvQ}) (:fv {fvP}) (:mv {state3_1goal0}))")
|
||||||
addTest $ LSpec.check "· apply Or.inr" (state3_1.goals.length = 1)
|
addTest $ LSpec.check "· apply Or.inr" (state3_1.goals.length = 1)
|
||||||
let state4_1 ← match ← state3_1.tacticOn (goalId := 0) (tactic := "assumption") with
|
let state4_1 ← match ← state3_1.tacticOn (goalId := 0) (tactic := "assumption") with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
addTest $ LSpec.check " assumption" state4_1.goals.isEmpty
|
addTest $ LSpec.check " assumption" state4_1.goals.isEmpty
|
||||||
let state4_1parent ← instantiateAll state4_1.parentExpr!
|
let state4_1parent ← instantiateAll state4_1.parentExpr?.get!
|
||||||
addTest $ LSpec.test "(4_1 parent)" state4_1parent.isFVar
|
addTest $ LSpec.test "(4_1 parent)" state4_1parent.isFVar
|
||||||
checkTrue "(4_1 root)" $ ¬ state4_1.isSolved
|
addTest $ LSpec.check "(4_1 root)" state4_1.rootExpr?.isNone
|
||||||
let state3_2 ← match ← state2.tacticOn (goalId := 1) (tactic := "apply Or.inl") with
|
let state3_2 ← match ← state2.tacticOn (goalId := 1) (tactic := "apply Or.inl") with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
addTest $ LSpec.check "· apply Or.inl" (state3_2.goals.length = 1)
|
addTest $ LSpec.check "· apply Or.inl" (state3_2.goals.length = 1)
|
||||||
let state4_2 ← match ← state3_2.tacticOn (goalId := 0) (tactic := "assumption") with
|
let state4_2 ← match ← state3_2.tacticOn (goalId := 0) (tactic := "assumption") with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
addTest $ LSpec.check " assumption" state4_2.goals.isEmpty
|
addTest $ LSpec.check " assumption" state4_2.goals.isEmpty
|
||||||
checkTrue "(4_2 root)" $ ¬ state4_2.isSolved
|
addTest $ LSpec.check "(4_2 root)" state4_2.rootExpr?.isNone
|
||||||
-- Ensure the proof can continue from `state4_2`.
|
-- Ensure the proof can continue from `state4_2`.
|
||||||
let state2b ← match state4_2.continue state2 with
|
let state2b ← match state4_2.continue state2 with
|
||||||
| .error msg => do
|
| .error msg => do
|
||||||
|
@ -277,18 +330,18 @@ def test_or_comm: TestM Unit := do
|
||||||
| .ok state => pure state
|
| .ok state => pure state
|
||||||
addTest $ LSpec.test "(resume)" (state2b.goals == [state2.goals[0]!])
|
addTest $ LSpec.test "(resume)" (state2b.goals == [state2.goals[0]!])
|
||||||
let state3_1 ← match ← state2b.tacticOn (goalId := 0) (tactic := "apply Or.inr") with
|
let state3_1 ← match ← state2b.tacticOn (goalId := 0) (tactic := "apply Or.inr") with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
addTest $ LSpec.check "· apply Or.inr" (state3_1.goals.length = 1)
|
addTest $ LSpec.check "· apply Or.inr" (state3_1.goals.length = 1)
|
||||||
let state4_1 ← match ← state3_1.tacticOn (goalId := 0) (tactic := "assumption") with
|
let state4_1 ← match ← state3_1.tacticOn (goalId := 0) (tactic := "assumption") with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
addTest $ LSpec.check " assumption" state4_1.goals.isEmpty
|
addTest $ LSpec.check " assumption" state4_1.goals.isEmpty
|
||||||
checkTrue "(4_1 root)" $ ¬ state4_1.rootExpr?.get!.hasExprMVar
|
addTest $ LSpec.check "(4_1 root)" state4_1.rootExpr?.isSome
|
||||||
|
|
||||||
return ()
|
return ()
|
||||||
where
|
where
|
||||||
|
@ -303,6 +356,193 @@ def test_or_comm: TestM Unit := do
|
||||||
]
|
]
|
||||||
}
|
}
|
||||||
|
|
||||||
|
example : ∀ (a b c1 c2: Nat), (b + a) + c1 = (b + a) + c2 → (a + b) + c1 = (b + a) + c2 := by
|
||||||
|
intro a b c1 c2 h
|
||||||
|
conv =>
|
||||||
|
lhs
|
||||||
|
congr
|
||||||
|
. rw [Nat.add_comm]
|
||||||
|
. rfl
|
||||||
|
exact h
|
||||||
|
|
||||||
|
def test_conv: TestM Unit := do
|
||||||
|
let state? ← startProof (.expr "∀ (a b c1 c2: Nat), (b + a) + c1 = (b + a) + c2 → (a + b) + c1 = (b + a) + c2")
|
||||||
|
let state0 ← match state? with
|
||||||
|
| .some state => pure state
|
||||||
|
| .none => do
|
||||||
|
addTest $ assertUnreachable "Goal could not parse"
|
||||||
|
return ()
|
||||||
|
|
||||||
|
let tactic := "intro a b c1 c2 h"
|
||||||
|
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check tactic ((← state1.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[interiorGoal [] "a + b + c1 = b + a + c2"])
|
||||||
|
|
||||||
|
let state2 ← match ← state1.conv (state1.get! 0) with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check "conv => ..." ((← state2.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[{ interiorGoal [] "a + b + c1 = b + a + c2" with isConversion := true }])
|
||||||
|
|
||||||
|
let convTactic := "rhs"
|
||||||
|
let state3R ← match ← state2.tacticOn (goalId := 0) convTactic with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check s!" {convTactic} (discard)" ((← state3R.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[{ interiorGoal [] "b + a + c2" with isConversion := true }])
|
||||||
|
|
||||||
|
let convTactic := "lhs"
|
||||||
|
let state3L ← match ← state2.tacticOn (goalId := 0) convTactic with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check s!" {convTactic}" ((← state3L.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[{ interiorGoal [] "a + b + c1" with isConversion := true }])
|
||||||
|
|
||||||
|
let convTactic := "congr"
|
||||||
|
let state4 ← match ← state3L.tacticOn (goalId := 0) convTactic with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check s!" {convTactic}" ((← state4.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[
|
||||||
|
{ interiorGoal [] "a + b" with isConversion := true, userName? := .some "a" },
|
||||||
|
{ interiorGoal [] "c1" with isConversion := true, userName? := .some "a" }
|
||||||
|
])
|
||||||
|
|
||||||
|
let convTactic := "rw [Nat.add_comm]"
|
||||||
|
let state5_1 ← match ← state4.tacticOn (goalId := 0) convTactic with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check s!" · {convTactic}" ((← state5_1.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[{ interiorGoal [] "b + a" with isConversion := true, userName? := .some "a" }])
|
||||||
|
|
||||||
|
let convTactic := "rfl"
|
||||||
|
let state6_1 ← match ← state5_1.tacticOn (goalId := 0) convTactic with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check s!" {convTactic}" ((← state6_1.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[])
|
||||||
|
|
||||||
|
let state4_1 ← match state6_1.continue state4 with
|
||||||
|
| .ok state => pure state
|
||||||
|
| .error e => do
|
||||||
|
addTest $ expectationFailure "continue" e
|
||||||
|
return ()
|
||||||
|
|
||||||
|
let convTactic := "rfl"
|
||||||
|
let state6 ← match ← state4_1.tacticOn (goalId := 0) convTactic with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check s!" · {convTactic}" ((← state6.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[])
|
||||||
|
|
||||||
|
let state1_1 ← match ← state6.convExit with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
|
||||||
|
let tactic := "exact h"
|
||||||
|
let stateF ← match ← state1_1.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check tactic ((← stateF.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[])
|
||||||
|
|
||||||
|
where
|
||||||
|
h := "b + a + c1 = b + a + c2"
|
||||||
|
interiorGoal (free: List (String × String)) (target: String) :=
|
||||||
|
let free := [("a", "Nat"), ("b", "Nat"), ("c1", "Nat"), ("c2", "Nat"), ("h", h)] ++ free
|
||||||
|
buildGoal free target
|
||||||
|
|
||||||
|
example : ∀ (a b c d: Nat), a + b = b + c → b + c = c + d → a + b = c + d := by
|
||||||
|
intro a b c d h1 h2
|
||||||
|
calc a + b = b + c := by apply h1
|
||||||
|
_ = c + d := by apply h2
|
||||||
|
|
||||||
|
def test_calc: TestM Unit := do
|
||||||
|
let state? ← startProof (.expr "∀ (a b c d: Nat), a + b = b + c → b + c = c + d → a + b = c + d")
|
||||||
|
let state0 ← match state? with
|
||||||
|
| .some state => pure state
|
||||||
|
| .none => do
|
||||||
|
addTest $ assertUnreachable "Goal could not parse"
|
||||||
|
return ()
|
||||||
|
let tactic := "intro a b c d h1 h2"
|
||||||
|
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check tactic ((← state1.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[interiorGoal [] "a + b = c + d"])
|
||||||
|
let pred := "a + b = b + c"
|
||||||
|
let state2 ← match ← state1.tryCalc (state1.get! 0) (pred := pred) with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check s!"calc {pred} := _" ((← state2.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[
|
||||||
|
interiorGoal [] "a + b = b + c" (.some "calc"),
|
||||||
|
interiorGoal [] "b + c = c + d"
|
||||||
|
])
|
||||||
|
addTest $ LSpec.test "(2.0 prev rhs)" (state2.calcPrevRhsOf? (state2.get! 0) |>.isNone)
|
||||||
|
addTest $ LSpec.test "(2.1 prev rhs)" (state2.calcPrevRhsOf? (state2.get! 1) |>.isSome)
|
||||||
|
|
||||||
|
let tactic := "apply h1"
|
||||||
|
let state2m ← match ← state2.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
let state3 ← match state2m.continue state2 with
|
||||||
|
| .ok state => pure state
|
||||||
|
| .error e => do
|
||||||
|
addTest $ expectationFailure "continue" e
|
||||||
|
return ()
|
||||||
|
let pred := "_ = c + d"
|
||||||
|
let state4 ← match ← state3.tryCalc (state3.get! 0) (pred := pred) with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.check s!"calc {pred} := _" ((← state4.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||||
|
#[
|
||||||
|
interiorGoal [] "b + c = c + d" (.some "calc")
|
||||||
|
])
|
||||||
|
addTest $ LSpec.test "(4.0 prev rhs)" (state4.calcPrevRhsOf? (state4.get! 0) |>.isNone)
|
||||||
|
let tactic := "apply h2"
|
||||||
|
let state4m ← match ← state4.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
|
| .success state => pure state
|
||||||
|
| other => do
|
||||||
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
return ()
|
||||||
|
addTest $ LSpec.test "(4m root)" state4m.rootExpr?.isSome
|
||||||
|
where
|
||||||
|
interiorGoal (free: List (String × String)) (target: String) (userName?: Option String := .none) :=
|
||||||
|
let free := [("a", "Nat"), ("b", "Nat"), ("c", "Nat"), ("d", "Nat"),
|
||||||
|
("h1", "a + b = b + c"), ("h2", "b + c = c + d")] ++ free
|
||||||
|
buildGoal free target userName?
|
||||||
|
|
||||||
def test_tactic_failure_unresolved_goals : TestM Unit := do
|
def test_tactic_failure_unresolved_goals : TestM Unit := do
|
||||||
let state? ← startProof (.expr "∀ (p : Nat → Prop), ∃ (x : Nat), p (0 + x + 0)")
|
let state? ← startProof (.expr "∀ (p : Nat → Prop), ∃ (x : Nat), p (0 + x + 0)")
|
||||||
let state0 ← match state? with
|
let state0 ← match state? with
|
||||||
|
@ -313,16 +553,14 @@ def test_tactic_failure_unresolved_goals : TestM Unit := do
|
||||||
|
|
||||||
let tactic := "intro p"
|
let tactic := "intro p"
|
||||||
let state1 ← match ← state0.tacticOn 0 tactic with
|
let state1 ← match ← state0.tacticOn 0 tactic with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
|
||||||
let tactic := "exact ⟨0, by simp⟩"
|
let tactic := "exact ⟨0, by simp⟩"
|
||||||
let .failure #[message] ← state1.tacticOn 0 tactic
|
let .failure messages ← state1.tacticOn 0 tactic | addTest $ assertUnreachable s!"{tactic} should fail"
|
||||||
| fail s!"{tactic} should fail with 1 message"
|
checkEq s!"{tactic} fails" messages #[s!"{← getFileName}:0:12: error: unsolved goals\np : Nat → Prop\n⊢ p 0\n"]
|
||||||
checkEq s!"{tactic} fails" (← message.toString)
|
|
||||||
s!"{← getFileName}:0:12: error: unsolved goals\np : Nat → Prop\n⊢ p 0\n"
|
|
||||||
|
|
||||||
def test_tactic_failure_synthesize_placeholder : TestM Unit := do
|
def test_tactic_failure_synthesize_placeholder : TestM Unit := do
|
||||||
let state? ← startProof (.expr "∀ (p q r : Prop) (h : p → q), q ∧ r")
|
let state? ← startProof (.expr "∀ (p q r : Prop) (h : p → q), q ∧ r")
|
||||||
|
@ -334,7 +572,7 @@ def test_tactic_failure_synthesize_placeholder : TestM Unit := do
|
||||||
|
|
||||||
let tactic := "intro p q r h"
|
let tactic := "intro p q r h"
|
||||||
let state1 ← match ← state0.tacticOn 0 tactic with
|
let state1 ← match ← state0.tacticOn 0 tactic with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -352,10 +590,9 @@ def test_tactic_failure_synthesize_placeholder : TestM Unit := do
|
||||||
-- buildGoal [("p", "Prop"), ("q", "Prop"), ("r", "Prop"), ("h", "p → q")] "p ∧ r"
|
-- buildGoal [("p", "Prop"), ("q", "Prop"), ("r", "Prop"), ("h", "p → q")] "p ∧ r"
|
||||||
--]
|
--]
|
||||||
|
|
||||||
let .failure #[message] ← state1.tacticOn 0 tactic
|
let .failure messages ← state1.tacticOn 0 tactic | addTest $ assertUnreachable s!"{tactic} should fail"
|
||||||
| 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" (← message.toString)
|
checkEq s!"{tactic} fails" messages #[message]
|
||||||
s!"{← getFileName}:0:31: error: don't know how to synthesize placeholder\ncontext:\np q r : Prop\nh : p → q\n⊢ p ∧ r\n"
|
|
||||||
|
|
||||||
def test_deconstruct : TestM Unit := do
|
def test_deconstruct : TestM Unit := do
|
||||||
let state? ← startProof (.expr "∀ (p q : Prop) (h : And p q), And q p")
|
let state? ← startProof (.expr "∀ (p q : Prop) (h : And p q), And q p")
|
||||||
|
@ -367,7 +604,7 @@ def test_deconstruct : TestM Unit := do
|
||||||
|
|
||||||
let tactic := "intro p q ⟨hp, hq⟩"
|
let tactic := "intro p q ⟨hp, hq⟩"
|
||||||
let state1 ← match ← state0.tacticOn 0 tactic with
|
let state1 ← match ← state0.tacticOn 0 tactic with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
fail other.toString
|
fail other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -382,8 +619,11 @@ def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
|
||||||
("identity", test_identity),
|
("identity", test_identity),
|
||||||
("Nat.add_comm", test_nat_add_comm false),
|
("Nat.add_comm", test_nat_add_comm false),
|
||||||
("Nat.add_comm manual", test_nat_add_comm true),
|
("Nat.add_comm manual", test_nat_add_comm true),
|
||||||
|
("Nat.add_comm delta", test_delta_variable),
|
||||||
("arithmetic", test_arith),
|
("arithmetic", test_arith),
|
||||||
("Or.comm", test_or_comm),
|
("Or.comm", test_or_comm),
|
||||||
|
("conv", test_conv),
|
||||||
|
("calc", test_calc),
|
||||||
("tactic failure with unresolved goals", test_tactic_failure_unresolved_goals),
|
("tactic failure with unresolved goals", test_tactic_failure_unresolved_goals),
|
||||||
("tactic failure with synthesize placeholder", test_tactic_failure_synthesize_placeholder),
|
("tactic failure with synthesize placeholder", test_tactic_failure_synthesize_placeholder),
|
||||||
("deconstruct", test_deconstruct),
|
("deconstruct", test_deconstruct),
|
||||||
|
|
|
@ -7,6 +7,9 @@ open Lean
|
||||||
|
|
||||||
namespace Pantograph.Test.Serial
|
namespace Pantograph.Test.Serial
|
||||||
|
|
||||||
|
def tempPath : IO System.FilePath := do
|
||||||
|
Prod.snd <$> IO.FS.createTempFile
|
||||||
|
|
||||||
structure MultiState where
|
structure MultiState where
|
||||||
coreContext : Core.Context
|
coreContext : Core.Context
|
||||||
env: Environment
|
env: Environment
|
||||||
|
@ -27,22 +30,22 @@ def runCoreM { α } (state : Core.State) (testCoreM : TestT CoreM α) : TestM (
|
||||||
set $ (← getThe LSpec.TestSeq) ++ tests
|
set $ (← getThe LSpec.TestSeq) ++ tests
|
||||||
return (a, state')
|
return (a, state')
|
||||||
|
|
||||||
def test_pickling_environment : TestM Unit := do
|
def test_environment_pickling : TestM Unit := do
|
||||||
let coreSrc : Core.State := { env := ← getEnv }
|
let coreSrc : Core.State := { env := ← getEnv }
|
||||||
let coreDst : Core.State := { env := ← getEnv }
|
let coreDst : Core.State := { env := ← getEnv }
|
||||||
|
|
||||||
let name := `mystery
|
let name := `mystery
|
||||||
IO.FS.withTempFile λ _ envPicklePath => do
|
let envPicklePath ← tempPath
|
||||||
let ((), _) ← runCoreM coreSrc do
|
let ((), _) ← runCoreM coreSrc do
|
||||||
let type: Expr := .forallE `p (.sort 0) (.forallE `h (.bvar 0) (.bvar 1) .default) .default
|
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 value: Expr := .lam `p (.sort 0) (.lam `h (.bvar 0) (.bvar 0) .default) .default
|
||||||
let c := Declaration.defnDecl <| mkDefinitionValEx
|
let c := Lean.Declaration.defnDecl <| Lean.mkDefinitionValEx
|
||||||
(name := name)
|
(name := name)
|
||||||
(levelParams := [])
|
(levelParams := [])
|
||||||
(type := type)
|
(type := type)
|
||||||
(value := value)
|
(value := value)
|
||||||
(hints := mkReducibilityHintsRegularEx 1)
|
(hints := Lean.mkReducibilityHintsRegularEx 1)
|
||||||
(safety := .safe)
|
(safety := Lean.DefinitionSafety.safe)
|
||||||
(all := [])
|
(all := [])
|
||||||
addDecl c
|
addDecl c
|
||||||
environmentPickle (← getEnv) envPicklePath
|
environmentPickle (← getEnv) envPicklePath
|
||||||
|
@ -53,10 +56,13 @@ def test_pickling_environment : TestM Unit := do
|
||||||
let anotherName := `mystery2
|
let anotherName := `mystery2
|
||||||
checkTrue s!"Doesn't have symbol {anotherName}" (env'.find? anotherName).isNone
|
checkTrue s!"Doesn't have symbol {anotherName}" (env'.find? anotherName).isNone
|
||||||
|
|
||||||
def test_goal_state_simple : TestM Unit := do
|
IO.FS.removeFile envPicklePath
|
||||||
|
|
||||||
|
def test_goal_state_pickling_simple : TestM Unit := do
|
||||||
let coreSrc : Core.State := { env := ← getEnv }
|
let coreSrc : Core.State := { env := ← getEnv }
|
||||||
let coreDst : Core.State := { env := ← getEnv }
|
let coreDst : Core.State := { env := ← getEnv }
|
||||||
IO.FS.withTempFile λ _ statePath => do
|
let statePath ← tempPath
|
||||||
|
|
||||||
let type: Expr := .forallE `p (.sort 0) (.forallE `h (.bvar 0) (.bvar 1) .default) .default
|
let type: Expr := .forallE `p (.sort 0) (.forallE `h (.bvar 0) (.bvar 1) .default) .default
|
||||||
let stateGenerate : MetaM GoalState := runTermElabMInMeta do
|
let stateGenerate : MetaM GoalState := runTermElabMInMeta do
|
||||||
GoalState.create type
|
GoalState.create type
|
||||||
|
@ -72,35 +78,13 @@ def test_goal_state_simple : TestM Unit := do
|
||||||
let types ← metaM.run'
|
let types ← metaM.run'
|
||||||
checkTrue "Goals" $ types[0]!.equal type
|
checkTrue "Goals" $ types[0]!.equal type
|
||||||
|
|
||||||
def test_pickling_env_extensions : TestM Unit := do
|
IO.FS.removeFile statePath
|
||||||
let coreSrc : Core.State := { env := ← getEnv }
|
|
||||||
let coreDst : Core.State := { env := ← getEnv }
|
|
||||||
IO.FS.withTempFile λ _ statePath => do
|
|
||||||
let ((), _) ← runCoreM coreSrc $ transformTestT runTermElabMInCore do
|
|
||||||
let .ok e ← elabTerm (← `(term|(2: Nat) ≤ 3 ∧ (3: Nat) ≤ 5)) .none | unreachable!
|
|
||||||
let state ← GoalState.create e
|
|
||||||
let .success state _ ← state.tacticOn' 0 (← `(tactic|apply And.intro)) | unreachable!
|
|
||||||
|
|
||||||
let goal := state.goals[0]!
|
|
||||||
let type ← goal.withContext do
|
|
||||||
let .ok type ← elabTerm (← `(term|(2: Nat) ≤ 3)) (.some $ .sort 0) | unreachable!
|
|
||||||
instantiateMVars type
|
|
||||||
let .success state1 _ ← state.tryTacticM goal (Tactic.assignWithAuxLemma type) | unreachable!
|
|
||||||
let parentExpr := state1.parentExpr!
|
|
||||||
checkTrue "src has aux lemma" $ parentExpr.getUsedConstants.any isAuxLemma
|
|
||||||
goalStatePickle state1 statePath
|
|
||||||
let ((), _) ← runCoreM coreDst $ transformTestT runTermElabMInCore do
|
|
||||||
let (state1, _) ← goalStateUnpickle statePath (← getEnv)
|
|
||||||
let parentExpr := state1.parentExpr!
|
|
||||||
checkTrue "dst has aux lemma" $ parentExpr.getUsedConstants.any isAuxLemma
|
|
||||||
|
|
||||||
return ()
|
|
||||||
|
|
||||||
structure Test where
|
structure Test where
|
||||||
name : String
|
name : String
|
||||||
routine: TestM Unit
|
routine: TestM Unit
|
||||||
|
|
||||||
protected def Test.run (test: Test) (env: Environment) : IO LSpec.TestSeq := do
|
protected def Test.run (test: Test) (env: Lean.Environment) : IO LSpec.TestSeq := do
|
||||||
-- Create the state
|
-- Create the state
|
||||||
let state : MultiState := {
|
let state : MultiState := {
|
||||||
coreContext := ← createCoreContext #[],
|
coreContext := ← createCoreContext #[],
|
||||||
|
@ -111,11 +95,10 @@ protected def Test.run (test: Test) (env: Environment) : IO LSpec.TestSeq := do
|
||||||
| .error e =>
|
| .error e =>
|
||||||
return LSpec.check s!"Emitted exception: {e.toString}" (e.toString == "")
|
return LSpec.check s!"Emitted exception: {e.toString}" (e.toString == "")
|
||||||
|
|
||||||
def suite (env : Environment): List (String × IO LSpec.TestSeq) :=
|
def suite (env : Lean.Environment): List (String × IO LSpec.TestSeq) :=
|
||||||
let tests: List Test := [
|
let tests: List Test := [
|
||||||
{ name := "environment", routine := test_pickling_environment, },
|
{ name := "environment_pickling", routine := test_environment_pickling, },
|
||||||
{ name := "goal simple", routine := test_goal_state_simple, },
|
{ name := "goal_state_pickling_simple", routine := test_goal_state_pickling_simple, },
|
||||||
{ name := "extensions", routine := test_pickling_env_extensions, },
|
|
||||||
]
|
]
|
||||||
tests.map (fun test => (test.name, test.run env))
|
tests.map (fun test => (test.name, test.run env))
|
||||||
|
|
||||||
|
|
|
@ -1,4 +1,2 @@
|
||||||
import Test.Tactic.Assign
|
import Test.Tactic.Assign
|
||||||
import Test.Tactic.Fragment
|
|
||||||
import Test.Tactic.Prograde
|
import Test.Tactic.Prograde
|
||||||
import Test.Tactic.Special
|
|
||||||
|
|
|
@ -1,313 +0,0 @@
|
||||||
import Pantograph.Goal
|
|
||||||
import Test.Common
|
|
||||||
|
|
||||||
open Lean
|
|
||||||
|
|
||||||
namespace Pantograph.Test.Tactic.Fragment
|
|
||||||
|
|
||||||
private def buildGoal (nameType: List (String × String)) (target: String):
|
|
||||||
Protocol.Goal :=
|
|
||||||
{
|
|
||||||
target := { pp? := .some target},
|
|
||||||
vars := (nameType.map fun x => ({
|
|
||||||
userName := x.fst,
|
|
||||||
type? := .some { pp? := .some x.snd },
|
|
||||||
})).toArray
|
|
||||||
}
|
|
||||||
|
|
||||||
abbrev TestM := TestT $ Elab.TermElabM
|
|
||||||
|
|
||||||
example : ∀ (a b c1 c2: Nat), (b + a) + c1 = (b + a) + c2 → (a + b) + c1 = (b + a) + c2 := by
|
|
||||||
intro a b c1 c2 h
|
|
||||||
conv =>
|
|
||||||
lhs
|
|
||||||
congr
|
|
||||||
. rw [Nat.add_comm]
|
|
||||||
. rfl
|
|
||||||
exact h
|
|
||||||
|
|
||||||
def test_conv_simple: TestM Unit := do
|
|
||||||
let rootTarget ← parseSentence "∀ (a b c1 c2: Nat), (b + a) + c1 = (b + a) + c2 → (a + b) + c1 = (b + a) + c2"
|
|
||||||
let state0 ← GoalState.create rootTarget
|
|
||||||
|
|
||||||
let tactic := "intro a b c1 c2 h"
|
|
||||||
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check tactic ((← state1.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[interiorGoal [] "a + b + c1 = b + a + c2"])
|
|
||||||
|
|
||||||
let goalConv := state1.goals[0]!
|
|
||||||
let state2 ← match ← state1.convEnter (state1.get! 0) with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check "conv => ..." ((← state2.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[{ interiorGoal [] "a + b + c1 = b + a + c2" with fragment := .conv }])
|
|
||||||
|
|
||||||
let convTactic := "rhs"
|
|
||||||
let state3R ← match ← state2.tacticOn (goalId := 0) convTactic with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check s!" {convTactic} (discard)" ((← state3R.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[{ interiorGoal [] "b + a + c2" with fragment := .conv }])
|
|
||||||
|
|
||||||
let convTactic := "lhs"
|
|
||||||
let state3L ← match ← state2.tacticOn (goalId := 0) convTactic with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check s!" {convTactic}" ((← state3L.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[{ interiorGoal [] "a + b + c1" with fragment := .conv }])
|
|
||||||
|
|
||||||
let convTactic := "congr"
|
|
||||||
let state4 ← match ← state3L.tacticOn (goalId := 0) convTactic with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check s!" {convTactic}" ((← state4.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[
|
|
||||||
{ interiorGoal [] "a + b" with fragment := .conv, userName? := .some "a" },
|
|
||||||
{ interiorGoal [] "c1" with fragment := .conv, userName? := .some "a" }
|
|
||||||
])
|
|
||||||
|
|
||||||
let convTactic := "rw [Nat.add_comm]"
|
|
||||||
let state5_1 ← match ← state4.tacticOn (goalId := 0) convTactic with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check s!" · {convTactic}" ((← state5_1.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[{ interiorGoal [] "b + a" with fragment := .conv, userName? := .some "a" }])
|
|
||||||
|
|
||||||
let convTactic := "rfl"
|
|
||||||
let state6_1 ← match ← state5_1.tacticOn (goalId := 0) convTactic with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check s!" {convTactic}" ((← state6_1.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[])
|
|
||||||
|
|
||||||
let state4_1 ← match state6_1.continue state4 with
|
|
||||||
| .ok state => pure state
|
|
||||||
| .error e => do
|
|
||||||
addTest $ expectationFailure "continue" e
|
|
||||||
return ()
|
|
||||||
|
|
||||||
let convTactic := "rfl"
|
|
||||||
let state1_1 ← match ← state4_1.tacticOn (goalId := 0) convTactic with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check s!" · {convTactic}" ((← state1_1.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[interiorGoal [] "b + a + c1 = b + a + c2"])
|
|
||||||
checkEq "(fragments)" state1_1.fragments.size 0
|
|
||||||
|
|
||||||
/-
|
|
||||||
let state1_1 ← match ← state6.fragmentExit goalConv with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
-/
|
|
||||||
|
|
||||||
let tactic := "exact h"
|
|
||||||
let stateF ← match ← state1_1.tacticOn (goalId := 0) (tactic := tactic) with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
checkEq tactic ((← stateF.serializeGoals).map (·.devolatilize)) #[]
|
|
||||||
checkEq "fragments" stateF.fragments.size 0
|
|
||||||
|
|
||||||
where
|
|
||||||
h := "b + a + c1 = b + a + c2"
|
|
||||||
interiorGoal (free: List (String × String)) (target: String) :=
|
|
||||||
let free := [("a", "Nat"), ("b", "Nat"), ("c1", "Nat"), ("c2", "Nat"), ("h", h)] ++ free
|
|
||||||
buildGoal free target
|
|
||||||
|
|
||||||
example (p : Prop) (x y z : Nat) : p → (p → x = y) → x + z = y + z ∧ p := by
|
|
||||||
intro hp hi
|
|
||||||
apply And.intro
|
|
||||||
conv =>
|
|
||||||
rhs
|
|
||||||
arg 1
|
|
||||||
rw [←hi]
|
|
||||||
rfl
|
|
||||||
tactic => exact hp
|
|
||||||
exact hp
|
|
||||||
|
|
||||||
def test_conv_unshielded : TestM Unit := do
|
|
||||||
let rootTarget ← parseSentence "∀ (p : Prop) (x y z : Nat), p → (p → x = y) → x + z = y + z ∧ p"
|
|
||||||
let state ← GoalState.create rootTarget
|
|
||||||
let tactic := "intro p x y z hp hi"
|
|
||||||
let .success state _ ← state.tacticOn 0 tactic | fail "intro failed"
|
|
||||||
let tactic := "apply And.intro"
|
|
||||||
let .success state _ ← state.tacticOn 0 tactic | fail "apply failed"
|
|
||||||
let .success state _ ← state.convEnter (.prefer state.goals[0]!) | fail "Cannot enter conversion tactic mode"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus "rhs" | fail "rhs failed"
|
|
||||||
let tactic := "arg 1"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize))
|
|
||||||
#[
|
|
||||||
{ interiorGoal [] "y" with fragment := .conv },
|
|
||||||
{ interiorGoal [] "p" with userName? := "right", },
|
|
||||||
]
|
|
||||||
let tactic := "rw [←hi]"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
checkEq s!" {tactic}" state.goals.length 3
|
|
||||||
let tactic := "rfl"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize))
|
|
||||||
#[
|
|
||||||
interiorGoal [] "p",
|
|
||||||
{ interiorGoal [] "p" with userName? := "right", },
|
|
||||||
]
|
|
||||||
checkEq "(n goals)" state.goals.length 2
|
|
||||||
checkEq "(fragments)" state.fragments.size 0
|
|
||||||
let tactic := "exact hp"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
let tactic := "exact hp"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
let root? := state.rootExpr?
|
|
||||||
checkTrue "root" root?.isSome
|
|
||||||
checkEq "fragments" state.fragments.size 0
|
|
||||||
where
|
|
||||||
interiorGoal (free: List (String × String)) (target: String) :=
|
|
||||||
let free := [("p", "Prop"), ("x", "Nat"), ("y", "Nat"), ("z", "Nat"), ("hp", "p"), ("hi", "p → x = y")] ++ free
|
|
||||||
buildGoal free target
|
|
||||||
|
|
||||||
example : ∀ (x y z w : Nat), y = z → x + z = w → x + y = w := by
|
|
||||||
intro x y z w hyz hxzw
|
|
||||||
conv =>
|
|
||||||
lhs
|
|
||||||
arg 2
|
|
||||||
rw [hyz]
|
|
||||||
rfl
|
|
||||||
exact hxzw
|
|
||||||
|
|
||||||
def test_conv_unfinished : TestM Unit := do
|
|
||||||
let rootTarget ← parseSentence "∀ (x y z w : Nat), y = z → x + z = w → x + y = w"
|
|
||||||
let state ← GoalState.create rootTarget
|
|
||||||
let tactic := "intro x y z w hyz hxzw"
|
|
||||||
let .success state _ ← state.tacticOn 0 tactic | fail "intro failed"
|
|
||||||
let convParent := state.goals[0]!
|
|
||||||
let .success state _ ← state.convEnter (.prefer convParent) | fail "Cannot enter conversion tactic mode"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus "lhs" | fail "rhs failed"
|
|
||||||
let tactic := "arg 2"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize))
|
|
||||||
#[
|
|
||||||
{ interiorGoal [] "y" with fragment := .conv },
|
|
||||||
]
|
|
||||||
let tactic := "rw [hyz]"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize))
|
|
||||||
#[
|
|
||||||
{ interiorGoal [] "z" with fragment := .conv },
|
|
||||||
]
|
|
||||||
checkTrue " (fragment)" $ state.fragments.contains state.mainGoal?.get!
|
|
||||||
checkTrue " (fragment parent)" $ state.fragments.contains convParent
|
|
||||||
checkTrue " (main goal)" state.mainGoal?.isSome
|
|
||||||
let tactic := "rfl"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize))
|
|
||||||
#[
|
|
||||||
interiorGoal [] "x + z = w",
|
|
||||||
]
|
|
||||||
checkEq "(fragments)" state.fragments.size 0
|
|
||||||
checkEq s!" {tactic}" state.goals.length 1
|
|
||||||
let tactic := "exact hxzw"
|
|
||||||
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
|
|
||||||
let root? := state.rootExpr?
|
|
||||||
checkTrue "root" root?.isSome
|
|
||||||
where
|
|
||||||
interiorGoal (free: List (String × String)) (target: String) :=
|
|
||||||
let free := [("x", "Nat"), ("y", "Nat"), ("z", "Nat"), ("w", "Nat"), ("hyz", "y = z"), ("hxzw", "x + z = w")] ++ free
|
|
||||||
buildGoal free target
|
|
||||||
|
|
||||||
example : ∀ (a b c d: Nat), a + b = b + c → b + c = c + d → a + b = c + d := by
|
|
||||||
intro a b c d h1 h2
|
|
||||||
calc a + b = b + c := by apply h1
|
|
||||||
_ = c + d := by apply h2
|
|
||||||
|
|
||||||
def test_calc: TestM Unit := do
|
|
||||||
let rootTarget ← parseSentence "∀ (a b c d: Nat), a + b = b + c → b + c = c + d → a + b = c + d"
|
|
||||||
let state0 ← GoalState.create rootTarget
|
|
||||||
let tactic := "intro a b c d h1 h2"
|
|
||||||
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check tactic ((← state1.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[interiorGoal [] "a + b = c + d"])
|
|
||||||
let pred := "a + b = b + c"
|
|
||||||
let .success state1 _ ← state1.calcEnter state1.mainGoal?.get! | fail "Could not enter calc"
|
|
||||||
let state2 ← match ← state1.tacticOn 0 pred with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check s!"calc {pred} := _" ((← state2.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[
|
|
||||||
{ interiorGoal [] "a + b = b + c" with userName? := .some "calc" },
|
|
||||||
{ interiorGoal [] "b + c = c + d" with fragment := .calc },
|
|
||||||
])
|
|
||||||
addTest $ LSpec.test "(2.0 prev rhs)" (state2.calcPrevRhsOf? (state2.get! 0) |>.isNone)
|
|
||||||
addTest $ LSpec.test "(2.1 prev rhs)" (state2.calcPrevRhsOf? (state2.get! 1) |>.isSome)
|
|
||||||
|
|
||||||
let tactic := "apply h1"
|
|
||||||
let state2m ← match ← state2.tacticOn (goalId := 0) (tactic := tactic) with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
let state3 ← match state2m.continue state2 with
|
|
||||||
| .ok state => pure state
|
|
||||||
| .error e => do
|
|
||||||
addTest $ expectationFailure "continue" e
|
|
||||||
return ()
|
|
||||||
let pred := "_ = c + d"
|
|
||||||
let state4 ← match ← state3.tacticOn 0 pred with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
addTest $ LSpec.check s!"calc {pred} := _" ((← state4.serializeGoals).map (·.devolatilize) =
|
|
||||||
#[
|
|
||||||
{ interiorGoal [] "b + c = c + d" with userName? := .some "calc" },
|
|
||||||
])
|
|
||||||
addTest $ LSpec.test "(4.0 prev rhs)" (state4.calcPrevRhsOf? (state4.get! 0) |>.isNone)
|
|
||||||
let tactic := "apply h2"
|
|
||||||
let state4m ← match ← state4.tacticOn (goalId := 0) (tactic := tactic) with
|
|
||||||
| .success state _ => pure state
|
|
||||||
| other => do
|
|
||||||
addTest $ assertUnreachable $ other.toString
|
|
||||||
return ()
|
|
||||||
checkEq "(fragments)" state4m.fragments.size 0
|
|
||||||
addTest $ LSpec.test "(4m root)" state4m.rootExpr?.isSome
|
|
||||||
where
|
|
||||||
interiorGoal (free: List (String × String)) (target: String) :=
|
|
||||||
let free := [("a", "Nat"), ("b", "Nat"), ("c", "Nat"), ("d", "Nat"),
|
|
||||||
("h1", "a + b = b + c"), ("h2", "b + c = c + d")] ++ free
|
|
||||||
buildGoal free target
|
|
||||||
|
|
||||||
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
|
|
||||||
[
|
|
||||||
("conv simple", test_conv_simple),
|
|
||||||
("conv unshielded", test_conv_unshielded),
|
|
||||||
("conv unfinished", test_conv_unfinished),
|
|
||||||
("calc", test_calc),
|
|
||||||
] |>.map (λ (name, t) => (name, runTestTermElabM env t))
|
|
||||||
|
|
||||||
end Pantograph.Test.Tactic.Fragment
|
|
|
@ -56,7 +56,7 @@ def test_define_proof : TestT Elab.TermElabM Unit := do
|
||||||
let state0 ← GoalState.create rootExpr
|
let state0 ← GoalState.create rootExpr
|
||||||
let tactic := "intro p q h"
|
let tactic := "intro p q h"
|
||||||
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -65,7 +65,7 @@ def test_define_proof : TestT Elab.TermElabM Unit := do
|
||||||
|
|
||||||
let expr := "Or.inl (Or.inl h)"
|
let expr := "Or.inl (Or.inl h)"
|
||||||
let state2 ← match ← state1.tryAssign (state1.get! 0) (expr := expr) with
|
let state2 ← match ← state1.tryAssign (state1.get! 0) (expr := expr) with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -75,7 +75,7 @@ def test_define_proof : TestT Elab.TermElabM Unit := do
|
||||||
let evalBind := "y"
|
let evalBind := "y"
|
||||||
let evalExpr := "Or.inl h"
|
let evalExpr := "Or.inl h"
|
||||||
let state2 ← match ← state1.tryDefine (state1.get! 0) (binderName := evalBind) (expr := evalExpr) with
|
let state2 ← match ← state1.tryDefine (state1.get! 0) (binderName := evalBind) (expr := evalExpr) with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -95,7 +95,7 @@ def test_define_proof : TestT Elab.TermElabM Unit := do
|
||||||
|
|
||||||
let expr := "Or.inl y"
|
let expr := "Or.inl y"
|
||||||
let state3 ← match ← state2.tryAssign (state2.get! 0) (expr := expr) with
|
let state3 ← match ← state2.tryAssign (state2.get! 0) (expr := expr) with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -112,31 +112,35 @@ def fun_define_root_expr: ∀ (p: Prop), PProd (Nat → p) Unit → p := by
|
||||||
def test_define_root_expr : TestT Elab.TermElabM Unit := do
|
def test_define_root_expr : TestT Elab.TermElabM Unit := do
|
||||||
--let rootExpr ← parseSentence "Nat"
|
--let rootExpr ← parseSentence "Nat"
|
||||||
--let state0 ← GoalState.create rootExpr
|
--let state0 ← GoalState.create rootExpr
|
||||||
--let .success state1 _ ← state0.tacticOn (goalId := 0) "exact 5" | addTest $ assertUnreachable "exact 5"
|
--let .success state1 ← state0.tacticOn (goalId := 0) "exact 5" | addTest $ assertUnreachable "exact 5"
|
||||||
--let .some rootExpr := state1.rootExpr? | addTest $ assertUnreachable "Root expr"
|
--let .some rootExpr := state1.rootExpr? | addTest $ assertUnreachable "Root expr"
|
||||||
--addTest $ LSpec.check "root" ((toString $ ← Meta.ppExpr rootExpr) = "5")
|
--addTest $ LSpec.check "root" ((toString $ ← Meta.ppExpr rootExpr) = "5")
|
||||||
let rootExpr ← parseSentence "∀ (p: Prop), PProd (Nat → p) Unit → p"
|
let rootExpr ← parseSentence "∀ (p: Prop), PProd (Nat → p) Unit → p"
|
||||||
let state0 ← GoalState.create rootExpr
|
let state0 ← GoalState.create rootExpr
|
||||||
let tactic := "intro p x"
|
let tactic := "intro p x"
|
||||||
let .success state1 _ ← state0.tacticOn (goalId := 0) tactic | addTest $ assertUnreachable tactic
|
let .success state1 ← state0.tacticOn (goalId := 0) tactic | addTest $ assertUnreachable tactic
|
||||||
let binderName := `binder
|
let binderName := `binder
|
||||||
let value := "x.fst"
|
let value := "x.fst"
|
||||||
let expr ← state1.goals[0]!.withContext $ strToTermSyntax value
|
let expr ← state1.goals[0]!.withContext $ strToTermSyntax value
|
||||||
let tacticM := Tactic.evalDefine binderName expr
|
let tacticM := Tactic.evalDefine binderName expr
|
||||||
let .success state2 _ ← state1.tryTacticM (state1.get! 0) tacticM | addTest $ assertUnreachable s!"define {binderName} := {value}"
|
let .success state2 ← state1.tryTacticM (state1.get! 0) tacticM | addTest $ assertUnreachable s!"define {binderName} := {value}"
|
||||||
let tactic := s!"apply {binderName}"
|
let tactic := s!"apply {binderName}"
|
||||||
let .success state3 _ ← state2.tacticOn (goalId := 0) tactic | addTest $ assertUnreachable tactic
|
let .success state3 ← state2.tacticOn (goalId := 0) tactic | addTest $ assertUnreachable tactic
|
||||||
let tactic := s!"exact 5"
|
let tactic := s!"exact 5"
|
||||||
let .success state4 _ ← state3.tacticOn (goalId := 0) tactic | addTest $ assertUnreachable tactic
|
let .success state4 ← state3.tacticOn (goalId := 0) tactic | addTest $ assertUnreachable tactic
|
||||||
let .some rootExpr := state4.rootExpr? | addTest $ assertUnreachable "Root expr"
|
let .some rootExpr := state4.rootExpr? | addTest $ assertUnreachable "Root expr"
|
||||||
addTest $ LSpec.check "root" ((toString $ ← Meta.ppExpr rootExpr) = "fun p x =>\n let binder := x.fst;\n binder 5")
|
addTest $ LSpec.check "root" ((toString $ ← Meta.ppExpr rootExpr) = "fun p x =>\n let binder := x.fst;\n binder 5")
|
||||||
|
|
||||||
|
--set_option pp.all true
|
||||||
|
--#check @PSigma (α := Prop) (β := λ (p: Prop) => p)
|
||||||
|
--def test_define_root_expr : TestT Elab.TermElabM Unit := do
|
||||||
|
|
||||||
def test_have_proof : TestT Elab.TermElabM Unit := do
|
def test_have_proof : TestT Elab.TermElabM Unit := do
|
||||||
let rootExpr ← parseSentence "∀ (p q: Prop), p → ((p ∨ q) ∨ (p ∨ q))"
|
let rootExpr ← parseSentence "∀ (p q: Prop), p → ((p ∨ q) ∨ (p ∨ q))"
|
||||||
let state0 ← GoalState.create rootExpr
|
let state0 ← GoalState.create rootExpr
|
||||||
let tactic := "intro p q h"
|
let tactic := "intro p q h"
|
||||||
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -145,7 +149,7 @@ def test_have_proof : TestT Elab.TermElabM Unit := do
|
||||||
|
|
||||||
let expr := "Or.inl (Or.inl h)"
|
let expr := "Or.inl (Or.inl h)"
|
||||||
let state2 ← match ← state1.tryAssign (state1.get! 0) (expr := expr) with
|
let state2 ← match ← state1.tryAssign (state1.get! 0) (expr := expr) with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -155,7 +159,7 @@ def test_have_proof : TestT Elab.TermElabM Unit := do
|
||||||
let haveBind := "y"
|
let haveBind := "y"
|
||||||
let haveType := "p ∨ q"
|
let haveType := "p ∨ q"
|
||||||
let state2 ← match ← state1.tryHave (state1.get! 0) (binderName := haveBind) (type := haveType) with
|
let state2 ← match ← state1.tryHave (state1.get! 0) (binderName := haveBind) (type := haveType) with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -167,7 +171,7 @@ def test_have_proof : TestT Elab.TermElabM Unit := do
|
||||||
|
|
||||||
let expr := "Or.inl h"
|
let expr := "Or.inl h"
|
||||||
let state3 ← match ← state2.tryAssign (state2.get! 0) (expr := expr) with
|
let state3 ← match ← state2.tryAssign (state2.get! 0) (expr := expr) with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -181,11 +185,12 @@ def test_have_proof : TestT Elab.TermElabM Unit := do
|
||||||
return ()
|
return ()
|
||||||
let expr := "Or.inl y"
|
let expr := "Or.inl y"
|
||||||
let state4 ← match ← state2b.tryAssign (state2b.get! 0) (expr := expr) with
|
let state4 ← match ← state2b.tryAssign (state2b.get! 0) (expr := expr) with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
checkEq s!":= {expr}" ((← state4.serializeGoals).map (·.devolatilize)) #[]
|
addTest $ LSpec.check s!":= {expr}" ((← state4.serializeGoals).map (·.devolatilize) =
|
||||||
|
#[])
|
||||||
|
|
||||||
let .some rootExpr := state4.rootExpr? | addTest $ assertUnreachable "Root expr"
|
let .some rootExpr := state4.rootExpr? | addTest $ assertUnreachable "Root expr"
|
||||||
addTest $ LSpec.check "root" ((toString $ ← Meta.ppExpr rootExpr) = "fun p q h y => Or.inl y")
|
addTest $ LSpec.check "root" ((toString $ ← Meta.ppExpr rootExpr) = "fun p q h y => Or.inl y")
|
||||||
|
@ -195,7 +200,7 @@ def test_let (specialized: Bool): TestT Elab.TermElabM Unit := do
|
||||||
let state0 ← GoalState.create rootExpr
|
let state0 ← GoalState.create rootExpr
|
||||||
let tactic := "intro a p h"
|
let tactic := "intro a p h"
|
||||||
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -211,7 +216,7 @@ def test_let (specialized: Bool): TestT Elab.TermElabM Unit := do
|
||||||
| true => state1.tryLet (state1.get! 0) (binderName := "b") (type := letType)
|
| true => state1.tryLet (state1.get! 0) (binderName := "b") (type := letType)
|
||||||
| false => state1.tryAssign (state1.get! 0) (expr := expr)
|
| false => state1.tryAssign (state1.get! 0) (expr := expr)
|
||||||
let state2 ← match result2 with
|
let state2 ← match result2 with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -236,7 +241,7 @@ def test_let (specialized: Bool): TestT Elab.TermElabM Unit := do
|
||||||
|
|
||||||
let tactic := "exact 1"
|
let tactic := "exact 1"
|
||||||
let state3 ← match ← state2.tacticOn (goalId := 0) (tactic := tactic) with
|
let state3 ← match ← state2.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
@ -263,15 +268,13 @@ def test_let (specialized: Bool): TestT Elab.TermElabM Unit := do
|
||||||
let tactic := "exact h"
|
let tactic := "exact h"
|
||||||
match ← state3r.tacticOn (goalId := 0) (tactic := tactic) with
|
match ← state3r.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
| .failure #[message] =>
|
| .failure #[message] =>
|
||||||
checkEq tactic
|
addTest $ LSpec.check tactic (message = s!"type mismatch\n h\nhas type\n a : Prop\nbut is expected to have type\n {mainTarget} : Prop")
|
||||||
(← message.toString)
|
|
||||||
s!"{← getFileName}:0:0: error: type mismatch\n h\nhas type\n a : Prop\nbut is expected to have type\n {mainTarget} : Prop\n"
|
|
||||||
| other => do
|
| other => do
|
||||||
fail s!"Should be a failure: {other.toString}"
|
addTest $ assertUnreachable $ other.toString
|
||||||
|
|
||||||
let tactic := "exact Or.inl (Or.inl h)"
|
let tactic := "exact Or.inl (Or.inl h)"
|
||||||
let state4 ← match ← state3r.tacticOn (goalId := 0) (tactic := tactic) with
|
let state4 ← match ← state3r.tacticOn (goalId := 0) (tactic := tactic) with
|
||||||
| .success state _ => pure state
|
| .success state => pure state
|
||||||
| other => do
|
| other => do
|
||||||
addTest $ assertUnreachable $ other.toString
|
addTest $ assertUnreachable $ other.toString
|
||||||
return ()
|
return ()
|
||||||
|
|
|
@ -1,25 +0,0 @@
|
||||||
import LSpec
|
|
||||||
import Lean
|
|
||||||
import Test.Common
|
|
||||||
|
|
||||||
open Lean
|
|
||||||
open Pantograph
|
|
||||||
|
|
||||||
namespace Pantograph.Test.Tactic.Special
|
|
||||||
|
|
||||||
def test_exact_q : TestT Elab.TermElabM Unit := do
|
|
||||||
let rootExpr ← parseSentence "1 + 2 = 2 + 3"
|
|
||||||
let state0 ← GoalState.create rootExpr
|
|
||||||
let tactic := "exact?"
|
|
||||||
let state1? ← state0.tacticOn (goalId := 0) (tactic := tactic)
|
|
||||||
let .failure messages := state1? | fail "Must fail"
|
|
||||||
checkEq "messages"
|
|
||||||
(← messages.mapM (·.toString))
|
|
||||||
#[s!"{← getFileName}:0:0: error: `exact?` could not close the goal. Try `apply?` to see partial suggestions.\n"]
|
|
||||||
|
|
||||||
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
|
|
||||||
[
|
|
||||||
("exact?", test_exact_q),
|
|
||||||
] |>.map (λ (name, t) => (name, runTestTermElabM env t))
|
|
||||||
|
|
||||||
end Pantograph.Test.Tactic.Special
|
|
|
@ -1,41 +0,0 @@
|
||||||
/- A tool for analysing Lean source code. -/
|
|
||||||
import Pantograph.Frontend
|
|
||||||
import Pantograph.Library
|
|
||||||
|
|
||||||
open Lean
|
|
||||||
|
|
||||||
namespace Pantograph
|
|
||||||
|
|
||||||
def fail (s : String) : IO UInt32 := do
|
|
||||||
IO.eprintln s
|
|
||||||
return 2
|
|
||||||
|
|
||||||
def dissect (args : List String) : IO UInt32 := do
|
|
||||||
let fileName :: _args := args | fail s!"Must supply a file name"
|
|
||||||
let file ← IO.FS.readFile fileName
|
|
||||||
let (context, state) ← do Frontend.createContextStateFromFile file fileName (env? := .none) {}
|
|
||||||
let frontendM: Elab.Frontend.FrontendM _ :=
|
|
||||||
Frontend.mapCompilationSteps λ step => do
|
|
||||||
IO.println s!"🐈 {step.stx.getKind.toString}"
|
|
||||||
for (tree, i) in step.trees.zipIdx do
|
|
||||||
IO.println s!"🌲[{i}] {← tree.toString}"
|
|
||||||
for (msg, i) in step.msgs.zipIdx do
|
|
||||||
IO.println s!"🔈[{i}] {← msg.toString}"
|
|
||||||
let (_, _) ← frontendM.run context |>.run state
|
|
||||||
return 0
|
|
||||||
|
|
||||||
end Pantograph
|
|
||||||
|
|
||||||
open Pantograph
|
|
||||||
|
|
||||||
def help : IO UInt32 := do
|
|
||||||
IO.println "Usage: tomograph dissect FILE_NAME"
|
|
||||||
return 1
|
|
||||||
|
|
||||||
def main (args : List String) : IO UInt32 := do
|
|
||||||
let command :: args := args | help
|
|
||||||
unsafe do
|
|
||||||
Pantograph.initSearch ""
|
|
||||||
match command with
|
|
||||||
| "dissect" => dissect args
|
|
||||||
| _ => fail s!"Unknown command {command}"
|
|
|
@ -37,22 +37,21 @@ differently from Lean in some times, but never at the sacrifice of soundness.
|
||||||
- When Lean LSP says "unresolved goals", that means a proof cannot finish where
|
- 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
|
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.
|
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
|
Pantograph cannot perform things that are inherently constrained by Lean. These
|
||||||
include:
|
include:
|
||||||
|
|
||||||
- If a tactic loses track of metavariables, it will not be caught until the end
|
- 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.
|
of the proof search. This is a bug in the tactic itself.
|
||||||
- Although a timeout feature exists in Pantograph, it relies on the coöperative
|
- Timeouts for executing tactics is not available. Maybe this will change in the
|
||||||
multitasking from the tactic implementation. There is nothing preventing a
|
future.
|
||||||
buggy tactic from stalling Lean if it does not check for cancellation often.
|
|
||||||
- For the same reason as above, there is no graceful way to stop a tactic which
|
|
||||||
leaks infinite memory. Users who wish to have this behaviour should run
|
|
||||||
Pantograph in a controlled environment with limited allocations. e.g.
|
|
||||||
Linux control groups.
|
|
||||||
- Interceptions of parsing errors generally cannot be turned into goals (e.g.
|
- Interceptions of parsing errors generally cannot be turned into goals (e.g.
|
||||||
`def mystery : Nat := :=`) due to Lean's parsing system. This question is also
|
`def mystery : Nat := :=`) due to Lean's parsing system.
|
||||||
not well-defined.
|
|
||||||
|
|
||||||
## References
|
## References
|
||||||
|
|
||||||
|
|
28
doc/repl.md
28
doc/repl.md
|
@ -13,7 +13,7 @@ See `Pantograph/Protocol.lean` for a description of the parameters and return va
|
||||||
only the values of definitions are printed.
|
only the values of definitions are printed.
|
||||||
* `env.save { "path": <fileName> }`, `env.load { "path": <fileName> }`: Save/Load the
|
* `env.save { "path": <fileName> }`, `env.load { "path": <fileName> }`: Save/Load the
|
||||||
current environment to/from a file
|
current environment to/from a file
|
||||||
* `env.module_read { "module": <name> }`: Reads a list of symbols from a module
|
* `env.module_read { "module": <name }`: Reads a list of symbols from a module
|
||||||
* `env.describe {}`: Describes the imports and modules in the current environment
|
* `env.describe {}`: Describes the imports and modules in the current environment
|
||||||
* `options.set { key: value, ... }`: Set one or more options (not Lean options; those
|
* `options.set { key: value, ... }`: Set one or more options (not Lean options; those
|
||||||
have to be set via command line arguments.), for options, see `Pantograph/Protocol.lean`
|
have to be set via command line arguments.), for options, see `Pantograph/Protocol.lean`
|
||||||
|
@ -28,21 +28,17 @@ See `Pantograph/Protocol.lean` for a description of the parameters and return va
|
||||||
* `options.print`: Display the current set of options
|
* `options.print`: Display the current set of options
|
||||||
* `goal.start {["name": <name>], ["expr": <expr>], ["levels": [<levels>]], ["copyFrom": <symbol>]}`:
|
* `goal.start {["name": <name>], ["expr": <expr>], ["levels": [<levels>]], ["copyFrom": <symbol>]}`:
|
||||||
Start a new proof from a given expression or symbol
|
Start a new proof from a given expression or symbol
|
||||||
* `goal.tactic {"stateId": <id>, ["goalId": <id>], ["autoResume": <bool>], ...}`:
|
* `goal.tactic {"stateId": <id>, "goalId": <id>, ...}`: Execute a tactic string on a
|
||||||
Execute a tactic string on a given goal site. The tactic is supplied as additional
|
given goal. The tactic is supplied as additional key-value pairs in one of the following formats:
|
||||||
key-value pairs in one of the following formats:
|
- `{ "tactic": <tactic> }`: Execute an ordinary tactic
|
||||||
- `{ "tactic": <tactic> }`: Executes a tactic in the current mode
|
|
||||||
- `{ "mode": <mode> }`: Enter a different tactic mode. The permitted values
|
|
||||||
are `tactic` (default), `conv`, `calc`. In case of `calc`, each step must
|
|
||||||
be of the form `lhs op rhs`. An `lhs` of `_` indicates that it should be set
|
|
||||||
to the previous `rhs`.
|
|
||||||
- `{ "expr": <expr> }`: Assign the given proof term to the current goal
|
- `{ "expr": <expr> }`: Assign the given proof term to the current goal
|
||||||
- `{ "have": <expr>, "binderName": <name> }`: Execute `have` and creates a branch goal
|
- `{ "have": <expr>, "binderName": <name> }`: Execute `have` and creates a branch goal
|
||||||
- `{ "let": <expr>, "binderName": <name> }`: Execute `let` and creates a branch goal
|
- `{ "calc": <expr> }`: Execute one step of a `calc` tactic. Each step must
|
||||||
- `{ "draft": <expr> }`: Draft an expression with `sorry`s, turning them into
|
be of the form `lhs op rhs`. An `lhs` of `_` indicates that it should be set
|
||||||
goals. Coupling is not allowed.
|
to the previous `rhs`.
|
||||||
If the `goals` field does not exist, the tactic execution has failed. Read
|
- `{ "conv": <bool> }`: Enter or exit conversion tactic mode. In the case of
|
||||||
`messages` to find the reason.
|
exit, the goal id is ignored.
|
||||||
|
- `{ "draft": <expr> }`: Draft an expression with `sorry`s, turning them into goals. Coupling is not allowed.
|
||||||
* `goal.continue {"stateId": <id>, ["branch": <id>], ["goals": <names>]}`:
|
* `goal.continue {"stateId": <id>, ["branch": <id>], ["goals": <names>]}`:
|
||||||
Execute continuation/resumption
|
Execute continuation/resumption
|
||||||
- `{ "branch": <id> }`: Continue on branch state. The current state must have no goals.
|
- `{ "branch": <id> }`: Continue on branch state. The current state must have no goals.
|
||||||
|
@ -54,9 +50,9 @@ 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>], readHeader: <bool>, inheritEnv: <bool>, invocations:
|
* `frontend.process { ["fileName": <fileName>,] ["file": <str>], readHeader: <bool>, inheritEnv: <bool>, invocations:
|
||||||
<string>, sorrys: <bool>, typeErrorsAsGoals: <bool>, newConstants: <bool> }`:
|
<bool>, sorrys: <bool>, typeErrorsAsGoals: <bool>, newConstants: <bool> }`:
|
||||||
Executes the Lean frontend on a file, collecting the tactic invocations
|
Executes the Lean frontend on a file, collecting the tactic invocations
|
||||||
(`"invocations": output-path`), the sorrys and type errors into goal states
|
(`"invocations": true`), the sorrys and type errors into goal states
|
||||||
(`"sorrys": true`), and new constants (`"newConstants": true`). In the case of
|
(`"sorrys": true`), and new constants (`"newConstants": true`). In the case of
|
||||||
`sorrys`, this command additionally outputs the position of each captured
|
`sorrys`, this command additionally outputs the position of each captured
|
||||||
`sorry`. Conditionally inherit the environment from executing the file.
|
`sorry`. Conditionally inherit the environment from executing the file.
|
||||||
|
|
48
flake.lock
48
flake.lock
|
@ -5,11 +5,11 @@
|
||||||
"nixpkgs-lib": "nixpkgs-lib"
|
"nixpkgs-lib": "nixpkgs-lib"
|
||||||
},
|
},
|
||||||
"locked": {
|
"locked": {
|
||||||
"lastModified": 1749398372,
|
"lastModified": 1743550720,
|
||||||
"narHash": "sha256-tYBdgS56eXYaWVW3fsnPQ/nFlgWi/Z2Ymhyu21zVM98=",
|
"narHash": "sha256-hIshGgKZCgWh6AYJpJmRgFdR3WUbkY04o82X05xqQiY=",
|
||||||
"owner": "hercules-ci",
|
"owner": "hercules-ci",
|
||||||
"repo": "flake-parts",
|
"repo": "flake-parts",
|
||||||
"rev": "9305fe4e5c2a6fcf5ba6a3ff155720fbe4076569",
|
"rev": "c621e8422220273271f52058f618c94e405bb0f5",
|
||||||
"type": "github"
|
"type": "github"
|
||||||
},
|
},
|
||||||
"original": {
|
"original": {
|
||||||
|
@ -39,14 +39,16 @@
|
||||||
"lean4-nix": {
|
"lean4-nix": {
|
||||||
"inputs": {
|
"inputs": {
|
||||||
"flake-parts": "flake-parts_2",
|
"flake-parts": "flake-parts_2",
|
||||||
"nixpkgs": "nixpkgs"
|
"nixpkgs": [
|
||||||
|
"nixpkgs"
|
||||||
|
]
|
||||||
},
|
},
|
||||||
"locked": {
|
"locked": {
|
||||||
"lastModified": 1750369222,
|
"lastModified": 1743534244,
|
||||||
"narHash": "sha256-KFFTVbciXUaHgeGN1yiaUtY88OLGU0gElXx5SfICDKg=",
|
"narHash": "sha256-WnoYs2iyrfgh35eXErCOyos8E2YbW3LT1xm/EtT88/k=",
|
||||||
"owner": "lenianiva",
|
"owner": "lenianiva",
|
||||||
"repo": "lean4-nix",
|
"repo": "lean4-nix",
|
||||||
"rev": "015ecd25206734d582a1b15dd11eb10be35ca555",
|
"rev": "5eb7f03be257e327fdb3cca9465392e68dc28a4d",
|
||||||
"type": "github"
|
"type": "github"
|
||||||
},
|
},
|
||||||
"original": {
|
"original": {
|
||||||
|
@ -57,27 +59,27 @@
|
||||||
},
|
},
|
||||||
"nixpkgs": {
|
"nixpkgs": {
|
||||||
"locked": {
|
"locked": {
|
||||||
"lastModified": 1743095683,
|
"lastModified": 1743975612,
|
||||||
"narHash": "sha256-gWd4urRoLRe8GLVC/3rYRae1h+xfQzt09xOfb0PaHSk=",
|
"narHash": "sha256-o4FjFOUmjSRMK7dn0TFdAT0RRWUWD+WsspPHa+qEQT8=",
|
||||||
"owner": "nixos",
|
"owner": "nixos",
|
||||||
"repo": "nixpkgs",
|
"repo": "nixpkgs",
|
||||||
"rev": "5e5402ecbcb27af32284d4a62553c019a3a49ea6",
|
"rev": "a880f49904d68b5e53338d1e8c7bf80f59903928",
|
||||||
"type": "github"
|
"type": "github"
|
||||||
},
|
},
|
||||||
"original": {
|
"original": {
|
||||||
"owner": "nixos",
|
"owner": "nixos",
|
||||||
"ref": "nixos-unstable",
|
"ref": "nixos-24.11",
|
||||||
"repo": "nixpkgs",
|
"repo": "nixpkgs",
|
||||||
"type": "github"
|
"type": "github"
|
||||||
}
|
}
|
||||||
},
|
},
|
||||||
"nixpkgs-lib": {
|
"nixpkgs-lib": {
|
||||||
"locked": {
|
"locked": {
|
||||||
"lastModified": 1748740939,
|
"lastModified": 1743296961,
|
||||||
"narHash": "sha256-rQaysilft1aVMwF14xIdGS3sj1yHlI6oKQNBRTF40cc=",
|
"narHash": "sha256-b1EdN3cULCqtorQ4QeWgLMrd5ZGOjLSLemfa00heasc=",
|
||||||
"owner": "nix-community",
|
"owner": "nix-community",
|
||||||
"repo": "nixpkgs.lib",
|
"repo": "nixpkgs.lib",
|
||||||
"rev": "656a64127e9d791a334452c6b6606d17539476e2",
|
"rev": "e4822aea2a6d1cdd36653c134cacfd64c97ff4fa",
|
||||||
"type": "github"
|
"type": "github"
|
||||||
},
|
},
|
||||||
"original": {
|
"original": {
|
||||||
|
@ -98,27 +100,11 @@
|
||||||
"url": "https://github.com/NixOS/nixpkgs/archive/fb192fec7cc7a4c26d51779e9bab07ce6fa5597a.tar.gz"
|
"url": "https://github.com/NixOS/nixpkgs/archive/fb192fec7cc7a4c26d51779e9bab07ce6fa5597a.tar.gz"
|
||||||
}
|
}
|
||||||
},
|
},
|
||||||
"nixpkgs_2": {
|
|
||||||
"locked": {
|
|
||||||
"lastModified": 1751048012,
|
|
||||||
"narHash": "sha256-MYbotu4UjWpTsq01wglhN5xDRfZYLFtNk7SBY0BcjkU=",
|
|
||||||
"owner": "nixos",
|
|
||||||
"repo": "nixpkgs",
|
|
||||||
"rev": "a684c58d46ebbede49f280b653b9e56100aa3877",
|
|
||||||
"type": "github"
|
|
||||||
},
|
|
||||||
"original": {
|
|
||||||
"owner": "nixos",
|
|
||||||
"ref": "nixos-24.11",
|
|
||||||
"repo": "nixpkgs",
|
|
||||||
"type": "github"
|
|
||||||
}
|
|
||||||
},
|
|
||||||
"root": {
|
"root": {
|
||||||
"inputs": {
|
"inputs": {
|
||||||
"flake-parts": "flake-parts",
|
"flake-parts": "flake-parts",
|
||||||
"lean4-nix": "lean4-nix",
|
"lean4-nix": "lean4-nix",
|
||||||
"nixpkgs": "nixpkgs_2"
|
"nixpkgs": "nixpkgs"
|
||||||
}
|
}
|
||||||
}
|
}
|
||||||
},
|
},
|
||||||
|
|
|
@ -4,7 +4,10 @@
|
||||||
inputs = {
|
inputs = {
|
||||||
nixpkgs.url = "github:nixos/nixpkgs/nixos-24.11";
|
nixpkgs.url = "github:nixos/nixpkgs/nixos-24.11";
|
||||||
flake-parts.url = "github:hercules-ci/flake-parts";
|
flake-parts.url = "github:hercules-ci/flake-parts";
|
||||||
lean4-nix.url = "github:lenianiva/lean4-nix";
|
lean4-nix = {
|
||||||
|
url = "github:lenianiva/lean4-nix";
|
||||||
|
inputs.nixpkgs.follows = "nixpkgs";
|
||||||
|
};
|
||||||
};
|
};
|
||||||
|
|
||||||
outputs = inputs @ {
|
outputs = inputs @ {
|
||||||
|
|
|
@ -10,7 +10,6 @@ lean_lib Pantograph {
|
||||||
|
|
||||||
lean_lib Repl {
|
lean_lib Repl {
|
||||||
}
|
}
|
||||||
|
|
||||||
@[default_target]
|
@[default_target]
|
||||||
lean_exe repl {
|
lean_exe repl {
|
||||||
root := `Main
|
root := `Main
|
||||||
|
@ -18,13 +17,6 @@ lean_exe repl {
|
||||||
supportInterpreter := true
|
supportInterpreter := true
|
||||||
}
|
}
|
||||||
|
|
||||||
@[default_target]
|
|
||||||
lean_exe tomograph {
|
|
||||||
root := `Tomograph
|
|
||||||
-- Solves the native symbol not found problem
|
|
||||||
supportInterpreter := true
|
|
||||||
}
|
|
||||||
|
|
||||||
require LSpec from git
|
require LSpec from git
|
||||||
"https://github.com/argumentcomputer/LSpec.git" @ "a6652a48b5c67b0d8dd3930fad6390a97d127e8d"
|
"https://github.com/argumentcomputer/LSpec.git" @ "a6652a48b5c67b0d8dd3930fad6390a97d127e8d"
|
||||||
lean_lib Test {
|
lean_lib Test {
|
||||||
|
|
|
@ -1 +1 @@
|
||||||
leanprover/lean4:v4.21.0
|
leanprover/lean4:v4.18.0
|
||||||
|
|
Loading…
Reference in New Issue