Skip to content

Commit 9dad7b3

Browse files
Implement timeout option
1 parent 506a6f3 commit 9dad7b3

File tree

4 files changed

+101
-14
lines changed

4 files changed

+101
-14
lines changed

REPL/JSON.lean

Lines changed: 13 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -11,7 +11,14 @@ open Lean Elab InfoTree
1111

1212
namespace REPL
1313

14-
structure CommandOptions where
14+
/-- Base structure with timeout that all command types can inherit from -/
15+
structure BaseOptions where
16+
/--
17+
Optional timeout in milliseconds. If none, no timeout will be applied.
18+
-/
19+
timeout : Option Nat := none
20+
21+
structure CommandOptions extends BaseOptions where
1522
allTactics : Option Bool := none
1623
rootGoals : Option Bool := none
1724
/--
@@ -37,7 +44,7 @@ deriving FromJson
3744
/--
3845
Run a tactic in a proof state.
3946
-/
40-
structure ProofStep where
47+
structure ProofStep extends BaseOptions where
4148
proofState : Nat
4249
tactic : String
4350
deriving ToJson, FromJson
@@ -169,21 +176,21 @@ structure Error where
169176
message : String
170177
deriving ToJson, FromJson
171178

172-
structure PickleEnvironment where
179+
structure PickleEnvironment extends BaseOptions where
173180
env : Nat
174181
pickleTo : System.FilePath
175182
deriving ToJson, FromJson
176183

177-
structure UnpickleEnvironment where
184+
structure UnpickleEnvironment extends BaseOptions where
178185
unpickleEnvFrom : System.FilePath
179186
deriving ToJson, FromJson
180187

181-
structure PickleProofState where
188+
structure PickleProofState extends BaseOptions where
182189
proofState : Nat
183190
pickleTo : System.FilePath
184191
deriving ToJson, FromJson
185192

186-
structure UnpickleProofState where
193+
structure UnpickleProofState extends BaseOptions where
187194
unpickleProofStateFrom : System.FilePath
188195
env : Option Nat
189196
deriving ToJson, FromJson

REPL/Main.lean

Lines changed: 27 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -361,6 +361,23 @@ def runProofStep (s : ProofStep) : M IO (ProofStepResponse ⊕ Error) := do
361361
catch ex =>
362362
return .inr ⟨"Lean error:\n" ++ ex.toString⟩
363363

364+
/--
365+
Run a task with an optional timeout in milliseconds.
366+
-/
367+
def runWithTimeout (timeout? : Option Nat) (task : M IO α) : M IO α :=
368+
match timeout? with
369+
| none => task
370+
| some timeout => do
371+
let jobTask ← IO.asTask (prio := .dedicated) (StateT.run task (← get))
372+
let timeoutTask : IO (α × State) := do
373+
IO.sleep timeout.toUInt32
374+
throw <| IO.userError s!"Operation timed out after {timeout}ms"
375+
match ← IO.waitAny [jobTask, ← IO.asTask timeoutTask] with
376+
| .ok (a, state) => do
377+
modify fun _ => state
378+
return a
379+
| .error e => throw e
380+
364381
end REPL
365382

366383
open REPL
@@ -425,14 +442,16 @@ where loop : M IO Unit := do
425442
if query = "" then
426443
return ()
427444
if query.startsWith "#" || query.startsWith "--" then loop else
428-
IO.println <| toString <| ← match ← parse query with
429-
| .command r => return toJson (← runCommand r)
430-
| .file r => return toJson (← processFile r)
431-
| .proofStep r => return toJson (← runProofStep r)
432-
| .pickleEnvironment r => return toJson (← pickleCommandSnapshot r)
433-
| .unpickleEnvironment r => return toJson (← unpickleCommandSnapshot r)
434-
| .pickleProofSnapshot r => return toJson (← pickleProofSnapshot r)
435-
| .unpickleProofSnapshot r => return toJson (← unpickleProofSnapshot r)
445+
IO.println <| toString <| ← try
446+
match ← parse query with
447+
| .command r => return toJson (← runWithTimeout r.timeout (runCommand r))
448+
| .file r => return toJson (← runWithTimeout r.timeout (processFile r))
449+
| .proofStep r => return toJson (← runWithTimeout r.timeout (runProofStep r))
450+
| .pickleEnvironment r => return toJson (← runWithTimeout r.timeout (pickleCommandSnapshot r))
451+
| .unpickleEnvironment r => return toJson (← runWithTimeout r.timeout (unpickleCommandSnapshot r))
452+
| .pickleProofSnapshot r => return toJson (← runWithTimeout r.timeout (pickleProofSnapshot r))
453+
| .unpickleProofSnapshot r => return toJson (← runWithTimeout r.timeout (unpickleProofSnapshot r))
454+
catch e => return toJson (⟨e.toString⟩ : Error)
436455
printFlush "\n" -- easier to parse the output if there are blank lines
437456
loop
438457

test/timeout.expected.out

Lines changed: 46 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,46 @@
1+
{"sorries":
2+
[{"proofState": 0,
3+
"pos": {"line": 1, "column": 49},
4+
"goal": "x : Nat\nh1 : x = 2\n⊢ x = 2",
5+
"endPos": {"line": 1, "column": 54}}],
6+
"messages":
7+
[{"severity": "warning",
8+
"pos": {"line": 1, "column": 8},
9+
"endPos": {"line": 1, "column": 10},
10+
"data": "declaration uses 'sorry'"}],
11+
"env": 0}
12+
13+
{"proofStatus": "Completed", "proofState": 1, "goals": []}
14+
15+
{"sorries":
16+
[{"proofState": 2,
17+
"pos": {"line": 1, "column": 49},
18+
"goal": "x : Nat\nh1 : x = 2\n⊢ x = 2",
19+
"endPos": {"line": 1, "column": 54}}],
20+
"messages":
21+
[{"severity": "warning",
22+
"pos": {"line": 1, "column": 8},
23+
"endPos": {"line": 1, "column": 10},
24+
"data": "declaration uses 'sorry'"}],
25+
"env": 1}
26+
27+
{"proofStatus": "Completed", "proofState": 3, "goals": []}
28+
29+
{"message": "Operation timed out after 1ms"}
30+
31+
{"message": "Unknown proof state."}
32+
33+
{"sorries":
34+
[{"proofState": 4,
35+
"pos": {"line": 1, "column": 49},
36+
"goal": "x : Nat\nh1 : x = 2\n⊢ x = 2",
37+
"endPos": {"line": 1, "column": 54}}],
38+
"messages":
39+
[{"severity": "warning",
40+
"pos": {"line": 1, "column": 8},
41+
"endPos": {"line": 1, "column": 10},
42+
"data": "declaration uses 'sorry'"}],
43+
"env": 2}
44+
45+
{"proofStatus": "Completed", "proofState": 5, "goals": []}
46+

test/timeout.in

Lines changed: 15 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,15 @@
1+
{"cmd": "theorem aa (x : Nat) (h1 : x = 2) : x = 2 := by sorry"}
2+
3+
{"tactic": "assumption", "proofState": 0}
4+
5+
{"cmd": "theorem aa (x : Nat) (h1 : x = 2) : x = 2 := by sorry", "timeout": 1000}
6+
7+
{"tactic": "assumption", "proofState": 2}
8+
9+
{"cmd": "theorem aa (x : Nat) (h1 : x = 2) : x = 2 := by sorry", "timeout": 1}
10+
11+
{"tactic": "assumption", "proofState": 4}
12+
13+
{"cmd": "theorem aa (x : Nat) (h1 : x = 2) : x = 2 := by sorry"}
14+
15+
{"tactic": "assumption", "proofState": 4}

0 commit comments

Comments
 (0)