author | Christian Urban <urbanc@in.tum.de> |
Sun, 01 Mar 2009 21:48:59 +0000 | |
changeset 154 | e81ebb37aa83 |
child 155 | b6fca043a796 |
permissions | -rw-r--r-- |
154
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1 |
theory Timing |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
2 |
imports "../Base" |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
3 |
begin |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
4 |
|
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
5 |
section {* Measuring the Time of an Operation\label{rec:timing} *} |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
6 |
|
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
7 |
text {* |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
8 |
{\bf Problem:} |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
9 |
You want to measure the running time of a tactic or function.\smallskip |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
10 |
|
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
11 |
{\bf Solution:} Time can be measured using the function |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
12 |
@{ML start_timing} and @{ML end_timing}.\smallskip |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
13 |
|
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
14 |
*} |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
15 |
|
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
16 |
ML{*fun ackermann (0, n) = n + 1 |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
17 |
| ackermann (m, 0) = ackermann (m - 1, 1) |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
18 |
| ackermann (m, n) = ackermann (m - 1, ackermann (m, n - 1)) *} |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
19 |
|
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
20 |
ML {* |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
21 |
let |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
22 |
val start = start_timing () |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
23 |
val res = ackermann (3,12) |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
24 |
val time = (end_timing start) |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
25 |
in |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
26 |
(res,time) |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
27 |
end |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
28 |
*} |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
29 |
|
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
30 |
text {* |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
31 |
check what @{text "#all"} does? |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
32 |
*} |
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
33 |
|
e81ebb37aa83
updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
34 |
end |