Pantograph/Pantograph.lean

194 lines
8.5 KiB
Plaintext
Raw Normal View History

2023-10-15 12:31:22 -07:00
import Pantograph.Goal
import Pantograph.Protocol
import Pantograph.Serial
import Pantograph.Environment
2024-03-08 23:50:44 -08:00
import Pantograph.Library
import Lean.Data.HashMap
2023-08-13 21:19:06 -07:00
namespace Pantograph
structure Context where
imports: List String
2023-08-13 21:19:06 -07:00
/-- Stores state of the REPL -/
structure State where
options: Protocol.Options := {}
nextId: Nat := 0
goalStates: Lean.HashMap Nat GoalState := Lean.HashMap.empty
2023-08-13 21:19:06 -07:00
/-- Main state monad for executing commands -/
abbrev MainM := ReaderT Context (StateT State Lean.CoreM)
-- HACK: For some reason writing `CommandM α := MainM (Except ... α)` disables
-- certain monadic features in `MainM`
abbrev CR α := Except Protocol.InteractionError α
2023-08-13 21:19:06 -07:00
def execute (command: Protocol.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
| .ok args => do
match (← comm args) with
| .ok result => return Lean.toJson result
| .error ierror => return Lean.toJson ierror
| .error error => return Lean.toJson $ errorCommand s!"Unable to parse json: {error}"
match command.cmd with
2023-11-06 12:20:08 -08:00
| "reset" => run reset
| "stat" => run stat
| "expr.echo" => run expr_echo
| "env.catalog" => run env_catalog
| "env.inspect" => run env_inspect
2023-12-13 19:35:32 -08:00
| "env.add" => run env_add
2023-11-06 12:20:08 -08:00
| "options.set" => run options_set
| "options.print" => run options_print
| "goal.start" => run goal_start
| "goal.tactic" => run goal_tactic
| "goal.continue" => run goal_continue
| "goal.delete" => run goal_delete
| "goal.print" => run goal_print
2023-08-13 21:19:06 -07:00
| cmd =>
let error: Protocol.InteractionError :=
errorCommand s!"Unknown command {cmd}"
2023-08-13 21:19:06 -07:00
return Lean.toJson error
where
2024-03-10 08:13:10 -07:00
errorCommand := errorI "command"
errorIndex := errorI "index"
-- Command Functions
reset (_: Protocol.Reset): MainM (CR Protocol.StatResult) := do
let state ← get
2023-08-27 19:58:52 -07:00
let nGoals := state.goalStates.size
set { state with nextId := 0, goalStates := Lean.HashMap.empty }
2023-08-27 19:58:52 -07:00
return .ok { nGoals }
stat (_: Protocol.Stat): MainM (CR Protocol.StatResult) := do
2023-08-27 19:58:52 -07:00
let state ← get
let nGoals := state.goalStates.size
return .ok { nGoals }
env_catalog (args: Protocol.EnvCatalog): MainM (CR Protocol.EnvCatalogResult) := do
2024-03-09 16:50:36 -08:00
let result ← Environment.catalog args
return .ok result
env_inspect (args: Protocol.EnvInspect): MainM (CR Protocol.EnvInspectResult) := do
let state ← get
Environment.inspect args state.options
2023-12-13 19:35:32 -08:00
env_add (args: Protocol.EnvAdd): MainM (CR Protocol.EnvAddResult) := do
Environment.addDecl args
expr_echo (args: Protocol.ExprEcho): MainM (CR Protocol.ExprEchoResult) := do
2023-08-14 21:43:40 -07:00
let state ← get
2024-03-31 16:43:30 -07:00
exprEcho args.expr args.type? state.options
options_set (args: Protocol.OptionsSet): MainM (CR Protocol.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,
noRepeat := args.noRepeat?.getD options.noRepeat,
printAuxDecls := args.printAuxDecls?.getD options.printAuxDecls,
printImplementationDetailHyps := args.printImplementationDetailHyps?.getD options.printImplementationDetailHyps
}
}
return .ok { }
2024-05-04 23:36:42 -07:00
options_print (_: Protocol.OptionsPrint): MainM (CR Protocol.Options) := do
return .ok (← get).options
goal_start (args: Protocol.GoalStart): MainM (CR Protocol.GoalStartResult) := do
2023-08-13 21:19:06 -07:00
let state ← get
let env ← Lean.MonadEnv.getEnv
let expr?: Except _ GoalState ← runTermElabM (match args.expr, args.copyFrom with
| .some expr, .none => goalStartExpr expr
2023-08-13 21:19:06 -07:00
| .none, .some copyFrom =>
(match env.find? <| copyFrom.toName with
2023-08-13 21:19:06 -07:00
| .none => return .error <| errorIndex s!"Symbol not found: {copyFrom}"
2023-12-08 17:31:25 -08:00
| .some cInfo => return .ok (← GoalState.create cInfo.type))
| _, _ =>
return .error <| errorI "arguments" "Exactly one of {expr, copyFrom} must be supplied")
2023-08-13 21:19:06 -07:00
match expr? with
| .error error => return .error error
2023-12-08 17:31:25 -08:00
| .ok goalState =>
let stateId := state.nextId
set { state with
goalStates := state.goalStates.insert stateId goalState,
nextId := state.nextId + 1
}
2023-11-06 11:51:31 -08:00
return .ok { stateId, root := goalState.root.name.toString }
goal_tactic (args: Protocol.GoalTactic): MainM (CR Protocol.GoalTacticResult) := do
2023-08-13 21:19:06 -07:00
let state ← get
match state.goalStates.find? args.stateId with
| .none => return .error $ errorIndex s!"Invalid state index {args.stateId}"
| .some goalState => do
2024-04-11 15:11:10 -07:00
let nextGoalState?: Except _ GoalState ←
match args.tactic?, args.expr?, args.have?, args.calc?, args.conv? with
| .some tactic, .none, .none, .none, .none => do
2024-03-10 08:13:10 -07:00
pure ( Except.ok (← goalTactic goalState args.goalId tactic))
2024-04-11 15:11:10 -07:00
| .none, .some expr, .none, .none, .none => do
2024-04-06 16:40:22 -07:00
pure ( Except.ok (← goalAssign goalState args.goalId expr))
2024-04-11 15:11:10 -07:00
| .none, .none, .some type, .none, .none => do
2024-04-06 16:40:22 -07:00
let binderName := args.binderName?.getD ""
pure ( Except.ok (← goalHave goalState args.goalId binderName type))
2024-04-11 15:11:10 -07:00
| .none, .none, .none, .some pred, .none => do
pure ( Except.ok (← goalCalc goalState args.goalId pred))
| .none, .none, .none, .none, .some true => do
pure ( Except.ok (← goalConv goalState args.goalId))
| .none, .none, .none, .none, .some false => do
pure ( Except.ok (← goalConvExit goalState))
| _, _, _, _, _ => pure (Except.error <|
errorI "arguments" "Exactly one of {tactic, expr, have, calc, conv} must be supplied")
match nextGoalState? with
| .error error => return .error error
| .ok (.success nextGoalState) =>
let nextStateId := state.nextId
set { state with
goalStates := state.goalStates.insert state.nextId nextGoalState,
nextId := state.nextId + 1,
}
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 (.indexError goalId) =>
return .error $ errorIndex s!"Invalid goal id index {goalId}"
2024-04-08 12:26:22 -07:00
| .ok (.invalidAction message) =>
return .error $ errorI "invalid" message
| .ok (.failure messages) =>
return .ok { tacticErrors? := .some messages }
2023-11-04 15:51:09 -07:00
goal_continue (args: Protocol.GoalContinue): MainM (CR Protocol.GoalContinueResult) := do
let state ← get
match state.goalStates.find? args.target with
2023-11-04 15:51:09 -07:00
| .none => return .error $ errorIndex s!"Invalid state index {args.target}"
| .some target => do
let nextState? ← match args.branch?, args.goals? with
| .some branchId, .none => do
match state.goalStates.find? branchId with
2023-11-04 15:51:09 -07:00
| .none => return .error $ errorIndex s!"Invalid state index {branchId}"
| .some branch => pure $ goalContinue target branch
2023-11-04 15:51:09 -07:00
| .none, .some goals =>
2024-03-10 08:13:10 -07:00
pure $ goalResume target goals
2023-11-04 15:51:09 -07:00
| _, _ => return .error <| errorI "arguments" "Exactly one of {branch, goals} must be supplied"
match nextState? with
2023-11-09 22:24:17 -08:00
| .error error => return .error <| errorI "structure" error
2023-11-04 15:51:09 -07:00
| .ok nextGoalState =>
let nextStateId := state.nextId
set { state with
goalStates := state.goalStates.insert nextStateId nextGoalState,
nextId := state.nextId + 1
}
2024-03-10 08:13:10 -07:00
let goals ← goalSerialize nextGoalState (options := state.options)
2023-11-04 15:51:09 -07:00
return .ok {
2023-11-09 22:24:17 -08:00
nextStateId,
goals,
2023-11-04 15:51:09 -07:00
}
goal_delete (args: Protocol.GoalDelete): MainM (CR Protocol.GoalDeleteResult) := do
let state ← get
let goalStates := args.stateIds.foldl (λ map id => map.erase id) state.goalStates
set { state with goalStates }
return .ok {}
goal_print (args: Protocol.GoalPrint): MainM (CR Protocol.GoalPrintResult) := do
let state ← get
match state.goalStates.find? args.stateId with
| .none => return .error $ errorIndex s!"Invalid state index {args.stateId}"
| .some goalState => runMetaM <| do
2024-03-14 16:34:01 -07:00
return .ok (← goalPrint goalState state.options)
2023-08-13 21:19:06 -07:00
end Pantograph