CookBook/Recipes/Timing.thy
changeset 168 009ca4807baa
parent 167 3e30ea95c7aa
child 175 7c09bd3227c5
equal deleted inserted replaced
167:3e30ea95c7aa 168:009ca4807baa
     9   You want to measure the running time of a tactic or function.\smallskip
     9   You want to measure the running time of a tactic or function.\smallskip
    10 
    10 
    11   {\bf Solution:} Time can be measured using the function 
    11   {\bf Solution:} Time can be measured using the function 
    12   @{ML start_timing} and @{ML end_timing}.\smallskip
    12   @{ML start_timing} and @{ML end_timing}.\smallskip
    13 
    13 
    14   Assume the following function defined in Isabelle. 
    14   Suppose you defined the Ackermann function inside Isabelle. 
    15 *}
    15 *}
    16 
    16 
    17 fun 
    17 fun 
    18  ackermann:: "(nat * nat) \<Rightarrow> nat"
    18  ackermann:: "(nat \<times> nat) \<Rightarrow> nat"
    19 where
    19 where
    20     "ackermann (0, n) = n + 1"
    20     "ackermann (0, n) = n + 1"
    21   | "ackermann (m, 0) = ackermann (m - 1, 1)"
    21   | "ackermann (m, 0) = ackermann (m - 1, 1)"
    22   | "ackermann (m, n) = ackermann (m - 1, ackermann (m, n - 1))"
    22   | "ackermann (m, n) = ackermann (m - 1, ackermann (m, n - 1))"
    23 
    23 
    24 text {* 
    24 text {* 
    25   We can now measure how long the simplifier takes to verify a datapoint
    25   You can measure how long the simplifier takes to verify a datapoint
    26   of this function. The timing can be done using the following wrapper function:
    26   of this function. The timing can be done using the following wrapper function:
    27 *}
    27 *}
    28 
    28 
    29 ML{*fun timing_wrapper tac st =
    29 ML{*fun timing_wrapper tac st =
    30 let 
    30 let 
    34 in
    34 in
    35   (warning (#message t_end); res)
    35   (warning (#message t_end); res)
    36 end*}
    36 end*}
    37 
    37 
    38 text {*
    38 text {*
    39   Note that this function also takes a state @{text "st"} as an argument and
    39   Note that this function, in addition to a tactic for which it measures the
    40   applies this state to the tactic. This is because tactics are lazy functions
    40   time, also takes a state @{text "st"} as argument and applies this state to
    41   and we need to force them to run, otherwise the timing will be meaningless.
    41   the tactic. The reason is that tactics are lazy functions and you need to force
    42   The used time will be calculated as the end time minus the start time.
    42   them to run, otherwise the timing will be meaningless.  The time between start
    43   The wrapper can now be used in the proof
    43   and finish of the tactic will be calculated as the end time minus the start time.  
       
    44   An example for the wrapper is the proof
       
    45 
    44 *}
    46 *}
    45 
    47 
    46 lemma "ackermann (3, 4) = 125"
    48 lemma "ackermann (3, 4) = 125"
    47 apply(tactic {* 
    49 apply(tactic {* 
    48   timing_wrapper (simp_tac (@{simpset} addsimps @{thms "nat_number"}) 1) *})
    50   timing_wrapper (simp_tac (@{simpset} addsimps @{thms "nat_number"}) 1) *})
    49 done
    51 done
    50 
    52 
    51 text {*
    53 text {*
    52   where it returns something on the cale of 3 seconds.
    54   where it returns something on the scale of 3 seconds. We choose to return
       
    55   this information as a string, but the timing information is also accessible
       
    56   in number format.
    53 
    57 
    54   \begin{readmore}
    58   \begin{readmore}
    55   Basic functions regarding timing are defined in @{ML_file 
    59   Basic functions regarding timing are defined in @{ML_file 
    56   "Pure/ML-Systems/polyml_common.ML"} (for the PolyML compiler). Some more
    60   "Pure/ML-Systems/polyml_common.ML"} (for the PolyML compiler). Some more
    57   advanced functions are defined in @{ML_file "Pure/General/output.ML"}.
    61   advanced functions are defined in @{ML_file "Pure/General/output.ML"}.