CookBook/Recipes/Timing.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 12 Mar 2009 18:39:10 +0000
changeset 173 d820cb5873ea
parent 168 009ca4807baa
child 175 7c09bd3227c5
permissions -rw-r--r--
used latex package boxedminipage
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
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
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 
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
168
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    14
  Suppose you defined the Ackermann function inside Isabelle. 
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
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    26
  of this function. The timing can be done using the following wrapper function:
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    27
*}
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
ML{*fun timing_wrapper tac st =
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    30
let 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    31
  val t_start = start_timing ();
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    32
  val res = tac st;
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    33
  val t_end = end_timing t_start;
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    34
in
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    35
  (warning (#message t_end); res)
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    36
end*}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    37
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    38
text {*
168
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    39
  Note that this function, in addition to a tactic for which it measures the
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    40
  time, also takes a state @{text "st"} as argument and applies this state to
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    41
  the tactic. The reason is that tactics are lazy functions and you need to force
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    42
  them to run, otherwise the timing will be meaningless.  The time between start
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    43
  and finish of the tactic will be calculated as the end time minus the start time.  
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    44
  An example for the wrapper is the proof
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    45
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    46
*}
154
e81ebb37aa83 updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
155
Christian Urban <urbanc@in.tum.de>
parents: 154
diff changeset
    48
lemma "ackermann (3, 4) = 125"
Christian Urban <urbanc@in.tum.de>
parents: 154
diff changeset
    49
apply(tactic {* 
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    50
  timing_wrapper (simp_tac (@{simpset} addsimps @{thms "nat_number"}) 1) *})
155
Christian Urban <urbanc@in.tum.de>
parents: 154
diff changeset
    51
done
154
e81ebb37aa83 updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    53
text {*
168
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    54
  where it returns something on the scale of 3 seconds. We choose to return
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    55
  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
    56
  in number format.
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    57
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    58
  \begin{readmore}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    59
  Basic functions regarding timing are defined in @{ML_file 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    60
  "Pure/ML-Systems/polyml_common.ML"} (for the PolyML compiler). Some more
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    61
  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
    62
  \end{readmore}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    63
*}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    64
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    65
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
    66
154
e81ebb37aa83 updated to repository version; added a section about timing
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
end