61
|
1 |
theory TimeLimit
|
63
|
2 |
imports "../Base"
|
61
|
3 |
begin
|
|
4 |
|
|
5 |
section {* Restricting the Runtime of a Function *}
|
|
6 |
|
|
7 |
text {*
|
|
8 |
{\bf Problem:}
|
63
|
9 |
Your tool should run only a specified amount of time.\smallskip
|
61
|
10 |
|
63
|
11 |
{\bf Solution:} This can be achieved using the function
|
|
12 |
@{ML timeLimit in TimeLimit}.\smallskip
|
61
|
13 |
|
63
|
14 |
Assume the following well-known function:
|
61
|
15 |
|
|
16 |
*}
|
|
17 |
|
63
|
18 |
ML {*
|
61
|
19 |
fun ackermann (0, n) = n + 1
|
|
20 |
| ackermann (m, 0) = ackermann (m - 1, 1)
|
|
21 |
| ackermann (m, n) = ackermann (m - 1, ackermann (m, n - 1))
|
|
22 |
*}
|
|
23 |
|
63
|
24 |
text {*
|
|
25 |
|
|
26 |
Now the call
|
|
27 |
|
|
28 |
@{ML_response_fake [display] "ackermann (4, 12)" "\<dots>"}
|
61
|
29 |
|
63
|
30 |
takes a bit long. To avoid this, the call can be encapsulated
|
|
31 |
in a time limit of five seconds. For this you have to write:
|
|
32 |
|
|
33 |
@{ML_response [display]
|
|
34 |
"TimeLimit.timeLimit (Time.fromSeconds 5) ackermann (4, 12)
|
|
35 |
handle TimeOut => ~1"
|
|
36 |
"~1"}
|
61
|
37 |
|
63
|
38 |
where @{ML_text TimeOut} is the exception raised when the time limit
|
|
39 |
is reached.
|
|
40 |
|
|
41 |
Note that @{ML "timeLimit" in TimeLimit} is only meaningful when you use PolyML,
|
|
42 |
because PolyML has a rich infrastructure for multithreading programming on
|
|
43 |
which @{ML "timeLimit" in TimeLimit} relies.
|
61
|
44 |
|
63
|
45 |
\begin{readmore}
|
|
46 |
The function @{ML "timeLimit" in TimeLimit} is defined in the structure
|
|
47 |
@{ML_struct TimeLimit} which can be found in the file
|
|
48 |
@{ML_file "Pure/ML-Systems/multithreading_polyml.ML"}.
|
|
49 |
\end{readmore}
|
|
50 |
|
|
51 |
|
61
|
52 |
*}
|
|
53 |
|
|
54 |
end |