pass lemma/tactic status to file worker

pull/43/head
Alexander Bentkamp 2 years ago
parent 4f93dbf928
commit bf7c68d4f9

@ -65,21 +65,23 @@ section Elab
-- TODO: use HashSet for allowed tactics? -- TODO: use HashSet for allowed tactics?
/-- Find all tactics in syntax object that are forbidden according to a /-- Find all tactics in syntax object that are forbidden according to a
set `allowed` of allowed tactics. -/ set `allowed` of allowed tactics. -/
partial def findForbiddenTactics (inputCtx : Parser.InputContext) (level : GameLevel) (stx : Syntax) partial def findForbiddenTactics (inputCtx : Parser.InputContext) (levelParams : Game.DidOpenLevelParams) (stx : Syntax)
: Elab.Command.CommandElabM Unit := do : Elab.Command.CommandElabM Unit := do
match stx with match stx with
| .missing => return () | .missing => return ()
| .node info kind args => | .node info kind args =>
for arg in args do for arg in args do
findForbiddenTactics inputCtx level arg findForbiddenTactics inputCtx levelParams arg
| .atom info val => | .atom info val =>
-- ignore syntax elements that do not start with a letter -- ignore syntax elements that do not start with a letter
-- and ignore "with" keyword -- and ignore "with" keyword
if 0 < val.length ∧ val.data[0]!.isAlpha ∧ val != "with" ∧ val != "at" then if 0 < val.length ∧ val.data[0]!.isAlpha ∧ val != "with" ∧ val != "at" then
if ¬ ((level.tactics.map (·.name.toString))).contains val then match levelParams.tactics.find? (fun t => t.name.toString == val) with
| none => addErrorMessage info s!"You have not unlocked the tactic '{val}' yet!"
| some tac =>
if tac.locked then
addErrorMessage info s!"You have not unlocked the tactic '{val}' yet!" addErrorMessage info s!"You have not unlocked the tactic '{val}' yet!"
else if level.disabledTactics.contains val else if tac.disabled then
(¬ level.onlyTactics.isEmpty ∧ ¬ level.onlyTactics.contains val)then
addErrorMessage info s!"The tactic '{val}' is disabled in this level!" addErrorMessage info s!"The tactic '{val}' is disabled in this level!"
| .ident info rawVal val preresolved => | .ident info rawVal val preresolved =>
let ns ← let ns ←
@ -88,10 +90,12 @@ partial def findForbiddenTactics (inputCtx : Parser.InputContext) (level : GameL
for n in ns do for n in ns do
let some (.thmInfo ..) := (← getEnv).find? n let some (.thmInfo ..) := (← getEnv).find? n
| return () -- not a theroem -> ignore | return () -- not a theroem -> ignore
if ¬ ((level.lemmas.map (·.name))).contains n then match levelParams.lemmas.find? (fun l => l.name.toString == n) with
| none => addErrorMessage info s!"You have not unlocked the lemma '{n}' yet!"
| some lem =>
if lem.locked then
addErrorMessage info s!"You have not unlocked the lemma '{n}' yet!" addErrorMessage info s!"You have not unlocked the lemma '{n}' yet!"
else if level.disabledLemmas.contains n else if lem.disabled then
(¬ level.onlyLemmas.isEmpty ∧ ¬ level.onlyLemmas.contains n)then
addErrorMessage info s!"The lemma '{n}' is disabled in this level!" addErrorMessage info s!"The lemma '{n}' is disabled in this level!"
where addErrorMessage (info : SourceInfo) (s : MessageData) := where addErrorMessage (info : SourceInfo) (s : MessageData) :=
modify fun st => { st with modify fun st => { st with
@ -104,7 +108,8 @@ where addErrorMessage (info : SourceInfo) (s : MessageData) :=
} }
open Elab Meta Expr in open Elab Meta Expr in
def compileProof (inputCtx : Parser.InputContext) (snap : Snapshot) (hasWidgets : Bool) (couldBeEndSnap : Bool) : IO Snapshot := do def compileProof (inputCtx : Parser.InputContext) (snap : Snapshot) (hasWidgets : Bool)
(couldBeEndSnap : Bool) (levelParams : Game.DidOpenLevelParams) : IO Snapshot := do
-- Recognize end snap -- Recognize end snap
if inputCtx.input.atEnd snap.mpState.pos ∧ couldBeEndSnap then if inputCtx.input.atEnd snap.mpState.pos ∧ couldBeEndSnap then
let endSnap : Snapshot := { let endSnap : Snapshot := {
@ -157,7 +162,7 @@ def compileProof (inputCtx : Parser.InputContext) (snap : Snapshot) (hasWidgets
parseResultRef.set (tacticStx, cmdParserState) parseResultRef.set (tacticStx, cmdParserState)
-- Check for forbidden tactics -- Check for forbidden tactics
findForbiddenTactics inputCtx level tacticStx findForbiddenTactics inputCtx levelParams tacticStx
-- Insert invisible `skip` command to make sure we always display the initial goal -- Insert invisible `skip` command to make sure we always display the initial goal
let skip := Syntax.node (.original default 0 default endOfWhitespace) ``Lean.Parser.Tactic.skip #[] let skip := Syntax.node (.original default 0 default endOfWhitespace) ``Lean.Parser.Tactic.skip #[]
@ -235,7 +240,7 @@ where
hOut.writeLspNotification { method := "$/game/completed", param } hOut.writeLspNotification { method := "$/game/completed", param }
/-- Elaborates the next command after `parentSnap` and emits diagnostics into `hOut`. -/ /-- Elaborates the next command after `parentSnap` and emits diagnostics into `hOut`. -/
private def nextSnap (ctx : WorkerContext) (m : DocumentMeta) (cancelTk : CancelToken) private def nextSnap (ctx : WorkerContext) (m : DocumentMeta) (cancelTk : CancelToken) (levelParams : Game.DidOpenLevelParams)
: AsyncElabM (Option Snapshot) := do : AsyncElabM (Option Snapshot) := do
cancelTk.check cancelTk.check
let s ← get let s ← get
@ -252,7 +257,7 @@ where
-- Make sure that there is at least one snap after the head snap, so that -- Make sure that there is at least one snap after the head snap, so that
-- we can see the current goal even on an empty document -- we can see the current goal even on an empty document
let couldBeEndSnap := s.snaps.size > 1 let couldBeEndSnap := s.snaps.size > 1
let snap ← compileProof m.mkInputContext lastSnap ctx.clientHasWidgets couldBeEndSnap let snap ← compileProof m.mkInputContext lastSnap ctx.clientHasWidgets couldBeEndSnap levelParams
set { s with snaps := s.snaps.push snap } set { s with snaps := s.snaps.push snap }
-- TODO(MH): check for interrupt with increased precision -- TODO(MH): check for interrupt with increased precision
cancelTk.check cancelTk.check
@ -273,7 +278,8 @@ where
return some snap return some snap
/-- Elaborates all commands after the last snap (at least the header snap is assumed to exist), emitting the diagnostics into `hOut`. -/ /-- Elaborates all commands after the last snap (at least the header snap is assumed to exist), emitting the diagnostics into `hOut`. -/
def unfoldSnaps (m : DocumentMeta) (snaps : Array Snapshot) (cancelTk : CancelToken) (startAfterMs : UInt32) def unfoldSnaps (m : DocumentMeta) (snaps : Array Snapshot) (cancelTk : CancelToken)
(startAfterMs : UInt32) (levelParams : Game.DidOpenLevelParams)
: ReaderT WorkerContext IO (AsyncList ElabTaskError Snapshot) := do : ReaderT WorkerContext IO (AsyncList ElabTaskError Snapshot) := do
let ctx ← read let ctx ← read
let some headerSnap := snaps[0]? | panic! "empty snapshots" let some headerSnap := snaps[0]? | panic! "empty snapshots"
@ -289,14 +295,14 @@ where
publishIleanInfoUpdate m ctx.hOut snaps publishIleanInfoUpdate m ctx.hOut snaps
return AsyncList.ofList snaps.toList ++ AsyncList.delayed (← EIO.asTask (ε := ElabTaskError) (prio := .dedicated) do return AsyncList.ofList snaps.toList ++ AsyncList.delayed (← EIO.asTask (ε := ElabTaskError) (prio := .dedicated) do
IO.sleep startAfterMs IO.sleep startAfterMs
AsyncList.unfoldAsync (nextSnap ctx m cancelTk) { snaps }) AsyncList.unfoldAsync (nextSnap ctx m cancelTk levelParams) { snaps })
end Elab end Elab
section Updates section Updates
/-- Given the new document, updates editable doc state. -/ /-- Given the new document, updates editable doc state. -/
def updateDocument (newMeta : DocumentMeta) : WorkerM Unit := do def updateDocument (newMeta : DocumentMeta) (levelParams : Game.DidOpenLevelParams) : WorkerM Unit := do
let ctx ← read let ctx ← read
let oldDoc := (←get).doc let oldDoc := (←get).doc
oldDoc.cancelTk.set oldDoc.cancelTk.set
@ -338,7 +344,7 @@ section Updates
validSnaps := validSnaps.dropLast validSnaps := validSnaps.dropLast
-- wait for a bit, giving the initial `cancelTk.check` in `nextCmdSnap` time to trigger -- wait for a bit, giving the initial `cancelTk.check` in `nextCmdSnap` time to trigger
-- before kicking off any expensive elaboration (TODO: make expensive elaboration cancelable) -- before kicking off any expensive elaboration (TODO: make expensive elaboration cancelable)
unfoldSnaps newMeta validSnaps.toArray cancelTk ctx unfoldSnaps newMeta validSnaps.toArray cancelTk levelParams ctx
(startAfterMs := ctx.initParams.editDelay.toUInt32) (startAfterMs := ctx.initParams.editDelay.toUInt32)
modify fun st => { st with doc := { meta := newMeta, cmdSnaps := AsyncList.delayed newSnaps, cancelTk } } modify fun st => { st with doc := { meta := newMeta, cmdSnaps := AsyncList.delayed newSnaps, cancelTk } }
@ -414,9 +420,9 @@ section Initialization
return (headerSnap, srcSearchPath) return (headerSnap, srcSearchPath)
def initializeWorker (meta : DocumentMeta) (i o e : FS.Stream) (initParams : InitializeParams) (opts : Options) def initializeWorker (meta : DocumentMeta) (i o e : FS.Stream) (initParams : InitializeParams) (opts : Options)
(levelModule : Name) : IO (WorkerContext × WorkerState) := do (levelParams : Game.DidOpenLevelParams) : IO (WorkerContext × WorkerState) := do
let clientHasWidgets := initParams.initializationOptions?.bind (·.hasWidgets?) |>.getD false let clientHasWidgets := initParams.initializationOptions?.bind (·.hasWidgets?) |>.getD false
let (headerStx, headerTask) ← compileHeader meta o opts (hasWidgets := clientHasWidgets) levelModule let (headerStx, headerTask) ← compileHeader meta o opts (hasWidgets := clientHasWidgets) levelParams.levelModule
let cancelTk ← CancelToken.new let cancelTk ← CancelToken.new
let ctx := let ctx :=
{ hIn := i { hIn := i
@ -427,7 +433,7 @@ section Initialization
clientHasWidgets clientHasWidgets
} }
let cmdSnaps ← EIO.mapTask (t := headerTask) (match · with let cmdSnaps ← EIO.mapTask (t := headerTask) (match · with
| Except.ok (s, _) => unfoldSnaps meta #[s] cancelTk ctx (startAfterMs := 0) | Except.ok (s, _) => unfoldSnaps meta #[s] cancelTk levelParams ctx (startAfterMs := 0)
| Except.error e => throw (e : ElabTaskError)) | Except.error e => throw (e : ElabTaskError))
let doc : EditableDocument := { meta, cmdSnaps := AsyncList.delayed cmdSnaps, cancelTk } let doc : EditableDocument := { meta, cmdSnaps := AsyncList.delayed cmdSnaps, cancelTk }
return (ctx, return (ctx,
@ -441,7 +447,7 @@ end Initialization
section NotificationHandling section NotificationHandling
def handleDidChange (p : DidChangeTextDocumentParams) : WorkerM Unit := do def handleDidChange (levelParams : Game.DidOpenLevelParams) (p : DidChangeTextDocumentParams) : WorkerM Unit := do
let docId := p.textDocument let docId := p.textDocument
let changes := p.contentChanges let changes := p.contentChanges
let oldDoc := (←get).doc let oldDoc := (←get).doc
@ -452,17 +458,17 @@ section NotificationHandling
IO.eprintln s!"Got outdated version number: {newVersion} ≤ {oldDoc.meta.version}" IO.eprintln s!"Got outdated version number: {newVersion} ≤ {oldDoc.meta.version}"
else if ¬ changes.isEmpty then else if ¬ changes.isEmpty then
let newDocText := foldDocumentChanges changes oldDoc.meta.text let newDocText := foldDocumentChanges changes oldDoc.meta.text
updateDocument ⟨docId.uri, newVersion, newDocText⟩ updateDocument ⟨docId.uri, newVersion, newDocText⟩ levelParams
end NotificationHandling end NotificationHandling
section MessageHandling section MessageHandling
def handleNotification (method : String) (params : Json) : WorkerM Unit := do def handleNotification (method : String) (params : Json) (levelParams : Game.DidOpenLevelParams) : WorkerM Unit := do
let handle := fun paramType [FromJson paramType] (handler : paramType → WorkerM Unit) => let handle := fun paramType [FromJson paramType] (handler : paramType → WorkerM Unit) =>
parseParams paramType params >>= handler parseParams paramType params >>= handler
match method with match method with
| "textDocument/didChange" => handle DidChangeTextDocumentParams handleDidChange | "textDocument/didChange" => handle DidChangeTextDocumentParams (handleDidChange levelParams)
| "$/cancelRequest" => handle CancelParams handleCancelRequest | "$/cancelRequest" => handle CancelParams handleCancelRequest
| "$/lean/rpc/release" => handle RpcReleaseParams handleRpcRelease | "$/lean/rpc/release" => handle RpcReleaseParams handleRpcRelease
| "$/lean/rpc/keepAlive" => handle RpcKeepAliveParams handleRpcKeepAlive | "$/lean/rpc/keepAlive" => handle RpcKeepAliveParams handleRpcKeepAlive
@ -471,7 +477,7 @@ section MessageHandling
end MessageHandling end MessageHandling
section MainLoop section MainLoop
partial def mainLoop : WorkerM Unit := do partial def mainLoop (levelParams : Game.DidOpenLevelParams) : WorkerM Unit := do
let ctx ← read let ctx ← read
let mut st ← get let mut st ← get
let msg ← ctx.hIn.readLspMessage let msg ← ctx.hIn.readLspMessage
@ -497,14 +503,14 @@ section MainLoop
match msg with match msg with
| Message.request id method (some params) => | Message.request id method (some params) =>
handleRequest id method (toJson params) handleRequest id method (toJson params)
mainLoop mainLoop levelParams
| Message.notification "exit" none => | Message.notification "exit" none =>
let doc := st.doc let doc := st.doc
doc.cancelTk.set doc.cancelTk.set
return () return ()
| Message.notification method (some params) => | Message.notification method (some params) =>
handleNotification method (toJson params) handleNotification method (toJson params) levelParams
mainLoop mainLoop levelParams
| _ => throwServerError "Got invalid JSON-RPC message" | _ => throwServerError "Got invalid JSON-RPC message"
end MainLoop end MainLoop
@ -513,7 +519,7 @@ def initAndRunWorker (i o e : FS.Stream) (opts : Options) : IO UInt32 := do
let o ← maybeTee "fwOut.txt" true o let o ← maybeTee "fwOut.txt" true o
let initParams ← i.readLspRequestAs "initialize" InitializeParams let initParams ← i.readLspRequestAs "initialize" InitializeParams
let ⟨_, param⟩ ← i.readLspNotificationAs "textDocument/didOpen" DidOpenTextDocumentParams let ⟨_, param⟩ ← i.readLspNotificationAs "textDocument/didOpen" DidOpenTextDocumentParams
let ⟨_, levelParam⟩ ← i.readLspNotificationAs "$/game/didOpenLevel" Game.DidOpenLevelParams let ⟨_, levelParams⟩ ← i.readLspNotificationAs "$/game/didOpenLevel" Game.DidOpenLevelParams
let doc := param.textDocument let doc := param.textDocument
/- NOTE(WN): `toFileMap` marks line beginnings as immediately following /- NOTE(WN): `toFileMap` marks line beginnings as immediately following
@ -525,8 +531,8 @@ def initAndRunWorker (i o e : FS.Stream) (opts : Options) : IO UInt32 := do
let e := e.withPrefix s!"[{param.textDocument.uri}] " let e := e.withPrefix s!"[{param.textDocument.uri}] "
let _ ← IO.setStderr e let _ ← IO.setStderr e
try try
let (ctx, st) ← initializeWorker meta i o e initParams.param opts levelParam.levelModule let (ctx, st) ← initializeWorker meta i o e initParams.param opts levelParams
let _ ← StateRefT'.run (s := st) <| ReaderT.run (r := ctx) mainLoop let _ ← StateRefT'.run (s := st) <| ReaderT.run (r := ctx) (mainLoop levelParams)
return (0 : UInt32) return (0 : UInt32)
catch e => catch e =>
IO.eprintln e IO.eprintln e

@ -57,6 +57,8 @@ structure LoadLevelParams where
structure DidOpenLevelParams where structure DidOpenLevelParams where
uri : String uri : String
levelModule : Name levelModule : Name
tactics: Array Availability
lemmas: Array Availability
deriving ToJson, FromJson deriving ToJson, FromJson
structure Doc where structure Doc where
@ -103,6 +105,8 @@ def handleDidOpenLevel (params : Json) : GameServerM Unit := do
param := { param := {
uri := m.uri uri := m.uri
levelModule := lvl.module levelModule := lvl.module
tactics := lvl.tactics
lemmas := lvl.lemmas
: DidOpenLevelParams : DidOpenLevelParams
} }
} }

Loading…
Cancel
Save