| author | Christian Urban <urbanc@in.tum.de> |
| Fri, 13 Mar 2009 01:15:55 +0100 | |
| changeset 174 | a29b81d4fa88 |
| parent 168 | 009ca4807baa |
| child 175 | 7c09bd3227c5 |
| permissions | -rw-r--r-- |
| 61 | 1 |
theory TimeLimit |
| 63 | 2 |
imports "../Base" |
| 61 | 3 |
begin |
4 |
||
| 94 | 5 |
section {* Restricting the Runtime of a Function\label{rec:timeout} *}
|
| 61 | 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 |
|
|
168
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
155
diff
changeset
|
14 |
Assume you defined the Ackermann function on the ML-level. |
|
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
155
diff
changeset
|
15 |
*} |
| 61 | 16 |
|
|
69
19106a9975c1
highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents:
68
diff
changeset
|
17 |
ML{*fun ackermann (0, n) = n + 1
|
| 61 | 18 |
| ackermann (m, 0) = ackermann (m - 1, 1) |
|
69
19106a9975c1
highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents:
68
diff
changeset
|
19 |
| ackermann (m, n) = ackermann (m - 1, ackermann (m, n - 1)) *} |
| 61 | 20 |
|
| 63 | 21 |
text {*
|
22 |
||
23 |
Now the call |
|
24 |
||
|
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents:
69
diff
changeset
|
25 |
@{ML_response_fake [display,gray] "ackermann (4, 12)" "\<dots>"}
|
| 61 | 26 |
|
|
68
e7519207c2b7
added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents:
67
diff
changeset
|
27 |
takes a bit of time before it finishes. To avoid this, the call can be encapsulated |
|
168
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
155
diff
changeset
|
28 |
in a time limit of five seconds. For this you have to write |
| 63 | 29 |
|
|
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents:
69
diff
changeset
|
30 |
@{ML_response [display,gray]
|
| 63 | 31 |
"TimeLimit.timeLimit (Time.fromSeconds 5) ackermann (4, 12) |
| 94 | 32 |
handle TimeLimit.TimeOut => ~1" |
| 63 | 33 |
"~1"} |
| 61 | 34 |
|
|
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
94
diff
changeset
|
35 |
where @{text TimeOut} is the exception raised when the time limit
|
| 63 | 36 |
is reached. |
37 |
||
38 |
Note that @{ML "timeLimit" in TimeLimit} is only meaningful when you use PolyML,
|
|
|
168
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
155
diff
changeset
|
39 |
because PolyML has the infrastructure for multithreading programming on |
| 63 | 40 |
which @{ML "timeLimit" in TimeLimit} relies.
|
| 61 | 41 |
|
| 63 | 42 |
\begin{readmore}
|
43 |
The function @{ML "timeLimit" in TimeLimit} is defined in the structure
|
|
44 |
@{ML_struct TimeLimit} which can be found in the file
|
|
45 |
@{ML_file "Pure/ML-Systems/multithreading_polyml.ML"}.
|
|
46 |
\end{readmore}
|
|
47 |
||
48 |
||
| 61 | 49 |
*} |
50 |
end |