2023-05-12 16:12:21 -07:00
|
|
|
import Pantograph.Commands
|
2023-08-13 21:19:06 -07:00
|
|
|
import Pantograph.Serial
|
2023-05-12 16:12:21 -07:00
|
|
|
import Pantograph.Symbols
|
2023-08-14 21:43:40 -07:00
|
|
|
import Pantograph.Tactic
|
2023-08-13 21:19:06 -07:00
|
|
|
|
|
|
|
namespace Pantograph
|
|
|
|
|
|
|
|
structure Context where
|
2023-08-14 17:07:53 -07:00
|
|
|
imports: List String
|
2023-08-13 21:19:06 -07:00
|
|
|
|
|
|
|
/-- Stores state of the REPL -/
|
|
|
|
structure State where
|
2023-08-14 17:07:53 -07:00
|
|
|
options: Commands.Options := {}
|
2023-08-13 21:19:06 -07:00
|
|
|
--environments: Array Lean.Environment := #[]
|
|
|
|
proofTrees: Array ProofTree := #[]
|
|
|
|
|
|
|
|
-- State monad
|
|
|
|
abbrev Subroutine := ReaderT Context (StateT State Lean.Elab.TermElabM)
|
|
|
|
|
|
|
|
/-- Parse a command either in `{ "cmd": ..., "payload": ... }` form or `cmd { ... }` form. -/
|
|
|
|
def parse_command (s: String): Except String Commands.Command := do
|
|
|
|
let s := s.trim
|
|
|
|
match s.get? 0 with
|
|
|
|
| .some '{' => -- Parse in Json mode
|
|
|
|
Lean.fromJson? (← Lean.Json.parse s)
|
|
|
|
| .some _ => -- Parse in line mode
|
|
|
|
let offset := s.posOf ' ' |> s.offsetOfPos
|
|
|
|
if offset = s.length then
|
|
|
|
return { cmd := s.take offset, payload := Lean.Json.null }
|
|
|
|
else
|
|
|
|
let payload ← s.drop offset |> Lean.Json.parse
|
|
|
|
return { cmd := s.take offset, payload := payload }
|
|
|
|
| .none => throw "Command is empty"
|
|
|
|
|
|
|
|
def execute (command: Commands.Command): Subroutine Lean.Json := do
|
|
|
|
match command.cmd with
|
2023-08-14 17:07:53 -07:00
|
|
|
| "options.set" =>
|
|
|
|
match Lean.fromJson? command.payload with
|
|
|
|
| .ok args => options_set args
|
|
|
|
| .error x => return errorJson x
|
|
|
|
| "options.print" =>
|
2023-08-13 21:19:06 -07:00
|
|
|
match Lean.fromJson? command.payload with
|
2023-08-14 17:07:53 -07:00
|
|
|
| .ok args => options_print args
|
2023-08-13 21:19:06 -07:00
|
|
|
| .error x => return errorJson x
|
|
|
|
| "catalog" =>
|
|
|
|
match Lean.fromJson? command.payload with
|
|
|
|
| .ok args => catalog args
|
|
|
|
| .error x => return errorJson x
|
|
|
|
| "inspect" =>
|
|
|
|
match Lean.fromJson? command.payload with
|
|
|
|
| .ok args => inspect args
|
|
|
|
| .error x => return errorJson x
|
|
|
|
| "clear" => clear
|
2023-08-14 21:43:40 -07:00
|
|
|
| "expr.echo" =>
|
2023-08-13 21:19:06 -07:00
|
|
|
match Lean.fromJson? command.payload with
|
2023-08-14 21:43:40 -07:00
|
|
|
| .ok args => expr_echo args
|
2023-08-13 21:19:06 -07:00
|
|
|
| .error x => return errorJson x
|
|
|
|
| "proof.start" =>
|
|
|
|
match Lean.fromJson? command.payload with
|
|
|
|
| .ok args => proof_start args
|
|
|
|
| .error x => return errorJson x
|
|
|
|
| "proof.tactic" =>
|
|
|
|
match Lean.fromJson? command.payload with
|
|
|
|
| .ok args => proof_tactic args
|
|
|
|
| .error x => return errorJson x
|
|
|
|
| "proof.printTree" =>
|
|
|
|
match Lean.fromJson? command.payload with
|
|
|
|
| .ok args => proof_print_tree args
|
|
|
|
| .error x => return errorJson x
|
|
|
|
| cmd =>
|
|
|
|
let error: Commands.InteractionError :=
|
|
|
|
{ error := "unknown", desc := s!"Unknown command {cmd}" }
|
|
|
|
return Lean.toJson error
|
|
|
|
where
|
|
|
|
errorI (type desc: String) := Lean.toJson (
|
|
|
|
{ error := type, desc := desc }: Commands.InteractionError)
|
|
|
|
errorJson := errorI "json"
|
|
|
|
errorIndex := errorI "index"
|
2023-08-14 17:07:53 -07:00
|
|
|
-- Command Functions
|
|
|
|
options_set (args: Commands.OptionsSet): Subroutine Lean.Json := do
|
|
|
|
let state ← get
|
|
|
|
set { state with
|
|
|
|
options := {
|
|
|
|
printExprPretty := args.printExprPretty?.getD true,
|
|
|
|
printExprAST := args.printExprAST?.getD true,
|
|
|
|
proofVariableDelta := args.proofVariableDelta?.getD false
|
|
|
|
}
|
|
|
|
}
|
|
|
|
return Lean.toJson ({ }: Commands.OptionsSetResult)
|
|
|
|
options_print (_: Commands.OptionsPrint): Subroutine Lean.Json := do
|
|
|
|
return Lean.toJson (← get).options
|
2023-08-13 21:19:06 -07:00
|
|
|
catalog (_: Commands.Catalog): Subroutine Lean.Json := do
|
|
|
|
let env ← Lean.MonadEnv.getEnv
|
|
|
|
let names := env.constants.fold (init := #[]) (λ acc name info =>
|
|
|
|
match to_filtered_symbol name info with
|
|
|
|
| .some x => acc.push x
|
|
|
|
| .none => acc)
|
|
|
|
return Lean.toJson <| ({ symbols := names }: Commands.CatalogResult)
|
|
|
|
inspect (args: Commands.Inspect): Subroutine Lean.Json := do
|
2023-08-14 17:07:53 -07:00
|
|
|
let state ← get
|
2023-08-13 21:19:06 -07:00
|
|
|
let env ← Lean.MonadEnv.getEnv
|
|
|
|
let name := str_to_name args.name
|
|
|
|
let info? := env.find? name
|
|
|
|
match info? with
|
|
|
|
| none => return errorIndex s!"Symbol not found {args.name}"
|
|
|
|
| some info =>
|
|
|
|
let module? := env.getModuleIdxFor? name >>=
|
|
|
|
(λ idx => env.allImportedModuleNames.get? idx.toNat) |>.map toString
|
2023-08-14 17:07:53 -07:00
|
|
|
let value? := match args.value?, info with
|
|
|
|
| .some true, _ => info.value?
|
|
|
|
| .some false, _ => .none
|
|
|
|
| .none, .defnInfo _ => info.value?
|
|
|
|
| .none, _ => .none
|
2023-08-13 21:19:06 -07:00
|
|
|
return Lean.toJson ({
|
2023-08-14 17:07:53 -07:00
|
|
|
type := ← serialize_expression state.options info.type,
|
|
|
|
value? := ← value?.mapM (λ v => serialize_expression state.options v),
|
2023-08-13 21:19:06 -07:00
|
|
|
module? := module?
|
|
|
|
}: Commands.InspectResult)
|
|
|
|
clear : Subroutine Lean.Json := do
|
|
|
|
let state ← get
|
|
|
|
let nTrees := state.proofTrees.size
|
|
|
|
set { state with proofTrees := #[] }
|
|
|
|
return Lean.toJson ({ nTrees := nTrees }: Commands.ClearResult)
|
2023-08-14 21:43:40 -07:00
|
|
|
expr_echo (args: Commands.ExprEcho): Subroutine Lean.Json := do
|
|
|
|
let state ← get
|
2023-08-13 21:19:06 -07:00
|
|
|
let env ← Lean.MonadEnv.getEnv
|
|
|
|
match syntax_from_str env args.expr with
|
|
|
|
| .error str => return errorI "parsing" str
|
|
|
|
| .ok syn => do
|
|
|
|
match (← syntax_to_expr syn) with
|
|
|
|
| .error str => return errorI "elab" str
|
|
|
|
| .ok expr => do
|
|
|
|
try
|
2023-08-14 21:43:40 -07:00
|
|
|
let type ← Lean.Meta.inferType expr
|
2023-08-13 21:19:06 -07:00
|
|
|
return Lean.toJson <| ({
|
2023-08-14 21:43:40 -07:00
|
|
|
type := (← serialize_expression (options := state.options) type),
|
|
|
|
expr := (← serialize_expression (options := state.options) expr)
|
|
|
|
}: Commands.ExprEchoResult)
|
2023-08-13 21:19:06 -07:00
|
|
|
catch exception =>
|
|
|
|
return errorI "typing" (← exception.toMessageData.toString)
|
|
|
|
proof_start (args: Commands.ProofStart): Subroutine Lean.Json := do
|
|
|
|
let state ← get
|
|
|
|
let env ← Lean.MonadEnv.getEnv
|
|
|
|
let expr?: Except Lean.Json Lean.Expr ← (match args.expr, args.copyFrom with
|
|
|
|
| .some expr, .none =>
|
|
|
|
(match syntax_from_str env expr with
|
|
|
|
| .error str => return .error <| errorI "parsing" str
|
|
|
|
| .ok syn => do
|
|
|
|
(match (← syntax_to_expr syn) with
|
|
|
|
| .error str => return .error <| errorI "elab" str
|
|
|
|
| .ok expr => return .ok expr))
|
|
|
|
| .none, .some copyFrom =>
|
|
|
|
(match env.find? <| str_to_name copyFrom with
|
|
|
|
| .none => return .error <| errorIndex s!"Symbol not found: {copyFrom}"
|
|
|
|
| .some cInfo => return .ok cInfo.type)
|
|
|
|
| .none, .none =>
|
|
|
|
return .error <| errorI "arguments" "At least one of {expr, copyFrom} must be supplied"
|
|
|
|
| _, _ => return .error <| errorI "arguments" "Cannot populate both of {expr, copyFrom}")
|
|
|
|
match expr? with
|
|
|
|
| .error error => return error
|
|
|
|
| .ok expr =>
|
|
|
|
let tree ← ProofTree.create (str_to_name <| args.name.getD "Untitled") expr
|
|
|
|
-- Put the new tree in the environment
|
|
|
|
let nextTreeId := state.proofTrees.size
|
|
|
|
set { state with proofTrees := state.proofTrees.push tree }
|
|
|
|
return Lean.toJson ({ treeId := nextTreeId }: Commands.ProofStartResult)
|
|
|
|
proof_tactic (args: Commands.ProofTactic): Subroutine Lean.Json := do
|
|
|
|
let state ← get
|
|
|
|
match state.proofTrees.get? args.treeId with
|
|
|
|
| .none => return errorIndex "Invalid tree index {args.treeId}"
|
|
|
|
| .some tree =>
|
|
|
|
let (result, nextTree) ← ProofTree.execute
|
|
|
|
(stateId := args.stateId)
|
|
|
|
(goalId := args.goalId.getD 0)
|
2023-08-14 21:43:40 -07:00
|
|
|
(tactic := args.tactic) |>.run state.options |>.run tree
|
2023-08-13 21:19:06 -07:00
|
|
|
match result with
|
|
|
|
| .invalid message => return Lean.toJson <| errorIndex message
|
|
|
|
| .success nextId? goals =>
|
|
|
|
set { state with proofTrees := state.proofTrees.set! args.treeId nextTree }
|
|
|
|
return Lean.toJson (
|
|
|
|
{ nextId? := nextId?, goals := goals }: Commands.ProofTacticResultSuccess)
|
|
|
|
| .failure messages =>
|
|
|
|
return Lean.toJson (
|
|
|
|
{ tacticErrors := messages }: Commands.ProofTacticResultFailure)
|
|
|
|
proof_print_tree (args: Commands.ProofPrintTree): Subroutine Lean.Json := do
|
|
|
|
let state ← get
|
|
|
|
match state.proofTrees.get? args.treeId with
|
|
|
|
| .none => return errorIndex "Invalid tree index {args.treeId}"
|
|
|
|
| .some tree =>
|
|
|
|
return Lean.toJson ({parents := tree.structure_array}: Commands.ProofPrintTreeResult)
|
|
|
|
|
|
|
|
|
|
|
|
end Pantograph
|