Compare commits

..

No commits in common. "dev" and "v0.2.21" have entirely different histories.
dev ... v0.2.21

51 changed files with 2413 additions and 3459 deletions

6
.gitignore vendored
View File

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

View File

@ -8,68 +8,64 @@ 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 let s := s.trim
| .some '{' => match s.get? 0 with
-- Parse in Json mode | .some '{' => -- Parse in Json mode
Lean.fromJson? (← Lean.Json.parse s) Lean.fromJson? (← Lean.Json.parse s)
| .some _ => | .some _ => -- Parse in line mode
-- Parse in line mode
let offset := s.posOf ' ' |> s.offsetOfPos let offset := s.posOf ' ' |> s.offsetOfPos
if offset = s.length then if offset = s.length then
return { cmd := s.take offset, payload := Lean.Json.null } return { cmd := s.take offset, payload := Lean.Json.null }
else else
let payload ← s.drop offset |> Lean.Json.parse let payload ← s.drop offset |> Lean.Json.parse
return { cmd := s.take offset, payload := payload } return { cmd := s.take offset, payload := payload }
| .none => | .none => throw "Command is empty"
throw "Command is empty"
partial def loop : MainM Unit := do repeat do partial def loop : MainM Unit := do
let state ← get let state ← get
let command ← (← IO.getStdin).getLine let command ← (← IO.getStdin).getLine
-- Halt the program if empty line is given if command.trim.length = 0 then return ()
if command.trim.length = 0 then break
match parseCommand command with match parseCommand command with
| .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.toMessageData.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
loop
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.
-- Separate imports and options
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 let coreContext ← args.filterMap (λ s => if s.startsWith "--" then .some <| s.drop 2 else .none)
let (options, imports) := args.partition (·.startsWith "--") |>.toArray |> Pantograph.createCoreContext
let coreContext ← options.map (·.drop 2) |>.toArray |> Pantograph.createCoreContext let imports:= args.filter (λ s => ¬ (s.startsWith "--"))
let coreState ← Pantograph.createCoreState imports.toArray let coreState ← Pantograph.createCoreState imports.toArray
let context: Context := {
imports
}
try try
let mainM := loop.run { coreContext } |>.run' { env := coreState.env } let coreM := loop.run context |>.run' {}
printImmediate "ready." IO.println "ready."
mainM discard <| coreM.toIO coreContext coreState
catch ex => catch ex =>
let message := ex.toString let message := ex.toString
let error := Lean.toJson ({ error := "io", desc := message }: InteractionError) let error := Lean.toJson ({ error := "io", desc := message }: InteractionError)

View File

@ -12,76 +12,31 @@ open Lean
namespace Pantograph namespace Pantograph
inductive Projection where structure ProjectionApplication where
-- Normal field case projector: Name
| field (projector : Name) (numParams : Nat) numParams: Nat
-- Singular inductive case inner: Expr
| singular (recursor : Name) (numParams : Nat) (numFields : Nat)
/-- Converts a `.proj` expression to a form suitable for exporting/transpilation -/
@[export pantograph_analyze_projection]
def analyzeProjection (env: Environment) (e: Expr): Projection :=
let (typeName, idx, _) := match e with
| .proj typeName idx struct => (typeName, idx, struct)
| _ => panic! "Argument must be proj"
if (getStructureInfo? env typeName).isSome then
let ctor := getStructureCtor env typeName
let fieldName := (getStructureFields env typeName)[idx]!
let projector := getProjFnForField? env typeName fieldName |>.get!
.field projector ctor.numParams
else
let recursor := mkRecOnName typeName
let ctor := getStructureCtor env typeName
.singular recursor ctor.numParams ctor.numFields
def anonymousLevel : Level := .mvar ⟨.anonymous⟩
@[export pantograph_expr_proj_to_app] @[export pantograph_expr_proj_to_app]
def exprProjToApp (env : Environment) (e : Expr) : Expr := def exprProjToApp (env: Environment) (e: Expr): ProjectionApplication :=
let anon : Expr := .mvar ⟨.anonymous⟩ let (typeName, idx, inner) := match e with
match analyzeProjection env e with | .proj typeName idx inner => (typeName, idx, inner)
| .field projector numParams => | _ => panic! "Argument must be proj"
let info := match env.find? projector with let ctor := getStructureCtor env typeName
| .some info => info let fieldName := getStructureFields env typeName |>.get! idx
| _ => panic! "Illegal projector" let projector := getProjFnForField? env typeName fieldName |>.get!
let callee := .const projector $ List.replicate info.numLevelParams anonymousLevel {
let args := (List.replicate numParams anon) ++ [e.projExpr!] projector,
mkAppN callee args.toArray numParams := ctor.numParams,
| .singular recursor numParams numFields => inner,
let info := match env.find? recursor with }
| .some info => info
| _ => panic! "Illegal recursor"
let callee := .const recursor $ List.replicate info.numLevelParams anonymousLevel
let typeArgs := List.replicate numParams anon
-- Motive type can be inferred directly
let motive := .lam .anonymous anon anon .default
let major := e.projExpr!
-- Generate a lambda of `numFields` parameters, and returns the `e.projIdx!` one.
let induct := List.foldl
(λ acc _ => .lam .anonymous anon acc .default)
(.bvar $ (numFields - e.projIdx! - 1))
(List.range numFields)
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]
def unfoldAuxLemmas (e : Expr) : CoreM Expr := do def unfoldAuxLemmas (e : Expr) : CoreM Expr := do
Meta.deltaExpand e isAuxLemma Lean.Meta.deltaExpand e Lean.Name.isAuxLemma
/-- Unfold all matcher applications -/
@[export pantograph_unfold_matchers_m]
def unfoldMatchers (expr : Expr) : CoreM Expr :=
Core.transform expr λ e => do
let .some mapp ← Meta.matchMatcherApp? e | return .continue e
let .some matcherInfo := (← getEnv).find? mapp.matcherName | panic! "Matcher must exist"
let f ← Meta.instantiateValueLevelParams matcherInfo mapp.matcherLevels.toList
let mdata := KVMap.empty.insert `matcher (DataValue.ofName mapp.matcherName)
return .visit $ .mdata mdata (f.betaRev e.getAppRevArgs (useZeta := true))
/-- /--
Force the instantiation of delayed metavariables even if they cannot be fully Force the instantiation of delayed metavariables even if they cannot be fully
@ -98,12 +53,13 @@ This function ensures any metavariable in the result is either
1. Delayed assigned with its pending mvar not assigned in any form 1. Delayed assigned with its pending mvar not assigned in any form
2. Not assigned (delay or not) 2. Not assigned (delay or not)
-/ -/
partial def instantiateDelayedMVars (expr : Expr) : MetaM Expr := partial def instantiateDelayedMVars (eOrig: Expr) : MetaM Expr := do
withTraceNode `Pantograph.Delate (λ _ex => return m!":= {← Meta.ppExpr expr}") do --let padding := String.join $ List.replicate level "│ "
let mut result ← Meta.transform (← instantiateMVars expr) --IO.println s!"{padding}Starting {toString eOrig}"
λ e => e.withApp fun f args => do let mut result ← Meta.transform (← instantiateMVars eOrig)
(pre := fun e => e.withApp fun f args => do
let .mvar mvarId := f | return .continue let .mvar mvarId := f | return .continue
trace[Pantograph.Delate] "V {e}" --IO.println s!"{padding}├V {e}"
let mvarDecl ← mvarId.getDecl let mvarDecl ← mvarId.getDecl
-- This is critical to maintaining the interdependency of metavariables. -- This is critical to maintaining the interdependency of metavariables.
@ -121,63 +77,60 @@ partial def instantiateDelayedMVars (expr : Expr) : MetaM Expr :=
panic! s!"In the context of {mvarId.name}, there are local context variable violations: {violations}" panic! s!"In the context of {mvarId.name}, there are local context variable violations: {violations}"
if let .some assign ← getExprMVarAssignment? mvarId then if let .some assign ← getExprMVarAssignment? mvarId then
trace[Pantograph.Delate] "A ?{mvarId.name}" --IO.println s!"{padding}├A ?{mvarId.name}"
assert! !(← mvarId.isDelayedAssigned) assert! !(← mvarId.isDelayedAssigned)
return .visit (mkAppN assign args) return .visit (mkAppN assign args)
else if let some { fvars, mvarIdPending } ← getDelayedMVarAssignment? mvarId then else if let some { fvars, mvarIdPending } ← getDelayedMVarAssignment? mvarId then
if ← isTracingEnabledFor `Pantograph.Delate then --let substTableStr := String.intercalate ", " $ Array.zipWith fvars args (λ fvar assign => s!"{fvar.fvarId!.name} := {assign}") |>.toList
let substTableStr := ",".intercalate $ --IO.println s!"{padding}├MD ?{mvarId.name} := ?{mvarIdPending.name} [{substTableStr}]"
Array.zipWith (λ fvar assign => s!"{fvar.fvarId!.name} := {assign}") fvars args |>.toList
trace[Pantograph.Delate]"MD ?{mvarId.name} := ?{mvarIdPending.name} [{substTableStr}]"
if args.size < fvars.size then if args.size < fvars.size then
throwError "Not enough arguments to instantiate a delay assigned mvar. This is due to bad implementations of a tactic: {args.size} < {fvars.size}. Expr: {toString e}; Origin: {toString expr}" throwError "Not enough arguments to instantiate a delay assigned mvar. This is due to bad implementations of a tactic: {args.size} < {fvars.size}. Expr: {toString e}; Origin: {toString eOrig}"
if !args.isEmpty then --if !args.isEmpty then
trace[Pantograph.Delate] "─ Arguments Begin" --IO.println s!"{padding}├── Arguments Begin"
let args ← args.mapM self let args ← args.mapM self
if !args.isEmpty then --if !args.isEmpty then
trace[Pantograph.Delate] "─ Arguments End" --IO.println s!"{padding}├── Arguments End"
if !(← mvarIdPending.isAssignedOrDelayedAssigned) then if !(← mvarIdPending.isAssignedOrDelayedAssigned) then
trace[Pantograph.Delate] "T1" --IO.println s!"{padding}├T1"
let result := mkAppN f args let result := mkAppN f args
return .done result return .done result
let pending ← mvarIdPending.withContext do let pending ← mvarIdPending.withContext do
let inner ← instantiateDelayedMVars (.mvar mvarIdPending) let inner ← instantiateDelayedMVars (.mvar mvarIdPending) --(level := level + 1)
trace[Pantograph.Delate] "Pre: {inner}" --IO.println s!"{padding}├Pre: {inner}"
pure <| (← inner.abstractM fvars).instantiateRev args pure <| (← inner.abstractM fvars).instantiateRev args
-- Tail arguments -- Tail arguments
let result := mkAppRange pending fvars.size args.size args let result := mkAppRange pending fvars.size args.size args
trace[Pantograph.Delate] "MD {result}" --IO.println s!"{padding}├MD {result}"
return .done result return .done result
else else
assert! !(← mvarId.isAssigned) assert! !(← mvarId.isAssigned)
assert! !(← mvarId.isDelayedAssigned) assert! !(← mvarId.isDelayedAssigned)
if !args.isEmpty then --if !args.isEmpty then
trace[Pantograph.Delate] "─ Arguments Begin" -- IO.println s!"{padding}├── Arguments Begin"
let args ← args.mapM self let args ← args.mapM self
if !args.isEmpty then --if !args.isEmpty then
trace[Pantograph.Delate] "─ Arguments End" -- IO.println s!"{padding}├── Arguments End"
trace[Pantograph.Delate] "M ?{mvarId.name}" --IO.println s!"{padding}├M ?{mvarId.name}"
return .done (mkAppN f args) return .done (mkAppN f args))
trace[Pantoraph.Delate] "Result {result}" --IO.println s!"{padding}└Result {result}"
return result return result
where where
self e := instantiateDelayedMVars e self e := instantiateDelayedMVars e --(level := level + 1)
/-- /--
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
3. No assigned mvars 3. No assigned mvars
-/ -/
@[export pantograph_instantiate_all_m] @[export pantograph_instantiate_all_m]
def instantiateAll (e: Expr): MetaM Expr := do def instantiateAll (e: Expr): MetaM Expr := do
let e ← instantiateDelayedMVars e let e ← instantiateDelayedMVars e
let e ← unfoldAuxLemmas e let e ← unfoldAuxLemmas e
let e ← unfoldMatchers e
return e return e
structure DelayedMVarInvocation where structure DelayedMVarInvocation where
@ -311,36 +264,38 @@ def serializeName (name: Name) (sanitize: Bool := true): String :=
if n.contains Lean.idBeginEscape then s!"{quote}{n}{quote}" else n if n.contains Lean.idBeginEscape then s!"{quote}{n}{quote}" else n
/-- serialize a sort level. Expression is optimized to be compact e.g. `(+ u 2)` -/ /-- serialize a sort level. Expression is optimized to be compact e.g. `(+ u 2)` -/
partial def serializeSortLevel (level: Level) : String := partial def serializeSortLevel (level: Level) (sanitize: Bool): String :=
let k := level.getOffset let k := level.getOffset
let u := level.getLevelOffset let u := level.getLevelOffset
let u_str := match u with let u_str := match u with
| .zero => "0" | .zero => "0"
| .succ _ => panic! "getLevelOffset should not return .succ" | .succ _ => panic! "getLevelOffset should not return .succ"
| .max v w => | .max v w =>
let v := serializeSortLevel v let v := serializeSortLevel v sanitize
let w := serializeSortLevel w let w := serializeSortLevel w sanitize
s!"(:max {v} {w})" s!"(:max {v} {w})"
| .imax v w => | .imax v w =>
let v := serializeSortLevel v let v := serializeSortLevel v sanitize
let w := serializeSortLevel w let w := serializeSortLevel w sanitize
s!"(:imax {v} {w})" s!"(:imax {v} {w})"
| .param name => | .param name =>
let name := serializeName name sanitize
s!"{name}" s!"{name}"
| .mvar id => | .mvar id =>
let name := id.name let name := serializeName id.name sanitize
s!"(:mv {name})" s!"(:mv {name})"
match k, u with match k, u with
| 0, _ => u_str | 0, _ => u_str
| _, .zero => s!"{k}" | _, .zero => s!"{k}"
| _, _ => s!"(+ {u_str} {k})" | _, _ => s!"(+ {u_str} {k})"
/-- /--
Completely serializes an expression tree. Json not used due to compactness Completely serializes an expression tree. Json not used due to compactness
A `_` symbol in the AST indicates automatic deductions not present in the original expression. A `_` symbol in the AST indicates automatic deductions not present in the original expression.
-/ -/
partial def serializeExpressionSexp (expr: Expr) : MetaM String := do partial def serializeExpressionSexp (expr: Expr) (sanitize: Bool := true): MetaM String := do
self expr self expr
where where
delayedMVarToSexp (e: Expr): MetaM (Option String) := do delayedMVarToSexp (e: Expr): MetaM (Option String) := do
@ -379,10 +334,9 @@ partial def serializeExpressionSexp (expr: Expr) : MetaM String := do
let name := mvarId.name let name := mvarId.name
pure s!"(:{pref} {name})" pure s!"(:{pref} {name})"
| .sort level => | .sort level =>
let level := serializeSortLevel level let level := serializeSortLevel level sanitize
pure s!"(:sort {level})" pure s!"(:sort {level})"
| .const declName _ => | .const declName _ =>
let declName := serializeName declName (sanitize := false)
-- The universe level of the const expression is elided since it should be -- The universe level of the const expression is elided since it should be
-- inferrable from surrounding expression -- inferrable from surrounding expression
pure s!"(:c {declName})" pure s!"(:c {declName})"
@ -415,29 +369,24 @@ partial def serializeExpressionSexp (expr: Expr) : MetaM String := do
-- is wrapped in a :lit sexp. -- is wrapped in a :lit sexp.
let v' := match v with let v' := match v with
| .natVal val => toString val | .natVal val => toString val
| .strVal val => IR.EmitC.quoteString val | .strVal val => s!"\"{val}\""
pure s!"(:lit {v'})" pure s!"(:lit {v'})"
| .mdata _ inner => | .mdata _ inner =>
-- NOTE: Equivalent to expr itself, but mdata influences the prettyprinter -- NOTE: Equivalent to expr itself, but mdata influences the prettyprinter
-- It may become necessary to incorporate the metadata. -- It may become necessary to incorporate the metadata.
self inner self inner
| .proj typeName idx inner => do | .proj _ _ _ => do
let env ← getEnv let env ← getEnv
match analyzeProjection env e with let projApp := exprProjToApp env e
| .field projector numParams => let autos := String.intercalate " " (List.replicate projApp.numParams "_")
let autos := String.intercalate " " (List.replicate numParams "_") let inner ← self projApp.inner
let inner' ← self inner pure s!"((:c {projApp.projector}) {autos} {inner})"
pure s!"((:c {projector}) {autos} {inner'})"
| .singular _ _ _ =>
let typeName' := serializeName typeName (sanitize := false)
let e' ← self e
pure s!"(:proj {typeName'} {idx} {e'})"
-- Elides all unhygenic names -- Elides all unhygenic names
binderInfoSexp : Lean.BinderInfo → String binderInfoSexp : Lean.BinderInfo → String
| .default => "" | .default => ""
| .implicit => " :i" | .implicit => " :implicit"
| .strictImplicit => " :si" | .strictImplicit => " :strictImplicit"
| .instImplicit => " :ii" | .instImplicit => " :instImplicit"
def serializeExpression (options: @&Protocol.Options) (e: Expr): MetaM Protocol.Expression := do def serializeExpression (options: @&Protocol.Options) (e: Expr): MetaM Protocol.Expression := do
let pp?: Option String ← match options.printExprPretty with let pp?: Option String ← match options.printExprPretty with
@ -455,6 +404,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 +470,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 +480,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 -/
@ -584,7 +532,7 @@ protected def GoalState.diag (goalState: GoalState) (parent?: Option GoalState :
then instantiateAll decl.type then instantiateAll decl.type
else pure $ decl.type else pure $ decl.type
let type_sexp ← if options.printSexp then let type_sexp ← if options.printSexp then
let sexp ← serializeExpressionSexp type let sexp ← serializeExpressionSexp type (sanitize := false)
pure <| " " ++ sexp pure <| " " ++ sexp
else else
pure "" pure ""
@ -608,11 +556,6 @@ 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
registerTraceClass `Pantograph.Delate
end Pantograph end Pantograph

View File

@ -6,8 +6,7 @@ namespace Pantograph
-- Functions for creating contexts and states -- Functions for creating contexts and states
@[export pantograph_default_elab_context] @[export pantograph_default_elab_context]
def defaultElabContext: Elab.Term.Context := { def defaultElabContext: Elab.Term.Context := {
declName? := .some `mystery, errToSorry := false
errToSorry := false,
} }
/-- Read syntax object from string -/ /-- Read syntax object from string -/

View File

@ -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]
@ -26,10 +25,7 @@ def env_catalog (env: Environment): Array Name := env.constants.fold (init := #[
@[export pantograph_environment_module_of_name] @[export pantograph_environment_module_of_name]
def module_of_name (env: Environment) (name: Name): Option Name := do def module_of_name (env: Environment) (name: Name): Option Name := do
let moduleId ← env.getModuleIdxFor? name let moduleId ← env.getModuleIdxFor? name
if h : moduleId.toNat < env.allImportedModuleNames.size then return env.allImportedModuleNames.get! moduleId.toNat
return env.allImportedModuleNames[moduleId.toNat]
else
.none
def toCompactSymbolName (n: Name) (info: ConstantInfo): String := def toCompactSymbolName (n: Name) (info: ConstantInfo): String :=
let pref := match info with let pref := match info with
@ -47,22 +43,6 @@ def toFilteredSymbol (n: Lean.Name) (info: Lean.ConstantInfo): Option String :=
if isNameInternal n || info.isUnsafe if isNameInternal n || info.isUnsafe
then Option.none then Option.none
else Option.some <| toCompactSymbolName n info else Option.some <| toCompactSymbolName n info
def describe (_: Protocol.EnvDescribe): CoreM Protocol.EnvDescribeResult := do
let env ← Lean.MonadEnv.getEnv
return {
imports := env.header.imports.map toString,
modules := env.header.moduleNames.map (·.toString),
}
def moduleRead (args: Protocol.EnvModuleRead): CoreM Protocol.EnvModuleReadResult := do
let env ← Lean.MonadEnv.getEnv
let .some i := env.header.moduleNames.findIdx? (· == args.module.toName) |
throwError s!"Module not found {args.module}"
let data := env.header.moduleData[i]!
return {
imports := data.imports.map toString,
constNames := data.constNames.map (·.toString),
extraConstNames := data.extraConstNames.map (·.toString),
}
def catalog (_: Protocol.EnvCatalog): CoreM Protocol.EnvCatalogResult := do def catalog (_: Protocol.EnvCatalog): CoreM Protocol.EnvCatalogResult := do
let env ← Lean.MonadEnv.getEnv let env ← Lean.MonadEnv.getEnv
let names := env.constants.fold (init := #[]) (λ acc name info => let names := env.constants.fold (init := #[]) (λ acc name info =>
@ -70,13 +50,15 @@ def catalog (_: Protocol.EnvCatalog): CoreM Protocol.EnvCatalogResult := do
| .some x => acc.push x | .some x => acc.push x
| .none => acc) | .none => acc)
return { symbols := names } return { symbols := names }
def inspect (args: Protocol.EnvInspect) (options: @&Protocol.Options): Protocol.FallibleT CoreM Protocol.EnvInspectResult := do def inspect (args: Protocol.EnvInspect) (options: @&Protocol.Options): CoreM (Protocol.CR Protocol.EnvInspectResult) := do
let env ← Lean.MonadEnv.getEnv let env ← Lean.MonadEnv.getEnv
let name := args.name.toName let name := args.name.toName
let info? := env.find? name let info? := env.find? name
let .some info := info? | Protocol.throw $ Protocol.errorIndex s!"Symbol not found {args.name}" let info ← match info? with
| none => return .error $ Protocol.errorIndex s!"Symbol not found {args.name}"
| some info => pure info
let module? := env.getModuleIdxFor? name >>= let module? := env.getModuleIdxFor? name >>=
(λ idx => env.allImportedModuleNames[idx.toNat]?) (λ idx => env.allImportedModuleNames.get? idx.toNat) |>.map toString
let value? := match args.value?, info with let value? := match args.value?, info with
| .some true, _ => info.value? | .some true, _ => info.value?
| .some false, _ => .none | .some false, _ => .none
@ -90,6 +72,7 @@ def inspect (args: Protocol.EnvInspect) (options: @&Protocol.Options): Protocol.
isUnsafe := info.isUnsafe, isUnsafe := info.isUnsafe,
value? := ← value?.mapM (λ v => serializeExpression options v |>.run'), value? := ← value?.mapM (λ v => serializeExpression options v |>.run'),
publicName? := Lean.privateToUserName? name |>.map (·.toString), publicName? := Lean.privateToUserName? name |>.map (·.toString),
-- BUG: Warning: getUsedConstants here will not include projections. This is a known bug.
typeDependency? := if args.dependency?.getD false typeDependency? := if args.dependency?.getD false
then .some <| type.getUsedConstants.map (λ n => serializeName n) then .some <| type.getUsedConstants.map (λ n => serializeName n)
else .none, else .none,
@ -97,7 +80,7 @@ def inspect (args: Protocol.EnvInspect) (options: @&Protocol.Options): Protocol.
then value?.map (λ e => then value?.map (λ e =>
e.getUsedConstants.filter (!isNameInternal ·) |>.map (λ n => serializeName n) ) e.getUsedConstants.filter (!isNameInternal ·) |>.map (λ n => serializeName n) )
else .none, else .none,
module? := module?.map (·.toString) module? := module?
} }
let result ← match info with let result ← match info with
| .inductInfo induct => pure { core with inductInfo? := .some { | .inductInfo induct => pure { core with inductInfo? := .some {
@ -130,72 +113,44 @@ def inspect (args: Protocol.EnvInspect) (options: @&Protocol.Options): Protocol.
k := r.k, k := r.k,
} } } }
| _ => pure core | _ => pure core
let result ← if args.source?.getD false then return .ok result
try def addDecl (args: Protocol.EnvAdd): CoreM (Protocol.CR Protocol.EnvAddResult) := do
let sourceUri? ← module?.bindM (Server.documentUriFromModule? ·)
let declRange? ← findDeclarationRanges? name
let sourceStart? := declRange?.map (·.range.pos)
let sourceEnd? := declRange?.map (·.range.endPos)
.pure {
result with
sourceUri? := sourceUri?.map (toString ·),
sourceStart?,
sourceEnd?,
}
catch _e =>
.pure result
else
.pure result
return result
/-- Elaborates and adds a declaration to the `CoreM` environment. -/
@[export pantograph_env_add_m]
def addDecl (name: String) (levels: Array String := #[]) (type?: Option String) (value: String) (isTheorem: Bool)
: Protocol.FallibleT CoreM Protocol.EnvAddResult := do
let env ← Lean.MonadEnv.getEnv let env ← Lean.MonadEnv.getEnv
let levelParams := levels.toList.map (·.toName) let tvM: Elab.TermElabM (Except String (Expr × Expr)) := do
let tvM: Elab.TermElabM (Except String (Expr × Expr)) := let type ← match parseTerm env args.type with
Elab.Term.withLevelNames levelParams do do | .ok syn => do
let expectedType?? : Except String (Option Expr) ← ExceptT.run $ type?.mapM λ type => do match ← elabTerm syn with
match parseTerm env type with
| .ok syn => elabTerm syn
| .error e => MonadExceptOf.throw e
let expectedType? ← match expectedType?? with
| .ok t? => pure t?
| .error e => return .error e | .error e => return .error e
let value ← match parseTerm env value with | .ok expr => pure expr
| .error e => return .error e
let value ← match parseTerm env args.value with
| .ok syn => do | .ok syn => do
try try
let expr ← Elab.Term.elabTerm (stx := syn) (expectedType? := expectedType?) let expr ← Elab.Term.elabTerm (stx := syn) (expectedType? := .some type)
Lean.Elab.Term.synthesizeSyntheticMVarsNoPostponing Lean.Elab.Term.synthesizeSyntheticMVarsNoPostponing
let expr ← instantiateMVars expr let expr ← instantiateMVars expr
pure $ expr pure $ expr
catch ex => return .error (← ex.toMessageData.toString) catch ex => return .error (← ex.toMessageData.toString)
| .error e => return .error e | .error e => return .error e
Elab.Term.synthesizeSyntheticMVarsNoPostponing pure $ .ok (type, value)
let type ← match expectedType? with
| .some t => pure t
| .none => Meta.inferType value
pure $ .ok (← instantiateMVars type, ← instantiateMVars value)
let (type, value) ← match ← tvM.run' (ctx := {}) |>.run' with let (type, value) ← match ← tvM.run' (ctx := {}) |>.run' with
| .ok t => pure t | .ok t => pure t
| .error e => Protocol.throw $ Protocol.errorExpr e | .error e => return .error $ Protocol.errorExpr e
let decl := if isTheorem then let constant := Lean.Declaration.defnDecl <| Lean.mkDefinitionValEx
Lean.Declaration.thmDecl <| Lean.mkTheoremValEx (name := args.name.toName)
(name := name.toName) (levelParams := [])
(levelParams := levelParams)
(type := type)
(value := value)
(all := [])
else
Lean.Declaration.defnDecl <| Lean.mkDefinitionValEx
(name := name.toName)
(levelParams := levelParams)
(type := type) (type := type)
(value := value) (value := value)
(hints := Lean.mkReducibilityHintsRegularEx 1) (hints := Lean.mkReducibilityHintsRegularEx 1)
(safety := Lean.DefinitionSafety.safe) (safety := Lean.DefinitionSafety.safe)
(all := []) (all := [])
Lean.addDecl decl let env' ← match env.addDecl (← getOptions) constant with
return {} | .error e => do
let options ← Lean.MonadOptions.getOptions
let desc ← (e.toMessageData options).toString
return .error $ { error := "kernel", desc }
| .ok env' => pure env'
Lean.MonadEnv.modifyEnv (λ _ => env')
return .ok {}
end Pantograph.Environment end Pantograph.Environment

View File

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

View File

@ -30,17 +30,9 @@ end Lean.PersistentArray
namespace Pantograph.Frontend namespace Pantograph.Frontend
@[export pantograph_frontend_stx_byte_range]
def stxByteRange (stx : Syntax) : String.Pos × String.Pos :=
let pos := stx.getPos?.getD 0
let endPos := stx.getTailPos?.getD 0
(pos, endPos)
abbrev FrontendM := Elab.Frontend.FrontendM abbrev FrontendM := Elab.Frontend.FrontendM
structure CompilationStep where structure CompilationStep where
scope : Elab.Command.Scope
fileName : String fileName : String
fileMap : FileMap fileMap : FileMap
src : Substring src : Substring
@ -68,14 +60,14 @@ def processOneCommand: FrontendM (CompilationStep × Bool) := do
let s := (← get).commandState let s := (← get).commandState
let before := s.env let before := s.env
let done ← Elab.Frontend.processCommand let done ← Elab.Frontend.processCommand
let stx := (← get).commands.back! let stx := (← get).commands.back
let src := (← read).inputCtx.input.toSubstring.extract (← get).cmdPos (← get).parserState.pos let src := (← read).inputCtx.input.toSubstring.extract (← get).cmdPos (← get).parserState.pos
let s' := (← get).commandState let s' := (← get).commandState
let after := s'.env let after := s'.env
let msgs := s'.messages.toList.drop s.messages.toList.length let msgs := s'.messages.toList.drop s.messages.toList.length
let trees := s'.infoState.trees.drop s.infoState.trees.size let trees := s'.infoState.trees.drop s.infoState.trees.size
let ⟨_, fileName, fileMap⟩ := (← read).inputCtx let ⟨_, fileName, fileMap⟩ := (← read).inputCtx
return ({ scope := s.scopes.head!, fileName, fileMap, src, stx, before, after, msgs, trees }, done) return ({ fileName, fileMap, src, stx, before, after, msgs, trees }, done)
partial def mapCompilationSteps { α } (f: CompilationStep → IO α) : FrontendM (List α) := do partial def mapCompilationSteps { α } (f: CompilationStep → IO α) : FrontendM (List α) := do
let (cmd, done) ← processOneCommand let (cmd, done) ← processOneCommand

View File

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

View File

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

View File

@ -62,14 +62,13 @@ private partial def translateExpr (srcExpr: Expr) : MetaTranslateM Expr := do
let sourceMCtx ← getSourceMCtx let sourceMCtx ← getSourceMCtx
-- We want to create as few mvars as possible -- We want to create as few mvars as possible
let (srcExpr, _) := instantiateMVarsCore (mctx := sourceMCtx) srcExpr let (srcExpr, _) := instantiateMVarsCore (mctx := sourceMCtx) srcExpr
trace[Pantograph.Frontend.MetaTranslate] "Transform src: {srcExpr}" --IO.println s!"Transform src: {srcExpr}"
let result ← Core.transform srcExpr λ e => do let result ← Core.transform srcExpr λ e => do
let state ← get let state ← get
match e with match e with
| .fvar fvarId => | .fvar fvarId =>
let .some fvarId' := state.fvarMap[fvarId]? | panic! s!"FVar id not registered: {fvarId.name}" let .some fvarId' := state.fvarMap[fvarId]? | panic! s!"FVar id not registered: {fvarId.name}"
-- Delegating this to `Meta.check` later assert! (← getLCtx).contains fvarId'
--assert! (← getLCtx).contains fvarId'
return .done $ .fvar fvarId' return .done $ .fvar fvarId'
| .mvar mvarId => do | .mvar mvarId => do
-- Must not be assigned -- Must not be assigned
@ -100,10 +99,10 @@ partial def translateLocalDecl (srcLocalDecl: LocalDecl) : MetaTranslateM LocalD
addTranslatedFVar srcLocalDecl.fvarId fvarId addTranslatedFVar srcLocalDecl.fvarId fvarId
match srcLocalDecl with match srcLocalDecl with
| .cdecl index _ userName type bi kind => do | .cdecl index _ userName type bi kind => do
trace[Pantograph.Frontend.MetaTranslate] "[CD] {userName} {toString type}" --IO.println s!"[CD] {userName} {toString type}"
return .cdecl index fvarId userName (← translateExpr type) bi kind return .cdecl index fvarId userName (← translateExpr type) bi kind
| .ldecl index _ userName type value nonDep kind => do | .ldecl index _ userName type value nonDep kind => do
trace[Pantograph.Frontend.MetaTranslate] "[LD] {toString type} := {toString value}" --IO.println s!"[LD] {toString type} := {toString value}"
return .ldecl index fvarId userName (← translateExpr type) (← translateExpr value) nonDep kind return .ldecl index fvarId userName (← translateExpr type) (← translateExpr value) nonDep kind
partial def translateLCtx : MetaTranslateM LocalContext := do partial def translateLCtx : MetaTranslateM LocalContext := do
@ -118,7 +117,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 +133,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 +147,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!
@ -165,7 +161,4 @@ end MetaTranslate
export MetaTranslate (MetaTranslateM) export MetaTranslate (MetaTranslateM)
initialize
registerTraceClass `Pantograph.Frontend.MetaTranslate
end Pantograph.Frontend end Pantograph.Frontend

View File

@ -10,72 +10,26 @@ import Lean
namespace Pantograph namespace Pantograph
open Lean open Lean
/-- The acting area of a tactic -/ def filename: String := "<pantograph>"
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 +44,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 +53,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,53 +71,67 @@ 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 :=
state.savedState.term.meta.core
protected def GoalState.withContext' (state: GoalState) (mvarId: MVarId) (m: MetaM α): MetaM α := do protected def GoalState.withContext (state: GoalState) (mvarId: MVarId) (m: MetaM α): MetaM α := do
mvarId.withContext m |>.run' (← read) state.metaState mvarId.withContext m |>.run' (← read) state.metaState
protected def GoalState.withContext { m } [MonadControlT MetaM m] [Monad m] (state: GoalState) (mvarId: MVarId) : m α → m α :=
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 α :=
Meta.mapMetaM <| state.withContext' state.parentMVars[0]!
protected def GoalState.withRootContext { n } [MonadControlT MetaM n] [Monad n] (state: GoalState): n α → n α :=
Meta.mapMetaM <| state.withContext' state.root
private def restoreCoreMExtra (state : Core.SavedState) : CoreM Unit := protected def GoalState.withParentContext { n } [MonadControlT MetaM n] [Monad n] (state: GoalState): n α → n α :=
let { nextMacroScope, ngen, .. } := state Meta.mapMetaM <| state.withContext state.parentMVar?.get!
modifyGetThe Core.State (fun st => ((), protected def GoalState.withRootContext { n } [MonadControlT MetaM n] [Monad n] (state: GoalState): n α → n α :=
{ st with nextMacroScope, ngen })) Meta.mapMetaM <| state.withContext state.root
-- Restore the name generator and macro scopes of the core state
protected def GoalState.restoreCoreMExtra (state: GoalState): CoreM Unit := private def GoalState.mvars (state: GoalState): SSet MVarId :=
restoreCoreMExtra state.coreState state.mctx.decls.foldl (init := .empty) fun acc k _ => acc.insert k
protected def GoalState.restoreMetaM (state: GoalState): MetaM Unit := do protected def GoalState.restoreMetaM (state: GoalState): MetaM Unit :=
state.restoreCoreMExtra
state.savedState.term.meta.restore state.savedState.term.meta.restore
protected def GoalState.restoreElabM (state: GoalState): Elab.TermElabM Unit := do protected def GoalState.restoreElabM (state: GoalState): Elab.TermElabM Unit :=
state.restoreCoreMExtra
state.savedState.term.restore state.savedState.term.restore
private def GoalState.restoreTacticM (state: GoalState) (goal: MVarId): Elab.Tactic.TacticM Unit := do private def GoalState.restoreTacticM (state: GoalState) (goal: MVarId): Elab.Tactic.TacticM Unit := do
state.restoreElabM state.savedState.restore
Elab.Tactic.setGoals [goal] Elab.Tactic.setGoals [goal]
@[export pantograph_goal_state_focus]
protected def GoalState.focus (state: GoalState) (goalId: Nat): Option GoalState := do
let goal ← state.savedState.tactic.goals.get? goalId
return {
state with
savedState := {
state.savedState with
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. User must ensure `goals` are distinct. Brings into scope a list of goals
-/ -/
@[export pantograph_goal_state_resume] @[export pantograph_goal_state_resume]
protected def GoalState.resume (state : GoalState) (goals : List MVarId) : Except String GoalState := do protected def GoalState.resume (state: GoalState) (goals: List MVarId): Except String GoalState :=
if ¬ (goals.all (state.mctx.decls.contains ·)) then if ¬ (goals.all (λ goal => state.mvars.contains goal)) then
let invalid_goals := goals.filter (λ goal => ¬ state.mctx.decls.contains goal) |>.map (·.name.toString) let invalid_goals := goals.filter (λ goal => ¬ state.mvars.contains goal) |>.map (·.name.toString)
.error s!"Goals {invalid_goals} are not in scope" .error s!"Goals {invalid_goals} are not in scope"
else
-- Set goals to the goals that have not been assigned yet, similar to the `focus` tactic. -- Set goals to the goals that have not been assigned yet, similar to the `focus` tactic.
let unassigned := goals.filter λ goal => let unassigned := goals.filter (λ goal =>
let isSolved := state.mctx.eAssignment.contains goal || state.mctx.dAssignment.contains goal let mctx := state.mctx
¬ isSolved ¬(mctx.eAssignment.contains goal || mctx.dAssignment.contains goal))
return { .ok {
state with state with
savedState := { savedState := {
term := state.savedState.term, term := state.savedState.term,
@ -193,286 +156,45 @@ protected def GoalState.rootExpr? (goalState : GoalState) : Option Expr := do
.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 ---
/-- protected def GoalState.step (state: GoalState) (goal: MVarId) (tacticM: Elab.Tactic.TacticM Unit)
These descendants serve as "seed" mvars. If a MVarError's mvar is related to one : Elab.TermElabM GoalState := do
of these seed mvars, it means an error has occurred when a tactic was executing unless (← getMCtx).decls.contains goal do
on `src`. `evalTactic`, will not capture these mvars, so we need to manually throwError s!"Goal is not in context: {goal.name}"
find them and save them into the goal list. See the rationales document for the goal.checkNotAssigned `GoalState.step
inspiration of this function. let (_, newGoals) ← tacticM { elaborator := .anonymous } |>.run { goals := [goal] }
-/
private def collectAllErroredMVars (src : MVarId) : Elab.TermElabM (List MVarId) := do
-- Mimics `Elab.Term.logUnassignedUsingErrorInfos`
let descendants ← Meta.getMVars (.mvar src)
--let _ ← Elab.Term.logUnassignedUsingErrorInfos descendants
let mut alreadyVisited : MVarIdSet := {}
let mut result : MVarIdSet := {}
for { mvarId, .. } in (← get).mvarErrorInfos do
unless alreadyVisited.contains mvarId do
alreadyVisited := alreadyVisited.insert mvarId
/- The metavariable `mvarErrorInfo.mvarId` may have been assigned or
delayed assigned to another metavariable that is unassigned. -/
let mvarDeps ← Meta.getMVars (.mvar mvarId)
if mvarDeps.any descendants.contains then do
result := mvarDeps.foldl (·.insert ·) result
return result.toList
/-- Merger of two unique lists -/
private def mergeMVarLists (li1 li2 : List MVarId) : List MVarId :=
let li2' := li2.filter (¬ li1.contains ·)
li1 ++ li2'
/--
Set `guardMVarErrors` to true to capture mvar errors. Lean will not
automatically collect mvars from text tactics (vide
`test_tactic_failure_synthesize_placeholder`)
-/
protected def GoalState.step' { α } (state : GoalState) (site : Site) (tacticM : Elab.Tactic.TacticM α) (guardMVarErrors : Bool := false)
: Elab.TermElabM (α × GoalState) := do
let ((a, parentMVars), { goals }) ← site.runTacticM tacticM
|>.run { elaborator := .anonymous }
|>.run state.savedState.tactic
let nextElabState ← MonadBacktrack.saveState let nextElabState ← MonadBacktrack.saveState
--Elab.Term.synthesizeSyntheticMVarsNoPostponing return {
let goals ← if guardMVarErrors then
parentMVars.foldlM (init := goals) λ goals parent => do
let errors ← collectAllErroredMVars parent
return mergeMVarLists goals errors
else
pure goals
let state' := {
state with state with
savedState := { term := nextElabState, tactic := { goals }, }, savedState := { term := nextElabState, tactic := newGoals },
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 String) | success (state: GoalState)
-- Tactic failed with messages -- Tactic failed with messages
| failure (messages: Array String) | failure (messages: Array String)
-- Could not parse tactic -- Could not parse tactic
@ -480,145 +202,196 @@ inductive TacticResult where
-- The given action cannot be executed in the state -- The given action cannot be executed in the state
| invalidAction (message: String) | invalidAction (message: String)
private def dumpMessageLog (prevMessageLength : Nat := 0) : CoreM (Bool × Array String) := do /-- Executes a `TacticM` monads on this `GoalState`, collecting the errors as necessary -/
let newMessages := (← Core.getMessageLog).toList.drop prevMessageLength protected def GoalState.tryTacticM (state: GoalState) (goal: MVarId) (tacticM: Elab.Tactic.TacticM Unit):
let hasErrors := newMessages.any (·.severity == .error) Elab.TermElabM TacticResult := do
let newMessages ← newMessages.mapM λ m => m.toString
Core.resetMessageLog
return (hasErrors, newMessages.toArray)
/-- Execute a `TermElabM` producing a goal state, capturing the error and turn it into a `TacticResult` -/
def withCapturingError (elabM : Elab.Term.TermElabM GoalState) : Elab.TermElabM TacticResult := do
let messageLog ← Core.getMessageLog
unless messageLog.toList.isEmpty do
IO.eprintln s!"{← messageLog.toList.mapM (·.toString)}"
assert! messageLog.toList.isEmpty
try try
let state ← elabM let nextState ← state.step goal tacticM
return .success nextState
-- Check if error messages have been generated in the core.
let (hasError, newMessages) ← dumpMessageLog
if hasError then
return .failure newMessages
else
return .success state newMessages
catch exception => catch exception =>
match exception with return .failure #[← exception.toMessageData.toString]
| .internal _ =>
let (_, messages) ← dumpMessageLog
return .failure messages
| _ => return .failure #[← exception.toMessageData.toString]
/-- Executes a `TacticM` monad on this `GoalState`, collecting the errors as necessary -/
protected def GoalState.tryTacticM
(state: GoalState) (site : Site)
(tacticM: Elab.Tactic.TacticM Unit)
(guardMVarErrors : Bool := false)
: Elab.TermElabM TacticResult := do
state.restoreElabM
withCapturingError do
state.step site tacticM guardMVarErrors
/-- Execute a string tactic on given state. Restores TermElabM -/ /-- Execute a string tactic on given state. Restores TermElabM -/
@[export pantograph_goal_state_try_tactic_m] protected def GoalState.tryTactic (state: GoalState) (goal: MVarId) (tactic: String):
protected def GoalState.tryTactic (state: GoalState) (site : Site) (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 := filename) 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 state.tryTacticM goal $ Elab.Tactic.evalTactic tactic
withCapturingError do
state.step site tacticM (guardMVarErrors := true) protected def GoalState.tryAssign (state: GoalState) (goal: MVarId) (expr: String):
Elab.TermElabM TacticResult := do
state.restoreElabM
let expr ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := expr)
(fileName := filename) with
| .ok syn => pure syn
| .error error => return .parseError error
state.tryTacticM goal $ Tactic.evalAssign expr
-- Specialized Tactics -- Specialized Tactics
protected def GoalState.tryAssign (state : GoalState) (site : Site) (expr : String) protected def GoalState.tryLet (state: GoalState) (goal: MVarId) (binderName: String) (type: String):
: Elab.TermElabM TacticResult := do Elab.TermElabM TacticResult := do
state.restoreElabM
let expr ← match Parser.runParserCategory
(env := ← getEnv)
(catName := `term)
(input := expr)
(fileName := ← getFileName) with
| .ok syn => pure syn
| .error error => return .parseError error
state.tryTacticM site $ Tactic.evalAssign expr
protected def GoalState.tryLet (state : GoalState) (site : Site) (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)
(catName := `term) (catName := `term)
(input := type) (input := type)
(fileName := ← getFileName) with (fileName := filename) 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.TermElabM TacticResult := do Elab.Tactic.replaceMainGoal [newGoal.mvarId!]
let .some goal := state.actingGoal? site | throwNoGoals pure rhs.mvarId!
let .some fragment := state.fragments[goal]? | return (← MonadBacktrack.saveState, convMVar)
return .invalidAction "Goal does not have a fragment" try
state.restoreElabM let (nextSavedState, convRhs) ← tacticM { elaborator := .anonymous } |>.run' state.savedState.tactic
withCapturingError do -- Other goals are now dormant
let (fragments, state') ← state.step' goal (fragment.exit goal state.fragments) let otherGoals := state.goals.filter $ λ g => g != goal
return { return .success {
state' with root := state.root,
fragments, 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
let (convRhs, convGoal, _) ← match state.convMVar? with
| .some mvar => pure mvar
| .none => return .invalidAction "Not in conv state"
let tacticM : Elab.Tactic.TacticM Elab.Tactic.SavedState:= do
-- Vide `Lean.Elab.Tactic.Conv.convert`
state.savedState.restore
-- Close all existing goals with `refl`
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 := filename) 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

View File

@ -3,7 +3,6 @@ import Pantograph.Goal
import Pantograph.Protocol import Pantograph.Protocol
import Pantograph.Delate import Pantograph.Delate
import Pantograph.Version import Pantograph.Version
import Lean import Lean
namespace Lean namespace Lean
@ -41,6 +40,8 @@ namespace Pantograph
def runMetaM { α } (metaM: MetaM α): CoreM α := def runMetaM { α } (metaM: MetaM α): CoreM α :=
metaM.run' metaM.run'
def runTermElabM { α } (termElabM: Elab.TermElabM α): CoreM α :=
termElabM.run' (ctx := defaultElabContext) |>.run'
def errorI (type desc: String): Protocol.InteractionError := { error := type, desc := desc } def errorI (type desc: String): Protocol.InteractionError := { error := type, desc := desc }
@ -69,123 +70,135 @@ 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_env_add_m]
def envAdd (name: String) (type: String) (value: String) (isTheorem: Bool):
CoreM (Protocol.CR Protocol.EnvAddResult) :=
Environment.addDecl { name, type, value, isTheorem }
@[export pantograph_parse_elab_type_m] @[export pantograph_parse_elab_type_m]
def parseElabType (type: String): Protocol.FallibleT Elab.TermElabM Expr := do def parseElabType (type: String): Elab.TermElabM (Protocol.CR Expr) := do
let env ← MonadEnv.getEnv let env ← MonadEnv.getEnv
let syn ← match parseTerm env type with let syn ← match parseTerm env type with
| .error str => Protocol.throw $ errorI "parsing" str | .error str => return .error $ errorI "parsing" str
| .ok syn => pure syn | .ok syn => pure syn
match ← elabType syn with match ← elabType syn with
| .error str => Protocol.throw $ errorI "elab" str | .error str => return .error $ errorI "elab" str
| .ok expr => return (← instantiateMVars expr) | .ok expr => return .ok (← instantiateMVars expr)
/-- This must be a TermElabM since the parsed expr contains extra information -/ /-- This must be a TermElabM since the parsed expr contains extra information -/
@[export pantograph_parse_elab_expr_m] @[export pantograph_parse_elab_expr_m]
def parseElabExpr (expr: String) (expectedType?: Option String := .none): Protocol.FallibleT Elab.TermElabM Expr := do def parseElabExpr (expr: String) (expectedType?: Option String := .none): Elab.TermElabM (Protocol.CR Expr) := do
let env ← MonadEnv.getEnv let env ← MonadEnv.getEnv
let expectedType? ← expectedType?.mapM parseElabType let expectedType? ← match ← expectedType?.mapM parseElabType with
| .none => pure $ .none
| .some (.ok t) => pure $ .some t
| .some (.error e) => return .error e
let syn ← match parseTerm env expr with let syn ← match parseTerm env expr with
| .error str => Protocol.throw $ errorI "parsing" str | .error str => return .error $ errorI "parsing" str
| .ok syn => pure syn | .ok syn => pure syn
match ← elabTerm syn expectedType? with match ← elabTerm syn expectedType? with
| .error str => Protocol.throw $ errorI "elab" str | .error str => return .error $ errorI "elab" str
| .ok expr => return (← instantiateMVars expr) | .ok expr => return .ok (← instantiateMVars expr)
@[export pantograph_expr_echo_m] @[export pantograph_expr_echo_m]
def exprEcho (expr: String) (expectedType?: Option String := .none) (options: @&Protocol.Options := {}): def exprEcho (expr: String) (expectedType?: Option String := .none) (levels: Array String := #[]) (options: @&Protocol.Options := {}):
Protocol.FallibleT Elab.TermElabM Protocol.ExprEchoResult := do CoreM (Protocol.CR Protocol.ExprEchoResult) :=
let expr ← parseElabExpr expr expectedType? runTermElabM $ Elab.Term.withLevelNames (levels.toList.map (·.toName)) do
let expr ← match ← parseElabExpr expr expectedType? with
| .error e => return .error e
| .ok expr => pure expr
try try
let type ← unfoldAuxLemmas (← Meta.inferType expr) let type ← unfoldAuxLemmas (← Meta.inferType expr)
return { return .ok {
type := (← serializeExpression options type), type := (← serializeExpression options type),
expr := (← serializeExpression options expr), expr := (← serializeExpression options expr)
} }
catch exception => catch exception =>
Protocol.throw $ errorI "typing" (← exception.toMessageData.toString) return .error $ errorI "typing" (← exception.toMessageData.toString)
@[export pantograph_goal_start_expr_m] @[export pantograph_goal_start_expr_m]
def goalStartExpr (expr: String) : Protocol.FallibleT Elab.TermElabM GoalState := do def goalStartExpr (expr: String) (levels: Array String): CoreM (Protocol.CR GoalState) :=
let t ← parseElabType expr runTermElabM $ Elab.Term.withLevelNames (levels.toList.map (·.toName)) do
GoalState.create t let expr ← match ← parseElabType expr with
| .error e => return .error e
| .ok expr => pure $ expr
return .ok $ ← GoalState.create expr
@[export pantograph_goal_resume]
def goalResume (target: GoalState) (goals: Array String): Except String GoalState :=
target.resume (goals.map (λ n => { name := n.toName }) |>.toList)
@[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) (options: @&Protocol.Options): CoreM Protocol.GoalPrintResult :=
: CoreM Protocol.GoalPrintResult := runMetaM do runMetaM do
state.restoreMetaM state.restoreMetaM
let rootExpr? := state.rootExpr?
let root? ← if rootExpr then
rootExpr?.mapM λ expr => state.withRootContext do
serializeExpression options (← instantiateAll expr)
else
pure .none
let parentExprs? ← if parentExprs then
.some <$> state.parentMVars.mapM λ parent => parent.withContext do
let val? := state.getMVarEAssignment parent
val?.mapM λ val => do
serializeExpression options (← instantiateAll val)
else
pure .none
let goals ← if goals then
goalSerialize state options
else
pure #[]
let extraMVars ← extraMVars.mapM λ mvarId => do
let mvarId: MVarId := { name := mvarId.toName }
let .some _ ← mvarId.findDecl? | return {}
state.withContext mvarId do
let .some expr ← getExprMVarAssignment? mvarId | return {}
serializeExpression options (← instantiateAll expr)
let env ← getEnv
return { return {
root?, root? := ← state.rootExpr?.mapM (λ expr =>
parentExprs?, state.withRootContext do
goals, serializeExpression options (← instantiateAll expr)),
extraMVars, parent? := ← state.parentExpr?.mapM (λ expr =>
rootHasSorry := rootExpr?.map (·.hasSorry) |>.getD false, state.withParentContext do
rootHasUnsafe := rootExpr?.map (env.hasUnsafe ·) |>.getD false, serializeExpression options (← instantiateAll expr)),
rootHasMVar := rootExpr?.map (·.hasExprMVar) |>.getD false,
} }
@[export pantograph_goal_tactic_m]
def goalTactic (state: GoalState) (goal: MVarId) (tactic: String): CoreM TacticResult :=
runTermElabM <| state.tryTactic goal tactic
@[export pantograph_goal_assign_m]
def goalAssign (state: GoalState) (goal: MVarId) (expr: String): CoreM TacticResult :=
runTermElabM <| state.tryAssign goal expr
@[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): CoreM 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
runTermElabM do
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): CoreM 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
runTermElabM do
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_motivated_apply_m]
protected def GoalState.tryDraft (state: GoalState) (site : Site) (expr: String): Elab.TermElabM TacticResult := do protected def GoalState.tryMotivatedApply (state: GoalState) (goal: MVarId) (recursor: String):
let expr ← match (← parseTermM expr) with Elab.TermElabM TacticResult := do
state.restoreElabM
let recursor ← match (← parseTermM recursor) with
| .ok syn => pure syn | .ok syn => pure syn
| .error error => return .parseError error | .error error => return .parseError error
state.tryTacticM goal (tacticM := Tactic.evalMotivatedApply recursor)
@[export pantograph_goal_try_no_confuse_m]
protected def GoalState.tryNoConfuse (state: GoalState) (goal: MVarId) (eq: String):
Elab.TermElabM TacticResult := do
state.restoreElabM state.restoreElabM
state.tryTacticM site $ Tactic.evalDraft expr let eq ← match (← parseTermM eq) with
| .ok syn => pure syn
-- Cancel the token after a timeout. | .error error => return .parseError error
@[export pantograph_run_cancel_token_with_timeout_m] state.tryTacticM goal (tacticM := Tactic.evalNoConfuse eq)
def runCancelTokenWithTimeout (cancelToken : IO.CancelToken) (timeout : UInt32) : IO Unit := do @[export pantograph_goal_let_m]
let _ ← EIO.asTask do def goalLet (state: GoalState) (goal: MVarId) (binderName: String) (type: String): CoreM TacticResult :=
IO.sleep timeout runTermElabM <| state.tryLet goal binderName type
cancelToken.set @[export pantograph_goal_conv_m]
return () def goalConv (state: GoalState) (goal: MVarId): CoreM TacticResult :=
runTermElabM <| state.conv goal
@[export pantograph_goal_conv_exit_m]
def goalConvExit (state: GoalState): CoreM TacticResult :=
runTermElabM <| state.convExit
@[export pantograph_goal_calc_m]
def goalCalc (state: GoalState) (goal: MVarId) (pred: String): CoreM TacticResult :=
runTermElabM <| state.tryCalc goal pred
end Pantograph end Pantograph

View File

@ -5,7 +5,6 @@ Note that no command other than `InteractionError` may have `error` as one of
its field names to avoid confusion with error messages generated by the REPL. 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
namespace Pantograph.Protocol namespace Pantograph.Protocol
@ -30,8 +29,6 @@ structure Options where
printImplementationDetailHyps: Bool := false printImplementationDetailHyps: Bool := false
-- If this is set to `true`, goals will never go dormant, so you don't have to manage resumption -- If this is set to `true`, goals will never go dormant, so you don't have to manage resumption
automaticMode: Bool := true automaticMode: Bool := true
-- Timeout for tactics and operations that could potentially execute a tactic
timeout: Nat := 0
deriving Lean.ToJson deriving Lean.ToJson
abbrev OptionsT := ReaderT Options abbrev OptionsT := ReaderT Options
@ -60,17 +57,12 @@ 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 of the metavariable -/
name: String := "" name: String := ""
/-- User-facing name -/ /-- Name of the metavariable -/
userName?: Option String := .none userName?: Option String := .none
fragment : Fragment := .tactic /-- Is the goal in conversion mode -/
isConversion: Bool := false
/-- target expression type -/ /-- target expression type -/
target: Expression target: Expression
/-- Variables -/ /-- Variables -/
@ -92,7 +84,6 @@ structure InteractionError where
deriving Lean.ToJson deriving Lean.ToJson
def errorIndex (desc: String): InteractionError := { error := "index", desc } def errorIndex (desc: String): InteractionError := { error := "index", desc }
def errorOperation (desc: String): InteractionError := { error := "operation", desc }
def errorExpr (desc: String): InteractionError := { error := "expr", desc } def errorExpr (desc: String): InteractionError := { error := "expr", desc }
@ -111,33 +102,15 @@ structure StatResult where
-- Return the type of an expression -- Return the type of an expression
structure ExprEcho where structure ExprEcho where
expr: String expr: String
type?: Option String := .none type?: Option String
-- universe levels -- universe levels
levels?: Option (Array String) := .none levels: Option (Array String) := .none
deriving Lean.FromJson deriving Lean.FromJson
structure ExprEchoResult where structure ExprEchoResult where
expr: Expression expr: Expression
type: Expression type: Expression
deriving Lean.ToJson deriving Lean.ToJson
-- Describe the current state of the environment
structure EnvDescribe where
deriving Lean.FromJson
structure EnvDescribeResult where
imports : Array String
modules : Array String
deriving Lean.ToJson
-- Describe a module
structure EnvModuleRead where
module : String
deriving Lean.FromJson
structure EnvModuleReadResult where
imports: Array String
constNames: Array String
extraConstNames: Array String
deriving Lean.ToJson
-- Print all symbols in environment -- Print all symbols in environment
structure EnvCatalog where structure EnvCatalog where
deriving Lean.FromJson deriving Lean.FromJson
@ -148,13 +121,11 @@ structure EnvCatalogResult where
-- Print the type of a symbol -- Print the type of a symbol
structure EnvInspect where structure EnvInspect where
name: String name: String
-- Show the value expressions; By default definitions values are shown and -- If true/false, show/hide the value expressions; By default definitions
-- theorem values are hidden. -- values are shown and theorem values are hidden.
value?: Option Bool := .some false value?: Option Bool := .some false
-- Show the type and value dependencies -- If true, show the type and value dependencies
dependency?: Option Bool := .some false dependency?: Option Bool := .some false
-- Show source location
source?: Option Bool := .some false
deriving Lean.FromJson deriving Lean.FromJson
-- See `InductiveVal` -- See `InductiveVal`
structure InductInfo where structure InductInfo where
@ -201,19 +172,13 @@ structure EnvInspectResult where
inductInfo?: Option InductInfo := .none inductInfo?: Option InductInfo := .none
constructorInfo?: Option ConstructorInfo := .none constructorInfo?: Option ConstructorInfo := .none
recursorInfo?: Option RecursorInfo := .none recursorInfo?: Option RecursorInfo := .none
-- Location in source
sourceUri?: Option String := .none
sourceStart?: Option Lean.Position := .none
sourceEnd?: Option Lean.Position := .none
deriving Lean.ToJson deriving Lean.ToJson
structure EnvAdd where structure EnvAdd where
name: String name: String
levels?: Option (Array String) := .none type: String
type?: Option String := .none
value: String value: String
isTheorem: Bool := false isTheorem: Bool
deriving Lean.FromJson deriving Lean.FromJson
structure EnvAddResult where structure EnvAddResult where
deriving Lean.ToJson deriving Lean.ToJson
@ -226,15 +191,14 @@ structure EnvSaveLoadResult where
/-- Set options; See `Options` struct above for meanings -/ /-- Set options; See `Options` struct above for meanings -/
structure OptionsSet where structure OptionsSet where
printJsonPretty?: Option Bool := .none printJsonPretty?: Option Bool
printExprPretty?: Option Bool := .none printExprPretty?: Option Bool
printExprAST?: Option Bool := .none printExprAST?: Option Bool
printDependentMVars?: Option Bool := .none printDependentMVars?: Option Bool
noRepeat?: Option Bool := .none noRepeat?: Option Bool
printAuxDecls?: Option Bool := .none printAuxDecls?: Option Bool
printImplementationDetailHyps?: Option Bool := .none printImplementationDetailHyps?: Option Bool
automaticMode?: Option Bool := .none automaticMode?: Option Bool
timeout?: Option Nat := .none
deriving Lean.FromJson deriving Lean.FromJson
structure OptionsSetResult where structure OptionsSetResult where
deriving Lean.ToJson deriving Lean.ToJson
@ -245,8 +209,8 @@ structure GoalStart where
-- Only one of the fields below may be populated. -- Only one of the fields below may be populated.
expr: Option String -- Directly parse in an expression expr: Option String -- Directly parse in an expression
-- universe levels -- universe levels
levels?: Option (Array String) := .none levels: Option (Array String) := .none
copyFrom: Option String := .none -- Copy the type from a theorem in the environment copyFrom: Option String -- Copy the type from a theorem in the environment
deriving Lean.FromJson deriving Lean.FromJson
structure GoalStartResult where structure GoalStartResult where
stateId: Nat := 0 stateId: Nat := 0
@ -254,18 +218,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
draft?: 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
-- 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
binderName?: Option String := .none binderName?: Option String := .none
@ -274,17 +237,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 String) := .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
@ -311,27 +271,12 @@ structure GoalDeleteResult where
structure GoalPrint where structure GoalPrint where
stateId: Nat stateId: Nat
-- Print root?
rootExpr?: Option Bool := .some False
-- Print the parent expressions
parentExprs?: Option Bool := .some False
-- Print goals?
goals?: Option Bool := .some False
-- Print values of extra mvars?
extraMVars?: Option (Array String) := .none
deriving Lean.FromJson deriving Lean.FromJson
structure GoalPrintResult where structure GoalPrintResult where
-- The root expression -- The root expression
root?: Option Expression := .none root?: Option Expression := .none
-- The filling expression of the parent goal -- The filling expression of the parent goal
parentExprs?: Option (List (Option Expression)) := .none parent?: Option Expression
goals: Array Goal := #[]
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,18 +308,10 @@ structure FrontendProcess where
-- One of these two must be supplied: Either supply the file name or the content. -- One of these two must be supplied: Either supply the file name or the content.
fileName?: Option String := .none fileName?: Option String := .none
file?: Option String := .none file?: Option String := .none
-- Whether to read the header -- If set to true, collect tactic invocations
readHeader : Bool := false invocations: Bool := false
-- Alter the REPL environment after the compilation units. -- If set to true, collect `sorry`s
inheritEnv : Bool := false
-- collect tactic invocations and output to a given file
invocations?: Option String := .none
-- collect `sorry`s
sorrys: Bool := false sorrys: Bool := false
-- collect type errors
typeErrorsAsGoals: Bool := false
-- list new constants from each compilation step
newConstants: Bool := false
deriving Lean.FromJson deriving Lean.FromJson
structure InvokedTactic where structure InvokedTactic where
goalBefore: String goalBefore: String
@ -388,30 +325,16 @@ structure InvokedTactic where
structure CompilationUnit where structure CompilationUnit where
-- String boundaries of compilation units -- String boundaries of compilation units
boundary: (Nat × Nat) boundary: (Nat × Nat)
messages: Array String := #[] -- Tactic invocations
-- Number of tactic invocations invocations?: Option (List InvokedTactic) := .none
nInvocations?: Option Nat := .none
goalStateId?: Option Nat := .none goalStateId?: Option Nat := .none
goals?: Option (Array Goal) := .none goals: Array Goal := #[]
-- Code segments which generated the goals messages: Array String := #[]
goalSrcBoundaries?: Option (Array (Nat × Nat)) := .none
-- New constants defined in compilation unit
newConstants?: Option (Array String) := .none
deriving Lean.ToJson deriving Lean.ToJson
structure FrontendProcessResult where structure FrontendProcessResult where
units: List CompilationUnit units: List CompilationUnit
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 CR α := Except InteractionError α
abbrev throw {m : Type v → Type w} [MonadExceptOf InteractionError m] {α : Type v} (e : InteractionError) : m α :=
throwThe InteractionError e
end Pantograph.Protocol end Pantograph.Protocol

View File

@ -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,9 @@ 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
(distilled : DistilledEnvironment)
(background? : Option Environment := .none)
: IO Environment := do
let (imports, constArray) := distilled
let env ← match background? with
| .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 +66,10 @@ 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 let env ← importModules imports {} 0
return (← resurrectEnvironment distilled background?, region) return (← env.replay (Std.HashMap.ofList map₂.toList), region)
open Lean.Core in open Lean.Core in
@ -112,15 +83,12 @@ 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 := {
meta := { meta := {
core := { core,
env, nextMacroScope, ngen, ..
},
meta, meta,
} }
«elab», «elab»,
@ -128,49 +96,47 @@ def goalStatePickle (goalState : GoalState) (path : System.FilePath) (background
tactic tactic
} }
root, root,
parentMVars, parentMVar?,
fragments, convMVar?,
calcPrevRhs?,
} := goalState } := goalState
pickle path ( --let env := core.env
distillEnvironment env background?, Pantograph.pickle path (
({ core with } : CompactCoreState),
({ nextMacroScope, ngen } : CompactCoreState),
meta, meta,
«elab», «elab»,
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,
compactCore, compactCore,
meta, meta,
«elab», «elab»,
tactic, tactic,
root, root,
parentMVars, parentMVar?,
fragments, convMVar?,
calcPrevRhs?,
), region) ← Pantograph.unpickle ( ), region) ← Pantograph.unpickle (
DistilledEnvironment ×
CompactCoreState × CompactCoreState ×
Meta.State × Meta.State ×
Elab.Term.State × Elab.Term.State ×
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 goalState := { let goalState := {
savedState := { savedState := {
term := { term := {
@ -187,8 +153,9 @@ def goalStateUnpickle (path : System.FilePath) (background? : Option Environment
tactic, tactic,
}, },
root, root,
parentMVars, parentMVar?,
fragments, convMVar?,
calcPrevRhs?,
} }
return (goalState, region) return (goalState, region)

View File

@ -1,3 +1,5 @@
import Pantograph.Tactic.Assign import Pantograph.Tactic.Assign
import Pantograph.Tactic.Fragment import Pantograph.Tactic.Congruence
import Pantograph.Tactic.MotivatedApply
import Pantograph.Tactic.NoConfuse
import Pantograph.Tactic.Prograde import Pantograph.Tactic.Prograde

View File

@ -27,39 +27,5 @@ def evalAssign : Elab.Tactic.Tactic := fun stx => Elab.Tactic.withMainContext do
goal.assign expr goal.assign expr
Elab.Tactic.replaceMainGoal nextGoals Elab.Tactic.replaceMainGoal nextGoals
def sorryToHole (src : Expr) : StateRefT (List MVarId) MetaM Expr := do
Meta.transform src λ expr =>
if expr.isSorry then do
let type ← instantiateMVars (expr.getArg! 0 |>.bindingBody!)
if type.hasSorry then
throwError s!"Coupling is not allowed in draft tactic: {← Meta.ppExpr type}"
let mvar ← Meta.mkFreshExprSyntheticOpaqueMVar type
modify (mvar.mvarId! :: .)
pure $ .done mvar
else
pure .continue
-- Given a complete (no holes) expression, extract the sorry's from it and convert them into goals.
def draft (goal : MVarId) (expr : Expr) : MetaM (List MVarId) := do
goal.checkNotAssigned `Pantograph.Tactic.draft
let exprType ← Meta.inferType expr
let goalType ← goal.getType
unless ← Meta.isDefEq goalType exprType do
throwError s!"{← Meta.ppExpr expr} : {← Meta.ppExpr exprType} ≠ {← Meta.ppExpr goalType}"
let (expr', holes) ← sorryToHole expr |>.run []
goal.assign expr'
return holes.reverse
def evalDraft : Elab.Tactic.Tactic := fun stx ↦ Elab.Tactic.withMainContext do
let target ← Elab.Tactic.getMainTarget
let goal ← Elab.Tactic.getMainGoal
let (expr, holeGoals) ← Elab.Tactic.elabTermWithHoles stx
(expectedType? := .some target)
(tagSuffix := .anonymous)
(allowNaturalHoles := true)
let draftGoals ← draft goal expr
Elab.Tactic.replaceMainGoal $ holeGoals ++ draftGoals
end Pantograph.Tactic end Pantograph.Tactic

View File

@ -0,0 +1,98 @@
import Lean
open Lean
namespace Pantograph.Tactic
def congruenceArg (mvarId: MVarId): MetaM (List MVarId) := mvarId.withContext do
mvarId.checkNotAssigned `Pantograph.Tactic.congruenceArg
let target ← mvarId.getType
let .some (β, _, _) := (← instantiateMVars target).eq? | throwError "Goal is not an Eq"
let userName := (← mvarId.getDecl).userName
let u ← Meta.mkFreshLevelMVar
let α ← Meta.mkFreshExprSyntheticOpaqueMVar (mkSort u)
(tag := userName ++ `α)
let f ← Meta.mkFreshExprSyntheticOpaqueMVar (.forallE .anonymous α β .default)
(tag := userName ++ `f)
let a₁ ← Meta.mkFreshExprSyntheticOpaqueMVar α
(tag := userName ++ `a₁)
let a₂ ← Meta.mkFreshExprSyntheticOpaqueMVar α
(tag := userName ++ `a₂)
let h ← Meta.mkFreshExprSyntheticOpaqueMVar (← Meta.mkEq a₁ a₂)
(tag := userName ++ `h)
let conduitType ← Meta.mkEq (← Meta.mkEq (.app f a₁) (.app f a₂)) target
let conduit ← Meta.mkFreshExprSyntheticOpaqueMVar conduitType
(tag := userName ++ `conduit)
mvarId.assign $ ← Meta.mkEqMP conduit (← Meta.mkCongrArg f h)
let result := [α, a₁, a₂, f, h, conduit]
return result.map (·.mvarId!)
def evalCongruenceArg: Elab.Tactic.TacticM Unit := do
let goal ← Elab.Tactic.getMainGoal
let nextGoals ← congruenceArg goal
Elab.Tactic.replaceMainGoal nextGoals
def congruenceFun (mvarId: MVarId): MetaM (List MVarId) := mvarId.withContext do
mvarId.checkNotAssigned `Pantograph.Tactic.congruenceFun
let target ← mvarId.getType
let .some (β, _, _) := (← instantiateMVars target).eq? | throwError "Goal is not an Eq"
let userName := (← mvarId.getDecl).userName
let u ← Meta.mkFreshLevelMVar
let α ← Meta.mkFreshExprSyntheticOpaqueMVar (mkSort u)
(tag := userName ++ `α)
let fType := .forallE .anonymous α β .default
let f₁ ← Meta.mkFreshExprSyntheticOpaqueMVar fType
(tag := userName ++ `f₁)
let f₂ ← Meta.mkFreshExprSyntheticOpaqueMVar fType
(tag := userName ++ `f₂)
let a ← Meta.mkFreshExprSyntheticOpaqueMVar α
(tag := userName ++ `a)
let h ← Meta.mkFreshExprSyntheticOpaqueMVar (← Meta.mkEq f₁ f₂)
(tag := userName ++ `h)
let conduitType ← Meta.mkEq (← Meta.mkEq (.app f₁ a) (.app f₂ a)) target
let conduit ← Meta.mkFreshExprSyntheticOpaqueMVar conduitType
(tag := userName ++ `conduit)
mvarId.assign $ ← Meta.mkEqMP conduit (← Meta.mkCongrFun h a)
let result := [α, f₁, f₂, h, a, conduit]
return result.map (·.mvarId!)
def evalCongruenceFun: Elab.Tactic.TacticM Unit := do
let goal ← Elab.Tactic.getMainGoal
let nextGoals ← congruenceFun goal
Elab.Tactic.replaceMainGoal nextGoals
def congruence (mvarId: MVarId): MetaM (List MVarId) := mvarId.withContext do
mvarId.checkNotAssigned `Pantograph.Tactic.congruence
let target ← mvarId.getType
let .some (β, _, _) := (← instantiateMVars target).eq? | throwError "Goal is not an Eq"
let userName := (← mvarId.getDecl).userName
let u ← Meta.mkFreshLevelMVar
let α ← Meta.mkFreshExprSyntheticOpaqueMVar (mkSort u)
(tag := userName ++ `α)
let fType := .forallE .anonymous α β .default
let f₁ ← Meta.mkFreshExprSyntheticOpaqueMVar fType
(tag := userName ++ `f₁)
let f₂ ← Meta.mkFreshExprSyntheticOpaqueMVar fType
(tag := userName ++ `f₂)
let a₁ ← Meta.mkFreshExprSyntheticOpaqueMVar α
(tag := userName ++ `a₁)
let a₂ ← Meta.mkFreshExprSyntheticOpaqueMVar α
(tag := userName ++ `a₂)
let h₁ ← Meta.mkFreshExprSyntheticOpaqueMVar (← Meta.mkEq f₁ f₂)
(tag := userName ++ `h₁)
let h₂ ← Meta.mkFreshExprSyntheticOpaqueMVar (← Meta.mkEq a₁ a₂)
(tag := userName ++ `h₂)
let conduitType ← Meta.mkEq (← Meta.mkEq (.app f₁ a₁) (.app f₂ a₂)) target
let conduit ← Meta.mkFreshExprSyntheticOpaqueMVar conduitType
(tag := userName ++ `conduit)
mvarId.assign $ ← Meta.mkEqMP conduit (← Meta.mkCongr h₁ h₂)
let result := [α, f₁, f₂, a₁, a₂, h₁, h₂, conduit]
return result.map (·.mvarId!)
def evalCongruence: Elab.Tactic.TacticM Unit := do
let goal ← Elab.Tactic.getMainGoal
let nextGoals ← congruence goal
Elab.Tactic.replaceMainGoal nextGoals
end Pantograph.Tactic

View File

@ -1,188 +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 goals := [ mvarBranch ] ++ remainder?.toList
Elab.Tactic.setGoals goals
match remainder? with
| .some goal => return map.erase goal |>.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

View File

@ -0,0 +1,106 @@
import Lean
open Lean
namespace Pantograph.Tactic
def getForallArgsBody: Expr → List Expr × Expr
| .forallE _ d b _ =>
let (innerArgs, innerBody) := getForallArgsBody b
(d :: innerArgs, innerBody)
| e => ([], e)
def replaceForallBody: Expr → Expr → Expr
| .forallE param domain body binderInfo, target =>
let body := replaceForallBody body target
.forallE param domain body binderInfo
| _, target => target
structure RecursorWithMotive where
args: List Expr
body: Expr
-- .bvar index for the motive and major from the body
iMotive: Nat
namespace RecursorWithMotive
protected def nArgs (info: RecursorWithMotive): Nat := info.args.length
protected def getMotiveType (info: RecursorWithMotive): Expr :=
let level := info.nArgs - info.iMotive - 1
let a := info.args.get! level
a
protected def surrogateMotiveType (info: RecursorWithMotive) (mvars: Array Expr) (resultant: Expr): MetaM Expr := do
let motiveType := Expr.instantiateRev info.getMotiveType mvars
let resultantType ← Meta.inferType resultant
return replaceForallBody motiveType resultantType
protected def conduitType (info: RecursorWithMotive) (mvars: Array Expr) (resultant: Expr): MetaM Expr := do
let motiveCall := Expr.instantiateRev info.body mvars
Meta.mkEq motiveCall resultant
end RecursorWithMotive
def getRecursorInformation (recursorType: Expr): Option RecursorWithMotive := do
let (args, body) := getForallArgsBody recursorType
if ¬ body.isApp then
.none
let iMotive ← match body.getAppFn with
| .bvar iMotive => pure iMotive
| _ => .none
return {
args,
body,
iMotive,
}
def collectMotiveArguments (forallBody: Expr): SSet Nat :=
match forallBody with
| .app (.bvar i) _ => SSet.empty.insert i
| _ => SSet.empty
/-- Applies a symbol of the type `∀ (motive: α → Sort u) (a: α)..., (motive α)` -/
def motivatedApply (mvarId: MVarId) (recursor: Expr) : MetaM (Array Meta.InductionSubgoal) := mvarId.withContext do
mvarId.checkNotAssigned `Pantograph.Tactic.motivatedApply
let recursorType ← Meta.inferType recursor
let resultant ← mvarId.getType
let tag ← mvarId.getTag
let info ← match getRecursorInformation recursorType with
| .some info => pure info
| .none => throwError "Recursor return type does not correspond with the invocation of a motive: {← Meta.ppExpr recursorType}"
let rec go (i: Nat) (prev: Array Expr): MetaM (Array Expr) := do
if i ≥ info.nArgs then
return prev
else
let argType := info.args.get! i
-- If `argType` has motive references, its goal needs to be placed in it
let argType := argType.instantiateRev prev
let bvarIndex := info.nArgs - i - 1
let argGoal ← if bvarIndex = info.iMotive then
let surrogateMotiveType ← info.surrogateMotiveType prev resultant
Meta.mkFreshExprSyntheticOpaqueMVar surrogateMotiveType (tag := tag ++ `motive)
else
Meta.mkFreshExprSyntheticOpaqueMVar argType (tag := .anonymous)
let prev := prev ++ [argGoal]
go (i + 1) prev
termination_by info.nArgs - i
let mut newMVars ← go 0 #[]
-- Create the conduit type which proves the result of the motive is equal to the goal
let conduitType ← info.conduitType newMVars resultant
let goalConduit ← Meta.mkFreshExprSyntheticOpaqueMVar conduitType (tag := `conduit)
mvarId.assign $ ← Meta.mkEqMP goalConduit (mkAppN recursor newMVars)
newMVars := newMVars ++ [goalConduit]
return newMVars.map (λ mvar => { mvarId := mvar.mvarId!})
def evalMotivatedApply : Elab.Tactic.Tactic := fun stx => Elab.Tactic.withMainContext do
let recursor ← Elab.Term.elabTerm (stx := stx) .none
let nextGoals ← motivatedApply (← Elab.Tactic.getMainGoal) recursor
Elab.Tactic.replaceMainGoal $ nextGoals.toList.map (·.mvarId)
end Pantograph.Tactic

View File

@ -0,0 +1,22 @@
import Lean
open Lean
namespace Pantograph.Tactic
def noConfuse (mvarId: MVarId) (h: Expr): MetaM Unit := mvarId.withContext do
mvarId.checkNotAssigned `Pantograph.Tactic.noConfuse
let target ← mvarId.getType
let noConfusion ← Meta.mkNoConfusion (target := target) (h := h)
unless ← Meta.isDefEq (← Meta.inferType noConfusion) target do
throwError "invalid noConfuse call: The resultant type {← Meta.ppExpr $ ← Meta.inferType noConfusion} cannot be unified with {← Meta.ppExpr target}"
mvarId.assign noConfusion
def evalNoConfuse: Elab.Tactic.Tactic := λ stx => do
let goal ← Elab.Tactic.getMainGoal
let h ← goal.withContext $ Elab.Term.elabTerm (stx := stx) .none
noConfuse goal h
Elab.Tactic.replaceMainGoal []
end Pantograph.Tactic

View File

@ -40,7 +40,7 @@ def «have» (mvarId: MVarId) (binderName: Name) (type: Expr): MetaM BranchResul
let fvarId ← mkFreshFVarId let fvarId ← mkFreshFVarId
let lctxUpstream := lctx.mkLocalDecl fvarId binderName type let lctxUpstream := lctx.mkLocalDecl fvarId binderName type
let mvarUpstream ← let mvarUpstream ←
Meta.withLCtx lctxUpstream #[] do withTheReader Meta.Context (fun ctx => { ctx with lctx := lctxUpstream }) do
Meta.withNewLocalInstances #[.fvar fvarId] 0 do Meta.withNewLocalInstances #[.fvar fvarId] 0 do
let mvarUpstream ← mkUpstreamMVar mvarId let mvarUpstream ← mkUpstreamMVar mvarId
--let expr: Expr := .app (.lam binderName type mvarBranch .default) mvarUpstream --let expr: Expr := .app (.lam binderName type mvarBranch .default) mvarUpstream

View File

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

View File

@ -7,8 +7,6 @@ A Machine-to-Machine interaction system for Lean 4.
Pantograph provides interfaces to execute proofs, construct expressions, and Pantograph provides interfaces to execute proofs, construct expressions, and
examine the symbol list of a Lean project for machine learning. examine the symbol list of a Lean project for machine learning.
See [documentations](doc/rationale.md) for design rationale and references.
## Installation ## Installation
For Nix users, run For Nix users, run
@ -17,9 +15,7 @@ nix build .#{sharedLib,executable}
``` ```
to build either the shared library or executable. to build either the shared library or executable.
Install `lake` and `lean` fixed to the version of the `lean-toolchain` file, and Install `elan` and `lake`, and run
run
``` sh ``` sh
lake build lake build
``` ```
@ -28,12 +24,9 @@ This builds the executable in `.lake/build/bin/pantograph-repl`.
## Executable Usage ## Executable Usage
``` sh ``` sh
pantograph-repl MODULES|LEAN_OPTIONS pantograph MODULES|LEAN_OPTIONS
``` ```
The `pantograph-repl` executable must be run with a list of modules to import.
It can also accept lean options of the form `--key=value` e.g. `--pp.raw=true`.
The REPL loop accepts commands as single-line JSON inputs and outputs either an The REPL loop accepts commands as single-line JSON inputs and outputs either an
`Error:` (indicating malformed command) or a JSON return value indicating the `Error:` (indicating malformed command) or a JSON return value indicating the
result of a command execution. The command can be passed in one of two formats result of a command execution. The command can be passed in one of two formats
@ -44,6 +37,8 @@ command { ... }
The list of available commands can be found in `Pantograph/Protocol.lean` and below. An The list of available commands can be found in `Pantograph/Protocol.lean` and below. An
empty command aborts the REPL. empty command aborts the REPL.
The `pantograph` executable must be run with a list of modules to import. It can
also accept lean options of the form `--key=value` e.g. `--pp.raw=true`.
Example: (~5k symbols) Example: (~5k symbols)
``` ```
@ -60,11 +55,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.
@ -80,8 +75,8 @@ the environment might be setup like this:
``` sh ``` sh
LIB="../lib" LIB="../lib"
LIB_MATHLIB="$LIB/mathlib4/.lake" LIB_MATHLIB="$LIB/mathlib4/lake-packages"
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 $@
``` ```

527
Repl.lean
View File

@ -3,35 +3,17 @@ import Pantograph
namespace Pantograph.Repl namespace Pantograph.Repl
open Lean
structure Context where structure Context where
coreContext : Core.Context imports: List String
-- 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 /-- Main state monad for executing commands -/
-- Parser state abbrev MainM := ReaderT Context (StateT State Lean.CoreM)
scope : Elab.Command.Scope := { header := "" }
/-- Main monad for executing commands -/
abbrev MainM := ReaderT Context $ StateRefT State IO
/-- Main with possible exception -/
abbrev EMainM := Protocol.FallibleT $ ReaderT Context $ StateRefT State IO
def getMainState : MainM State := get
instance : MonadEnv MainM where
getEnv := return (← get).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
withTheReader Context ({ · with inheritEnv := true }) z
def newGoalState (goalState: GoalState) : MainM Nat := do def newGoalState (goalState: GoalState) : MainM Nat := do
let state ← get let state ← get
@ -42,244 +24,30 @@ def newGoalState (goalState : GoalState) : MainM Nat := do
} }
return stateId return stateId
def runCoreM { α } (coreM : CoreM α) : EMainM α := do
let scope := (← get).scope
let options := (← get).options
let cancelTk? ← match options.timeout with
| 0 => pure .none
| _ => .some <$> IO.CancelToken.new
let coreCtx : Core.Context := {
(← read).coreContext with
currNamespace := scope.currNamespace,
openDecls := scope.openDecls,
options := scope.opts,
initHeartbeats := ← IO.getNumHeartbeats,
cancelTk?,
}
let coreState : Core.State := {
env := (← get).env
}
-- Remap the coreM to capture every exception
let coreM' : CoreM _ :=
try
Except.ok <$> coreM
catch ex =>
let desc ← ex.toMessageData.toString
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
runCancelTokenWithTimeout token (timeout := .ofBitVec options.timeout)
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.internal id _) => Protocol.throw $ { error := "internal", desc := (← id.getName).toString }
| Except.ok a => pure a
if (← read).inheritEnv && result matches .ok _ then
setEnv state'.env
liftExcept result
def runCoreM' { α } (coreM : Protocol.FallibleT CoreM α) : EMainM α := do -- HACK: For some reason writing `CommandM α := MainM (Except ... α)` disables
liftExcept $ ← runCoreM coreM.run -- certain monadic features in `MainM`
abbrev CR α := Except Protocol.InteractionError α
def liftMetaM { α } (metaM : MetaM α): EMainM α := def runMetaInMainM { α } (metaM: Lean.MetaM α): MainM α :=
runCoreM metaM.run' metaM.run'
def liftTermElabM { α } (termElabM : Elab.TermElabM α) (levelNames : List Name := []) def runTermElabInMainM { α } (termElabM: Lean.Elab.TermElabM α) : MainM α :=
: EMainM α := do termElabM.run' (ctx := defaultElabContext) |>.run'
let scope := (← get).scope
let context := {
errToSorry := false,
isNoncomputableSection := scope.isNoncomputable,
}
let state := {
levelNames := scope.levelNames ++ levelNames,
}
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'
return {
nextStateId? := .some nextStateId,
goals? := .some goals,
messages? := .some messages,
hasSorry,
hasUnsafe,
}
| .ok (.parseError message) =>
return { messages? := .none, parseError? := .some message }
| .ok (.invalidAction message) =>
Protocol.throw $ errorI "invalid" message
| .ok (.failure messages) =>
return { messages? := .some messages }
end Goal
section Frontend
structure CompilationUnit where
-- Environment immediately before the unit
env : Environment
boundary : Nat × Nat
invocations : List Protocol.InvokedTactic
sorrys : List Frontend.InfoWithContext
messages : Array String
newConstants : List Name
def frontend_process (args: Protocol.FrontendProcess): EMainM Protocol.FrontendProcessResult := do
let options := (← getMainState).options
let (fileName, file) ← match args.fileName?, args.file? with
| .some fileName, .none => do
let file ← IO.FS.readFile fileName
pure (fileName, file)
| .none, .some file =>
pure ("<anonymous>", file)
| _, _ => Protocol.throw $ errorI "arguments" "Exactly one of {fileName, file} must be supplied"
let env?: Option Environment ← if args.readHeader then
pure .none
else do
.some <$> getEnv
let (context, state) ← do Frontend.createContextStateFromFile file fileName env? {}
let frontendM: Elab.Frontend.FrontendM (List CompilationUnit) :=
Frontend.mapCompilationSteps λ step => do
let boundary := (step.src.startPos.byteIdx, step.src.stopPos.byteIdx)
let invocations: Option (List Protocol.InvokedTactic) ← if args.invocations?.isSome then
Frontend.collectTacticsFromCompilationStep step
else
pure []
let sorrys ← if args.sorrys then
Frontend.collectSorrys step (options := { collectTypeErrors := args.typeErrorsAsGoals })
else
pure []
let messages ← step.messageStrings
let newConstants ← if args.newConstants then
Frontend.collectNewDefinedConstants step
else
pure []
return {
env := step.before,
boundary,
invocations,
sorrys,
messages,
newConstants
}
let (li, state') ← frontendM.run context |>.run state
if args.inheritEnv then
setEnv state'.commandState.env
if let .some scope := state'.commandState.scopes.head? then
-- modify the 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 newConstants? := if args.newConstants then
.some $ step.newConstants.toArray.map λ name => name.toString
else
.none
let (goalStateId?, goals?, goalSrcBoundaries?) ← if step.sorrys.isEmpty then
pure (.none, .none, .none)
else do
let ({ state, srcBoundaries }, goals) ← liftMetaM do
let result@{state, .. } ← Frontend.sorrysToGoalState step.sorrys
let goals ← goalSerialize state options
pure (result, goals)
let stateId ← newGoalState state
let srcBoundaries := srcBoundaries.toArray.map (λ (b, e) => (b.byteIdx, e.byteIdx))
pure (.some stateId, .some goals, .some srcBoundaries)
let nInvocations? := if args.invocations?.isSome then .some step.invocations.length else .none
return {
boundary := step.boundary,
messages := step.messages,
nInvocations?,
goalStateId?,
goals?,
goalSrcBoundaries?,
newConstants?,
}
return { units }
end Frontend
/-- Main loop command of the REPL -/ /-- Main loop command of the REPL -/
def execute (command: Protocol.Command): MainM Json := do def execute (command: Protocol.Command): MainM Lean.Json := do
let run { α β: Type } [FromJson α] [ToJson β] (comm: α → EMainM β): MainM Json := let run { α β: Type } [Lean.FromJson α] [Lean.ToJson β] (comm: α → MainM (CR β)): MainM Lean.Json :=
try match Lean.fromJson? command.payload with
match fromJson? command.payload with
| .ok args => do | .ok args => do
match (← comm args |>.run) with match (← comm args) with
| .ok result => return toJson result | .ok result => return Lean.toJson result
| .error ierror => return toJson ierror | .error ierror => return Lean.toJson ierror
| .error error => return toJson $ errorCommand s!"Unable to parse json: {error}" | .error error => return Lean.toJson $ errorCommand s!"Unable to parse json: {error}"
catch ex : IO.Error => try
let error : Protocol.InteractionError := { error := "io", desc := ex.toString }
return toJson error
match command.cmd with match command.cmd with
| "reset" => run reset | "reset" => run reset
| "stat" => run stat | "stat" => run stat
| "expr.echo" => run expr_echo | "expr.echo" => run expr_echo
| "env.describe" => run env_describe
| "env.module_read" => run env_module_read
| "env.catalog" => run env_catalog | "env.catalog" => run env_catalog
| "env.inspect" => run env_inspect | "env.inspect" => run env_inspect
| "env.add" => run env_add | "env.add" => run env_add
@ -298,48 +66,45 @@ def execute (command: Protocol.Command): MainM Json := do
| cmd => | cmd =>
let error: Protocol.InteractionError := let error: Protocol.InteractionError :=
errorCommand s!"Unknown command {cmd}" errorCommand s!"Unknown command {cmd}"
return toJson error return Lean.toJson error
catch ex => do
let error ← ex.toMessageData.toString
return Lean.toJson $ errorIO 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): MainM (CR Protocol.StatResult) := do
let state ← getMainState let state ← get
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 .ok { nGoals }
stat (_: Protocol.Stat): EMainM Protocol.StatResult := do stat (_: Protocol.Stat): MainM (CR Protocol.StatResult) := do
let state ← getMainState let state ← get
let nGoals := state.goalStates.size let nGoals := state.goalStates.size
return { nGoals } return .ok { nGoals }
env_describe (args: Protocol.EnvDescribe): EMainM Protocol.EnvDescribeResult := do env_catalog (args: Protocol.EnvCatalog): MainM (CR Protocol.EnvCatalogResult) := do
let result ← runCoreM $ Environment.describe args let result ← Environment.catalog args
return result return .ok result
env_module_read (args: Protocol.EnvModuleRead): EMainM Protocol.EnvModuleReadResult := do env_inspect (args: Protocol.EnvInspect): MainM (CR Protocol.EnvInspectResult) := do
runCoreM $ Environment.moduleRead args let state ← get
env_catalog (args: Protocol.EnvCatalog): EMainM Protocol.EnvCatalogResult := do Environment.inspect args state.options
let result ← runCoreM $ Environment.catalog args env_add (args: Protocol.EnvAdd): MainM (CR Protocol.EnvAddResult) := do
return result Environment.addDecl args
env_inspect (args: Protocol.EnvInspect): EMainM Protocol.EnvInspectResult := do env_save (args: Protocol.EnvSaveLoad): MainM (CR Protocol.EnvSaveLoadResult) := do
let state ← getMainState let env ← Lean.MonadEnv.getEnv
runCoreM' $ Environment.inspect args state.options
env_add (args: Protocol.EnvAdd): EMainM Protocol.EnvAddResult := withInheritEnv do
runCoreM' $ Environment.addDecl args.name (args.levels?.getD #[]) args.type? args.value args.isTheorem
env_save (args: Protocol.EnvSaveLoad): EMainM Protocol.EnvSaveLoadResult := do
let env ← MonadEnv.getEnv
environmentPickle env args.path environmentPickle env args.path
return {} return .ok {}
env_load (args: Protocol.EnvSaveLoad): EMainM Protocol.EnvSaveLoadResult := do env_load (args: Protocol.EnvSaveLoad): MainM (CR Protocol.EnvSaveLoadResult) := do
let (env, _) ← environmentUnpickle args.path let (env, _) ← environmentUnpickle args.path
setEnv env Lean.setEnv env
return {} return .ok {}
expr_echo (args: Protocol.ExprEcho): EMainM Protocol.ExprEchoResult := do expr_echo (args: Protocol.ExprEcho): MainM (CR Protocol.ExprEchoResult) := do
let state ← getMainState let state ← get
let levelNames := (args.levels?.getD #[]).toList.map (·.toName) exprEcho args.expr (expectedType? := args.type?) (levels := args.levels.getD #[]) (options := state.options)
liftExcept $ ← liftTermElabM (levelNames := levelNames) do options_set (args: Protocol.OptionsSet): MainM (CR Protocol.OptionsSetResult) := do
(exprEcho args.expr (expectedType? := args.type?) (options := state.options)).run let state ← get
options_set (args: Protocol.OptionsSet): EMainM Protocol.OptionsSetResult := do
let state ← getMainState
let options := state.options let options := state.options
set { state with set { state with
options := { options := {
@ -352,76 +117,170 @@ def execute (command: Protocol.Command): MainM Json := do
printAuxDecls := args.printAuxDecls?.getD options.printAuxDecls, printAuxDecls := args.printAuxDecls?.getD options.printAuxDecls,
printImplementationDetailHyps := args.printImplementationDetailHyps?.getD options.printImplementationDetailHyps printImplementationDetailHyps := args.printImplementationDetailHyps?.getD options.printImplementationDetailHyps
automaticMode := args.automaticMode?.getD options.automaticMode, automaticMode := args.automaticMode?.getD options.automaticMode,
timeout := args.timeout?.getD options.timeout,
} }
} }
return { } return .ok { }
options_print (_: Protocol.OptionsPrint): EMainM Protocol.Options := do options_print (_: Protocol.OptionsPrint): MainM (CR Protocol.Options) := do
return (← getMainState).options return .ok (← get).options
goal_start (args: Protocol.GoalStart): EMainM Protocol.GoalStartResult := do goal_start (args: Protocol.GoalStart): MainM (CR Protocol.GoalStartResult) := do
let levelNames := (args.levels?.getD #[]).toList.map (·.toName) let env ← Lean.MonadEnv.getEnv
let expr?: Except _ GoalState ← liftTermElabM (levelNames := levelNames) do let expr?: Except _ GoalState ← runTermElabInMainM (match args.expr, args.copyFrom with
match args.expr, args.copyFrom with | .some expr, .none => goalStartExpr expr (args.levels.getD #[])
| .some expr, .none => goalStartExpr expr |>.run | .none, .some copyFrom =>
| .none, .some copyFrom => do (match env.find? <| copyFrom.toName with
(match (← getEnv).find? <| copyFrom.toName with | .none => return .error <| errorIndex s!"Symbol not found: {copyFrom}"
| .none => return .error <| Protocol.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")
match expr? with match expr? with
| .error error => Protocol.throw error | .error error => return .error error
| .ok goalState => | .ok goalState =>
let stateId ← newGoalState goalState let stateId ← newGoalState goalState
return { stateId, root := goalState.root.name.toString } return .ok { stateId, root := goalState.root.name.toString }
goal_continue (args: Protocol.GoalContinue): EMainM Protocol.GoalContinueResult := do goal_tactic (args: Protocol.GoalTactic): MainM (CR Protocol.GoalTacticResult) := do
let state ← getMainState let state ← get
let .some goalState := state.goalStates[args.stateId]? |
return .error $ errorIndex s!"Invalid state index {args.stateId}"
let .some goal := goalState.goals.get? args.goalId |
return .error $ errorIndex s!"Invalid goal index {args.goalId}"
let nextGoalState?: Except _ TacticResult ← runTermElabInMainM do
match args.tactic?, args.expr?, args.have?, args.let?, args.calc?, args.conv? with
| .some tactic, .none, .none, .none, .none, .none => do
pure <| Except.ok <| ← goalState.tryTactic goal tactic
| .none, .some expr, .none, .none, .none, .none => do
pure <| Except.ok <| ← goalState.tryAssign goal expr
| .none, .none, .some type, .none, .none, .none => do
let binderName := args.binderName?.getD ""
pure <| Except.ok <| ← goalState.tryHave goal binderName type
| .none, .none, .none, .some type, .none, .none => do
let binderName := args.binderName?.getD ""
pure <| Except.ok <| ← goalState.tryLet goal binderName type
| .none, .none, .none, .none, .some pred, .none => do
pure <| Except.ok <| ← goalState.tryCalc goal pred
| .none, .none, .none, .none, .none, .some true => do
pure <| Except.ok <| ← goalState.conv goal
| .none, .none, .none, .none, .none, .some false => do
pure <| Except.ok <| ← goalState.convExit
| _, _, _, _, _, _ =>
let error := errorI "arguments" "Exactly one of {tactic, expr, have, calc, conv} must be supplied"
pure $ Except.error $ error
match nextGoalState? with
| .error error => return .error 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) |
throwError "Resuming known goals"
pure result
| true, .some true => pure nextGoalState
| true, .some false => do
let .some (_, _, dormantGoals) := goalState.convMVar? |
throwError "If conv exit succeeded this should not fail"
let .ok result := nextGoalState.resume (nextGoalState.goals ++ dormantGoals) |
throwError "Resuming known goals"
pure result
| false, _ => pure nextGoalState
let nextStateId ← newGoalState nextGoalState
let goals ← nextGoalState.serializeGoals (parent := .some goalState) (options := state.options) |>.run'
return .ok {
nextStateId? := .some nextStateId,
goals? := .some goals,
}
| .ok (.parseError message) =>
return .ok { parseError? := .some message }
| .ok (.invalidAction message) =>
return .error $ errorI "invalid" message
| .ok (.failure messages) =>
return .ok { tacticErrors? := .some messages }
goal_continue (args: Protocol.GoalContinue): MainM (CR Protocol.GoalContinueResult) := do
let state ← get
let .some target := state.goalStates[args.target]? | let .some target := state.goalStates[args.target]? |
Protocol.throw $ Protocol.errorIndex s!"Invalid state index {args.target}" return .error $ errorIndex s!"Invalid state index {args.target}"
let nextGoalState? : GoalState ← match args.branch?, args.goals? with let nextState? ← 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 => return .error $ 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 }) pure $ goalResume target goals
pure $ target.resume goals | _, _ => return .error <| errorI "arguments" "Exactly one of {branch, goals} must be supplied"
| _, _ => Protocol.throw $ errorI "arguments" "Exactly one of {branch, goals} must be supplied" match nextState? with
match nextGoalState? with | .error error => return .error <| errorI "structure" error
| .error error => Protocol.throw $ errorI "structure" error
| .ok nextGoalState => | .ok nextGoalState =>
let nextStateId ← newGoalState nextGoalState let nextStateId ← newGoalState nextGoalState
let goals ← liftMetaM $ goalSerialize nextGoalState (options := state.options) let goals ← goalSerialize nextGoalState (options := state.options)
return { return .ok {
nextStateId, nextStateId,
goals, goals,
} }
goal_delete (args: Protocol.GoalDelete): EMainM Protocol.GoalDeleteResult := do goal_delete (args: Protocol.GoalDelete): MainM (CR Protocol.GoalDeleteResult) := do
let state ← getMainState let state ← get
let goalStates := args.stateIds.foldl (λ map id => map.erase id) state.goalStates let goalStates := args.stateIds.foldl (λ map id => map.erase id) state.goalStates
set { state with goalStates } set { state with goalStates }
return {} return .ok {}
goal_print (args: Protocol.GoalPrint): EMainM Protocol.GoalPrintResult := do goal_print (args: Protocol.GoalPrint): MainM (CR Protocol.GoalPrintResult) := do
let state ← getMainState let state ← get
let .some goalState := state.goalStates[args.stateId]? | let .some goalState := state.goalStates[args.stateId]? |
Protocol.throw $ Protocol.errorIndex s!"Invalid state index {args.stateId}" return .error $ errorIndex s!"Invalid state index {args.stateId}"
let result ← liftMetaM <| goalPrint let result ← runMetaInMainM <| goalPrint goalState state.options
goalState return .ok result
(rootExpr := args.rootExpr?.getD False) goal_save (args: Protocol.GoalSave): MainM (CR Protocol.GoalSaveResult) := do
(parentExprs := args.parentExprs?.getD False) let state ← get
(goals := args.goals?.getD False)
(extraMVars := args.extraMVars?.getD #[])
(options := state.options)
return result
goal_save (args: Protocol.GoalSave): EMainM Protocol.GoalSaveResult := do
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}" return .error $ errorIndex s!"Invalid state index {args.id}"
goalStatePickle goalState args.path (background? := .some $ ← getEnv) goalStatePickle goalState args.path
return {} return .ok {}
goal_load (args: Protocol.GoalLoad): EMainM Protocol.GoalLoadResult := do goal_load (args: Protocol.GoalLoad): MainM (CR Protocol.GoalLoadResult) := do
let (goalState, _) ← goalStateUnpickle args.path (background? := .some $ ← getEnv) let (goalState, _) ← goalStateUnpickle args.path (← Lean.MonadEnv.getEnv)
let id ← newGoalState goalState let id ← newGoalState goalState
return { id } return .ok { id }
frontend_process (args: Protocol.FrontendProcess): MainM (CR Protocol.FrontendProcessResult) := do
let options := (← get).options
try
let (fileName, file) ← match args.fileName?, args.file? with
| .some fileName, .none => do
let file ← IO.FS.readFile fileName
pure (fileName, file)
| .none, .some file =>
pure ("<anonymous>", file)
| _, _ => return .error <| errorI "arguments" "Exactly one of {fileName, file} must be supplied"
let env?: Option Lean.Environment ← if args.fileName?.isSome then
pure .none
else do
let env ← Lean.MonadEnv.getEnv
pure <| .some env
let (context, state) ← do Frontend.createContextStateFromFile file fileName env? {}
let frontendM := Frontend.mapCompilationSteps λ step => do
let boundary := (step.src.startPos.byteIdx, step.src.stopPos.byteIdx)
let invocations?: Option (List Protocol.InvokedTactic) ← if args.invocations then
let invocations ← Frontend.collectTacticsFromCompilationStep step
pure $ .some invocations
else
pure .none
let sorrys := if args.sorrys then
Frontend.collectSorrys step
else
[]
let messages ← step.messageStrings
return (step.before, boundary, invocations?, sorrys, messages)
let li ← frontendM.run context |>.run' state
let units ← li.mapM λ (env, boundary, invocations?, sorrys, messages) => Lean.withEnv env do
let (goalStateId?, goals) ← if sorrys.isEmpty then do
pure (.none, #[])
else do
let goalState ← runMetaInMainM $ Frontend.sorrysToGoalState sorrys
let stateId ← newGoalState goalState
let goals ← goalSerialize goalState options
pure (.some stateId, goals)
return {
boundary,
invocations?,
goalStateId?,
goals,
messages,
}
return .ok { units }
catch e =>
return .error $ errorI "frontend" (← e.toMessageData.toString)
end Pantograph.Repl end Pantograph.Repl

View File

@ -48,12 +48,6 @@ namespace Condensed
deriving instance BEq, Repr for LocalDecl deriving instance BEq, Repr for LocalDecl
deriving instance BEq, Repr for Goal deriving instance BEq, Repr for Goal
-- Enable string interpolation
instance : ToString FVarId where
toString id := id.name.toString
instance : ToString MVarId where
toString id := id.name.toString
protected def LocalDecl.devolatilize (decl: LocalDecl): LocalDecl := protected def LocalDecl.devolatilize (decl: LocalDecl): LocalDecl :=
{ {
decl with fvarId := { name := .anonymous } decl with fvarId := { name := .anonymous }
@ -67,13 +61,11 @@ protected def Goal.devolatilize (goal: Goal): Goal :=
end Condensed end Condensed
def GoalState.get! (state: GoalState) (i: Nat): MVarId := state.goals[i]! def GoalState.get! (state: GoalState) (i: Nat): MVarId := state.goals.get! 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.goals.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 =>
let messages := "\n".intercalate messages.toList let messages := "\n".intercalate messages.toList
s!".failure {messages}" s!".failure {messages}"
@ -96,31 +88,29 @@ 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 [Monad m] [MonadEnv m] (s: String): m Syntax := do
let .ok stx := Parser.runParserCategory let .ok stx := Parser.runParserCategory
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
(catName := `term) (catName := `term)
(input := s) (input := s)
(fileName := ← getFileName) | panic! s!"Failed to parse {s}" (fileName := filename) | panic! s!"Failed to parse {s}"
return stx return stx
def parseSentence (s : String) (expectedType? : Option Expr := .none) : Elab.TermElabM Expr := do def parseSentence (s: String): Elab.TermElabM Expr := do
let stx ← match Parser.runParserCategory let stx ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
(catName := `term) (catName := `term)
(input := s) (input := s)
(fileName := ← getFileName) with (fileName := filename) with
| .ok syn => pure syn | .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}" | .error error => throwError "Failed to parse: {error}"
Elab.Term.elabTerm (stx := stx) expectedType? Elab.Term.elabTerm (stx := stx) .none
def runTacticOnMVar (tacticM: Elab.Tactic.TacticM Unit) (goal: MVarId): Elab.TermElabM (List MVarId) := do def runTacticOnMVar (tacticM: Elab.Tactic.TacticM Unit) (goal: MVarId): Elab.TermElabM (List MVarId) := do
let (_, newGoals) ← tacticM { elaborator := .anonymous } |>.run { goals := [goal] } let (_, newGoals) ← tacticM { elaborator := .anonymous } |>.run { goals := [goal] }
@ -133,41 +123,26 @@ def mvarUserNameAndType (mvarId: MVarId): MetaM (Name × String) := do
-- Monadic testing -- Monadic testing
abbrev TestT := StateRefT' IO.RealWorld LSpec.TestSeq abbrev TestT := StateT LSpec.TestSeq
section Monadic def addTest [Monad m] (test: LSpec.TestSeq) : TestT m Unit := do
variable [Monad m] [MonadLiftT (ST IO.RealWorld) m]
def addTest (test: LSpec.TestSeq) : TestT m Unit := do
set $ (← get) ++ test set $ (← get) ++ test
def checkEq [DecidableEq α] [Repr α] (desc : String) (lhs rhs : α) : TestT m Unit := do def checkEq [Monad m] [DecidableEq α] (desc : String) (lhs rhs : α) : TestT m Unit := do
addTest $ LSpec.check desc (lhs = rhs) addTest $ LSpec.check desc (lhs == rhs)
def checkTrue (desc : String) (flag : Bool) : TestT m Unit := do def checkTrue [Monad m] (desc : String) (flag : Bool) : TestT m Unit := do
addTest $ LSpec.check desc flag addTest $ LSpec.check desc flag
def fail (desc : String) : TestT m Unit := do def fail [Monad m] (desc : String) : TestT m Unit := do
addTest $ LSpec.check desc false addTest $ LSpec.check desc false
def runTest (t: TestT m Unit): m LSpec.TestSeq := def runTest [Monad m] (t: TestT m Unit): m LSpec.TestSeq :=
Prod.snd <$> t.run LSpec.TestSeq.done Prod.snd <$> t.run LSpec.TestSeq.done
def runTestWithResult { α } (t: TestT m α): m (α × LSpec.TestSeq) := def runTestWithResult { α } [Monad m] (t: TestT m α): m (α × LSpec.TestSeq) :=
t.run LSpec.TestSeq.done t.run LSpec.TestSeq.done
def runTestCoreM (env: Environment) (coreM: TestT CoreM Unit) (options: Array String := #[]): IO LSpec.TestSeq := do
runCoreMSeq env (runTest coreM) options
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 }
@ -183,29 +158,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

View File

@ -3,10 +3,11 @@ import Pantograph.Delate
import Test.Common import Test.Common
import Lean import Lean
open Lean Pantograph open Lean
namespace Pantograph.Test.Delate namespace Pantograph.Test.Delate
open Pantograph
deriving instance Repr, DecidableEq for Protocol.BoundExpression deriving instance Repr, DecidableEq for Protocol.BoundExpression
def test_serializeName: LSpec.TestSeq := def test_serializeName: LSpec.TestSeq :=
@ -34,7 +35,7 @@ def test_sexp_of_symbol (env: Environment): IO LSpec.TestSeq := do
("Nat.add", "(:forall a (:c Nat) (:forall a (:c Nat) (:c Nat)))"), ("Nat.add", "(:forall a (:c Nat) (:forall a (:c Nat) (:c Nat)))"),
-- These ones are normal and easy -- These ones are normal and easy
("Nat.add_one", "(:forall n (:c Nat) ((:c Eq) (:c Nat) ((:c HAdd.hAdd) (:c Nat) (:c Nat) (:c Nat) ((:c instHAdd) (:c Nat) (:c instAddNat)) 0 ((:c OfNat.ofNat) (:c Nat) (:lit 1) ((:c instOfNatNat) (:lit 1)))) ((:c Nat.succ) 0)))"), ("Nat.add_one", "(:forall n (:c Nat) ((:c Eq) (:c Nat) ((:c HAdd.hAdd) (:c Nat) (:c Nat) (:c Nat) ((:c instHAdd) (:c Nat) (:c instAddNat)) 0 ((:c OfNat.ofNat) (:c Nat) (:lit 1) ((:c instOfNatNat) (:lit 1)))) ((:c Nat.succ) 0)))"),
("Nat.le_of_succ_le", "(:forall n (:c Nat) (:forall m (:c Nat) (:forall h ((:c LE.le) (:c Nat) (:c instLENat) ((:c Nat.succ) 1) 0) ((:c LE.le) (:c Nat) (:c instLENat) 2 1)) :i) :i)"), ("Nat.le_of_succ_le", "(:forall n (:c Nat) (:forall m (:c Nat) (:forall h ((:c LE.le) (:c Nat) (:c instLENat) ((:c Nat.succ) 1) 0) ((:c LE.le) (:c Nat) (:c instLENat) 2 1)) :implicit) :implicit)"),
-- Handling of higher order types -- Handling of higher order types
("Or", "(:forall a (:sort 0) (:forall b (:sort 0) (:sort 0)))"), ("Or", "(:forall a (:sort 0) (:forall b (:sort 0) (:sort 0)))"),
("List", "(:forall α (:sort (+ u 1)) (:sort (+ u 1)))") ("List", "(:forall α (:sort (+ u 1)) (:sort (+ u 1)))")
@ -48,9 +49,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))"),
("λ {α: Sort (u + 1)} => List α", [`u], "(:lambda α (:sort (+ u 1)) ((:c List) 0) :i)"), ("λ x: Array Nat => x.data", [], "(:lambda x ((:c Array) (:c Nat)) ((:c Array.data) (:c Nat) 0))"),
("λ {α} => List α", [], "(:lambda α (:sort (+ (:mv _uniq.4) 1)) ((:c List) 0) :i)"), ("λ {α: Sort (u + 1)} => List α", [`u], "(:lambda α (:sort (+ u 1)) ((:c List) 0) :implicit)"),
("(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)))"), ("λ {α} => List α", [], "(:lambda α (:sort (+ (:mv _uniq.4) 1)) ((:c List) 0) :implicit)"),
("(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
@ -75,7 +77,7 @@ def test_sexp_of_expr (env: Environment): IO LSpec.TestSeq := do
.default) .default)
.implicit) .implicit)
.implicit, .implicit,
"(:lambda p (:sort 0) (:lambda q (:sort 0) (:lambda k ((:c And) 1 0) ((:c And.right) _ _ 0)) :i) :i)" "(:lambda p (:sort 0) (:lambda q (:sort 0) (:lambda k ((:c And) 1 0) ((:c And.right) _ _ 0)) :implicit) :implicit)"
), ),
] ]
let termElabM: Elab.TermElabM LSpec.TestSeq := entries.foldlM (λ suites (expr, target) => do let termElabM: Elab.TermElabM LSpec.TestSeq := entries.foldlM (λ suites (expr, target) => do
@ -94,30 +96,6 @@ def test_instance (env: Environment): IO LSpec.TestSeq :=
let _expr := (← runTermElabMInMeta <| elabTerm s) |>.toOption |>.get! let _expr := (← runTermElabMInMeta <| elabTerm s) |>.toOption |>.get!
return LSpec.TestSeq.done return LSpec.TestSeq.done
def test_projection_prod (env: Environment) : IO LSpec.TestSeq:= runTest do
let struct := .app (.bvar 1) (.bvar 0)
let expr := .proj `Prod 1 struct
let .field projector numParams := analyzeProjection env expr |
fail "`Prod has fields"
checkEq "projector" projector `Prod.snd
checkEq "numParams" numParams 2
def test_projection_exists (env: Environment) : IO LSpec.TestSeq:= runTest do
let struct := .app (.bvar 1) (.bvar 0)
let expr := .proj `Exists 1 struct
let .singular recursor numParams numFields := analyzeProjection env expr |
fail "`Exists has no projectors"
checkEq "recursor" recursor `Exists.recOn
checkEq "numParams" numParams 2
checkEq "numFields" numFields 2
def test_matcher : TestT Elab.TermElabM Unit := do
let t ← parseSentence "Nat → Nat"
let e ← parseSentence "fun (n : Nat) => match n with | 0 => 0 | k => k" (.some t)
let .some _ ← Meta.matchMatcherApp? e.bindingBody! | fail "Must be a matcher app"
let e' ← instantiateAll e
checkTrue "ok" <| ← Meta.isTypeCorrect e'
def suite (env: Environment): List (String × IO LSpec.TestSeq) := def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
[ [
("serializeName", do pure test_serializeName), ("serializeName", do pure test_serializeName),
@ -126,9 +104,6 @@ def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
("Sexp from elaborated expr", test_sexp_of_elab env), ("Sexp from elaborated expr", test_sexp_of_elab env),
("Sexp from expr", test_sexp_of_expr env), ("Sexp from expr", test_sexp_of_expr env),
("Instance", test_instance env), ("Instance", test_instance env),
("Projection Prod", test_projection_prod env),
("Projection Exists", test_projection_exists env),
("Matcher", runTestTermElabM env test_matcher),
] ]
end Pantograph.Test.Delate end Pantograph.Test.Delate

View File

@ -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,11 @@ 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 suite: List (String × IO LSpec.TestSeq) :=
addTest $ ← runTestCoreM env do
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"
-- Extraction of source doesn't work for symbols in `Init` for some reason
checkTrue "file" result.sourceUri?.isNone
checkEq "sourceStart" (result.sourceStart?.map (·.column)) <| .some 0
checkEq "sourceEnd" (result.sourceEnd?.map (·.column)) <| .some 88
let { imports, constNames, .. } ← Environment.moduleRead ⟨"Init.Data.Nat.Basic"⟩
checkEq "imports" imports #["Init.SimpLemmas", "Init.Data.NeZero", "Init.Grind.Tactics"]
checkTrue "constNames" $ constNames.contains "Nat.succ_add"
def test_matcher (env : Environment) : TestT IO Unit := do
checkTrue "not matcher" $ ¬ Meta.isMatcherCore env `Nat.strongRecOn
def suite (env : Environment) : 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),
("Matcher", runTest $ test_matcher env),
] ]
end Pantograph.Test.Environment end Pantograph.Test.Environment

View File

@ -1,40 +1,22 @@
import LSpec
import Pantograph import Pantograph
import Repl import Repl
import Test.Common import Test.Common
import LSpec
open Lean Pantograph open Lean Pantograph
namespace Pantograph.Test.Frontend namespace Pantograph.Test.Frontend
def runFrontend { α } (source: String) (f : Frontend.CompilationStep → IO α) : MetaM (List α) := do def collectSorrysFromSource (source: String) : MetaM (List GoalState) := do
let filename := "<anonymous>"
let (context, state) ← do Frontend.createContextStateFromFile source filename (← getEnv) {}
let m := Frontend.mapCompilationSteps f
m.run context |>.run' state
def test_open : TestT MetaM Unit := do
let sketch := "
open Nat
example : ∀ (n : Nat), n + 1 = Nat.succ n := by
intro
apply add_one
"
let errors ← runFrontend sketch λ step => step.msgs.mapM (·.toString)
checkEq "errors" errors [[], []]
def collectSorrysFromSource (source: String) (options : Frontend.GoalCollectionOptions := {})
: MetaM (List GoalState) := do
let filename := "<anonymous>" let filename := "<anonymous>"
let (context, state) ← do Frontend.createContextStateFromFile source filename (← getEnv) {} let (context, state) ← do Frontend.createContextStateFromFile source filename (← getEnv) {}
let m := Frontend.mapCompilationSteps λ step => do let m := Frontend.mapCompilationSteps λ step => do
return (step.before, Frontend.collectSorrys step options) return (step.before, Frontend.collectSorrys step)
let li ← m.run context |>.run' state let li ← m.run context |>.run' state
let goalStates ← li.filterMapM λ (env, sorrys) => withEnv env do let goalStates ← li.filterMapM λ (env, sorrys) => withEnv env do
if sorrys.isEmpty then if sorrys.isEmpty then
return .none return .none
let { state, .. } ← Frontend.sorrysToGoalState sorrys let goalState ← Frontend.sorrysToGoalState sorrys
return .some state return .some goalState
return goalStates return goalStates
def test_multiple_sorrys_in_proof : TestT MetaM Unit := do def test_multiple_sorrys_in_proof : TestT MetaM Unit := do
@ -195,71 +177,14 @@ example (n: Nat) : mystery n + 1 = n + 2 := sorry
} }
]) ])
def test_capture_type_mismatch : TestT MetaM Unit := do
let input := "
def mystery (k: Nat) : Nat := true
"
let options := { collectTypeErrors := true }
let goalStates ← (collectSorrysFromSource input options).run' {}
let [goalState] := goalStates | panic! s!"Incorrect number of states: {goalStates.length}"
checkEq "goals" ((← goalState.serializeGoals).map (·.devolatilize)) #[
{
target := { pp? := "Nat" },
vars := #[{
userName := "k",
type? := .some { pp? := "Nat" },
}],
}
]
def test_capture_type_mismatch_in_binder : TestT MetaM Unit := do
let input := "
example (p: Prop) (h: (∀ (x: Prop), Nat) → p): p := h (λ (y: Nat) => 5)
"
let options := { collectTypeErrors := true }
let goalStates ← (collectSorrysFromSource input options).run' {}
let [goalState] := goalStates | panic! s!"Incorrect number of states: {goalStates.length}"
checkEq "goals" ((← goalState.serializeGoals (options := {})).map (·.devolatilize)) #[
]
def collectNewConstants (source: String) : MetaM (List (List Name)) := do
let filename := "<anonymous>"
let (context, state) ← do Frontend.createContextStateFromFile source filename (← getEnv) {}
let m := Frontend.mapCompilationSteps λ step => do
Frontend.collectNewDefinedConstants step
m.run context |>.run' state
def test_collect_one_constant : TestT MetaM Unit := do
let input := "
def mystery : Nat := 123
"
let names ← collectNewConstants input
checkEq "constants" names [[`mystery]]
def test_collect_one_theorem : TestT MetaM Unit := do
let input := "
theorem mystery [SizeOf α] (as : List α) (i : Fin as.length) : sizeOf (as.get i) < sizeOf as := by
match as, i with
| a::as, ⟨0, _⟩ => simp_arith [get]
| a::as, ⟨i+1, h⟩ =>
have ih := sizeOf_get as ⟨i, Nat.le_of_succ_le_succ h⟩
apply Nat.lt_trans ih
simp_arith
"
let names ← collectNewConstants input
checkEq "constants" names [[`mystery]]
def suite (env : Environment): List (String × IO LSpec.TestSeq) := def suite (env : Environment): List (String × IO LSpec.TestSeq) :=
let tests := [ let tests := [
("open", test_open),
("multiple_sorrys_in_proof", test_multiple_sorrys_in_proof), ("multiple_sorrys_in_proof", test_multiple_sorrys_in_proof),
("sorry_in_middle", test_sorry_in_middle), ("sorry_in_middle", test_sorry_in_middle),
("sorry_in_induction", test_sorry_in_induction), ("sorry_in_induction", test_sorry_in_induction),
("sorry_in_coupled", test_sorry_in_coupled), ("sorry_in_coupled", test_sorry_in_coupled),
("environment_capture", test_environment_capture), ("environment_capture", test_environment_capture),
("capture_type_mismatch", test_capture_type_mismatch),
--("capture_type_mismatch_in_binder", test_capture_type_mismatch_in_binder),
("collect_one_constant", test_collect_one_constant),
("collect_one_theorem", test_collect_one_theorem),
] ]
tests.map (fun (name, test) => (name, runMetaMSeq env $ runTest test)) tests.map (fun (name, test) => (name, runMetaMSeq env $ runTest test))

View File

@ -8,115 +8,76 @@ import Test.Common
namespace Pantograph.Test.Integration namespace Pantograph.Test.Integration
open Pantograph.Repl open Pantograph.Repl
deriving instance Lean.ToJson for Protocol.EnvInspect def step { α } [Lean.ToJson α] (cmd: String) (payload: List (String × Lean.Json))
deriving instance Lean.ToJson for Protocol.EnvAdd (expected: α) (name? : Option String := .none): MainM LSpec.TestSeq := do
deriving instance Lean.ToJson for Protocol.ExprEcho let payload := Lean.Json.mkObj payload
deriving instance Lean.ToJson for Protocol.OptionsSet
deriving instance Lean.ToJson for Protocol.OptionsPrint
deriving instance Lean.ToJson for Protocol.GoalStart
deriving instance Lean.ToJson for Protocol.GoalPrint
deriving instance Lean.ToJson for Protocol.GoalTactic
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 step { α β } [Lean.ToJson α] [Lean.ToJson β] (cmd: String) (payload: α)
(expected: β) (name? : Option String := .none) : TestM Unit := do
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 (toString result = toString (Lean.toJson expected))
def stepFile { α } [Lean.ToJson α] (name : String) (path : String)
(expected : α) : TestM Unit := do
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 := abbrev Test := List (MainM LSpec.TestSeq)
def test_elab : Test :=
[
step "expr.echo" step "expr.echo"
({ expr := "λ {α : Sort (u + 1)} => List α", levels? := .some #["u"]}: Protocol.ExprEcho) [("expr", .str "λ {α : Sort (u + 1)} => List α"), ("levels", .arr #["u"])]
({ (Lean.toJson ({
type := { pp? := .some "{α : Type u} → Type u" }, type := { pp? := .some "{α : Type u} → Type u" },
expr := { pp? := .some "fun {α} => List α" } expr := { pp? := .some "fun {α} => List α" }
}: Protocol.ExprEchoResult) }: 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", .str "Nat.add_one")]
step "options.set" ({ printExprAST? := .some true } : Protocol.OptionsSet) ({ type := { pp? }, module? }: Protocol.EnvInspectResult),
({ }: Protocol.OptionsSetResult) step "options.set" [("printExprAST", .bool true)]
step "env.inspect" ({ name := "Nat.add_one" } : Protocol.EnvInspect) ({ }: Protocol.OptionsSetResult),
({ type := { pp?, sexp? }, module? }: Protocol.EnvInspectResult) step "env.inspect" [("name", .str "Nat.add_one")]
step "options.print" ({} : Protocol.OptionsPrint) ({ type := { pp?, sexp? }, module? }: Protocol.EnvInspectResult),
({ options with printExprAST := true }: Protocol.Options) step "options.print" []
({ 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) [
step invalid [("name", .str "Nat.add_one")]
({ error := "command", desc := s!"Unknown command {invalid}" }: Protocol.InteractionError) ({ error := "command", desc := s!"Unknown command {invalid}" }: Protocol.InteractionError)
(name? := .some "Invalid Command") (name? := .some "Invalid Command"),
step "expr.echo" (Lean.Json.mkObj [(invalid, .str "Random garbage data")]) step "expr.echo" [(invalid, .str "Random garbage data")]
({ error := "command", desc := s!"Unable to parse json: Pantograph.Protocol.ExprEcho.expr: String expected" }: ({ error := "command", desc := s!"Unable to parse json: Pantograph.Protocol.ExprEcho.expr: String expected" }:
Protocol.InteractionError) (name? := .some "JSON Deserialization") Protocol.InteractionError)
def test_tactic : Test := do (name? := .some "JSON Deserialization")
let varX := { name := "_uniq.10", userName := "x", type? := .some { pp? := .some "Prop" }} ]
def test_tactic : Test :=
let goal1: Protocol.Goal := { let goal1: Protocol.Goal := {
name := "_uniq.11", name := "_uniq.11",
target := { pp? := .some "∀ (q : Prop), x q → q x" }, target := { pp? := .some "∀ (q : Prop), x q → q x" },
vars := #[varX], vars := #[{ name := "_uniq.10", userName := "x", type? := .some { pp? := .some "Prop" }}],
} }
let goal2: Protocol.Goal := { let goal2: Protocol.Goal := {
name := "_uniq.14", name := "_uniq.17",
target := { pp? := .some "x y → y x" }, target := { pp? := .some "x y → y x" },
vars := #[ vars := #[
varX, { name := "_uniq.10", userName := "x", type? := .some { pp? := .some "Prop" }},
{ name := "_uniq.13", userName := "y", type? := .some { pp? := .some "Prop" }} { name := "_uniq.16", 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", .str "∀ (p q: Prop), p q → q p")]
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", .num 0), ("goalId", .num 0), ("tactic", .str "intro x")]
step "goal.print" ({ stateId := 1, parentExprs? := .some true, rootExpr? := .some true }: Protocol.GoalPrint) ({ nextStateId? := .some 1, goals? := #[goal1], }: Protocol.GoalTacticResult),
({ step "goal.print" [("stateId", .num 1)]
root? := .some { pp? := "fun x => ?m.11"}, ({ parent? := .some { pp? := .some "fun x => ?m.12 x" }, }: Protocol.GoalPrintResult),
parentExprs? := .some [.some { pp? := .some "fun x => ?m.11" }], step "goal.tactic" [("stateId", .num 1), ("goalId", .num 0), ("tactic", .str "intro y")]
}: Protocol.GoalPrintResult) ({ nextStateId? := .some 2, goals? := #[goal2], }: Protocol.GoalTacticResult),
step "goal.tactic" ({ stateId := 1, tactic? := .some "intro y" }: Protocol.GoalTactic) ]
({ nextStateId? := .some 2, goals? := #[goal2], }: Protocol.GoalTacticResult) def test_automatic_mode (automatic: Bool): Test :=
step "goal.tactic" ({ stateId := 1, tactic? := .some "apply Nat.le_of_succ_le" }: Protocol.GoalTactic)
({ messages? := .some #["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
simp
def test_tactic_timeout : Test := do
step "goal.start" ({ expr := "(1 : Nat) + (2 * 3) = 1 + (4 - 3) + (6 - 4) + 3" }: Protocol.GoalStart)
({ stateId := 0, root := "_uniq.370" }: Protocol.GoalStartResult)
-- timeout of 10 milliseconds
step "options.set" ({ timeout? := .some 10 } : Protocol.OptionsSet)
({ }: Protocol.OptionsSetResult)
step "goal.tactic" ({ stateId := 0, expr? := .some "by\nsleep 1000; simp" }: Protocol.GoalTactic)
({ error := "internal", desc := "interrupt" }: Protocol.InteractionError)
-- ensure graceful recovery
step "options.set" ({ timeout? := .some 0 } : Protocol.OptionsSet)
({ }: Protocol.OptionsSetResult)
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
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" }}
@ -129,148 +90,94 @@ def test_automatic_mode (automatic: Bool): Test := do
], ],
} }
let goal2l: Protocol.Goal := { let goal2l: Protocol.Goal := {
name := "_uniq.61", name := "_uniq.59",
userName? := .some "inl", userName? := .some "inl",
target := { pp? := .some "q p" }, target := { pp? := .some "q p" },
vars := varsPQ ++ #[ vars := varsPQ ++ #[
{ name := "_uniq.49", userName := "h✝", type? := .some { pp? := .some "p" }, isInaccessible := true} { name := "_uniq.47", userName := "h✝", type? := .some { pp? := .some "p" }, isInaccessible := true}
], ],
} }
let goal2r: Protocol.Goal := { let goal2r: Protocol.Goal := {
name := "_uniq.74", name := "_uniq.72",
userName? := .some "inr", userName? := .some "inr",
target := { pp? := .some "q p" }, target := { pp? := .some "q p" },
vars := varsPQ ++ #[ vars := varsPQ ++ #[
{ name := "_uniq.62", userName := "h✝", type? := .some { pp? := .some "q" }, isInaccessible := true} { name := "_uniq.60", userName := "h✝", type? := .some { pp? := .some "q" }, isInaccessible := true}
], ],
} }
let goal3l: Protocol.Goal := { let goal3l: Protocol.Goal := {
name := "_uniq.80", name := "_uniq.78",
userName? := .some "inl.h", userName? := .some "inl.h",
target := { pp? := .some "p" }, target := { pp? := .some "p" },
vars := varsPQ ++ #[ vars := varsPQ ++ #[
{ name := "_uniq.49", userName := "h✝", type? := .some { pp? := .some "p" }, isInaccessible := true} { name := "_uniq.47", userName := "h✝", type? := .some { pp? := .some "p" }, isInaccessible := true}
], ],
} }
step "options.set" ({automaticMode? := .some automatic}: Protocol.OptionsSet) [
({}: Protocol.OptionsSetResult) step "options.set" [("automaticMode", .bool automatic)]
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", .str "∀ (p q: Prop), p q → q p")]
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", .num 0), ("goalId", .num 0), ("tactic", .str "intro p q h")]
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", .num 1), ("goalId", .num 0), ("tactic", .str "cases h")]
({ nextStateId? := .some 2, goals? := #[goal2l, goal2r], }: Protocol.GoalTacticResult),
let goals? := if automatic then #[goal3l, goal2r] else #[goal3l] let goals? := if automatic then #[goal3l, goal2r] else #[goal3l]
step "goal.tactic" ({ stateId := 2, tactic? := .some "apply Or.inr" }: Protocol.GoalTactic) step "goal.tactic" [("stateId", .num 2), ("goalId", .num 0), ("tactic", .str "apply Or.inr")]
({ nextStateId? := .some 3, goals?, }: Protocol.GoalTacticResult) ({ 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" [
step "env.add" step "env.add"
({ [
name := name1, ("name", .str name1),
value := "λ (a b: Prop) => Or a b", ("type", .str "Prop → Prop → Prop"),
isTheorem := false ("value", .str "λ (a b: Prop) => Or a b"),
}: Protocol.EnvAdd) ("isTheorem", .bool false)
({}: Protocol.EnvAddResult) ]
step "env.inspect" ({name := name1, value? := .some true} : Protocol.EnvInspect) ({}: Protocol.EnvAddResult),
step "env.inspect" [("name", .str name1)]
({ ({
value? := .some { pp? := .some "fun a b => a b" }, value? := .some { pp? := .some "fun a b => a b" },
type := { pp? := .some "Prop → Prop → Prop" }, type := { pp? := .some "Prop → Prop → Prop" },
}: Protocol.EnvInspectResult) }:
Protocol.EnvInspectResult),
step "env.add" step "env.add"
({ [
name := name2, ("name", .str name2),
type? := "Nat → Int", ("type", .str "Nat → Int"),
value := "λ (a: Nat) => a + 1", ("value", .str "λ (a: Nat) => a + 1"),
isTheorem := false ("isTheorem", .bool false)
}: Protocol.EnvAdd) ]
({}: Protocol.EnvAddResult) ({}: Protocol.EnvAddResult),
step "env.inspect" ({name := name2, value? := .some true} : Protocol.EnvInspect) step "env.inspect" [("name", .str name2)]
({ ({
value? := .some { pp? := .some "fun a => ↑a + 1" }, value? := .some { pp? := .some "fun a => ↑a + 1" },
type := { pp? := .some "Nat → Int" }, type := { pp? := .some "Nat → Int" },
}: Protocol.EnvInspectResult) }:
step "env.add" Protocol.EnvInspectResult)
({ ]
name := name3,
levels? := .some #["u"]
type? := "(α : Type u) → α → (α × α)",
value := "λ (α : Type u) (x : α) => (x, x)",
isTheorem := false
}: Protocol.EnvAdd)
({}: Protocol.EnvAddResult)
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", .str file),
invocations? := .some filename, ("invocations", .bool true),
}: Protocol.FrontendProcess) ("sorrys", .bool false),
]
({ ({
units := [{ units := [{
boundary := (0, file.utf8ByteSize), boundary := (0, file.utf8ByteSize),
nInvocations? := .some 2,
}],
}: Protocol.FrontendProcessResult)
stepFile (α := Protocol.FrontendData) "invocations" filename
{ units := [{
invocations? := .some [ invocations? := .some [
{ {
goalBefore := "⊢ ∀ (p q : Prop), p → p q", goalBefore := "⊢ ∀ (p q : Prop), p → p q",
@ -285,14 +192,17 @@ def test_frontend_process : Test := do
usedConstants := #["Or.inl"], usedConstants := #["Or.inl"],
}, },
] ]
} ] } }],
}: Protocol.FrontendProcessResult),
]
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 file := s!"{solved}{withSorry}"
let goal1: Protocol.Goal := { let goal1: Protocol.Goal := {
name := "_uniq.6", name := "_uniq.6",
@ -300,99 +210,49 @@ def test_frontend_process_sorry : Test := do
vars := #[{ name := "_uniq.4", userName := "p", type? := .some { pp? := .some "Prop" }}], vars := #[{ name := "_uniq.4", userName := "p", type? := .some { pp? := .some "Prop" }}],
} }
step "frontend.process" step "frontend.process"
({ [
file? := .some file, ("file", .str file),
sorrys := true, ("invocations", .bool false),
}: Protocol.FrontendProcess) ("sorrys", .bool true),
]
({ ({
units := [{ units := [{
boundary := (0, solved.utf8ByteSize), boundary := (0, solved.utf8ByteSize),
}, { }, {
boundary := (solved.utf8ByteSize, solved.utf8ByteSize + withSorry.utf8ByteSize), boundary := (solved.utf8ByteSize, solved.utf8ByteSize + withSorry.utf8ByteSize),
goalStateId? := .some 0, goalStateId? := .some 0,
goals? := .some #[goal1], goals := #[goal1],
goalSrcBoundaries? := .some #[(57, 62)],
messages := #["<anonymous>:2:0: warning: declaration uses 'sorry'\n"], messages := #["<anonymous>:2:0: warning: declaration uses 'sorry'\n"],
}], }],
}: Protocol.FrontendProcessResult) }: Protocol.FrontendProcessResult),
]
def test_import_open : Test := do
let header := "import Init\nopen Nat\nuniverse u"
let goal1: Protocol.Goal := {
name := "_uniq.81",
target := { pp? := .some "n + 1 = n.succ" },
vars := #[{ name := "_uniq.80", userName := "n", type? := .some { pp? := .some "Nat" }}],
}
step "frontend.process"
({
file? := .some header,
readHeader := true,
inheritEnv := true,
}: Protocol.FrontendProcess)
({
units := [
{ boundary := (12, 21) },
{ boundary := (21, header.utf8ByteSize) },
],
}: Protocol.FrontendProcessResult)
step "goal.start" ({ expr := "∀ (n : Nat), n + 1 = Nat.succ n"} : Protocol.GoalStart)
({ stateId := 0, root := "_uniq.79" }: Protocol.GoalStartResult)
step "goal.tactic" ({ stateId := 0, tactic? := .some "intro n" }: Protocol.GoalTactic)
({ nextStateId? := .some 1, goals? := #[goal1], }: Protocol.GoalTacticResult)
step "goal.tactic" ({ stateId := 1, tactic? := .some "apply add_one" }: Protocol.GoalTactic)
({ nextStateId? := .some 2, goals? := .some #[], }: Protocol.GoalTacticResult)
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 := #["<anonymous>:1:8: warning: declaration uses 'sorry'\n"],
}],
}: Protocol.FrontendProcessResult)
step "goal.tactic" ({ stateId := 0, tactic? := .some "exact?" }: Protocol.GoalTactic)
({ messages? := .some #["`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 context: Context := {
let mainM : MainM LSpec.TestSeq := runTest steps imports := ["Init"]
mainM.run { coreContext } |>.run' { env } }
let commands: MainM LSpec.TestSeq :=
steps.foldlM (λ suite step => do
let result ← step
return suite ++ result) LSpec.TestSeq.done
runCoreMSeq env <| commands.run context |>.run' {}
def suite (env : Lean.Environment): List (String × IO LSpec.TestSeq) := def suite (env : Lean.Environment): List (String × IO LSpec.TestSeq) :=
let tests := [ let tests := [
("expr.echo", test_expr_echo), ("expr.echo", test_elab),
("options.set options.print", test_option_modify), ("options.set options.print", test_option_modify),
("Malformed command", test_malformed_command), ("Malformed command", test_malformed_command),
("Tactic", test_tactic), ("Tactic", test_tactic),
("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 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

View File

@ -8,18 +8,15 @@ open Pantograph
namespace Pantograph.Test.Library namespace Pantograph.Test.Library
def runTermElabM { α } (termElabM: Elab.TermElabM α): CoreM α :=
termElabM.run' (ctx := defaultElabContext) |>.run'
def test_expr_echo (env: Environment): IO LSpec.TestSeq := do def test_expr_echo (env: Environment): IO LSpec.TestSeq := do
let inner: CoreM LSpec.TestSeq := do let inner: CoreM LSpec.TestSeq := do
let prop_and_proof := "⟨∀ (x: Prop), x → x, λ (x: Prop) (h: x) => h⟩" let prop_and_proof := "⟨∀ (x: Prop), x → x, λ (x: Prop) (h: x) => h⟩"
let tests := LSpec.TestSeq.done let tests := LSpec.TestSeq.done
let echoResult ← runTermElabM $ exprEcho prop_and_proof (options := {}) let echoResult ← exprEcho prop_and_proof (options := {})
let tests := tests.append (LSpec.test "fail" (echoResult.toOption == .some { let tests := tests.append (LSpec.test "fail" (echoResult.toOption == .some {
type := { pp? := "?m.2" }, expr := { pp? := "?m.3" } type := { pp? := "?m.2" }, expr := { pp? := "?m.3" }
})) }))
let echoResult ← runTermElabM $ exprEcho prop_and_proof (expectedType? := .some "Σ' p:Prop, p") (options := { printExprAST := true }) let echoResult ← exprEcho prop_and_proof (expectedType? := .some "Σ' p:Prop, p") (options := { printExprAST := true })
let tests := tests.append (LSpec.test "fail" (echoResult.toOption == .some { let tests := tests.append (LSpec.test "fail" (echoResult.toOption == .some {
type := { type := {
pp? := "(p : Prop) ×' p", pp? := "(p : Prop) ×' p",

View File

@ -17,9 +17,9 @@ def addPrefix (pref: String) (tests: List (String × α)): List (String × α)
tests.map (λ (name, x) => (pref ++ "/" ++ name, x)) tests.map (λ (name, x) => (pref ++ "/" ++ name, x))
/-- Runs test in parallel. Filters test name if given -/ /-- Runs test in parallel. Filters test name if given -/
def runTestGroup (nameFilter?: Option String) (tests: List (String × IO LSpec.TestSeq)): IO LSpec.TestSeq := do def runTestGroup (filter: Option String) (tests: List (String × IO LSpec.TestSeq)): IO LSpec.TestSeq := do
let tests: List (String × IO LSpec.TestSeq) := match nameFilter? with let tests: List (String × IO LSpec.TestSeq) := match filter with
| .some nameFilter => tests.filter (λ (name, _) => nameFilter.isPrefixOf name) | .some filter => tests.filter (λ (name, _) => filter.isPrefixOf name)
| .none => tests | .none => tests
let tasks: List (String × Task _) ← tests.mapM (λ (name, task) => do let tasks: List (String × Task _) ← tests.mapM (λ (name, task) => do
return (name, ← EIO.asTask task)) return (name, ← EIO.asTask task))
@ -37,30 +37,26 @@ open Pantograph.Test
/-- Main entry of tests; Provide an argument to filter tests by prefix -/ /-- Main entry of tests; Provide an argument to filter tests by prefix -/
def main (args: List String) := do def main (args: List String) := do
let nameFilter? := args.head? let name_filter := 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/Congruence", Tactic.Congruence.suite env_default),
("Tactic/Fragment", Tactic.Fragment.suite), ("Tactic/Motivated Apply", Tactic.MotivatedApply.suite env_default),
("Tactic/Prograde", Tactic.Prograde.suite), ("Tactic/No Confuse", Tactic.NoConfuse.suite env_default),
("Tactic/Special", Tactic.Special.suite), ("Tactic/Prograde", Tactic.Prograde.suite env_default),
] ]
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 LSpec.lspecIO (← runTestGroup name_filter tests)
let tests : List (String × IO LSpec.TestSeq) := suites.foldl (init := []) λ acc (name, suite) =>
acc ++ (addPrefix name $ suiterunner suite)
LSpec.lspecEachIO [()] (λ () => runTestGroup nameFilter? tests)

View File

@ -8,7 +8,10 @@ namespace Pantograph.Test.Metavar
open Pantograph open Pantograph
open Lean open Lean
abbrev TestM := TestT $ ReaderT Protocol.Options Elab.TermElabM abbrev TestM := StateRefT LSpec.TestSeq (ReaderT Protocol.Options Elab.TermElabM)
def addTest (test: LSpec.TestSeq): TestM Unit := do
set $ (← get) ++ test
-- Tests that all delay assigned mvars are instantiated -- Tests that all delay assigned mvars are instantiated
def test_instantiate_mvar: TestM Unit := do def test_instantiate_mvar: TestM Unit := do
@ -25,10 +28,12 @@ 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
let env ← Lean.MonadEnv.getEnv let env ← Lean.MonadEnv.getEnv
let syn? := parseTerm env expr let syn? := parseTerm env expr
@ -79,20 +84,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 +105,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 +116,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 +139,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 +151,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 +163,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 +178,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 +189,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.29 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 +221,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 +229,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 +243,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 ≤ ?m.succ", .some "?m.succ ≤ 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 ≤ ?m.succ", .some "?m.succ ≤ 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 +278,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))

View File

@ -11,16 +11,19 @@ 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 := StateRefT LSpec.TestSeq (ReaderT Protocol.Options Elab.TermElabM)
def addTest (test: LSpec.TestSeq): TestM Unit := do
set $ (← get) ++ test
def startProof (start: Start): TestM (Option GoalState) := do def startProof (start: Start): TestM (Option GoalState) := do
let env ← Lean.MonadEnv.getEnv let env ← Lean.MonadEnv.getEnv
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 +32,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 +60,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 +70,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 +83,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!) (sanitize := false)
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 +109,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 +119,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 ()
@ -110,13 +133,48 @@ def test_nat_add_comm (manual: Bool): TestM Unit := 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
@ -132,27 +190,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
@ -177,24 +235,22 @@ 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 ()
let [state1g0] := state1.goals | fail "Should have 1 goal" let fvP := "_uniq.10"
let (fvP, fvQ, fvH) ← state1.withContext state1g0 do let fvQ := "_uniq.13"
let lctx ← getLCtx let fvH := "_uniq.16"
let #[fvP, fvQ, fvH] := lctx.getFVarIds.map (toString ·.name) | let state1g0 := "_uniq.17"
panic! "Incorrect number of decls"
pure (fvP, fvQ, fvH)
addTest $ LSpec.check tactic ((← state1.serializeGoals (options := ← read)) = addTest $ LSpec.check tactic ((← state1.serializeGoals (options := ← read)) =
#[{ #[{
name := state1g0.name.toString, name := state1g0,
target := { pp? := .some "q p" }, target := { pp? := .some "q p" },
vars := #[ vars := #[
{ name := fvP, userName := "p", type? := .some { pp? := .some "Prop" } }, { name := fvP, userName := "p", type? := .some { pp? := .some "Prop" } },
@ -202,30 +258,28 @@ 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!) (sanitize := false)
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 ()
addTest $ LSpec.check tactic ((← state2.serializeGoals (options := ← read)).map (·.devolatilize) = addTest $ LSpec.check tactic ((← state2.serializeGoals (options := ← read)).map (·.devolatilize) =
#[branchGoal "inl" "p", branchGoal "inr" "q"]) #[branchGoal "inl" "p", branchGoal "inr" "q"])
let [state2g0, state2g1] := state2.goals | let (caseL, caseR) := ("_uniq.64", "_uniq.77")
fail s!"Should have 2 goals, but it has {state2.goals.length}"
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!) (sanitize := false)
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}))"
@ -236,57 +290,56 @@ 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!) (sanitize := false)
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 _uniq.91))")
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
addTest $ assertUnreachable $ msg addTest $ assertUnreachable $ msg
return () return ()
| .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.get! 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
@ -301,87 +354,368 @@ def test_or_comm: TestM Unit := do
] ]
} }
def test_tactic_failure_unresolved_goals : TestM Unit := do example : ∀ (a b c1 c2: Nat), (b + a) + c1 = (b + a) + c2 → (a + b) + c1 = (b + a) + c2 := by
let state? ← startProof (.expr "∀ (p : Nat → Prop), ∃ (x : Nat), p (0 + x + 0)") 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 let state0 ← match state? with
| .some state => pure state | .some state => pure state
| .none => do | .none => do
addTest $ assertUnreachable "Goal could not parse" addTest $ assertUnreachable "Goal could not parse"
return () return ()
let tactic := "intro p" let tactic := "intro a b c1 c2 h"
let state1 ← match ← state0.tacticOn 0 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.serializeGoals (options := ← read)).map (·.devolatilize) =
#[interiorGoal [] "a + b + c1 = b + a + c2"])
let tactic := "exact ⟨0, by simp⟩" let state2 ← match ← state1.conv (state1.get! 0) with
let .failure messages ← state1.tacticOn 0 tactic | addTest $ assertUnreachable s!"{tactic} should fail" | .success state => pure state
checkEq s!"{tactic} fails" messages #[s!"{← getFileName}:0:12: error: unsolved goals\np : Nat → Prop\n⊢ p 0\n"]
def test_tactic_failure_synthesize_placeholder : TestM Unit := do
let state? ← startProof (.expr "∀ (p q r : Prop) (h : p → q), q ∧ r")
let state0 ← match state? with
| .some state => pure state
| .none => do
addTest $ assertUnreachable "Goal could not parse"
return ()
let tactic := "intro p q r h"
let state1 ← match ← state0.tacticOn 0 tactic with
| .success state _ => pure state
| other => do | other => do
addTest $ assertUnreachable $ other.toString addTest $ assertUnreachable $ other.toString
return () return ()
addTest $ LSpec.check "conv => ..." ((← state2.serializeGoals (options := ← read)).map (·.devolatilize) =
#[{ interiorGoal [] "a + b + c1 = b + a + c2" with isConversion := true }])
let tactic := "simpa [h] using And.imp_left h _" let convTactic := "rhs"
--let state2 ← match ← state1.tacticOn 0 tactic with let state3R ← match ← state2.tacticOn (goalId := 0) convTactic with
-- | .success state => pure state | .success state => pure state
-- | other => do
-- addTest $ assertUnreachable $ other.toString
-- return ()
-- Volatile behaviour. This easily changes across Lean versions
--checkEq tactic ((← state2.serializeGoals).map (·.devolatilize)) #[
-- buildGoal [("p", "Prop"), ("q", "Prop"), ("r", "Prop"), ("h", "p → q")] "p ∧ r"
--]
let .failure messages ← state1.tacticOn 0 tactic | addTest $ assertUnreachable s!"{tactic} should fail"
let message := s!"<Pantograph>:0:31: error: don't know how to synthesize placeholder\ncontext:\np q r : Prop\nh : p → q\n⊢ p ∧ r\n"
checkEq s!"{tactic} fails" messages #[message]
def test_deconstruct : TestM Unit := do
let state? ← startProof (.expr "∀ (p q : Prop) (h : And p q), And q p")
let state0 ← match state? with
| .some state => pure state
| .none => do
addTest $ assertUnreachable "Goal could not parse"
return ()
let tactic := "intro p q ⟨hp, hq⟩"
let state1 ← match ← state0.tacticOn 0 tactic with
| .success state _ => pure state
| other => do | other => do
fail other.toString addTest $ assertUnreachable $ other.toString
return () return ()
checkEq tactic ((← state1.serializeGoals (options := ← read)).map (·.devolatilize)) 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) =
#[ #[
buildGoal [("p", "Prop"), ("q", "Prop"), ("hp", "p"), ("hq", "q")] "q ∧ p" { 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_nat_zero_add: TestM Unit := do
let state? ← startProof (.expr "∀ (n: Nat), n + 0 = n")
let state0 ← match state? with
| .some state => pure state
| .none => do
addTest $ assertUnreachable "Goal could not parse"
return ()
let tactic := "intro n"
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) =
#[buildGoal [("n", "Nat")] "n + 0 = n"])
let recursor := "@Nat.brecOn"
let state2 ← match ← state1.tryMotivatedApply (state1.get! 0) (recursor := recursor) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check s!"mapply {recursor}" ((← state2.serializeGoals (options := ← read)).map (·.devolatilizeVars) =
#[
buildNamedGoal "_uniq.67" [("n", "Nat")] "Nat → Prop" (.some "motive"),
buildNamedGoal "_uniq.68" [("n", "Nat")] "Nat",
buildNamedGoal "_uniq.69" [("n", "Nat")] "∀ (t : Nat), Nat.below t → ?motive t",
buildNamedGoal "_uniq.70" [("n", "Nat")] "?motive ?m.68 = (n + 0 = n)" (.some "conduit")
])
let tactic := "exact n"
let state3b ← match ← state2.tacticOn (goalId := 1) (tactic := tactic) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check tactic ((← state3b.serializeGoals (options := ← read)).map (·.devolatilize) =
#[])
let state2b ← match state3b.continue state2 with
| .ok state => pure state
| .error e => do
addTest $ assertUnreachable e
return ()
let tactic := "exact (λ x => x + 0 = x)"
let state3c ← match ← state2b.tacticOn (goalId := 0) (tactic := tactic) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check tactic ((← state3c.serializeGoals (options := ← read)).map (·.devolatilize) =
#[])
let state2c ← match state3c.continue state2b with
| .ok state => pure state
| .error e => do
addTest $ assertUnreachable e
return ()
let tactic := "intro t h"
let state3 ← match ← state2c.tacticOn (goalId := 0) (tactic := tactic) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check tactic ((← state3.serializeGoals (options := ← read)).map (·.devolatilize) =
#[buildGoal [("n", "Nat"), ("t", "Nat"), ("h", "Nat.below t")] "t + 0 = t"])
let tactic := "simp"
let state3d ← match ← state3.tacticOn (goalId := 0) (tactic := tactic) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
let state2d ← match state3d.continue state2c with
| .ok state => pure state
| .error e => do
addTest $ assertUnreachable e
return ()
let tactic := "rfl"
let stateF ← match ← state2d.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)) =
#[])
let expr := stateF.mctx.eAssignment.find! stateF.root
let (expr, _) := instantiateMVarsCore (mctx := stateF.mctx) (e := expr)
addTest $ LSpec.check "(F root)" stateF.rootExpr?.isSome
def test_nat_zero_add_alt: TestM Unit := do
let state? ← startProof (.expr "∀ (n: Nat), n + 0 = n")
let state0 ← match state? with
| .some state => pure state
| .none => do
addTest $ assertUnreachable "Goal could not parse"
return ()
let tactic := "intro n"
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) =
#[buildGoal [("n", "Nat")] "n + 0 = n"])
let recursor := "@Nat.brecOn"
let state2 ← match ← state1.tryMotivatedApply (state1.get! 0) (recursor := recursor) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
let major := "_uniq.68"
addTest $ LSpec.check s!"mapply {recursor}" ((← state2.serializeGoals (options := ← read)).map (·.devolatilizeVars) =
#[
buildNamedGoal "_uniq.67" [("n", "Nat")] "Nat → Prop" (.some "motive"),
buildNamedGoal major [("n", "Nat")] "Nat",
buildNamedGoal "_uniq.69" [("n", "Nat")] "∀ (t : Nat), Nat.below t → ?motive t",
buildNamedGoal "_uniq.70" [("n", "Nat")] "?motive ?m.68 = (n + 0 = n)" (.some "conduit")
])
let tactic := "intro x"
let state3m ← match ← state2.tacticOn (goalId := 0) (tactic := tactic) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check tactic ((← state3m.serializeGoals (options := ← read)).map (·.devolatilize) =
#[buildGoal [("n", "Nat"), ("x", "Nat")] "Prop" (.some "motive")])
let tactic := "apply Eq"
let state3m2 ← match ← state3m.tacticOn (goalId := 0) (tactic := tactic) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
let (eqL, eqR, eqT) := ("_uniq.88", "_uniq.89", "_uniq.87")
addTest $ LSpec.check tactic $ state3m2.goals.map (·.name.toString) = [eqL, eqR, eqT]
let [_motive, _major, _step, conduit] := state2.goals | panic! "Goals conflict"
let state2b ← match state3m2.resume [conduit] with
| .ok state => pure state
| .error e => do
addTest $ assertUnreachable e
return ()
let cNatAdd := "(:c HAdd.hAdd) (:c Nat) (:c Nat) (:c Nat) ((:c instHAdd) (:c Nat) (:c instAddNat))"
let cNat0 := "((:c OfNat.ofNat) (:c Nat) (:lit 0) ((:c instOfNatNat) (:lit 0)))"
let fvN := "_uniq.63"
let conduitRight := s!"((:c Eq) (:c Nat) ({cNatAdd} (:fv {fvN}) {cNat0}) (:fv {fvN}))"
let substOf (mv: String) := s!"(:subst (:mv {mv}) (:fv {fvN}) (:mv {major}))"
addTest $ LSpec.check "resume" ((← state2b.serializeGoals (options := { ← read with printExprAST := true })) =
#[
{
name := "_uniq.70",
userName? := .some "conduit",
target := {
pp? := .some "(?m.92 ?m.68 = ?m.94 ?m.68) = (n + 0 = n)",
sexp? := .some s!"((:c Eq) (:sort 0) ((:c Eq) {substOf eqT} {substOf eqL} {substOf eqR}) {conduitRight})",
},
vars := #[{
name := fvN,
userName := "n",
type? := .some { pp? := .some "Nat", sexp? := .some "(:c Nat)" },
}],
}
])
def suite (env: Environment): List (String × IO LSpec.TestSeq) := def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
let tests := [ let tests := [
("identity", test_identity), ("identity", test_identity),
("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),
("tactic failure with unresolved goals", test_tactic_failure_unresolved_goals), ("conv", test_conv),
("tactic failure with synthesize placeholder", test_tactic_failure_synthesize_placeholder), ("calc", test_calc),
("deconstruct", test_deconstruct), ("Nat.zero_add", test_nat_zero_add),
("Nat.zero_add alt", test_nat_zero_add_alt),
] ]
tests.map (fun (name, test) => (name, proofRunner env test)) tests.map (fun (name, test) => (name, proofRunner env test))

View File

@ -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,25 +30,29 @@ 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 let env' ← match (← getEnv).addDecl (← getOptions) c with
environmentPickle (← getEnv) envPicklePath | .error e => do
let error ← (e.toMessageData (← getOptions)).toString
throwError error
| .ok env' => pure env'
environmentPickle env' envPicklePath
let _ ← runCoreM coreDst do let _ ← runCoreM coreDst do
let (env', _) ← environmentUnpickle envPicklePath let (env', _) ← environmentUnpickle envPicklePath
@ -53,10 +60,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 +82,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 +99,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))

View File

@ -1,4 +1,4 @@
import Test.Tactic.Assign import Test.Tactic.Congruence
import Test.Tactic.Fragment import Test.Tactic.MotivatedApply
import Test.Tactic.NoConfuse
import Test.Tactic.Prograde import Test.Tactic.Prograde
import Test.Tactic.Special

View File

@ -1,33 +0,0 @@
import Lean.Meta
import Lean.Elab
import LSpec
import Test.Common
open Lean
namespace Pantograph.Test.Tactic.Assign
def test_draft : TestT Elab.TermElabM Unit := do
let expr := "forall (p : Prop), (p p) p"
let skeleton := "by\nhave a : p p := sorry\nsorry"
let expr ← parseSentence expr
Meta.forallTelescope expr $ λ _ body => do
let skeleton' ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := skeleton)
(fileName := ← getFileName) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let tactic := Tactic.evalDraft skeleton'
let newGoals ← runTacticOnMVar tactic target.mvarId!
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) = ["p p", "(p p) p"])
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
[
("draft", test_draft),
] |>.map (λ (name, t) => (name, runTestTermElabM env t))
end Pantograph.Test.Tactic.Assign

View File

@ -0,0 +1,88 @@
import LSpec
import Lean
import Test.Common
open Lean
open Pantograph
namespace Pantograph.Test.Tactic.Congruence
def test_congr_arg_list : TestT Elab.TermElabM Unit := do
let expr := "λ {α} (l1 l2 : List α) (h: l1 = l2) => l1.reverse = l2.reverse"
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let newGoals ← runTacticOnMVar Tactic.evalCongruenceArg target.mvarId!
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
[
(`α, "Sort ?u.30"),
(`a₁, "?α"),
(`a₂, "?α"),
(`f, "?α → List α"),
(`h, "?a₁ = ?a₂"),
(`conduit, "(?f ?a₁ = ?f ?a₂) = (l1.reverse = l2.reverse)"),
])
let f := newGoals.get! 3
let h := newGoals.get! 4
let c := newGoals.get! 5
let results ← Meta.withAssignableSyntheticOpaque do f.apply (← parseSentence "List.reverse")
addTest $ LSpec.check "apply" (results.length = 0)
addTest $ LSpec.check "h" ((← exprToStr $ ← h.getType) = "?a₁ = ?a₂")
addTest $ LSpec.check "conduit" ((← exprToStr $ ← c.getType) = "(?a₁.reverse = ?a₂.reverse) = (l1.reverse = l2.reverse)")
def test_congr_arg : TestT Elab.TermElabM Unit := do
let expr := "λ (n m: Nat) (h: n = m) => n * n = m * m"
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let newGoals ← runTacticOnMVar Tactic.evalCongruenceArg target.mvarId!
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
[
(`α, "Sort ?u.70"),
(`a₁, "?α"),
(`a₂, "?α"),
(`f, "?α → Nat"),
(`h, "?a₁ = ?a₂"),
(`conduit, "(?f ?a₁ = ?f ?a₂) = (n * n = m * m)"),
])
def test_congr_fun : TestT Elab.TermElabM Unit := do
let expr := "λ (n m: Nat) => (n + m) + (n + m) = (n + m) * 2"
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let newGoals ← runTacticOnMVar Tactic.evalCongruenceFun target.mvarId!
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
[
(`α, "Sort ?u.159"),
(`f₁, "?α → Nat"),
(`f₂, "?α → Nat"),
(`h, "?f₁ = ?f₂"),
(`a, "?α"),
(`conduit, "(?f₁ ?a = ?f₂ ?a) = (n + m + (n + m) = (n + m) * 2)"),
])
def test_congr : TestT Elab.TermElabM Unit := do
let expr := "λ (a b: Nat) => a = b"
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let newGoals ← runTacticOnMVar Tactic.evalCongruence target.mvarId!
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
[
(`α, "Sort ?u.10"),
(`f₁, "?α → Nat"),
(`f₂, "?α → Nat"),
(`a₁, "?α"),
(`a₂, "?α"),
(`h₁, "?f₁ = ?f₂"),
(`h₂, "?a₁ = ?a₂"),
(`conduit, "(?f₁ ?a₁ = ?f₂ ?a₂) = (a = b)"),
])
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
[
("congrArg List.reverse", test_congr_arg_list),
("congrArg", test_congr_arg),
("congrFun", test_congr_fun),
("congr", test_congr),
] |>.map (λ (name, t) => (name, runTestTermElabM env t))
end Pantograph.Test.Tactic.Congruence

View File

@ -1,312 +0,0 @@
import Pantograph.Goal
import Test.Common
open Lean
namespace Pantograph.Test.Tactic.Fragment
private def buildGoal (nameType: List (String × String)) (target: String):
Protocol.Goal :=
{
target := { pp? := .some target},
vars := (nameType.map fun x => ({
userName := x.fst,
type? := .some { pp? := .some x.snd },
})).toArray
}
abbrev TestM := TestT $ Elab.TermElabM
example : ∀ (a b c1 c2: Nat), (b + a) + c1 = (b + a) + c2 → (a + b) + c1 = (b + a) + c2 := by
intro a b c1 c2 h
conv =>
lhs
congr
. rw [Nat.add_comm]
. rfl
exact h
def test_conv_simple: TestM Unit := do
let rootTarget ← parseSentence "∀ (a b c1 c2: Nat), (b + a) + c1 = (b + a) + c2 → (a + b) + c1 = (b + a) + c2"
let state0 ← GoalState.create rootTarget
let tactic := "intro a b c1 c2 h"
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
| .success state _ => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check tactic ((← state1.serializeGoals).map (·.devolatilize) =
#[interiorGoal [] "a + b + c1 = b + a + c2"])
let goalConv := state1.goals[0]!
let state2 ← match ← state1.convEnter (state1.get! 0) with
| .success state _ => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check "conv => ..." ((← state2.serializeGoals).map (·.devolatilize) =
#[{ interiorGoal [] "a + b + c1 = b + a + c2" with fragment := .conv }])
let convTactic := "rhs"
let state3R ← match ← state2.tacticOn (goalId := 0) convTactic with
| .success state _ => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check s!" {convTactic} (discard)" ((← state3R.serializeGoals).map (·.devolatilize) =
#[{ interiorGoal [] "b + a + c2" with fragment := .conv }])
let convTactic := "lhs"
let state3L ← match ← state2.tacticOn (goalId := 0) convTactic with
| .success state _ => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check s!" {convTactic}" ((← state3L.serializeGoals).map (·.devolatilize) =
#[{ interiorGoal [] "a + b + c1" with fragment := .conv }])
let convTactic := "congr"
let state4 ← match ← state3L.tacticOn (goalId := 0) convTactic with
| .success state _ => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check s!" {convTactic}" ((← state4.serializeGoals).map (·.devolatilize) =
#[
{ interiorGoal [] "a + b" with fragment := .conv, userName? := .some "a" },
{ interiorGoal [] "c1" with fragment := .conv, userName? := .some "a" }
])
let convTactic := "rw [Nat.add_comm]"
let state5_1 ← match ← state4.tacticOn (goalId := 0) convTactic with
| .success state _ => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check s!" · {convTactic}" ((← state5_1.serializeGoals).map (·.devolatilize) =
#[{ interiorGoal [] "b + a" with fragment := .conv, userName? := .some "a" }])
let convTactic := "rfl"
let state6_1 ← match ← state5_1.tacticOn (goalId := 0) convTactic with
| .success state _ => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check s!" {convTactic}" ((← state6_1.serializeGoals).map (·.devolatilize) =
#[])
let state4_1 ← match state6_1.continue state4 with
| .ok state => pure state
| .error e => do
addTest $ expectationFailure "continue" e
return ()
let convTactic := "rfl"
let state1_1 ← match ← state4_1.tacticOn (goalId := 0) convTactic with
| .success state _ => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check s!" · {convTactic}" ((← state1_1.serializeGoals).map (·.devolatilize) =
#[interiorGoal [] "b + a + c1 = b + a + c2"])
checkEq "(fragments)" state1_1.fragments.size 0
/-
let state1_1 ← match ← state6.fragmentExit goalConv with
| .success state _ => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
-/
let tactic := "exact h"
let stateF ← match ← state1_1.tacticOn (goalId := 0) (tactic := tactic) with
| .success state _ => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check tactic ((← stateF.serializeGoals).map (·.devolatilize) =
#[])
where
h := "b + a + c1 = b + a + c2"
interiorGoal (free: List (String × String)) (target: String) :=
let free := [("a", "Nat"), ("b", "Nat"), ("c1", "Nat"), ("c2", "Nat"), ("h", h)] ++ free
buildGoal free target
example (p : Prop) (x y z : Nat) : p → (p → x = y) → x + z = y + z ∧ p := by
intro hp hi
apply And.intro
conv =>
rhs
arg 1
rw [←hi]
rfl
tactic => exact hp
exact hp
def test_conv_unshielded : TestM Unit := do
let rootTarget ← parseSentence "∀ (p : Prop) (x y z : Nat), p → (p → x = y) → x + z = y + z ∧ p"
let state ← GoalState.create rootTarget
let tactic := "intro p x y z hp hi"
let .success state _ ← state.tacticOn 0 tactic | fail "intro failed"
let tactic := "apply And.intro"
let .success state _ ← state.tacticOn 0 tactic | fail "apply failed"
let .success state _ ← state.convEnter (.prefer state.goals[0]!) | fail "Cannot enter conversion tactic mode"
let .success state _ ← state.tryTactic .unfocus "rhs" | fail "rhs failed"
let tactic := "arg 1"
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize))
#[
{ interiorGoal [] "y" with fragment := .conv },
{ interiorGoal [] "p" with userName? := "right", },
]
let tactic := "rw [←hi]"
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
checkEq s!" {tactic}" state.goals.length 3
let tactic := "rfl"
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize))
#[
interiorGoal [] "p",
{ interiorGoal [] "p" with userName? := "right", },
]
checkEq "(n goals)" state.goals.length 2
checkEq "(fragments)" state.fragments.size 0
let tactic := "exact hp"
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
let tactic := "exact hp"
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
let root? := state.rootExpr?
checkTrue "root" root?.isSome
where
interiorGoal (free: List (String × String)) (target: String) :=
let free := [("p", "Prop"), ("x", "Nat"), ("y", "Nat"), ("z", "Nat"), ("hp", "p"), ("hi", "p → x = y")] ++ free
buildGoal free target
example : ∀ (x y z w : Nat), y = z → x + z = w → x + y = w := by
intro x y z w hyz hxzw
conv =>
lhs
arg 2
rw [hyz]
rfl
exact hxzw
def test_conv_unfinished : TestM Unit := do
let rootTarget ← parseSentence "∀ (x y z w : Nat), y = z → x + z = w → x + y = w"
let state ← GoalState.create rootTarget
let tactic := "intro x y z w hyz hxzw"
let .success state _ ← state.tacticOn 0 tactic | fail "intro failed"
let convParent := state.goals[0]!
let .success state _ ← state.convEnter (.prefer convParent) | fail "Cannot enter conversion tactic mode"
let .success state _ ← state.tryTactic .unfocus "lhs" | fail "rhs failed"
let tactic := "arg 2"
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize))
#[
{ interiorGoal [] "y" with fragment := .conv },
]
let tactic := "rw [hyz]"
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize))
#[
{ interiorGoal [] "z" with fragment := .conv },
]
checkTrue " (fragment)" $ state.fragments.contains state.mainGoal?.get!
checkTrue " (fragment parent)" $ state.fragments.contains convParent
checkTrue " (main goal)" state.mainGoal?.isSome
let tactic := "rfl"
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize))
#[
interiorGoal [] "x + z = w",
]
checkEq "(fragments)" state.fragments.size 0
checkEq s!" {tactic}" state.goals.length 1
let tactic := "exact hxzw"
let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed"
let root? := state.rootExpr?
checkTrue "root" root?.isSome
where
interiorGoal (free: List (String × String)) (target: String) :=
let free := [("x", "Nat"), ("y", "Nat"), ("z", "Nat"), ("w", "Nat"), ("hyz", "y = z"), ("hxzw", "x + z = w")] ++ free
buildGoal free target
example : ∀ (a b c d: Nat), a + b = b + c → b + c = c + d → a + b = c + d := by
intro a b c d h1 h2
calc a + b = b + c := by apply h1
_ = c + d := by apply h2
def test_calc: TestM Unit := do
let rootTarget ← parseSentence "∀ (a b c d: Nat), a + b = b + c → b + c = c + d → a + b = c + d"
let state0 ← GoalState.create rootTarget
let tactic := "intro a b c d h1 h2"
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
| .success state _ => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check tactic ((← state1.serializeGoals).map (·.devolatilize) =
#[interiorGoal [] "a + b = c + d"])
let pred := "a + b = b + c"
let .success state1 _ ← state1.calcEnter state1.mainGoal?.get! | fail "Could not enter calc"
let state2 ← match ← state1.tacticOn 0 pred with
| .success state _ => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check s!"calc {pred} := _" ((← state2.serializeGoals).map (·.devolatilize) =
#[
{ interiorGoal [] "a + b = b + c" with userName? := .some "calc" },
{ interiorGoal [] "b + c = c + d" with fragment := .calc },
])
addTest $ LSpec.test "(2.0 prev rhs)" (state2.calcPrevRhsOf? (state2.get! 0) |>.isNone)
addTest $ LSpec.test "(2.1 prev rhs)" (state2.calcPrevRhsOf? (state2.get! 1) |>.isSome)
let tactic := "apply h1"
let state2m ← match ← state2.tacticOn (goalId := 0) (tactic := tactic) with
| .success state _ => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
let state3 ← match state2m.continue state2 with
| .ok state => pure state
| .error e => do
addTest $ expectationFailure "continue" e
return ()
let pred := "_ = c + d"
let state4 ← match ← state3.tacticOn 0 pred with
| .success state _ => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check s!"calc {pred} := _" ((← state4.serializeGoals).map (·.devolatilize) =
#[
{ interiorGoal [] "b + c = c + d" with userName? := .some "calc" },
])
addTest $ LSpec.test "(4.0 prev rhs)" (state4.calcPrevRhsOf? (state4.get! 0) |>.isNone)
let tactic := "apply h2"
let state4m ← match ← state4.tacticOn (goalId := 0) (tactic := tactic) with
| .success state _ => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
checkEq "(fragments)" state4m.fragments.size 0
addTest $ LSpec.test "(4m root)" state4m.rootExpr?.isSome
where
interiorGoal (free: List (String × String)) (target: String) :=
let free := [("a", "Nat"), ("b", "Nat"), ("c", "Nat"), ("d", "Nat"),
("h1", "a + b = b + c"), ("h2", "b + c = c + d")] ++ free
buildGoal free target
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
[
("conv simple", test_conv_simple),
("conv unshielded", test_conv_unshielded),
("conv unfinished", test_conv_unfinished),
("calc", test_calc),
] |>.map (λ (name, t) => (name, runTestTermElabM env t))
end Pantograph.Test.Tactic.Fragment

View File

@ -0,0 +1,113 @@
import LSpec
import Lean
import Test.Common
open Lean
open Pantograph
namespace Pantograph.Test.Tactic.MotivatedApply
def test_type_extract : TestT Elab.TermElabM Unit := do
let recursor ← parseSentence "@Nat.brecOn"
let recursorType ← Meta.inferType recursor
addTest $ LSpec.check "recursorType" ("{motive : Nat → Sort ?u.1} → (t : Nat) → ((t : Nat) → Nat.below t → motive t) → motive t" =
(← exprToStr recursorType))
let info ← match Tactic.getRecursorInformation recursorType with
| .some info => pure info
| .none => throwError "Failed to extract recursor info"
addTest $ LSpec.check "iMotive" (info.iMotive = 2)
let motiveType := info.getMotiveType
addTest $ LSpec.check "motiveType" ("Nat → Sort ?u.1" =
(← exprToStr motiveType))
def test_nat_brec_on : TestT Elab.TermElabM Unit := do
let expr := "λ (n t: Nat) => n + 0 = n"
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "@Nat.brecOn")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let tactic := Tactic.evalMotivatedApply recursor
let newGoals ← runTacticOnMVar tactic target.mvarId!
let test := LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
[
"Nat → Prop",
"Nat",
"∀ (t : Nat), Nat.below t → ?motive t",
"?motive ?m.67 = (n + 0 = n)",
])
addTest test
def test_list_brec_on : TestT Elab.TermElabM Unit := do
let expr := "λ {α : Type} (l: List α) => l ++ [] = [] ++ l"
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "@List.brecOn")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let tactic := Tactic.evalMotivatedApply recursor
let newGoals ← runTacticOnMVar tactic target.mvarId!
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
[
"Type ?u.90",
"List ?m.92 → Prop",
"List ?m.92",
"∀ (t : List ?m.92), List.below t → ?motive t",
"?motive ?m.94 = (l ++ [] = [] ++ l)",
])
def test_partial_motive_instantiation : TestT Elab.TermElabM Unit := do
let expr := "λ (n t: Nat) => n + 0 = n"
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "@Nat.brecOn")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let tactic := Tactic.evalMotivatedApply recursor
let newGoals ← runTacticOnMVar tactic target.mvarId!
let majorId := 67
addTest $ (LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
[
"Nat → Prop",
"Nat",
"∀ (t : Nat), Nat.below t → ?motive t",
s!"?motive ?m.{majorId} = (n + 0 = n)",
]))
let [motive, major, step, conduit] := newGoals | panic! "Incorrect goal number"
addTest $ (LSpec.check "goal name" (major.name.toString = s!"_uniq.{majorId}"))
-- Assign motive to `λ x => x + _`
let motive_assign ← parseSentence "λ (x: Nat) => @Nat.add x + 0 = _"
motive.assign motive_assign
addTest $ ← conduit.withContext do
let t := toString (← Meta.ppExpr $ ← conduit.getType)
return LSpec.check "conduit" (t = s!"(?m.{majorId}.add + 0 = ?m.138 ?m.{majorId}) = (n + 0 = n)")
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
[
("type_extract", test_type_extract),
("Nat.brecOn", test_nat_brec_on),
("List.brecOn", test_list_brec_on),
("Nat.brecOn partial motive instantiation", test_partial_motive_instantiation),
] |>.map (λ (name, t) => (name, runTestTermElabM env t))
end Pantograph.Test.Tactic.MotivatedApply

View File

@ -0,0 +1,72 @@
import LSpec
import Lean
import Test.Common
open Lean
open Pantograph
namespace Pantograph.Test.Tactic.NoConfuse
def test_nat : TestT Elab.TermElabM Unit := do
let expr := "λ (n: Nat) (h: 0 = n + 1) => False"
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "h")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let tactic := Tactic.evalNoConfuse recursor
let newGoals ← runTacticOnMVar tactic target.mvarId!
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) = [])
def test_nat_fail : TestT Elab.TermElabM Unit := do
let expr := "λ (n: Nat) (h: n = n) => False"
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "h")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
try
let tactic := Tactic.evalNoConfuse recursor
let _ ← runTacticOnMVar tactic target.mvarId!
addTest $ assertUnreachable "Tactic should fail"
catch _ =>
addTest $ LSpec.check "Tactic should fail" true
def test_list : TestT Elab.TermElabM Unit := do
let expr := "λ (l: List Nat) (h: [] = 1 :: l) => False"
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "h")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let tactic := Tactic.evalNoConfuse recursor
let newGoals ← runTacticOnMVar tactic target.mvarId!
addTest $ LSpec.check "goals"
((← newGoals.mapM (λ g => do exprToStr (← g.getType))) = [])
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
[
("Nat", test_nat),
("Nat fail", test_nat_fail),
("List", test_list),
] |>.map (λ (name, t) => (name, runTestTermElabM env t))
end Pantograph.Test.Tactic.NoConfuse

View File

@ -15,7 +15,7 @@ def test_define : TestT Elab.TermElabM Unit := do
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
(catName := `term) (catName := `term)
(input := "Or.inl h") (input := "Or.inl h")
(fileName := ← getFileName) with (fileName := filename) with
| .ok syn => pure syn | .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}" | .error error => throwError "Failed to parse: {error}"
-- Apply the tactic -- Apply the tactic
@ -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 ()
@ -87,7 +87,7 @@ def test_define_proof : TestT Elab.TermElabM Unit := do
{ userName := "q", type? := .some { pp? := .some "Prop" } }, { userName := "q", type? := .some { pp? := .some "Prop" } },
{ userName := "h", type? := .some { pp? := .some "p" } }, { userName := "h", type? := .some { pp? := .some "p" } },
{ userName := "y", { userName := "y",
type? := .some { pp? := .some "p ?m.19" }, type? := .some { pp? := .some "p ?m.25" },
value? := .some { pp? := .some "Or.inl h" }, value? := .some { pp? := .some "Or.inl h" },
} }
] ]
@ -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,7 +185,7 @@ 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 ()
@ -196,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 ()
@ -212,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 ()
@ -237,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 ()
@ -270,7 +274,7 @@ def test_let (specialized: Bool): TestT Elab.TermElabM Unit := do
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 ()

View File

@ -1,23 +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 #["`exact?` could not close the goal. Try `apply?` to see partial suggestions."]
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

View File

@ -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}"

View File

@ -4,17 +4,17 @@
<svg <svg
width="256" width="256"
height="256" height="256"
viewBox="0 0 67.733332 67.733333" viewBox="0 0 55.900957 55.900957"
version="1.1" version="1.1"
id="svg1" id="svg21534"
xml:space="preserve"
inkscape:version="1.2.2 (b0a8486541, 2022-12-01)"
sodipodi:docname="icon.svg" sodipodi:docname="icon.svg"
inkscape:version="1.3.2 (091e20ef0f, 2023-11-25, custom)"
xmlns:inkscape="http://www.inkscape.org/namespaces/inkscape" xmlns:inkscape="http://www.inkscape.org/namespaces/inkscape"
xmlns:sodipodi="http://sodipodi.sourceforge.net/DTD/sodipodi-0.dtd" xmlns:sodipodi="http://sodipodi.sourceforge.net/DTD/sodipodi-0.dtd"
xmlns="http://www.w3.org/2000/svg" xmlns="http://www.w3.org/2000/svg"
xmlns:svg="http://www.w3.org/2000/svg"> xmlns:svg="http://www.w3.org/2000/svg"><sodipodi:namedview
<sodipodi:namedview id="namedview21536"
id="namedview1"
pagecolor="#ffffff" pagecolor="#ffffff"
bordercolor="#111111" bordercolor="#111111"
borderopacity="1" borderopacity="1"
@ -23,135 +23,51 @@
inkscape:pagecheckerboard="1" inkscape:pagecheckerboard="1"
inkscape:deskcolor="#d1d1d1" inkscape:deskcolor="#d1d1d1"
inkscape:document-units="px" inkscape:document-units="px"
showguides="true" showgrid="true"
inkscape:zoom="5.1882633" inkscape:zoom="5.1754899"
inkscape:cx="81.819286" inkscape:cx="158.82554"
inkscape:cy="132.22151" inkscape:cy="91.682142"
inkscape:window-width="3774" inkscape:window-width="3777"
inkscape:window-height="2126" inkscape:window-height="2093"
inkscape:window-x="0" inkscape:window-x="0"
inkscape:window-y="0" inkscape:window-y="0"
inkscape:window-maximized="1" inkscape:window-maximized="1"
inkscape:current-layer="layer2"> inkscape:current-layer="layer1"><inkscape:grid
<sodipodi:guide type="xygrid"
position="33.866666,69.8437" id="grid23833"
orientation="-1,0" spacingx="3.4938098"
id="guide1" spacingy="3.4938098"
inkscape:locked="false" empspacing="4" /></sodipodi:namedview><defs
inkscape:label="" id="defs21531" /><g
inkscape:color="rgb(0,134,229)" /> inkscape:label="Layer 1"
<sodipodi:guide
position="-27.673679,33.866666"
orientation="0,1"
id="guide2"
inkscape:locked="false"
inkscape:label=""
inkscape:color="rgb(0,134,229)" />
<sodipodi:guide
position="16.933333,29.94582"
orientation="-1,0"
id="guide3"
inkscape:locked="false"
inkscape:label=""
inkscape:color="rgb(0,134,229)" />
<sodipodi:guide
position="50.799999,37.44627"
orientation="-1,0"
id="guide4"
inkscape:locked="false"
inkscape:label=""
inkscape:color="rgb(0,134,229)" />
<sodipodi:guide
position="31.336956,16.933333"
orientation="0,1"
id="guide5"
inkscape:locked="false"
inkscape:label=""
inkscape:color="rgb(0,134,229)" />
<sodipodi:guide
position="24.528038,25.4"
orientation="0,1"
id="guide6"
inkscape:locked="false"
inkscape:label=""
inkscape:color="rgb(0,134,229)" />
<sodipodi:guide
position="33.866666,50.799999"
orientation="0,1"
id="guide7"
inkscape:locked="false"
inkscape:label=""
inkscape:color="rgb(0,134,229)" />
<sodipodi:guide
position="32.770414,55.033333"
orientation="0,1"
id="guide8"
inkscape:locked="false"
inkscape:label=""
inkscape:color="rgb(0,134,229)" />
<sodipodi:guide
position="25.347689,33.866666"
orientation="1,0"
id="guide9"
inkscape:locked="false" />
<sodipodi:guide
position="25.347689,42.333333"
orientation="0,1"
id="guide10"
inkscape:locked="false"
inkscape:label=""
inkscape:color="rgb(0,134,229)" />
</sodipodi:namedview>
<defs
id="defs1" />
<g
inkscape:groupmode="layer" inkscape:groupmode="layer"
id="layer4" id="layer1"><rect
inkscape:label="bg" /> style="fill:#3e3e3e;fill-opacity:1;fill-rule:evenodd;stroke:none;stroke-width:1.78013;stroke-miterlimit:3.4;stroke-dasharray:none"
<g id="rect26805"
inkscape:groupmode="layer" width="11.502316"
id="layer1" height="2.2512667"
inkscape:label="Circle"> x="33.344425"
<path y="7.6690259"
id="path1" ry="0.28140834"
style="fill:#666b98;fill-rule:evenodd;stroke:#ffffff;stroke-width:0.0191989;stroke-miterlimit:3.4;fill-opacity:1" rx="0.47926313" /><path
d="M 33.866666 0.009818522 A 33.857067 33.857067 0 0 0 0.009818522 33.866666 A 33.857067 33.857067 0 0 0 33.866666 67.723514 A 33.857067 33.857067 0 0 0 67.723514 33.866666 A 33.857067 33.857067 0 0 0 33.866666 0.009818522 z M 33.866666 4.2416015 A 29.624933 29.624933 0 0 1 63.491731 33.866666 A 29.624933 29.624933 0 0 1 33.866666 63.491731 A 29.624933 29.624933 0 0 1 4.2416015 33.866666 A 29.624933 29.624933 0 0 1 33.866666 4.2416015 z " /> style="fill:#3e3e3e;stroke:none;stroke-width:0.218363px;stroke-linecap:butt;stroke-linejoin:miter;stroke-opacity:1;fill-opacity:1"
</g> d="m 35.764667,9.9513013 c 0,0 -26.8581417,13.7987337 -28.0863506,14.9501437 -1.250042,1.171878 3.2347846,3.945325 3.2347846,3.945325 l 21.34979,14.934062 6.624567,0.453105 -27.599216,-17.304358 c 0,0 -0.603209,-0.08927 -0.600411,-0.762283 0.0028,-0.673015 27.32022,-16.4227356 27.32022,-16.4227356 z"
<g id="path27381"
inkscape:groupmode="layer" sodipodi:nodetypes="csccccscc" /><path
id="layer2" style="fill:#3e3e3e;stroke:none;stroke-width:0.218363px;stroke-linecap:butt;stroke-linejoin:miter;stroke-opacity:1;fill-opacity:1"
inkscape:label="Pantograph-Core"> d="M 10.97848,26.985751 40.537772,9.7943227 41.921795,9.7005084 11.210626,27.421377 Z"
<rect id="path27479" /><path
style="fill:#666b98;fill-opacity:1;fill-rule:evenodd;stroke:#ffffff;stroke-width:0.01905;stroke-miterlimit:3.4" style="fill:#3e3e3e;stroke:none;stroke-width:0.218363px;stroke-linecap:butt;stroke-linejoin:miter;stroke-opacity:1;fill-opacity:1"
id="rect8" d="m 7.0509847,25.522367 c -0.8266141,1.386819 -2.4011783,4.48805 -2.4706357,4.90223 -0.069458,0.414182 0.4434324,0.513474 0.8491061,0.757041 C 5.835129,31.425204 19.33424,43.917182 19.33424,43.917182 l 0.324562,-0.539228 c 0,0 -14.2055729,-12.369493 -14.0644435,-12.868167 0.1411292,-0.498672 3.544896,-3.777392 3.544896,-3.777392 L 7.4596884,25.117508 Z"
width="16.942858" id="path27481" /><rect
height="4.2257233" style="fill:#3e3e3e;fill-opacity:1;fill-rule:evenodd;stroke:none;stroke-width:2.11692;stroke-miterlimit:3.4;stroke-dasharray:none;stroke-opacity:1"
x="33.866665" id="rect27483"
y="12.7" width="36.38942"
rx="0.58070719" height="3.6217353"
ry="0.34097314" /> x="13.953447"
<rect y="43.009739"
style="fill:#666b98;fill-opacity:1;fill-rule:evenodd;stroke:#ffffff;stroke-width:0.01905;stroke-miterlimit:3.4" rx="0.43672624"
id="rect1" ry="0.43672624" /><path
width="33.885715" style="fill:none;stroke:#000000;stroke-width:0.218363px;stroke-linecap:butt;stroke-linejoin:miter;stroke-opacity:1"
height="8.4211359" d="M -2.1119274,7.666599 H 64.179192"
x="16.933332" id="path27487" /></g></svg>
y="42.333332"
rx="0.58070719"
ry="0.34097314" />
<path
style="fill:#666b98;fill-opacity:1;stroke:none;stroke-width:0.264583px;stroke-linecap:butt;stroke-linejoin:miter;stroke-opacity:1"
d="m 42.338095,16.925724 -16.990406,8.474275 13.121218,16.923808 -4.602241,0.0095 -4.254289,0.0015 -8.564029,-16.934789 17.310554,-8.464751 z"
id="path10"
sodipodi:nodetypes="cccccccc" />
<path
style="fill:none;stroke:#666b98;stroke-width:0.264583px;stroke-linecap:butt;stroke-linejoin:miter;stroke-opacity:1"
d="M 46.53445,16.925724 26.018901,26.73418"
id="path11" />
<path
style="fill:none;stroke:#666b98;stroke-width:0.264583px;stroke-linecap:butt;stroke-linejoin:miter;stroke-opacity:1"
d="m 21.048348,25.399999 4.352167,16.934808"
id="path12"
sodipodi:nodetypes="cc" />
</g>
</svg>

Before

Width:  |  Height:  |  Size: 5.2 KiB

After

Width:  |  Height:  |  Size: 3.5 KiB

View File

@ -1,60 +0,0 @@
# Design Rationale
A great problem in machine learning is to use ML agents to automatically prove
mathematical theorems. This sort of proof necessarily involves *search*.
Compatibility for search is the main reason for creating Pantograph. The Lean 4
LSP interface is not conducive to search. Pantograph is designed with this in
mind. It emphasizes the difference between 3 views of a proof:
- **Presentation View**: The view of a written, polished proof. e.g. Mathlib and
math papers are almost always written in this form.
- **Search View**: The view of a proof exploration trajectory. This is not
explicitly supported by Lean LSP.
- **Kernel View**: The proof viewed as a set of metavariables.
Pantograph enables proof agents to operate on the search view.
## Name
The name Pantograph is a pun. It means two things
- A pantograph is an instrument for copying down writing. As an agent explores
the vast proof search space, Pantograph records the current state to ensure
the proof is sound.
- A pantograph is also an equipment for an electric train. It supplies power to
a locomotive. In comparison the (relatively) simple Pantograph software powers
theorem proving projects.
## Caveats and Limitations
Pantograph does not exactly mimic Lean LSP's behaviour. That would not grant the
flexibility it offers. To support tree search means Pantograph has to act
differently from Lean in some times, but never at the sacrifice of soundness.
- When Lean LSP says "don't know how to synthesize placeholder", this indicates
the human operator needs to manually move the cursor to the placeholder and
type in the correct expression. This error therefore should not halt the proof
process, and the placeholder should be turned into a goal.
- When Lean LSP says "unresolved goals", that means a proof cannot finish where
it is supposed to finish at the end of a `by` block. Pantograph will raise the
error in this case, since it indicates the termination of a proof search branch.
Pantograph cannot perform things that are inherently constrained by Lean. These
include:
- If a tactic loses track of metavariables, it will not be caught until the end
of the proof search. This is a bug in the tactic itself.
- Although a timeout feature exists in Pantograph, it relies on the coöperative
multitasking from the tactic implementation. There is nothing preventing a
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.
`def mystery : Nat := :=`) due to Lean's parsing system. This question is also
not well-defined.
## References
* [Pantograph Paper](https://arxiv.org/abs/2410.16429)

View File

@ -11,10 +11,8 @@ See `Pantograph/Protocol.lean` for a description of the parameters and return va
* `env.inspect {"name": <name>, "value": <bool>}`: Show the type and package of a * `env.inspect {"name": <name>, "value": <bool>}`: Show the type and package of a
given symbol; If value flag is set, the value is printed or hidden. By default given symbol; If value flag is set, the value is printed or hidden. By default
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 }`, `env.load { path }`: Save/Load the current environment
current environment to/from a file to/from a file
* `env.module_read { "module": <name> }`: Reads a list of symbols from a module
* `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`
@ -22,46 +20,32 @@ See `Pantograph/Protocol.lean` for a description of the parameters and return va
automatic mode (flag: `"automaticMode"`). By default it is turned on, with automatic mode (flag: `"automaticMode"`). By default it is turned on, with
all goals automatically resuming. This makes Pantograph act like a gym, all goals automatically resuming. This makes Pantograph act like a gym,
with no resumption necessary to manage your goals. with no resumption necessary to manage your goals.
Set `timeout` to a non-zero number to specify timeout (milliseconds) for all `CoreM`
operations.
* `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.
* `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.
- `{ "goals": <names> }`: Resume the given goals - `{ "goals": <names> }`: Resume the given goals
* `goal.remove {"stateIds": [<id>]}"`: Drop the goal states specified in the list * `goal.remove {"stateIds": [<id>]}"`: Drop the goal states specified in the list
* `goal.print {"stateId": <id>}"`: Print a goal state * `goal.print {"stateId": <id>}"`: Print a goal state
* `goal.save { "id": <id>, "path": <fileName> }`, `goal.load { "path": <fileName> }`: * `goal.save{ id, path }`, `goal.load { path }`: Save/Load a goal state to/from a
Save/Load a goal state to/from a file. The environment is not carried with the file. The environment is not carried with the state. The user is responsible
state. The user is responsible to ensure the sender/receiver instances share to ensure the sender/receiver instances share the same environment.
the same environment. * `frontend.process { ["fileName": <fileName>",] ["file": <str>], invocations:
* `frontend.process { ["fileName": <fileName>,] ["file": <str>], readHeader: <bool>, inheritEnv: <bool>, invocations: <bool>, sorrys: <bool> }`: Executes the Lean frontend on a file, collecting
<string>, sorrys: <bool>, typeErrorsAsGoals: <bool>, newConstants: <bool> }`: either the tactic invocations (`"invocations": true`) or the sorrys into goal
Executes the Lean frontend on a file, collecting the tactic invocations states (`"sorrys": true`)
(`"invocations": output-path`), the sorrys and type errors into goal states
(`"sorrys": true`), and new constants (`"newConstants": true`). In the case of
`sorrys`, this command additionally outputs the position of each captured
`sorry`. Conditionally inherit the environment from executing the file.
Warning: Behaviour is unstable in case of multiple `sorry`s. Use the draft
tactic if possible.
## Errors ## Errors

View File

@ -5,11 +5,11 @@
"nixpkgs-lib": "nixpkgs-lib" "nixpkgs-lib": "nixpkgs-lib"
}, },
"locked": { "locked": {
"lastModified": 1749398372, "lastModified": 1730504689,
"narHash": "sha256-tYBdgS56eXYaWVW3fsnPQ/nFlgWi/Z2Ymhyu21zVM98=", "narHash": "sha256-hgmguH29K2fvs9szpq2r3pz2/8cJd2LPS+b4tfNFCwE=",
"owner": "hercules-ci", "owner": "hercules-ci",
"repo": "flake-parts", "repo": "flake-parts",
"rev": "9305fe4e5c2a6fcf5ba6a3ff155720fbe4076569", "rev": "506278e768c2a08bec68eb62932193e341f55c90",
"type": "github" "type": "github"
}, },
"original": { "original": {
@ -42,11 +42,11 @@
"nixpkgs": "nixpkgs" "nixpkgs": "nixpkgs"
}, },
"locked": { "locked": {
"lastModified": 1750369222, "lastModified": 1731711316,
"narHash": "sha256-KFFTVbciXUaHgeGN1yiaUtY88OLGU0gElXx5SfICDKg=", "narHash": "sha256-s5u+A2/Ea9gPveB5wwVM5dWW0NST6kamDsTeovGuLEs=",
"owner": "lenianiva", "owner": "lenianiva",
"repo": "lean4-nix", "repo": "lean4-nix",
"rev": "015ecd25206734d582a1b15dd11eb10be35ca555", "rev": "136fc6057c48de970579e960b62421e9c295b67d",
"type": "github" "type": "github"
}, },
"original": { "original": {
@ -55,35 +55,49 @@
"type": "github" "type": "github"
} }
}, },
"lspec": {
"flake": false,
"locked": {
"lastModified": 1728279187,
"narHash": "sha256-ZMqbvCqR/gHXRuIkuo7b0Yp9N1vOQR7xnrcy/SeIBoQ=",
"owner": "argumentcomputer",
"repo": "LSpec",
"rev": "504a8cecf8da601b9466ac727aebb6b511aae4ab",
"type": "github"
},
"original": {
"owner": "argumentcomputer",
"ref": "504a8cecf8da601b9466ac727aebb6b511aae4ab",
"repo": "LSpec",
"type": "github"
}
},
"nixpkgs": { "nixpkgs": {
"locked": { "locked": {
"lastModified": 1743095683, "lastModified": 1728500571,
"narHash": "sha256-gWd4urRoLRe8GLVC/3rYRae1h+xfQzt09xOfb0PaHSk=", "narHash": "sha256-dOymOQ3AfNI4Z337yEwHGohrVQb4yPODCW9MDUyAc4w=",
"owner": "nixos", "owner": "nixos",
"repo": "nixpkgs", "repo": "nixpkgs",
"rev": "5e5402ecbcb27af32284d4a62553c019a3a49ea6", "rev": "d51c28603def282a24fa034bcb007e2bcb5b5dd0",
"type": "github" "type": "github"
}, },
"original": { "original": {
"owner": "nixos", "owner": "nixos",
"ref": "nixos-unstable", "ref": "nixos-24.05",
"repo": "nixpkgs", "repo": "nixpkgs",
"type": "github" "type": "github"
} }
}, },
"nixpkgs-lib": { "nixpkgs-lib": {
"locked": { "locked": {
"lastModified": 1748740939, "lastModified": 1730504152,
"narHash": "sha256-rQaysilft1aVMwF14xIdGS3sj1yHlI6oKQNBRTF40cc=", "narHash": "sha256-lXvH/vOfb4aGYyvFmZK/HlsNsr/0CVWlwYvo2rxJk3s=",
"owner": "nix-community", "type": "tarball",
"repo": "nixpkgs.lib", "url": "https://github.com/NixOS/nixpkgs/archive/cc2f28000298e1269cea6612cd06ec9979dd5d7f.tar.gz"
"rev": "656a64127e9d791a334452c6b6606d17539476e2",
"type": "github"
}, },
"original": { "original": {
"owner": "nix-community", "type": "tarball",
"repo": "nixpkgs.lib", "url": "https://github.com/NixOS/nixpkgs/archive/cc2f28000298e1269cea6612cd06ec9979dd5d7f.tar.gz"
"type": "github"
} }
}, },
"nixpkgs-lib_2": { "nixpkgs-lib_2": {
@ -100,16 +114,16 @@
}, },
"nixpkgs_2": { "nixpkgs_2": {
"locked": { "locked": {
"lastModified": 1751048012, "lastModified": 1731386116,
"narHash": "sha256-MYbotu4UjWpTsq01wglhN5xDRfZYLFtNk7SBY0BcjkU=", "narHash": "sha256-lKA770aUmjPHdTaJWnP3yQ9OI1TigenUqVC3wweqZuI=",
"owner": "nixos", "owner": "nixos",
"repo": "nixpkgs", "repo": "nixpkgs",
"rev": "a684c58d46ebbede49f280b653b9e56100aa3877", "rev": "689fed12a013f56d4c4d3f612489634267d86529",
"type": "github" "type": "github"
}, },
"original": { "original": {
"owner": "nixos", "owner": "nixos",
"ref": "nixos-24.11", "ref": "nixos-24.05",
"repo": "nixpkgs", "repo": "nixpkgs",
"type": "github" "type": "github"
} }
@ -118,6 +132,7 @@
"inputs": { "inputs": {
"flake-parts": "flake-parts", "flake-parts": "flake-parts",
"lean4-nix": "lean4-nix", "lean4-nix": "lean4-nix",
"lspec": "lspec",
"nixpkgs": "nixpkgs_2" "nixpkgs": "nixpkgs_2"
} }
} }

View File

@ -2,9 +2,13 @@
description = "Pantograph"; description = "Pantograph";
inputs = { inputs = {
nixpkgs.url = "github:nixos/nixpkgs/nixos-24.11"; nixpkgs.url = "github:nixos/nixpkgs/nixos-24.05";
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";
lspec = {
url = "github:argumentcomputer/LSpec?ref=504a8cecf8da601b9466ac727aebb6b511aae4ab";
flake = false;
};
}; };
outputs = inputs @ { outputs = inputs @ {
@ -12,77 +16,46 @@
nixpkgs, nixpkgs,
flake-parts, flake-parts,
lean4-nix, lean4-nix,
lspec,
... ...
}: } : flake-parts.lib.mkFlake { inherit inputs; } {
flake-parts.lib.mkFlake {inherit inputs;} {
flake = { flake = {
}; };
systems = [ systems = [
"aarch64-linux"
"aarch64-darwin"
"x86_64-linux" "x86_64-linux"
"x86_64-darwin" "x86_64-darwin"
]; ];
perSystem = { perSystem = { system, pkgs, ... }: let
system,
pkgs,
...
}: let
pkgs = import nixpkgs { pkgs = import nixpkgs {
inherit system; inherit system;
overlays = [ (lean4-nix.readToolchainFile ./lean-toolchain) ]; overlays = [ (lean4-nix.readToolchainFile ./lean-toolchain) ];
}; };
manifest = pkgs.lib.importJSON ./lake-manifest.json;
manifest-lspec = builtins.head manifest.packages;
lspecLib = pkgs.lean.buildLeanPackage { lspecLib = pkgs.lean.buildLeanPackage {
name = "LSpec"; name = "LSpec";
roots = ["LSpec"]; roots = [ "Main" "LSpec" ];
src = builtins.fetchGit {inherit (manifest-lspec) url rev;}; src = "${lspec}";
};
inherit (pkgs.lib.fileset) unions toSource fileFilter;
src = ./.;
set-project = unions [
./Pantograph.lean
(fileFilter (file: file.hasExt "lean") ./Pantograph)
];
set-repl = unions [
./Main.lean
./Repl.lean
];
set-test = unions [
(fileFilter (file: file.hasExt "lean") ./Test)
];
src-project = toSource {
root = src;
fileset = unions [
set-project
];
};
src-repl = toSource {
root = src;
fileset = unions [
set-project
set-repl
];
};
src-test = toSource {
root = src;
fileset = unions [
set-project
set-repl
set-test
];
}; };
project = pkgs.lean.buildLeanPackage { project = pkgs.lean.buildLeanPackage {
name = "Pantograph"; name = "Pantograph";
roots = [ "Pantograph" ]; roots = [ "Pantograph" ];
src = src-project; src = pkgs.lib.cleanSource (pkgs.lib.cleanSourceWith {
src = ./.;
filter = path: type:
!(pkgs.lib.hasInfix "/Test/" path) &&
!(pkgs.lib.hasSuffix ".md" path) &&
!(pkgs.lib.hasSuffix "Repl.lean" path);
});
}; };
repl = pkgs.lean.buildLeanPackage { repl = pkgs.lean.buildLeanPackage {
name = "Repl"; name = "Repl";
roots = [ "Main" "Repl" ]; roots = [ "Main" "Repl" ];
deps = [ project ]; deps = [ project ];
src = src-repl; src = pkgs.lib.cleanSource (pkgs.lib.cleanSourceWith {
src = ./.;
filter = path: type:
!(pkgs.lib.hasInfix "/Test/" path) &&
!(pkgs.lib.hasSuffix ".md" path);
});
}; };
test = pkgs.lean.buildLeanPackage { test = pkgs.lean.buildLeanPackage {
name = "Test"; name = "Test";
@ -91,7 +64,11 @@
# Environment`) and thats where `lakefile.lean` resides. # Environment`) and thats where `lakefile.lean` resides.
roots = [ "Test.Main" ]; roots = [ "Test.Main" ];
deps = [ lspecLib repl ]; deps = [ lspecLib repl ];
src = src-test; src = pkgs.lib.cleanSource (pkgs.lib.cleanSourceWith {
src = ./.;
filter = path: type:
!(pkgs.lib.hasInfix "Pantograph" path);
});
}; };
in rec { in rec {
packages = { packages = {
@ -105,15 +82,13 @@
leanPkgs = pkgs.lean; leanPkgs = pkgs.lean;
}; };
checks = { checks = {
test = test = pkgs.runCommand "test" {
pkgs.runCommand "test" {
buildInputs = [ test.executable pkgs.lean.lean-all ]; buildInputs = [ test.executable pkgs.lean.lean-all ];
} '' } ''
#export LEAN_SRC_PATH="${./.}" #export LEAN_SRC_PATH="${./.}"
${test.executable}/bin/test > $out ${test.executable}/bin/test > $out
''; '';
}; };
formatter = pkgs.alejandra;
devShells.default = pkgs.mkShell { devShells.default = pkgs.mkShell {
buildInputs = [ pkgs.lean.lean-all pkgs.lean.lean ]; buildInputs = [ pkgs.lean.lean-all pkgs.lean.lean ];
}; };

View File

@ -1,15 +1,15 @@
{"version": "1.1.0", {"version": "1.1.0",
"packagesDir": ".lake/packages", "packagesDir": ".lake/packages",
"packages": "packages":
[{"url": "https://github.com/argumentcomputer/LSpec.git", [{"url": "https://github.com/lenianiva/LSpec.git",
"type": "git", "type": "git",
"subDir": null, "subDir": null,
"scope": "", "scope": "",
"rev": "a6652a48b5c67b0d8dd3930fad6390a97d127e8d", "rev": "c492cecd0bc473e2f9c8b94d545d02cc0056034f",
"name": "LSpec", "name": "LSpec",
"manifestFile": "lake-manifest.json", "manifestFile": "lake-manifest.json",
"inputRev": "a6652a48b5c67b0d8dd3930fad6390a97d127e8d", "inputRev": "c492cecd0bc473e2f9c8b94d545d02cc0056034f",
"inherited": false, "inherited": false,
"configFile": "lakefile.toml"}], "configFile": "lakefile.lean"}],
"name": "pantograph", "name": "pantograph",
"lakeDir": ".lake"} "lakeDir": ".lake"}

View File

@ -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,15 +17,8 @@ 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/lenianiva/LSpec.git" @ "c492cecd0bc473e2f9c8b94d545d02cc0056034f"
lean_lib Test { lean_lib Test {
} }
@[test_driver] @[test_driver]

View File

@ -1 +1 @@
leanprover/lean4:v4.21.0 leanprover/lean4:v4.12.0