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-20 14:04:09 -07:00
|
|
|
|
import Pantograph.IO
|
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
|
|
|
|
|
coreContext: Lean.Core.Context
|
|
|
|
|
|
2023-05-20 14:04:09 -07:00
|
|
|
|
/-- Stores state of the REPL -/
|
|
|
|
|
structure State where
|
2023-05-20 15:58:38 -07:00
|
|
|
|
environments: Array Lean.Environment := #[]
|
2023-05-20 14:04:09 -07:00
|
|
|
|
|
|
|
|
|
-- State monad
|
2023-05-20 15:58:38 -07:00
|
|
|
|
abbrev Subroutine := ReaderT Context (StateT State IO)
|
2023-05-20 14:04:09 -07:00
|
|
|
|
|
|
|
|
|
def nextId (s: State): Nat := s.environments.size
|
|
|
|
|
|
|
|
|
|
def State.getEnv (state: State) (id: Nat): Except String Lean.Environment :=
|
|
|
|
|
match state.environments.get? id with
|
|
|
|
|
| .some env => return env
|
|
|
|
|
| .none => throw s!"Invalid environment id {id}"
|
|
|
|
|
|
|
|
|
|
|
2023-05-09 22:51:19 -07:00
|
|
|
|
-- Utilities
|
|
|
|
|
def option_expect (o: Option α) (error: String): Except String α :=
|
|
|
|
|
match o with
|
|
|
|
|
| .some value => return value
|
|
|
|
|
| .none => throw error
|
|
|
|
|
|
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-20 15:58:38 -07:00
|
|
|
|
unsafe def execute (command: Command): Subroutine Lean.Json := do
|
2023-05-09 18:01:09 -07:00
|
|
|
|
match command.cmd with
|
|
|
|
|
| "create" =>
|
2023-05-20 15:58:38 -07:00
|
|
|
|
match Lean.fromJson? command.payload with
|
|
|
|
|
| .ok args => create args
|
|
|
|
|
| .error x => return errorJson x
|
2023-05-09 18:01:09 -07:00
|
|
|
|
| "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-09 22:51:19 -07:00
|
|
|
|
| "clear" =>
|
|
|
|
|
-- Delete all the environments
|
2023-05-17 21:58:03 -07:00
|
|
|
|
let ret ← clear
|
|
|
|
|
return Lean.toJson ret
|
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-17 21:58:03 -07:00
|
|
|
|
| "proof.trace" =>
|
2023-05-20 15:58:38 -07:00
|
|
|
|
match Lean.fromJson? command.payload with
|
|
|
|
|
| .ok args => proof_trace args
|
|
|
|
|
| .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)
|
|
|
|
|
create (args: Create): Subroutine Lean.Json := do
|
2023-05-17 21:58:03 -07:00
|
|
|
|
let state ← get
|
|
|
|
|
let id := nextId state
|
|
|
|
|
let env ← Lean.importModules
|
|
|
|
|
(imports := args.imports.map (λ str => { module := strToName str, runtimeOnly := false }))
|
|
|
|
|
(opts := {})
|
|
|
|
|
(trustLevel := 1)
|
|
|
|
|
modify fun s => { environments := s.environments.push env }
|
|
|
|
|
let num_filtered_symbols := env.constants.fold (init := 0) (λ acc name info =>
|
|
|
|
|
acc + if is_symbol_unsafe_or_internal name info then 0 else 1)
|
2023-05-20 15:58:38 -07:00
|
|
|
|
return Lean.toJson ({
|
2023-05-17 21:58:03 -07:00
|
|
|
|
id := id,
|
|
|
|
|
symbols := env.constants.size,
|
2023-05-20 15:58:38 -07:00
|
|
|
|
filtered_symbols := num_filtered_symbols }: CreateResult)
|
|
|
|
|
catalog (args: Catalog): Subroutine Lean.Json := do
|
2023-05-17 21:58:03 -07:00
|
|
|
|
let state ← get
|
2023-05-20 15:58:38 -07:00
|
|
|
|
match state.getEnv args.id with
|
|
|
|
|
| .error error => return Lean.toJson <| errorIndex error
|
|
|
|
|
| .ok env =>
|
|
|
|
|
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 <| ({ theorems := names }: CatalogResult)
|
|
|
|
|
clear: Subroutine Lean.Json := do
|
2023-05-17 21:58:03 -07:00
|
|
|
|
let state ← get
|
2023-05-20 15:58:38 -07:00
|
|
|
|
let nEnv := state.environments.size
|
2023-05-17 21:58:03 -07:00
|
|
|
|
for env in state.environments do
|
|
|
|
|
env.freeRegions
|
2023-05-20 15:58:38 -07:00
|
|
|
|
set { state with environments := #[] }
|
|
|
|
|
return Lean.toJson ({ nEnv := nEnv }: ClearResult)
|
|
|
|
|
inspect (args: Inspect): Subroutine Lean.Json := do
|
|
|
|
|
let context ← read
|
2023-05-20 14:04:09 -07:00
|
|
|
|
let state ← get
|
2023-05-20 15:58:38 -07:00
|
|
|
|
match state.getEnv args.id with
|
|
|
|
|
| .error error => return Lean.toJson <| errorIndex error
|
|
|
|
|
| .ok env =>
|
|
|
|
|
let info? := env.find? <| strToName args.symbol
|
|
|
|
|
match info? with
|
|
|
|
|
| none => return Lean.toJson <| errorIndex s!"Symbol not found {args.symbol}"
|
|
|
|
|
| some info =>
|
|
|
|
|
let format ← IO.exprToStr
|
|
|
|
|
(env := env)
|
|
|
|
|
(coreContext := context.coreContext)
|
|
|
|
|
(expr := info.toConstantVal.type)
|
|
|
|
|
return Lean.toJson ({ type := format }: InspectResult)
|
|
|
|
|
proof_trace (args: ProofTrace): Subroutine Lean.Json := do
|
2023-05-17 21:58:03 -07:00
|
|
|
|
-- Step 1: Create tactic state
|
|
|
|
|
-- Step 2: Execute tactic
|
|
|
|
|
-- Step 3: ??
|
2023-05-20 15:58:38 -07:00
|
|
|
|
return Lean.toJson ({ expr := "test" }: ProofTraceResult)
|
2023-05-17 21:58:03 -07:00
|
|
|
|
|
|
|
|
|
|
|
|
|
|
end Pantograph
|
|
|
|
|
|
|
|
|
|
open Pantograph
|
|
|
|
|
|
2023-05-09 16:39:24 -07:00
|
|
|
|
|
2023-05-09 22:51:19 -07:00
|
|
|
|
|
2023-05-09 16:39:24 -07:00
|
|
|
|
-- Main IO functions
|
|
|
|
|
|
|
|
|
|
unsafe def getLines : IO String := do
|
|
|
|
|
match (← (← IO.getStdin).getLine) with
|
|
|
|
|
| "" => pure ""
|
|
|
|
|
| "\n" => pure "\n"
|
|
|
|
|
| line => pure <| line ++ (← getLines)
|
|
|
|
|
|
2023-05-20 15:58:38 -07:00
|
|
|
|
unsafe def loop : Subroutine Unit := do
|
|
|
|
|
let command ← (← IO.getStdin).getLine
|
|
|
|
|
match parse_command command with
|
|
|
|
|
| .error _ =>
|
|
|
|
|
-- Halt execution if command is empty
|
|
|
|
|
return ()
|
|
|
|
|
| .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-12 01:08:36 -07:00
|
|
|
|
unsafe def main : 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-20 15:58:38 -07:00
|
|
|
|
let context: Context := {
|
|
|
|
|
coreContext := {
|
|
|
|
|
currNamespace := strToName "Aniva",
|
|
|
|
|
openDecls := [], -- No 'open' directives needed
|
|
|
|
|
fileName := "<Pantograph>",
|
|
|
|
|
fileMap := { source := "", positions := #[0], lines := #[1] }
|
|
|
|
|
}
|
|
|
|
|
}
|
|
|
|
|
loop.run context |>.run' {}
|