Skip to content

Commit

Permalink
feat: report sorries in tactic mode
Browse files Browse the repository at this point in the history
  • Loading branch information
kim-em committed Jan 13, 2024
1 parent 7df9e30 commit 5adf21a
Show file tree
Hide file tree
Showing 5 changed files with 54 additions and 3 deletions.
14 changes: 12 additions & 2 deletions REPL/JSON.lean
Original file line number Diff line number Diff line change
Expand Up @@ -68,7 +68,15 @@ structure Sorry where
You can use the `ProofStep` instruction to run a tactic at this state.
-/
proofState : Option Nat
deriving ToJson, FromJson
deriving FromJson

instance : ToJson Sorry where
toJson r := Json.mkObj <| .join [
[("goal", r.goal)],
[("proofState", toJson r.proofState)],
if r.pos.line ≠ 0 then [("pos", toJson r.pos)] else [],
if r.endPos.line ≠ 0 then [("endPos", toJson r.endPos)] else [],
]

/-- Construct the JSON representation of a Lean sorry. -/
def Sorry.of (goal : String) (pos endPos : Lean.Position) (proofState : Option Nat) : Sorry :=
Expand Down Expand Up @@ -124,13 +132,15 @@ structure ProofStepResponse where
proofState : Nat
goals : List String
messages : List Message := []
sorries : List Sorry := []
deriving ToJson, FromJson

instance : ToJson ProofStepResponse where
toJson r := Json.mkObj <| .join [
[("proofState", r.proofState)],
[("goals", toJson r.goals)],
Json.nonemptyList "messages" r.messages
Json.nonemptyList "messages" r.messages,
Json.nonemptyList "sorries" r.sorries
]

/-- Json wrapper for an error. -/
Expand Down
20 changes: 19 additions & 1 deletion REPL/Main.lean
Original file line number Diff line number Diff line change
Expand Up @@ -128,8 +128,26 @@ def createProofStepReponse (proofState : ProofSnapshot) (old? : Option ProofSnap
M m ProofStepResponse := do
let messages := proofState.newMessages old?
let messages ← messages.mapM fun m => Message.of m
let trees := proofState.newInfoTrees old?
let trees := match old? with
| some old =>
let ctx : ContextInfo :=
{ env := old.coreState.env
ngen := old.coreState.ngen
fileMap := old.coreContext.fileMap
options := old.coreContext.options
currNamespace := old.coreContext.currNamespace
openDecls := old.coreContext.openDecls
mctx := old.metaState.mctx }
trees.map fun t => InfoTree.context ctx t
| none => trees
let sorries ← sorries trees
let id ← recordProofSnapshot proofState
return { proofState := id, goals := (← proofState.ppGoals).map fun s => s!"{s}", messages }
return {
proofState := id
goals := (← proofState.ppGoals).map fun s => s!"{s}"
messages
sorries }

/-- Pickle an `CommandSnapshot`, generating a JSON response. -/
def pickleCommandSnapshot (n : PickleEnvironment) : M m (CommandResponse ⊕ Error) := do
Expand Down
6 changes: 6 additions & 0 deletions REPL/Snapshots.lean
Original file line number Diff line number Diff line change
Expand Up @@ -118,6 +118,12 @@ def newMessages (new : ProofSnapshot) (old? : Option ProofSnapshot := none) : Li
| none => new.coreState.messages.msgs.toList
| some old => new.coreState.messages.msgs.toList.drop (old.coreState.messages.msgs.size)

/-- New info trees in a `ProofSnapshot`, relative to an optional previous `ProofSnapshot`. -/
def newInfoTrees (new : ProofSnapshot) (old? : Option ProofSnapshot := none) : List InfoTree :=
match old? with
| none => new.coreState.infoState.trees.toList
| some old => new.coreState.infoState.trees.toList.drop (old.coreState.infoState.trees.size)

/-- Run a `CoreM` monadic function in the current `ProofSnapshot`, updating the `Core.State`. -/
def runCoreM (p : ProofSnapshot) (t : CoreM α) : IO (α × ProofSnapshot) := do
let (a, coreState) ← (Lean.Core.CoreM.toIO · p.coreContext p.coreState) do t
Expand Down
14 changes: 14 additions & 0 deletions test/tactic_mode_sorry.expected.out
Original file line number Diff line number Diff line change
@@ -0,0 +1,14 @@
{"sorries":
[{"proofState": 0,
"pos": {"line": 1, "column": 18},
"goal": "⊢ Nat",
"endPos": {"line": 1, "column": 23}}],
"messages":
[{"severity": "warning",
"pos": {"line": 1, "column": 4},
"endPos": {"line": 1, "column": 5},
"data": "declaration uses 'sorry'"}],
"env": 0}

{"sorries": [{"proofState": 1, "goal": "⊢ Nat"}], "proofState": 2, "goals": []}

3 changes: 3 additions & 0 deletions test/tactic_mode_sorry.in
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
{"cmd": "def f : Nat := by sorry"}

{"proofState": 0, "tactic": "sorry"}

0 comments on commit 5adf21a

Please sign in to comment.