author | nipkow |
Sat, 12 Oct 2024 12:45:29 +0900 | |
changeset 81147 | 503e5280ba72 |
parent 81088 | 28ef01901650 |
child 81356 | 9c47740e974a |
permissions | -rw-r--r-- |
79490 | 1 |
(* |
2 |
Author: Jonas Stahl |
|
3 |
||
79973
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
4 |
Automatic definition of step-counting running-time functions from HOL functions |
79490 | 5 |
following the translation described in Section 1.5, Running Time, of the book |
6 |
Functional Data Structures and Algorithms. A Proof Assistant Approach. |
|
79973
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
7 |
See https://functional-algorithms-verified.org |
79490 | 8 |
*) |
9 |
||
10 |
theory Define_Time_Function |
|
11 |
imports Main |
|
79969 | 12 |
keywords "time_fun" :: thy_decl |
81088 | 13 |
and "time_function" :: thy_decl |
14 |
and "time_definition" :: thy_decl |
|
79490 | 15 |
and "equations" |
79969 | 16 |
and "time_fun_0" :: thy_decl |
79490 | 17 |
begin |
18 |
||
19 |
ML_file Define_Time_0.ML |
|
20 |
ML_file Define_Time_Function.ML |
|
21 |
||
22 |
declare [[time_prefix = "T_"]] |
|
23 |
||
79973
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
24 |
text \<open> |
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
25 |
This theory provides commands for the automatic definition of step-counting running-time functions |
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
26 |
from HOL functions following the translation described in Section 1.5, Running Time, of the book |
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
27 |
"Functional Data Structures and Algorithms. A Proof Assistant Approach." |
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
28 |
See @{url "https://functional-algorithms-verified.org"} |
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
29 |
|
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
30 |
Command \<open>time_fun f\<close> retrieves the definition of \<open>f\<close> and defines a corresponding step-counting |
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
31 |
running-time function \<open>T_f\<close>. For all auxiliary functions used by \<open>f\<close> (excluding constructors), |
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
32 |
running time functions must already have been defined. |
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
33 |
If the definition of the function requires a manual termination proof, |
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
34 |
use \<open>time_function\<close> accompanied by a \<open>termination\<close> command. |
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
35 |
Limitation: The commands do not work properly in locales yet. |
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
36 |
|
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
37 |
The pre-defined functions below are assumed to have constant running time. |
79490 | 38 |
In fact, we make that constant 0. |
39 |
This does not change the asymptotic running time of user-defined functions using the |
|
40 |
pre-defined functions because 1 is added for every user-defined function call. |
|
41 |
||
79973
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
42 |
Many of the functions below are polymorphic and reside in type classes. |
79490 | 43 |
The constant-time assumption is justified only for those types where the hardware offers |
44 |
suitable support, e.g. numeric types. The argument size is implicitly bounded, too. |
|
45 |
||
46 |
The constant-time assumption for \<open>(=)\<close> is justified for recursive data types such as lists and trees |
|
79973
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
47 |
as long as the comparison is of the form \<open>t = c\<close> where \<open>c\<close> is a constant term, for example \<open>xs = []\<close>. |
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
48 |
|
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
49 |
Users of this running time framework need to ensure that 0-time functions are used only |
7bbb0d65ce72
documented running time function framework by Jonas Stahl
nipkow
parents:
79969
diff
changeset
|
50 |
within the above restrictions.\<close> |
79490 | 51 |
|
79969 | 52 |
time_fun_0 "(+)" |
53 |
time_fun_0 "(-)" |
|
54 |
time_fun_0 "(*)" |
|
55 |
time_fun_0 "(/)" |
|
56 |
time_fun_0 "(div)" |
|
57 |
time_fun_0 "(<)" |
|
58 |
time_fun_0 "(\<le>)" |
|
59 |
time_fun_0 "Not" |
|
60 |
time_fun_0 "(\<and>)" |
|
61 |
time_fun_0 "(\<or>)" |
|
62 |
time_fun_0 "Num.numeral_class.numeral" |
|
63 |
time_fun_0 "(=)" |
|
79490 | 64 |
|
65 |
end |