You cannot select more than 25 topics Topics must start with a letter or number, can include dashes ('-') and can be up to 35 characters long.
lean4game/server/GameServer/WasmServer.lean

92 lines
2.9 KiB
Plaintext

This file contains ambiguous Unicode characters!

This file contains ambiguous Unicode characters that may be confused with others in your current locale. If your use case is intentional and legitimate, you can safely ignore this warning. Use the Escape button to highlight these characters.

import Lean.Server.Watchdog
import GameServer.Game
namespace WasmServer.Watchdog
open Lean
open Server
open Watchdog
open IO
open Lsp
open JsonRpc
open System.Uri
def ServerState := GameServerState
@[export game_make_state]
def makeState : IO (Ref ServerState) := do
let e ← IO.getStderr
try
searchPathRef.set ["/lib", "/gamelib"]
let state : GameServerState := {
env := ← importModules #[
{ module := `Init : Import }
-- { module := `GameServer : Import }
] {} 0 --← createEnv gameDir module,
game := "TEST",
gameDir := "test",
inventory := #[]
difficulty := 0
}
return ← IO.mkRef state
catch err =>
e.putStrLn s!"Import error: {err}"
throw err
def readMessage (s : String) : IO JsonRpc.Message := do
let j ← ofExcept (Json.parse s)
let m ← match fromJson? j with
| Except.ok (m : JsonRpc.Message) => pure m
| Except.error inner => throw $ userError s!"JSON '{j.compress}' did not have the format of a JSON-RPC message.\n{inner}"
return m
def readLspRequestAs (s : String) (expectedMethod : String) (α : Type) [FromJson α] : IO (Request α) := do
let m ← readMessage s
match m with
| Message.request id method params? =>
if method = expectedMethod then
let j := toJson params?
match fromJson? j with
| Except.ok v => pure $ JsonRpc.Request.mk id expectedMethod (v : α)
| Except.error inner => throw $ userError s!"Unexpected param '{j.compress}' for method '{expectedMethod}'\n{inner}"
else
throw $ userError s!"Expected method '{expectedMethod}', got method '{method}'"
| _ => throw $ userError s!"Expected JSON-RPC request, got: '{(toJson m).compress}'"
def initializeServer (id : RequestID) : IO Unit := do
let o ← IO.getStdout
o.writeLspResponse {
id := id
result := {
capabilities := mkLeanServerCapabilities
serverInfo? := some {
name := "Lean 4 Game Server"
version? := "0.1.1"
}
: InitializeResult
}
}
return ()
@[export game_send_message]
def sendMessage (s : String) (state : Ref ServerState) : IO Unit := do
let o ← IO.getStdout
let e ← IO.getStderr
try
let m ← readMessage s
match m with
| Message.request id "initialize" params? =>
initializeServer id
| Message.request id "info" _ =>
let some game := (gameExt.getState (← state.get).env).find? `TestGame
| throwServerError "Game not found"
let gameJson : Json := toJson game
-- Add world sizes to Json object
let worldSize := game.worlds.nodes.toList.map (fun (n, w) => (n.toString, w.levels.size))
let gameJson := gameJson.mergeObj (Json.mkObj [("worldSize", Json.mkObj worldSize)])
o.writeLspResponse ⟨id, gameJson⟩
| _ => throw $ userError s!"Expected JSON-RPC request, got: '{(toJson m).compress}'"
catch err =>
e.putStrLn s!"Server error: {err}"
return ()