CookBook/Recipes/TimeLimit.thy
author Christian Urban <urbanc@in.tum.de>
Sat, 14 Feb 2009 16:09:04 +0000
changeset 119 4536782969fa
parent 102 5e309df58557
child 155 b6fca043a796
permissions -rw-r--r--
added an acknowledgement section
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     1
theory TimeLimit
63
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
     2
imports "../Base"
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     3
begin
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     4
94
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 72
diff changeset
     5
section {* Restricting the Runtime of a Function\label{rec:timeout} *} 
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     6
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     7
text {*
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     8
  {\bf Problem:}
63
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
     9
  Your tool should run only a specified amount of time.\smallskip
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    10
63
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    11
  {\bf Solution:} This can be achieved using the function 
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    12
  @{ML timeLimit in TimeLimit}.\smallskip
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    13
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
    14
  Assume you defined the Ackermann function:
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    15
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    16
  *}
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    17
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    18
ML{*fun ackermann (0, n) = n + 1
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    19
  | 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
    20
  | ackermann (m, n) = ackermann (m - 1, ackermann (m, n - 1)) *}
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    21
63
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    22
text {*
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    23
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    24
  Now the call 
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    25
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
    26
  @{ML_response_fake [display,gray] "ackermann (4, 12)" "\<dots>"}
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    27
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
    28
  takes a bit of time before it finishes. To avoid this, the call can be encapsulated 
63
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    29
  in a time limit of five seconds. For this you have to write:
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    30
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
    31
@{ML_response [display,gray]
63
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    32
"TimeLimit.timeLimit (Time.fromSeconds 5) ackermann (4, 12) 
94
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 72
diff changeset
    33
  handle TimeLimit.TimeOut => ~1"
63
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    34
"~1"}
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    35
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
    36
  where @{text TimeOut} is the exception raised when the time limit
63
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    37
  is reached.
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    38
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    39
  Note that @{ML  "timeLimit" in TimeLimit} is only meaningful when you use PolyML, 
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    40
  because PolyML has a rich infrastructure for multithreading programming on 
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    41
  which @{ML "timeLimit" in TimeLimit} relies.
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    42
63
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    43
\begin{readmore}
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    44
   The function @{ML "timeLimit" in TimeLimit} is defined in the structure
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    45
  @{ML_struct TimeLimit} which can be found in the file 
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    46
  @{ML_file "Pure/ML-Systems/multithreading_polyml.ML"}.
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    47
\end{readmore}
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    48
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    49
 
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    50
*}
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    51
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    52
end