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
|
|
|
|
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
|
|
|
|
2023-08-22 09:54:37 -07:00
|
|
|
/-- Parse a command either in `{ "cmd": ..., "payload": ... }` form or `cmd { ... }` form. -/
|
|
|
|
def parse_command (s: String): Except String Commands.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-08-15 15:40:54 -07:00
|
|
|
unsafe def loop : MainM Unit := do
|
2023-08-16 19:25:32 -07:00
|
|
|
let state ← get
|
2023-05-20 15:58:38 -07:00
|
|
|
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 =>
|
2023-08-23 13:00:11 -07:00
|
|
|
let error := Lean.toJson ({ error := "command", desc := error }: Commands.InteractionError)
|
2023-08-16 19:25:32 -07:00
|
|
|
-- Using `Lean.Json.compress` here to prevent newline
|
|
|
|
IO.println error.compress
|
2023-05-20 15:58:38 -07:00
|
|
|
| .ok command =>
|
|
|
|
let ret ← execute command
|
2023-08-16 19:25:32 -07:00
|
|
|
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
|
|
|
|
2023-08-14 17:07:53 -07:00
|
|
|
namespace Lean
|
|
|
|
|
|
|
|
/-- This is better than the default version since it handles `.` and doesn't
|
|
|
|
crash the program when it fails. -/
|
|
|
|
def setOptionFromString' (opts : Options) (entry : String) : ExceptT String IO Options := do
|
|
|
|
let ps := (entry.splitOn "=").map String.trim
|
|
|
|
let [key, val] ← pure ps | throw "invalid configuration option entry, it must be of the form '<key> = <value>'"
|
|
|
|
let key := Pantograph.str_to_name key
|
|
|
|
let defValue ← getOptionDefaultValue key
|
|
|
|
match defValue with
|
|
|
|
| DataValue.ofString _ => pure $ opts.setString key val
|
|
|
|
| DataValue.ofBool _ =>
|
|
|
|
match val with
|
|
|
|
| "true" => pure $ opts.setBool key true
|
|
|
|
| "false" => pure $ opts.setBool key false
|
|
|
|
| _ => throw s!"invalid Bool option value '{val}'"
|
|
|
|
| DataValue.ofName _ => pure $ opts.setName key val.toName
|
|
|
|
| DataValue.ofNat _ =>
|
|
|
|
match val.toNat? with
|
|
|
|
| none => throw s!"invalid Nat option value '{val}'"
|
|
|
|
| some v => pure $ opts.setNat key v
|
|
|
|
| DataValue.ofInt _ =>
|
|
|
|
match val.toInt? with
|
|
|
|
| none => throw s!"invalid Int option value '{val}'"
|
|
|
|
| some v => pure $ opts.setInt key v
|
|
|
|
| DataValue.ofSyntax _ => throw s!"invalid Syntax option value"
|
|
|
|
|
|
|
|
end Lean
|
|
|
|
|
|
|
|
|
2023-05-23 05:12:46 -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
|
|
|
|
|
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
|
|
|
|
2023-08-13 21:19:06 -07:00
|
|
|
let options? ← args.filterMap (λ s => if s.startsWith "--" then .some <| s.drop 2 else .none)
|
2023-08-14 17:07:53 -07:00
|
|
|
|>.foldlM Lean.setOptionFromString' Lean.Options.empty
|
2023-08-13 21:19:06 -07:00
|
|
|
|>.run
|
|
|
|
let options ← match options? with
|
|
|
|
| .ok options => pure options
|
|
|
|
| .error e => throw $ IO.userError s!"Options cannot be parsed: {e}"
|
2023-05-24 00:54:48 -07:00
|
|
|
let imports:= args.filter (λ s => ¬ (s.startsWith "--"))
|
|
|
|
|
2023-05-23 05:12:46 -07:00
|
|
|
let env ← Lean.importModules
|
2023-05-24 00:54:48 -07:00
|
|
|
(imports := imports.map (λ str => { module := str_to_name str, runtimeOnly := false }))
|
2023-05-23 05:12:46 -07:00
|
|
|
(opts := {})
|
|
|
|
(trustLevel := 1)
|
2023-05-20 15:58:38 -07:00
|
|
|
let context: Context := {
|
2023-08-14 17:07:53 -07:00
|
|
|
imports
|
2023-05-23 05:12:46 -07:00
|
|
|
}
|
|
|
|
let coreContext: Lean.Core.Context := {
|
2023-08-13 21:19:06 -07:00
|
|
|
currNamespace := Lean.Name.str .anonymous "Aniva"
|
2023-05-23 05:12:46 -07:00
|
|
|
openDecls := [], -- No 'open' directives needed
|
|
|
|
fileName := "<Pantograph>",
|
2023-05-24 00:54:48 -07:00
|
|
|
fileMap := { source := "", positions := #[0], lines := #[1] },
|
2023-08-13 21:19:06 -07:00
|
|
|
options := options
|
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'
|
2023-10-01 21:58:58 -07:00
|
|
|
IO.println "ready."
|
2023-05-23 05:12:46 -07:00
|
|
|
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
|