Pantograph/Main.lean

72 lines
2.3 KiB
Plaintext
Raw Normal View History

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-08-13 21:19:06 -07:00
import Pantograph
import Repl
2023-05-09 22:51:19 -07:00
2023-05-09 16:39:24 -07:00
-- Main IO functions
open Pantograph.Repl
2024-09-09 18:43:34 -07:00
open Pantograph.Protocol
2023-05-09 16:39:24 -07:00
/-- Parse a command either in `{ "cmd": ..., "payload": ... }` form or `cmd { ... }` form. -/
2024-09-09 18:43:34 -07:00
def parseCommand (s: String): Except String Command := do
2025-01-24 19:22:58 -08:00
match s.trim.get? 0 with
| .some '{' =>
-- Parse in Json mode
Lean.fromJson? (← Lean.Json.parse s)
2025-01-24 19:22:58 -08:00
| .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 }
2025-01-24 19:22:58 -08:00
| .none =>
throw "Command is empty"
2025-01-24 19:11:05 -08:00
partial def loop : MainM Unit := do repeat do
let state ← get
let command ← (← IO.getStdin).getLine
2025-01-24 19:22:58 -08:00
-- Halt the program if empty line is given
2025-01-24 19:11:05 -08:00
if command.trim.length = 0 then break
match parseCommand command with
| .error error =>
2024-09-09 18:43:34 -07:00
let error := Lean.toJson ({ error := "command", desc := error }: InteractionError)
-- Using `Lean.Json.compress` here to prevent newline
IO.println error.compress
| .ok command =>
2024-10-08 00:17:31 -07:00
try
let ret ← execute command
let str := match state.options.printJsonPretty with
| true => ret.pretty
| false => ret.compress
IO.println str
catch e =>
let message ← e.toMessageData.toString
2024-10-08 00:45:45 -07:00
let error := Lean.toJson ({ error := "main", desc := message }: InteractionError)
2024-10-08 00:17:31 -07:00
IO.println error.compress
2023-05-09 16:39:24 -07:00
unsafe def main (args: List String): IO Unit := do
2023-08-13 21:19:06 -07:00
-- NOTE: A more sophisticated scheme of command line argument handling is needed.
if args == ["--version"] then do
2024-10-08 00:45:45 -07:00
IO.println s!"{Pantograph.version}"
2023-08-13 21:19:06 -07:00
return
2024-09-09 18:43:34 -07:00
Pantograph.initSearch ""
2025-01-24 19:22:58 -08:00
-- Separate imports and options
let (options, imports) := args.partition (·.startsWith "--")
let coreContext ← options.map (·.drop 2) |>.toArray |> Pantograph.createCoreContext
2024-09-09 18:43:34 -07:00
let coreState ← Pantograph.createCoreState imports.toArray
2025-01-24 19:22:58 -08:00
let context: Context := {}
2023-05-21 17:41:39 -07:00
try
let coreM := loop.run context |>.run' {}
IO.println "ready."
2024-03-09 16:50:36 -08:00
discard <| coreM.toIO coreContext coreState
2023-05-21 17:41:39 -07:00
catch ex =>
2024-10-08 00:45:45 -07:00
let message := ex.toString
let error := Lean.toJson ({ error := "io", desc := message }: InteractionError)
IO.println error.compress