| author | wenzelm | 
| Mon, 14 Jul 2025 11:18:10 +0200 | |
| changeset 82864 | 2703f19d323e | 
| parent 82081 | 50dd4fc40fcb | 
| child 83194 | c2566f1548f4 | 
| 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  | 
|
| 82081 | 15  | 
and "time_partial_function" :: thy_decl  | 
| 79490 | 16  | 
and "equations"  | 
| 79969 | 17  | 
and "time_fun_0" :: thy_decl  | 
| 79490 | 18  | 
begin  | 
19  | 
||
20  | 
ML_file Define_Time_0.ML  | 
|
21  | 
ML_file Define_Time_Function.ML  | 
|
22  | 
||
23  | 
declare [[time_prefix = "T_"]]  | 
|
24  | 
||
| 
79973
 
7bbb0d65ce72
documented running time function framework by Jonas Stahl
 
nipkow 
parents: 
79969 
diff
changeset
 | 
25  | 
text \<open>  | 
| 
 
7bbb0d65ce72
documented running time function framework by Jonas Stahl
 
nipkow 
parents: 
79969 
diff
changeset
 | 
26  | 
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
 | 
27  | 
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
 | 
28  | 
"Functional Data Structures and Algorithms. A Proof Assistant Approach."  | 
| 
 
7bbb0d65ce72
documented running time function framework by Jonas Stahl
 
nipkow 
parents: 
79969 
diff
changeset
 | 
29  | 
See @{url "https://functional-algorithms-verified.org"}
 | 
| 
 
7bbb0d65ce72
documented running time function framework by Jonas Stahl
 
nipkow 
parents: 
79969 
diff
changeset
 | 
30  | 
|
| 
 
7bbb0d65ce72
documented running time function framework by Jonas Stahl
 
nipkow 
parents: 
79969 
diff
changeset
 | 
31  | 
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
 | 
32  | 
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
 | 
33  | 
running time functions must already have been defined.  | 
| 
 
7bbb0d65ce72
documented running time function framework by Jonas Stahl
 
nipkow 
parents: 
79969 
diff
changeset
 | 
34  | 
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
 | 
35  | 
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
 | 
36  | 
Limitation: The commands do not work properly in locales yet.  | 
| 
 
7bbb0d65ce72
documented running time function framework by Jonas Stahl
 
nipkow 
parents: 
79969 
diff
changeset
 | 
37  | 
|
| 
 
7bbb0d65ce72
documented running time function framework by Jonas Stahl
 
nipkow 
parents: 
79969 
diff
changeset
 | 
38  | 
The pre-defined functions below are assumed to have constant running time.  | 
| 79490 | 39  | 
In fact, we make that constant 0.  | 
40  | 
This does not change the asymptotic running time of user-defined functions using the  | 
|
41  | 
pre-defined functions because 1 is added for every user-defined function call.  | 
|
42  | 
||
| 
79973
 
7bbb0d65ce72
documented running time function framework by Jonas Stahl
 
nipkow 
parents: 
79969 
diff
changeset
 | 
43  | 
Many of the functions below are polymorphic and reside in type classes.  | 
| 79490 | 44  | 
The constant-time assumption is justified only for those types where the hardware offers  | 
45  | 
suitable support, e.g. numeric types. The argument size is implicitly bounded, too.  | 
|
46  | 
||
47  | 
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
 | 
48  | 
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
 | 
49  | 
|
| 
 
7bbb0d65ce72
documented running time function framework by Jonas Stahl
 
nipkow 
parents: 
79969 
diff
changeset
 | 
50  | 
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
 | 
51  | 
within the above restrictions.\<close>  | 
| 79490 | 52  | 
|
| 81356 | 53  | 
time_fun_0 "min"  | 
54  | 
time_fun_0 "max"  | 
|
| 79969 | 55  | 
time_fun_0 "(+)"  | 
56  | 
time_fun_0 "(-)"  | 
|
57  | 
time_fun_0 "(*)"  | 
|
58  | 
time_fun_0 "(/)"  | 
|
59  | 
time_fun_0 "(div)"  | 
|
60  | 
time_fun_0 "(<)"  | 
|
61  | 
time_fun_0 "(\<le>)"  | 
|
62  | 
time_fun_0 "Not"  | 
|
63  | 
time_fun_0 "(\<and>)"  | 
|
64  | 
time_fun_0 "(\<or>)"  | 
|
65  | 
time_fun_0 "Num.numeral_class.numeral"  | 
|
66  | 
time_fun_0 "(=)"  | 
|
| 79490 | 67  | 
|
68  | 
end  |