2023-05-09 16:39:24 -07:00
|
|
|
import Lean.Data.Json
|
2023-05-09 22:51:19 -07:00
|
|
|
import Lean.Environment
|
2023-05-07 15:19:45 -07:00
|
|
|
|
2023-05-09 22:51:19 -07:00
|
|
|
import Pantograph.Commands
|
2023-05-21 17:41:39 -07:00
|
|
|
import Pantograph.Serial
|
|
|
|
import Pantograph.Meta
|
2023-05-12 16:12:21 -07:00
|
|
|
import Pantograph.Symbols
|
2023-05-09 16:39:24 -07:00
|
|
|
|
|
|
|
namespace Pantograph
|
|
|
|
|
2023-05-20 15:58:38 -07:00
|
|
|
|
|
|
|
structure Context where
|
|
|
|
|
2023-05-20 14:04:09 -07:00
|
|
|
/-- Stores state of the REPL -/
|
|
|
|
structure State where
|
2023-05-23 05:12:46 -07:00
|
|
|
--environments: Array Lean.Environment := #[]
|
2023-05-21 17:41:39 -07:00
|
|
|
proofTrees: Array Meta.ProofTree := #[]
|
2023-05-20 14:04:09 -07:00
|
|
|
|
|
|
|
-- State monad
|
2023-05-23 05:12:46 -07:00
|
|
|
abbrev Subroutine := ReaderT Context (StateT State Lean.Elab.TermElabM)
|
2023-05-20 15:58:38 -07:00
|
|
|
|
|
|
|
open Commands
|
2023-05-20 13:03:12 -07:00
|
|
|
|
|
|
|
/-- Parse a command either in `{ "cmd": ..., "payload": ... }` form or `cmd { ... }` form. -/
|
|
|
|
def parse_command (s: String): Except String 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"
|
|
|
|
|
2023-05-17 21:58:03 -07:00
|
|
|
|
2023-05-09 18:01:09 -07:00
|
|
|
|
2023-05-23 05:12:46 -07:00
|
|
|
def execute (command: Command): Subroutine Lean.Json := do
|
2023-05-09 18:01:09 -07:00
|
|
|
match command.cmd with
|
|
|
|
| "catalog" =>
|
2023-05-20 15:58:38 -07:00
|
|
|
match Lean.fromJson? command.payload with
|
|
|
|
| .ok args => catalog args
|
|
|
|
| .error x => return errorJson x
|
2023-05-20 14:04:09 -07:00
|
|
|
| "inspect" =>
|
2023-05-20 15:58:38 -07:00
|
|
|
match Lean.fromJson? command.payload with
|
|
|
|
| .ok args => inspect args
|
|
|
|
| .error x => return errorJson x
|
2023-05-21 17:41:39 -07:00
|
|
|
| "proof.start" =>
|
2023-05-20 15:58:38 -07:00
|
|
|
match Lean.fromJson? command.payload with
|
2023-05-21 17:41:39 -07:00
|
|
|
| .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
|
2023-05-20 15:58:38 -07:00
|
|
|
| .error x => return errorJson x
|
|
|
|
| cmd =>
|
|
|
|
let error: InteractionError := { error := "unknown", desc := s!"Unknown command {cmd}" }
|
|
|
|
return Lean.toJson error
|
2023-05-17 21:58:03 -07:00
|
|
|
where
|
2023-05-20 15:58:38 -07:00
|
|
|
errorJson (s: String) := Lean.toJson ({ error := "json", desc := s }: InteractionError)
|
|
|
|
errorIndex (s: String) := Lean.toJson ({ error := "index", desc := s }: InteractionError)
|
2023-05-23 05:12:46 -07:00
|
|
|
catalog (_: Catalog): Subroutine Lean.Json := do
|
|
|
|
let env ← Lean.MonadEnv.getEnv
|
|
|
|
let names := env.constants.fold (init := []) (λ es name info =>
|
|
|
|
match to_filtered_symbol name info with
|
|
|
|
| .some x => x::es
|
|
|
|
| .none => es)
|
|
|
|
return Lean.toJson <| ({ symbols := names }: CatalogResult)
|
2023-05-20 15:58:38 -07:00
|
|
|
inspect (args: Inspect): Subroutine Lean.Json := do
|
2023-05-23 05:12:46 -07:00
|
|
|
let env ← Lean.MonadEnv.getEnv
|
|
|
|
let name := str_to_name args.name
|
|
|
|
let info? := env.find? name
|
|
|
|
match info? with
|
|
|
|
| none => return Lean.toJson <| errorIndex s!"Symbol not found {args.name}"
|
|
|
|
| some info =>
|
|
|
|
let format ← Lean.Meta.ppExpr info.toConstantVal.type
|
|
|
|
let module? := env.getModuleIdxFor? name >>=
|
|
|
|
(λ idx => env.allImportedModuleNames.get? idx.toNat) |>.map toString
|
|
|
|
return Lean.toJson ({
|
|
|
|
type := toString format,
|
|
|
|
module? := module?
|
|
|
|
}: InspectResult)
|
2023-05-21 17:41:39 -07:00
|
|
|
proof_start (args: ProofStart): Subroutine Lean.Json := do
|
|
|
|
let state ← get
|
2023-05-23 05:12:46 -07:00
|
|
|
let env ← Lean.MonadEnv.getEnv
|
|
|
|
let expr?: Except Lean.Json Lean.Expr ← (match args.expr, args.copyFrom with
|
|
|
|
| .some expr, .none =>
|
|
|
|
(match Serial.syntax_from_str env expr with
|
|
|
|
| .error str => return .error <| Lean.toJson ({ error := "parsing", desc := str }: InteractionError)
|
|
|
|
| .ok syn => do
|
|
|
|
(match (← Meta.syntax_to_expr syn) with
|
|
|
|
| .error str => return .error <| Lean.toJson ({ error := "elab", desc := str }: InteractionError)
|
|
|
|
| .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 <| Lean.toJson ({ error := "arguments", desc := "At least one of {expr, copyFrom} must be supplied" }: InteractionError)
|
|
|
|
| _, _ => return .error <| Lean.toJson ({ error := "arguments", desc := "Cannot populate both of {expr, copyFrom}" }: InteractionError))
|
|
|
|
match expr? with
|
2023-05-21 17:41:39 -07:00
|
|
|
| .error error => return error
|
2023-05-23 05:12:46 -07:00
|
|
|
| .ok expr =>
|
|
|
|
let tree ← Meta.ProofTree.create (str_to_name <| args.name.getD "Untitled") expr
|
2023-05-21 17:41:39 -07:00
|
|
|
-- Put the new tree in the environment
|
2023-05-22 14:56:43 -07:00
|
|
|
let nextTreeId := state.proofTrees.size
|
2023-05-21 17:41:39 -07:00
|
|
|
set { state with proofTrees := state.proofTrees.push tree }
|
2023-05-22 14:56:43 -07:00
|
|
|
return Lean.toJson ({ treeId := nextTreeId }: ProofStartResult)
|
2023-05-21 17:41:39 -07:00
|
|
|
proof_tactic (args: ProofTactic): Subroutine Lean.Json := do
|
|
|
|
let state ← get
|
|
|
|
match state.proofTrees.get? args.treeId with
|
|
|
|
| .none => return Lean.toJson <| errorIndex "Invalid tree index {args.treeId}"
|
|
|
|
| .some tree =>
|
2023-05-23 05:12:46 -07:00
|
|
|
let (result, nextTree) ← Meta.ProofTree.execute
|
2023-05-22 00:49:37 -07:00
|
|
|
(stateId := args.stateId)
|
|
|
|
(goalId := args.goalId.getD 0)
|
|
|
|
(tactic := args.tactic) |>.run tree
|
2023-05-21 17:41:39 -07:00
|
|
|
match result with
|
|
|
|
| .invalid message => return Lean.toJson <| errorIndex message
|
2023-05-22 14:56:43 -07:00
|
|
|
| .success nextId? goals =>
|
2023-05-21 17:41:39 -07:00
|
|
|
set { state with proofTrees := state.proofTrees.set! args.treeId nextTree }
|
2023-05-22 14:56:43 -07:00
|
|
|
return Lean.toJson ({ nextId? := nextId?, goals := goals }: ProofTacticResultSuccess)
|
2023-05-21 17:41:39 -07:00
|
|
|
| .failure messages =>
|
|
|
|
return Lean.toJson ({ errorMessages := messages }: ProofTacticResultFailure)
|
|
|
|
proof_print_tree (args: ProofPrintTree): Subroutine Lean.Json := do
|
|
|
|
let state ← get
|
|
|
|
match state.proofTrees.get? args.treeId with
|
|
|
|
| .none => return Lean.toJson <| errorIndex "Invalid tree index {args.treeId}"
|
|
|
|
| .some tree =>
|
2023-05-22 22:48:48 -07:00
|
|
|
return Lean.toJson ({parents := tree.structure_array}: ProofPrintTreeResult)
|
2023-05-17 21:58:03 -07:00
|
|
|
|
|
|
|
|
|
|
|
end Pantograph
|
|
|
|
|
2023-05-09 22:51:19 -07:00
|
|
|
|
2023-05-09 16:39:24 -07:00
|
|
|
-- Main IO functions
|
2023-05-21 17:41:39 -07:00
|
|
|
open Pantograph
|
2023-05-09 16:39:24 -07:00
|
|
|
|
2023-05-20 15:58:38 -07:00
|
|
|
unsafe def loop : Subroutine Unit := do
|
|
|
|
let command ← (← IO.getStdin).getLine
|
2023-05-22 14:49:56 -07:00
|
|
|
if command.trim.length = 0 then return ()
|
2023-05-20 15:58:38 -07:00
|
|
|
match parse_command command with
|
2023-05-22 14:49:56 -07:00
|
|
|
| .error error =>
|
|
|
|
let error := Lean.toJson ({ error := "json", desc := error }: Commands.InteractionError)
|
|
|
|
IO.println (toString error)
|
2023-05-20 15:58:38 -07:00
|
|
|
| .ok command =>
|
|
|
|
let ret ← execute command
|
|
|
|
IO.println <| toString <| ret
|
2023-05-09 22:51:19 -07:00
|
|
|
loop
|
2023-05-09 16:39:24 -07:00
|
|
|
|
2023-05-23 05:12:46 -07:00
|
|
|
unsafe def main (args: List String): IO Unit := do
|
2023-05-12 16:12:21 -07:00
|
|
|
Lean.enableInitializersExecution
|
2023-05-12 01:08:36 -07:00
|
|
|
Lean.initSearchPath (← Lean.findSysroot)
|
2023-05-23 05:12:46 -07:00
|
|
|
|
|
|
|
let env ← Lean.importModules
|
|
|
|
(imports := args.map (λ str => { module := str_to_name str, runtimeOnly := false }))
|
|
|
|
(opts := {})
|
|
|
|
(trustLevel := 1)
|
2023-05-20 15:58:38 -07:00
|
|
|
let context: Context := {
|
2023-05-23 05:12:46 -07:00
|
|
|
}
|
|
|
|
let coreContext: Lean.Core.Context := {
|
|
|
|
currNamespace := str_to_name "Aniva",
|
|
|
|
openDecls := [], -- No 'open' directives needed
|
|
|
|
fileName := "<Pantograph>",
|
|
|
|
fileMap := { source := "", positions := #[0], lines := #[1] }
|
2023-05-20 15:58:38 -07:00
|
|
|
}
|
2023-05-21 17:41:39 -07:00
|
|
|
try
|
2023-05-23 05:12:46 -07:00
|
|
|
let termElabM := loop.run context |>.run' {}
|
|
|
|
let metaM := termElabM.run' (ctx := {
|
|
|
|
declName? := some "_pantograph",
|
|
|
|
errToSorry := false
|
|
|
|
})
|
|
|
|
let coreM := metaM.run'
|
|
|
|
discard <| coreM.toIO coreContext { env := env }
|
2023-05-21 17:41:39 -07:00
|
|
|
catch ex =>
|
|
|
|
IO.println "Uncaught IO exception"
|
|
|
|
IO.println ex.toString
|