ProgTutorial/Recipes/Timing.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Wed, 28 May 2014 12:41:09 +0100
changeset 556 3c214b215f7e
parent 544 501491d56798
child 565 cecd7a941885
permissions -rw-r--r--
some small updates for Isabelle and corrections in the Parsing chapter
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
154
e81ebb37aa83 updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Timing
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
     2
imports "../Appendix"
154
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
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
     5
section {* Measuring Time\label{rec:timing} *} 
154
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 
460
5c33c4b52ad7 Updated to changes in timing structure
Christian Urban <urbanc@in.tum.de>
parents: 456
diff changeset
    12
  @{ML start in Timing} and @{ML result in Timing}.\smallskip
154
e81ebb37aa83 updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
191
0150cf5982ae polished some recipies
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    14
  Suppose you defined the Ackermann function on the Isabelle level. 
154
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
155
Christian Urban <urbanc@in.tum.de>
parents: 154
diff changeset
    17
fun 
168
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    18
 ackermann:: "(nat \<times> nat) \<Rightarrow> nat"
155
Christian Urban <urbanc@in.tum.de>
parents: 154
diff changeset
    19
where
Christian Urban <urbanc@in.tum.de>
parents: 154
diff changeset
    20
    "ackermann (0, n) = n + 1"
Christian Urban <urbanc@in.tum.de>
parents: 154
diff changeset
    21
  | "ackermann (m, 0) = ackermann (m - 1, 1)"
Christian Urban <urbanc@in.tum.de>
parents: 154
diff changeset
    22
  | "ackermann (m, n) = ackermann (m - 1, ackermann (m, n - 1))"
Christian Urban <urbanc@in.tum.de>
parents: 154
diff changeset
    23
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    24
text {* 
168
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    25
  You can measure how long the simplifier takes to verify a datapoint
191
0150cf5982ae polished some recipies
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    26
  of this function. The actual timing is done inside the
175
7c09bd3227c5 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
    27
  wrapper function:
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    28
*}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    29
185
043ef82000b4 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 175
diff changeset
    30
ML %linenosgray{*fun timing_wrapper tac st =
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    31
let 
460
5c33c4b52ad7 Updated to changes in timing structure
Christian Urban <urbanc@in.tum.de>
parents: 456
diff changeset
    32
  val t_start = Timing.start ();
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    33
  val res = tac st;
460
5c33c4b52ad7 Updated to changes in timing structure
Christian Urban <urbanc@in.tum.de>
parents: 456
diff changeset
    34
  val t_end = Timing.result t_start;
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    35
in
460
5c33c4b52ad7 Updated to changes in timing structure
Christian Urban <urbanc@in.tum.de>
parents: 456
diff changeset
    36
  (writeln (Timing.message t_end); res)
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    37
end*}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    38
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    39
text {*
175
7c09bd3227c5 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
    40
  Note that this function, in addition to a tactic, also takes a state @{text
185
043ef82000b4 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 175
diff changeset
    41
  "st"} as argument and applies this state to the tactic (Line 4). The reason is that
175
7c09bd3227c5 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
    42
  tactics are lazy functions and you need to force them to run, otherwise the
191
0150cf5982ae polished some recipies
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    43
  timing will be meaningless. The simplifier tactic, amongst others,  can be 
0150cf5982ae polished some recipies
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    44
  forced to run by just applying the state to it. But ``fully'' lazy tactics,
0150cf5982ae polished some recipies
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    45
  such as @{ML "resolve_tac"}, need even more ``standing-on-ones-head'' to force
0150cf5982ae polished some recipies
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    46
  them to run. 
0150cf5982ae polished some recipies
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    47
0150cf5982ae polished some recipies
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    48
  The time between start and finish of the simplifier will be calculated 
0150cf5982ae polished some recipies
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    49
  as the end time minus the start time.  An example of the
175
7c09bd3227c5 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
    50
  wrapper is the proof
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    51
*}
154
e81ebb37aa83 updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
155
Christian Urban <urbanc@in.tum.de>
parents: 154
diff changeset
    53
lemma "ackermann (3, 4) = 125"
Christian Urban <urbanc@in.tum.de>
parents: 154
diff changeset
    54
apply(tactic {* 
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 460
diff changeset
    55
  timing_wrapper (simp_tac (@{context} addsimps @{thms "eval_nat_numeral"}) 1) *})
155
Christian Urban <urbanc@in.tum.de>
parents: 154
diff changeset
    56
done
154
e81ebb37aa83 updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    58
text {*
185
043ef82000b4 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 175
diff changeset
    59
  where it returns something on the scale of 3 seconds. We chose to return
168
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    60
  this information as a string, but the timing information is also accessible
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    61
  in number format.
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    62
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    63
  \begin{readmore}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    64
  Basic functions regarding timing are defined in @{ML_file 
460
5c33c4b52ad7 Updated to changes in timing structure
Christian Urban <urbanc@in.tum.de>
parents: 456
diff changeset
    65
  "Pure/General/timing.ML"} (for the PolyML compiler). Some more
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    66
  advanced functions are defined in @{ML_file "Pure/General/output.ML"}.
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    67
  \end{readmore}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    68
*}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    69
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    70
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    71
456
89fccd3d5055 'nat_number' was renamed to 'eval_nat_numeral'
griff
parents: 346
diff changeset
    72
end