Skip to content

Commit

Permalink
plug-in variables in hints client-side
Browse files Browse the repository at this point in the history
  • Loading branch information
joneugster committed Feb 16, 2024
1 parent 3775ad9 commit 2649f98
Show file tree
Hide file tree
Showing 4 changed files with 66 additions and 6 deletions.
24 changes: 21 additions & 3 deletions client/src/components/hints.tsx
Original file line number Diff line number Diff line change
Expand Up @@ -5,15 +5,33 @@ import { DeletedChatContext, ProofContext } from "./infoview/context";
import { lastStepHasErrors } from "./infoview/goals";
import { Button } from "./button";

/** Plug-in the variable names in a hint. We do this client-side to prepare
* for i18n in the future. i.e. one should be able translate the `rawText`
* and have the variables substituted just before displaying.
*/
function getHintText(hint: GameHint): string {
if (hint.rawText) {
// Replace the variable names used in the hint with the ones used by the player
// variable names are marked like `«{g}»` inside the text.
return hint.rawText.replace(/«\{(.*?)\}»/, ((_, v) =>
// `hint.varNames` contains tuples `[oldName, newName]`
(hint.varNames.find(x => x[0] == v))[1]))
} else {
// hints created in the frontend do not have a `rawText`
// TODO: `hint.text` could be removed in theory.
return hint.text
}
}

export function Hint({hint, step, selected, toggleSelection, lastLevel} : {hint: GameHint, step: number, selected: number, toggleSelection: any, lastLevel?: boolean}) {
return <div className={`message information step-${step}` + (step == selected ? ' selected' : '') + (lastLevel ? ' recent' : '')} onClick={toggleSelection}>
<Markdown>{hint.text}</Markdown>
<Markdown>{getHintText(hint)}</Markdown>
</div>
}

export function HiddenHint({hint, step, selected, toggleSelection, lastLevel} : {hint: GameHint, step: number, selected: number, toggleSelection: any, lastLevel?: boolean}) {
return <div className={`message warning step-${step}` + (step == selected ? ' selected' : '') + (lastLevel ? ' recent' : '')} onClick={toggleSelection}>
<Markdown>{hint.text}</Markdown>
<Markdown>{getHintText(hint)}</Markdown>
</div>
}

Expand All @@ -31,7 +49,7 @@ export function Hints({hints, showHidden, step, selected, toggleSelection, lastL

export function DeletedHint({hint} : {hint: GameHint}) {
return <div className="message information deleted-hint">
<Markdown>{hint.text}</Markdown>
<Markdown>{getHintText(hint)}</Markdown>
</div>
}

Expand Down
2 changes: 2 additions & 0 deletions client/src/components/infoview/rpc_api.ts
Original file line number Diff line number Diff line change
Expand Up @@ -49,6 +49,8 @@ export interface InteractiveTermGoal extends InteractiveGoalCore {
export interface GameHint {
text: string;
hidden: boolean;
rawText: string;
varNames: string[][]; // in Lean: `Array (Name × Name)`
}

export interface InteractiveGoalWithHints {
Expand Down
37 changes: 34 additions & 3 deletions server/GameServer/RpcHandlers.lean
Original file line number Diff line number Diff line change
Expand Up @@ -121,14 +121,45 @@ def findHints (goal : MVarId) (m : DocumentMeta) (initParams : Lsp.InitializePar
openAbstractCtxResult hint.goal fun hintFVars hintGoal => do
if let some fvarBij := matchExpr (← instantiateMVars $ hintGoal) (← instantiateMVars $ ← inferType $ mkMVar goal)
then
let lctx := (← goal.getDecl).lctx
if let some bij ← matchDecls hintFVars lctx.getFVars (strict := hint.strict) (initBij := fvarBij)

-- NOTE: This is a bit a hack of somebody who does not know how meta-programming works.
-- All we want here is a list of `userNames` for the `FVarId`s in `hintFVars`...
-- and we wrap them in `«{}»` here since I don't know how to do it later.
let mut hintFVarsNames : Array Expr := #[]
for fvar in hintFVars do
let name₁ ← fvar.fvarId!.getUserName
hintFVarsNames := hintFVarsNames.push <| Expr.fvar ⟨s!"«\{{name₁}}»"

let lctx := (← goal.getDecl).lctx -- the player's local context
if let some bij ← matchDecls hintFVars lctx.getFVars
(strict := hint.strict) (initBij := fvarBij)
then
let userFVars := hintFVars.map fun v => bij.forward.findD v.fvarId! v.fvarId!
-- Evaluate the text in the player's context to get the new variable names.
let text := (← evalHintMessage hint.text) (userFVars.map Expr.fvar)
let ctx := {env := ← getEnv, mctx := ← getMCtx, lctx := lctx, opts := {}}
let text ← (MessageData.withContext ctx text).toString
return some { text := text, hidden := hint.hidden }

-- Evaluate the text in the `Hint`'s context to get the old variable names.
let rawText := (← evalHintMessage hint.text) hintFVarsNames
let ctx₂ := {env := ← getEnv, mctx := ← getMCtx, lctx := ← getLCtx, opts := {}}
let rawText ← (MessageData.withContext ctx₂ rawText).toString

-- Here we map the goal's variable names to the player's variable names.
let mut varNames : Array <| Name × Name := #[]
for (fvar₁, fvar₂) in bij.forward.toArray do
-- get the `userName` of the fvar in the opened local context of the hint.
let name₁ ← fvar₁.getUserName
-- get the `userName` in the player's local context.
let name₂ := (lctx.get! fvar₂).userName
varNames := varNames.push (name₁, name₂)

return some {
text := text,
hidden := hint.hidden,
rawText := rawText,
varNames := varNames }

else return none
else
return none
Expand Down
9 changes: 9 additions & 0 deletions server/GameServer/Structures.lean
Original file line number Diff line number Diff line change
Expand Up @@ -54,8 +54,17 @@ deriving RpcEncodable

/-- A hint in the game at the corresponding goal. -/
structure GameHint where
/-- The text with the variable names already inserted.
Note: This is in theory superfluous and will be completely replaced by `rawText`. We just left
it in for debugging for now. -/
text : String
/-- Flag whether the hint should be hidden initially. -/
hidden : Bool
/-- The text with the variables not inserted yet. -/
rawText : String
/-- The assignment of variable names in the `rawText` to the ones the player used. -/
varNames : Array <| Name × Name
deriving FromJson, ToJson

/-- Bundled `InteractiveGoal` together with an array of hints that apply at this stage. -/
Expand Down

0 comments on commit 2649f98

Please sign in to comment.