author | wenzelm |
Tue, 27 Jul 2021 15:31:54 +0200 | |
changeset 74072 | dc98bb7a439b |
parent 71847 | da12452c9be2 |
child 76987 | 4c275405faae |
permissions | -rw-r--r-- |
39307 | 1 |
(* Title: HOL/Imperative_HOL/Overview.thy |
2 |
Author: Florian Haftmann, TU Muenchen |
|
3 |
*) |
|
4 |
||
5 |
(*<*) |
|
6 |
theory Overview |
|
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63167
diff
changeset
|
7 |
imports Imperative_HOL "HOL-Library.LaTeXsugar" |
39307 | 8 |
begin |
9 |
||
10 |
(* type constraints with spacing *) |
|
61143 | 11 |
no_syntax (output) |
12 |
"_constrain" :: "logic => type => logic" ("_::_" [4, 0] 3) |
|
13 |
"_constrain" :: "prop' => type => prop'" ("_::_" [4, 0] 3) |
|
14 |
||
15 |
syntax (output) |
|
16 |
"_constrain" :: "logic => type => logic" ("_ :: _" [4, 0] 3) |
|
17 |
"_constrain" :: "prop' => type => prop'" ("_ :: _" [4, 0] 3) |
|
18 |
(*>*) |
|
39307 | 19 |
|
63167 | 20 |
text \<open> |
71847 | 21 |
\<open>Imperative HOL\<close> is a lightweight framework for reasoning |
63167 | 22 |
about imperative data structures in \<open>Isabelle/HOL\<close> |
58622 | 23 |
@{cite "Nipkow-et-al:2002:tutorial"}. Its basic ideas are described in |
24 |
@{cite "Bulwahn-et-al:2008:imp_HOL"}. However their concrete |
|
39307 | 25 |
realisation has changed since, due to both extensions and |
26 |
refinements. Therefore this overview wants to present the framework |
|
27 |
\qt{as it is} by now. It focusses on the user-view, less on matters |
|
40358 | 28 |
of construction. For details study of the theory sources is |
39307 | 29 |
encouraged. |
63167 | 30 |
\<close> |
39307 | 31 |
|
32 |
||
63167 | 33 |
section \<open>A polymorphic heap inside a monad\<close> |
39307 | 34 |
|
63167 | 35 |
text \<open> |
71847 | 36 |
Heaps (\<^type>\<open>heap\<close>) can be populated by values of class \<^class>\<open>heap\<close>; HOL's default types are |
37 |
already instantiated to class \<^class>\<open>heap\<close>. Class \<^class>\<open>heap\<close> is a subclass of \<^class>\<open>countable\<close>; |
|
38 |
see theory \<open>Countable\<close> for ways to instantiate types as \<^class>\<open>countable\<close>. |
|
39307 | 39 |
|
69597 | 40 |
The heap is wrapped up in a monad \<^typ>\<open>'a Heap\<close> by means of the |
39307 | 41 |
following specification: |
42 |
||
43 |
\begin{quote} |
|
69597 | 44 |
\<^datatype>\<open>Heap\<close> |
39307 | 45 |
\end{quote} |
46 |
||
47 |
Unwrapping of this monad type happens through |
|
48 |
||
49 |
\begin{quote} |
|
69597 | 50 |
\<^term_type>\<open>execute\<close> \\ |
39307 | 51 |
@{thm execute.simps [no_vars]} |
52 |
\end{quote} |
|
53 |
||
54 |
This allows for equational reasoning about monadic expressions; the |
|
63167 | 55 |
fact collection \<open>execute_simps\<close> contains appropriate rewrites |
39307 | 56 |
for all fundamental operations. |
57 |
||
39610 | 58 |
Primitive fine-granular control over heaps is available through rule |
63167 | 59 |
\<open>Heap_cases\<close>: |
39307 | 60 |
|
61 |
\begin{quote} |
|
62 |
@{thm [break] Heap_cases [no_vars]} |
|
63 |
\end{quote} |
|
64 |
||
65 |
Monadic expression involve the usual combinators: |
|
66 |
||
67 |
\begin{quote} |
|
69597 | 68 |
\<^term_type>\<open>return\<close> \\ |
69 |
\<^term_type>\<open>bind\<close> \\ |
|
70 |
\<^term_type>\<open>raise\<close> |
|
39307 | 71 |
\end{quote} |
72 |
||
71847 | 73 |
This is also associated with nice monad do-syntax. The \<^typ>\<open>string\<close> argument to \<^const>\<open>raise\<close> is |
74 |
just a codified comment. |
|
39307 | 75 |
|
76 |
Among a couple of generic combinators the following is helpful for |
|
77 |
establishing invariants: |
|
78 |
||
79 |
\begin{quote} |
|
69597 | 80 |
\<^term_type>\<open>assert\<close> \\ |
39307 | 81 |
@{thm assert_def [no_vars]} |
82 |
\end{quote} |
|
63167 | 83 |
\<close> |
39307 | 84 |
|
85 |
||
69597 | 86 |
section \<open>Relational reasoning about \<^type>\<open>Heap\<close> expressions\<close> |
39307 | 87 |
|
63167 | 88 |
text \<open> |
39307 | 89 |
To establish correctness of imperative programs, predicate |
90 |
||
91 |
\begin{quote} |
|
69597 | 92 |
\<^term_type>\<open>effect\<close> |
39307 | 93 |
\end{quote} |
94 |
||
71847 | 95 |
provides a simple relational calculus. Primitive rules are \<open>effectI\<close> and \<open>effectE\<close>, rules |
96 |
appropriate for reasoning about imperative operations are available in the \<open>effect_intros\<close> and |
|
63167 | 97 |
\<open>effect_elims\<close> fact collections. |
39307 | 98 |
|
99 |
Often non-failure of imperative computations does not depend |
|
100 |
on the heap at all; reasoning then can be easier using predicate |
|
101 |
||
102 |
\begin{quote} |
|
69597 | 103 |
\<^term_type>\<open>success\<close> |
39307 | 104 |
\end{quote} |
105 |
||
69597 | 106 |
Introduction rules for \<^const>\<open>success\<close> are available in the |
63167 | 107 |
\<open>success_intro\<close> fact collection. |
39307 | 108 |
|
69597 | 109 |
\<^const>\<open>execute\<close>, \<^const>\<open>effect\<close>, \<^const>\<open>success\<close> and \<^const>\<open>bind\<close> |
71847 | 110 |
are related by rules \<open>execute_bind_success\<close>, \<open>success_bind_executeI\<close>, \<open>success_bind_effectI\<close>, |
111 |
\<open>effect_bindI\<close>, \<open>effect_bindE\<close> and \<open>execute_bind_eq_SomeI\<close>. |
|
63167 | 112 |
\<close> |
39307 | 113 |
|
114 |
||
63167 | 115 |
section \<open>Monadic data structures\<close> |
39307 | 116 |
|
63167 | 117 |
text \<open> |
39307 | 118 |
The operations for monadic data structures (arrays and references) |
119 |
come in two flavours: |
|
120 |
||
121 |
\begin{itemize} |
|
122 |
||
123 |
\item Operations on the bare heap; their number is kept minimal |
|
124 |
to facilitate proving. |
|
125 |
||
126 |
\item Operations on the heap wrapped up in a monad; these are designed |
|
127 |
for executing. |
|
128 |
||
129 |
\end{itemize} |
|
130 |
||
131 |
Provided proof rules are such that they reduce monad operations to |
|
132 |
operations on bare heaps. |
|
39717 | 133 |
|
134 |
Note that HOL equality coincides with reference equality and may be |
|
135 |
used as primitive executable operation. |
|
63167 | 136 |
\<close> |
39307 | 137 |
|
63167 | 138 |
subsection \<open>Arrays\<close> |
39307 | 139 |
|
63167 | 140 |
text \<open> |
39307 | 141 |
Heap operations: |
142 |
||
143 |
\begin{quote} |
|
69597 | 144 |
\<^term_type>\<open>Array.alloc\<close> \\ |
145 |
\<^term_type>\<open>Array.present\<close> \\ |
|
146 |
\<^term_type>\<open>Array.get\<close> \\ |
|
147 |
\<^term_type>\<open>Array.set\<close> \\ |
|
148 |
\<^term_type>\<open>Array.length\<close> \\ |
|
149 |
\<^term_type>\<open>Array.update\<close> \\ |
|
150 |
\<^term_type>\<open>Array.noteq\<close> |
|
39307 | 151 |
\end{quote} |
152 |
||
153 |
Monad operations: |
|
154 |
||
155 |
\begin{quote} |
|
69597 | 156 |
\<^term_type>\<open>Array.new\<close> \\ |
157 |
\<^term_type>\<open>Array.of_list\<close> \\ |
|
158 |
\<^term_type>\<open>Array.make\<close> \\ |
|
159 |
\<^term_type>\<open>Array.len\<close> \\ |
|
160 |
\<^term_type>\<open>Array.nth\<close> \\ |
|
161 |
\<^term_type>\<open>Array.upd\<close> \\ |
|
162 |
\<^term_type>\<open>Array.map_entry\<close> \\ |
|
163 |
\<^term_type>\<open>Array.swap\<close> \\ |
|
164 |
\<^term_type>\<open>Array.freeze\<close> |
|
39307 | 165 |
\end{quote} |
63167 | 166 |
\<close> |
39307 | 167 |
|
63167 | 168 |
subsection \<open>References\<close> |
39307 | 169 |
|
63167 | 170 |
text \<open> |
39307 | 171 |
Heap operations: |
172 |
||
173 |
\begin{quote} |
|
69597 | 174 |
\<^term_type>\<open>Ref.alloc\<close> \\ |
175 |
\<^term_type>\<open>Ref.present\<close> \\ |
|
176 |
\<^term_type>\<open>Ref.get\<close> \\ |
|
177 |
\<^term_type>\<open>Ref.set\<close> \\ |
|
178 |
\<^term_type>\<open>Ref.noteq\<close> |
|
39307 | 179 |
\end{quote} |
180 |
||
181 |
Monad operations: |
|
182 |
||
183 |
\begin{quote} |
|
69597 | 184 |
\<^term_type>\<open>Ref.ref\<close> \\ |
185 |
\<^term_type>\<open>Ref.lookup\<close> \\ |
|
186 |
\<^term_type>\<open>Ref.update\<close> \\ |
|
187 |
\<^term_type>\<open>Ref.change\<close> |
|
39307 | 188 |
\end{quote} |
63167 | 189 |
\<close> |
39307 | 190 |
|
191 |
||
63167 | 192 |
section \<open>Code generation\<close> |
39307 | 193 |
|
63167 | 194 |
text \<open> |
39307 | 195 |
Imperative HOL sets up the code generator in a way that imperative |
196 |
operations are mapped to suitable counterparts in the target |
|
63167 | 197 |
language. For \<open>Haskell\<close>, a suitable \<open>ST\<close> monad is used; |
198 |
for \<open>SML\<close>, \<open>Ocaml\<close> and \<open>Scala\<close> unit values ensure |
|
39307 | 199 |
that the evaluation order is the same as you would expect from the |
200 |
original monadic expressions. These units may look cumbersome; the |
|
63167 | 201 |
target language variants \<open>SML_imp\<close>, \<open>Ocaml_imp\<close> and |
202 |
\<open>Scala_imp\<close> make some effort to optimize some of them away. |
|
203 |
\<close> |
|
39307 | 204 |
|
205 |
||
63167 | 206 |
section \<open>Some hints for using the framework\<close> |
39307 | 207 |
|
63167 | 208 |
text \<open> |
39307 | 209 |
Of course a framework itself does not by itself indicate how to make |
210 |
best use of it. Here some hints drawn from prior experiences with |
|
211 |
Imperative HOL: |
|
212 |
||
213 |
\begin{itemize} |
|
214 |
||
215 |
\item Proofs on bare heaps should be strictly separated from those |
|
216 |
for monadic expressions. The first capture the essence, while the |
|
217 |
latter just describe a certain wrapping-up. |
|
218 |
||
219 |
\item A good methodology is to gradually improve an imperative |
|
220 |
program from a functional one. In the extreme case this means |
|
221 |
that an original functional program is decomposed into suitable |
|
222 |
operations with exactly one corresponding imperative operation. |
|
223 |
Having shown suitable correspondence lemmas between those, the |
|
224 |
correctness prove of the whole imperative program simply |
|
225 |
consists of composing those. |
|
226 |
||
227 |
\item Whether one should prefer equational reasoning (fact |
|
63167 | 228 |
collection \<open>execute_simps\<close> or relational reasoning (fact |
229 |
collections \<open>effect_intros\<close> and \<open>effect_elims\<close>) depends |
|
39610 | 230 |
on the problems to solve. For complex expressions or |
71847 | 231 |
expressions involving binders, the relation style is usually |
39610 | 232 |
superior but requires more proof text. |
39307 | 233 |
|
234 |
\item Note that you can extend the fact collections of Imperative |
|
39610 | 235 |
HOL yourself whenever appropriate. |
39307 | 236 |
|
237 |
\end{itemize} |
|
63167 | 238 |
\<close> |
39307 | 239 |
|
62390 | 240 |
end |