Pantograph/Main.lean

67 lines
2.1 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.Version
2024-03-09 16:50:36 -08:00
import Pantograph.Library
2023-08-13 21:19:06 -07:00
import 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
/-- Parse a command either in `{ "cmd": ..., "payload": ... }` form or `cmd { ... }` form. -/
def parseCommand (s: String): Except String Protocol.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-11-06 11:43:57 -08:00
partial def loop : MainM Unit := do
let state ← get
let command ← (← IO.getStdin).getLine
if command.trim.length = 0 then return ()
match parseCommand command with
| .error error =>
let error := Lean.toJson ({ error := "command", desc := error }: Protocol.InteractionError)
-- Using `Lean.Json.compress` here to prevent newline
IO.println error.compress
| .ok command =>
let ret ← execute command
let str := match state.options.printJsonPretty with
| true => ret.pretty
| false => ret.compress
IO.println str
2023-05-09 22:51:19 -07:00
loop
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.
-- Separate imports and options
if args == ["--version"] then do
println! s!"{version}"
return
2024-03-09 19:36:25 -08:00
initSearch ""
2024-03-09 16:50:36 -08:00
let coreContext ← args.filterMap (λ s => if s.startsWith "--" then .some <| s.drop 2 else .none)
|>.toArray |> createCoreContext
let imports:= args.filter (λ s => ¬ (s.startsWith "--"))
2024-03-09 16:50:36 -08:00
let coreState ← createCoreState imports.toArray
let context: Context := {
imports
}
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 =>
IO.println "Uncaught IO exception"
IO.println ex.toString