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-27 19:53:09 -07:00
|
|
|
|
import Pantograph.SemihashMap
|
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-27 19:53:09 -07:00
|
|
|
|
goalStates: SemihashMap GoalState := SemihashMap.empty
|
2023-08-13 21:19:06 -07:00
|
|
|
|
|
|
|
|
|
-- State monad
|
2023-08-15 15:40:54 -07:00
|
|
|
|
abbrev MainM := ReaderT Context (StateT State Lean.Elab.TermElabM)
|
|
|
|
|
-- For some reason writing `CommandM α := MainM (Except ... α)` disables certain
|
|
|
|
|
-- monadic features in `MainM`
|
|
|
|
|
abbrev CR α := Except Commands.InteractionError α
|
2023-08-13 21:19:06 -07:00
|
|
|
|
|
2023-08-15 15:40:54 -07:00
|
|
|
|
def execute (command: Commands.Command): MainM Lean.Json := do
|
|
|
|
|
let run { α β: Type } [Lean.FromJson α] [Lean.ToJson β] (comm: α → MainM (CR β)): MainM Lean.Json :=
|
2023-08-13 21:19:06 -07:00
|
|
|
|
match Lean.fromJson? command.payload with
|
2023-08-15 15:40:54 -07:00
|
|
|
|
| .ok args => do
|
|
|
|
|
match (← comm args) with
|
|
|
|
|
| .ok result => return Lean.toJson result
|
|
|
|
|
| .error ierror => return Lean.toJson ierror
|
2023-08-24 22:51:40 -07:00
|
|
|
|
| .error error => return Lean.toJson $ errorCommand s!"Unable to parse json: {error}"
|
2023-08-15 15:40:54 -07:00
|
|
|
|
match command.cmd with
|
2023-08-16 19:25:32 -07:00
|
|
|
|
| "reset" => run reset
|
|
|
|
|
| "expr.echo" => run expr_echo
|
|
|
|
|
| "lib.catalog" => run lib_catalog
|
|
|
|
|
| "lib.inspect" => run lib_inspect
|
2023-08-15 15:40:54 -07:00
|
|
|
|
| "options.set" => run options_set
|
|
|
|
|
| "options.print" => run options_print
|
|
|
|
|
| "proof.start" => run proof_start
|
|
|
|
|
| "proof.tactic" => run proof_tactic
|
|
|
|
|
| "proof.printTree" => run proof_print_tree
|
2023-08-13 21:19:06 -07:00
|
|
|
|
| cmd =>
|
|
|
|
|
let error: Commands.InteractionError :=
|
2023-08-24 22:51:40 -07:00
|
|
|
|
errorCommand s!"Unknown command {cmd}"
|
2023-08-13 21:19:06 -07:00
|
|
|
|
return Lean.toJson error
|
|
|
|
|
where
|
2023-08-15 15:40:54 -07:00
|
|
|
|
errorI (type desc: String): Commands.InteractionError := { error := type, desc := desc }
|
2023-08-24 22:51:40 -07:00
|
|
|
|
errorCommand := errorI "command"
|
2023-08-13 21:19:06 -07:00
|
|
|
|
errorIndex := errorI "index"
|
2023-08-14 17:07:53 -07:00
|
|
|
|
-- Command Functions
|
2023-08-16 19:25:32 -07:00
|
|
|
|
reset (_: Commands.Reset): MainM (CR Commands.ResetResult) := do
|
2023-08-14 17:07:53 -07:00
|
|
|
|
let state ← get
|
2023-08-27 19:53:09 -07:00
|
|
|
|
let nStates := state.goalStates.size
|
|
|
|
|
set { state with goalStates := SemihashMap.empty }
|
|
|
|
|
return .ok { nStates := nStates }
|
2023-08-16 19:25:32 -07:00
|
|
|
|
lib_catalog (_: Commands.LibCatalog): MainM (CR Commands.LibCatalogResult) := do
|
2023-08-13 21:19:06 -07:00
|
|
|
|
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)
|
2023-08-15 15:40:54 -07:00
|
|
|
|
return .ok { symbols := names }
|
2023-08-16 19:25:32 -07:00
|
|
|
|
lib_inspect (args: Commands.LibInspect): MainM (CR Commands.LibInspectResult) := 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
|
2023-08-15 15:40:54 -07:00
|
|
|
|
| none => return .error $ errorIndex s!"Symbol not found {args.name}"
|
2023-08-13 21:19:06 -07:00
|
|
|
|
| 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-15 15:40:54 -07:00
|
|
|
|
return .ok {
|
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?
|
2023-08-15 15:40:54 -07:00
|
|
|
|
}
|
|
|
|
|
expr_echo (args: Commands.ExprEcho): MainM (CR Commands.ExprEchoResult) := do
|
2023-08-14 21:43:40 -07:00
|
|
|
|
let state ← get
|
2023-08-13 21:19:06 -07:00
|
|
|
|
let env ← Lean.MonadEnv.getEnv
|
|
|
|
|
match syntax_from_str env args.expr with
|
2023-08-15 15:40:54 -07:00
|
|
|
|
| .error str => return .error $ errorI "parsing" str
|
2023-08-13 21:19:06 -07:00
|
|
|
|
| .ok syn => do
|
|
|
|
|
match (← syntax_to_expr syn) with
|
2023-08-15 15:40:54 -07:00
|
|
|
|
| .error str => return .error $ errorI "elab" str
|
2023-08-13 21:19:06 -07:00
|
|
|
|
| .ok expr => do
|
|
|
|
|
try
|
2023-08-14 21:43:40 -07:00
|
|
|
|
let type ← Lean.Meta.inferType expr
|
2023-08-15 15:40:54 -07:00
|
|
|
|
return .ok {
|
2023-08-14 21:43:40 -07:00
|
|
|
|
type := (← serialize_expression (options := state.options) type),
|
|
|
|
|
expr := (← serialize_expression (options := state.options) expr)
|
2023-08-15 15:40:54 -07:00
|
|
|
|
}
|
2023-08-13 21:19:06 -07:00
|
|
|
|
catch exception =>
|
2023-08-15 15:40:54 -07:00
|
|
|
|
return .error $ errorI "typing" (← exception.toMessageData.toString)
|
2023-08-16 19:25:32 -07:00
|
|
|
|
options_set (args: Commands.OptionsSet): MainM (CR Commands.OptionsSetResult) := do
|
|
|
|
|
let state ← get
|
|
|
|
|
let options := state.options
|
|
|
|
|
set { state with
|
|
|
|
|
options := {
|
|
|
|
|
-- FIXME: This should be replaced with something more elegant
|
|
|
|
|
printJsonPretty := args.printJsonPretty?.getD options.printJsonPretty,
|
|
|
|
|
printExprPretty := args.printExprPretty?.getD options.printExprPretty,
|
|
|
|
|
printExprAST := args.printExprAST?.getD options.printExprAST,
|
|
|
|
|
proofVariableDelta := args.proofVariableDelta?.getD options.proofVariableDelta,
|
|
|
|
|
printAuxDecls := args.printAuxDecls?.getD options.printAuxDecls,
|
|
|
|
|
printImplementationDetailHyps := args.printImplementationDetailHyps?.getD options.printImplementationDetailHyps
|
|
|
|
|
}
|
|
|
|
|
}
|
|
|
|
|
return .ok { }
|
|
|
|
|
options_print (_: Commands.OptionsPrint): MainM (CR Commands.OptionsPrintResult) := do
|
|
|
|
|
return .ok (← get).options
|
2023-08-15 15:40:54 -07:00
|
|
|
|
proof_start (args: Commands.ProofStart): MainM (CR Commands.ProofStartResult) := do
|
2023-08-13 21:19:06 -07:00
|
|
|
|
let state ← get
|
|
|
|
|
let env ← Lean.MonadEnv.getEnv
|
2023-08-15 15:40:54 -07:00
|
|
|
|
let expr?: Except _ Lean.Expr ← (match args.expr, args.copyFrom with
|
2023-08-13 21:19:06 -07:00
|
|
|
|
| .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
|
2023-08-15 15:40:54 -07:00
|
|
|
|
| .error error => return .error error
|
2023-08-13 21:19:06 -07:00
|
|
|
|
| .ok expr =>
|
2023-08-27 19:53:09 -07:00
|
|
|
|
let goalState ← GoalState.create expr
|
|
|
|
|
let (goalStates, goalId) := state.goalStates.insert goalState
|
|
|
|
|
set { state with goalStates }
|
|
|
|
|
return .ok { goalId }
|
2023-08-15 15:40:54 -07:00
|
|
|
|
proof_tactic (args: Commands.ProofTactic): MainM (CR Commands.ProofTacticResult) := do
|
2023-08-13 21:19:06 -07:00
|
|
|
|
let state ← get
|
2023-08-27 19:53:09 -07:00
|
|
|
|
match state.goalStates.get? args.goalId with
|
|
|
|
|
| .none => return .error $ errorIndex "Invalid goal index {args.goalId}"
|
|
|
|
|
| .some goalState =>
|
|
|
|
|
let result ← GoalState.execute goalState args.tactic |>.run state.options
|
2023-08-13 21:19:06 -07:00
|
|
|
|
match result with
|
2023-08-27 19:53:09 -07:00
|
|
|
|
| .success goals =>
|
|
|
|
|
if goals.isEmpty then
|
|
|
|
|
return .ok {}
|
|
|
|
|
else
|
|
|
|
|
-- Append all goals
|
|
|
|
|
let (goalStates, goalIds, sGoals) := Array.foldl (λ acc itr =>
|
|
|
|
|
let (map, indices, serializedGoals) := acc
|
|
|
|
|
let (goalState, sGoal) := itr
|
|
|
|
|
let (map, index) := map.insert goalState
|
|
|
|
|
(map, index :: indices, sGoal :: serializedGoals)
|
|
|
|
|
) (state.goalStates, [], []) goals
|
|
|
|
|
set { state with goalStates }
|
|
|
|
|
return .ok { goals? := .some sGoals.reverse.toArray, goalIds? := .some goalIds.reverse.toArray }
|
2023-08-13 21:19:06 -07:00
|
|
|
|
| .failure messages =>
|
2023-08-15 15:40:54 -07:00
|
|
|
|
return .ok { tacticErrors? := .some messages }
|
2023-08-27 19:53:09 -07:00
|
|
|
|
proof_print_tree (_: Commands.ProofPrintTree): MainM (CR Commands.ProofPrintTreeResult) := do
|
2023-08-13 21:19:06 -07:00
|
|
|
|
let state ← get
|
2023-08-27 19:53:09 -07:00
|
|
|
|
return .ok { nGoals := state.goalStates.size }
|
2023-08-13 21:19:06 -07:00
|
|
|
|
|
|
|
|
|
end Pantograph
|