| author | wenzelm | 
| Fri, 18 Feb 2022 23:10:33 +0100 | |
| changeset 75102 | 678fae02f9b3 | 
| parent 69697 | 4d95261fab5a | 
| child 76673 | 059a68d21f0f | 
| permissions | -rw-r--r-- | 
| 38405 | 1  | 
theory Foundations  | 
| 28419 | 2  | 
imports Introduction  | 
| 28213 | 3  | 
begin  | 
4  | 
||
| 59377 | 5  | 
section \<open>Code generation foundations \label{sec:foundations}\<close>
 | 
| 28419 | 6  | 
|
| 59377 | 7  | 
subsection \<open>Code generator architecture \label{sec:architecture}\<close>
 | 
| 28419 | 8  | 
|
| 59377 | 9  | 
text \<open>  | 
| 38437 | 10  | 
The code generator is actually a framework consisting of different  | 
11  | 
components which can be customised individually.  | 
|
12  | 
||
13  | 
Conceptually all components operate on Isabelle's logic framework  | 
|
| 68484 | 14  | 
Pure. Practically, the object logic HOL  | 
| 38437 | 15  | 
provides the necessary facilities to make use of the code generator,  | 
| 68484 | 16  | 
mainly since it is an extension of Isabelle/Pure.  | 
| 38437 | 17  | 
|
18  | 
The constellation of the different components is visualized in the  | 
|
19  | 
following picture.  | 
|
20  | 
||
21  | 
  \begin{figure}[h]
 | 
|
| 52742 | 22  | 
    \def\sys#1{\emph{#1}}
 | 
23  | 
    \begin{tikzpicture}[x = 4cm, y = 1cm]
 | 
|
24  | 
\tikzstyle positive=[color = black, fill = white];  | 
|
25  | 
\tikzstyle negative=[color = white, fill = black];  | 
|
26  | 
\tikzstyle entity=[rounded corners, draw, thick];  | 
|
27  | 
\tikzstyle process=[ellipse, draw, thick];  | 
|
28  | 
\tikzstyle arrow=[-stealth, semithick];  | 
|
29  | 
      \node (spec) at (0, 3) [entity, positive] {specification tools};
 | 
|
30  | 
      \node (user) at (1, 3) [entity, positive] {user proofs};
 | 
|
31  | 
      \node (spec_user_join) at (0.5, 3) [shape=coordinate] {};
 | 
|
32  | 
      \node (raw) at (0.5, 4) [entity, positive] {raw code equations};
 | 
|
33  | 
      \node (pre) at (1.5, 4) [process, positive] {preprocessing};
 | 
|
34  | 
      \node (eqn) at (2.5, 4) [entity, positive] {code equations};
 | 
|
35  | 
      \node (iml) at (0.5, 0) [entity, positive] {intermediate program};
 | 
|
36  | 
      \node (seri) at (1.5, 0) [process, positive] {serialisation};
 | 
|
37  | 
      \node (SML) at (2.5, 3) [entity, positive] {\sys{SML}};
 | 
|
38  | 
      \node (OCaml) at (2.5, 2) [entity, positive] {\sys{OCaml}};
 | 
|
39  | 
      \node (Haskell) at (2.5, 1) [entity, positive] {\sys{Haskell}};
 | 
|
40  | 
      \node (Scala) at (2.5, 0) [entity, positive] {\sys{Scala}};
 | 
|
41  | 
\draw [semithick] (spec) -- (spec_user_join);  | 
|
42  | 
\draw [semithick] (user) -- (spec_user_join);  | 
|
43  | 
\draw [-diamond, semithick] (spec_user_join) -- (raw);  | 
|
44  | 
\draw [arrow] (raw) -- (pre);  | 
|
45  | 
\draw [arrow] (pre) -- (eqn);  | 
|
46  | 
      \draw [arrow] (eqn) -- node (transl) [process, positive] {translation} (iml);
 | 
|
47  | 
\draw [arrow] (iml) -- (seri);  | 
|
48  | 
\draw [arrow] (seri) -- (SML);  | 
|
49  | 
\draw [arrow] (seri) -- (OCaml);  | 
|
50  | 
\draw [arrow] (seri) -- (Haskell);  | 
|
51  | 
\draw [arrow] (seri) -- (Scala);  | 
|
52  | 
    \end{tikzpicture}
 | 
|
| 38437 | 53  | 
    \caption{Code generator architecture}
 | 
54  | 
    \label{fig:arch}
 | 
|
55  | 
  \end{figure}
 | 
|
56  | 
||
57  | 
  \noindent Central to code generation is the notion of \emph{code
 | 
|
58  | 
equations}. A code equation as a first approximation is a theorem  | 
|
| 69505 | 59  | 
of the form \<open>f t\<^sub>1 t\<^sub>2 \<dots> t\<^sub>n \<equiv> t\<close> (an equation headed by a  | 
60  | 
constant \<open>f\<close> with arguments \<open>t\<^sub>1 t\<^sub>2 \<dots> t\<^sub>n\<close> and right  | 
|
61  | 
hand side \<open>t\<close>).  | 
|
| 38437 | 62  | 
|
63  | 
  \begin{itemize}
 | 
|
64  | 
||
65  | 
\item Starting point of code generation is a collection of (raw)  | 
|
66  | 
code equations in a theory. It is not relevant where they stem  | 
|
67  | 
from, but typically they were either produced by specification  | 
|
68  | 
tools or proved explicitly by the user.  | 
|
69  | 
||
70  | 
\item These raw code equations can be subjected to theorem  | 
|
71  | 
      transformations.  This \qn{preprocessor} (see
 | 
|
72  | 
      \secref{sec:preproc}) can apply the full expressiveness of
 | 
|
73  | 
ML-based theorem transformations to code generation. The result  | 
|
74  | 
of preprocessing is a structured collection of code equations.  | 
|
75  | 
||
76  | 
    \item These code equations are \qn{translated} to a program in an
 | 
|
77  | 
abstract intermediate language. Think of it as a kind of  | 
|
| 69505 | 78  | 
      \qt{Mini-Haskell} with four \qn{statements}: \<open>data\<close> (for
 | 
79  | 
datatypes), \<open>fun\<close> (stemming from code equations), also  | 
|
80  | 
\<open>class\<close> and \<open>inst\<close> (for type classes).  | 
|
| 38437 | 81  | 
|
82  | 
    \item Finally, the abstract program is \qn{serialised} into
 | 
|
83  | 
concrete source code of a target language. This step only  | 
|
84  | 
produces concrete syntax but does not change the program in  | 
|
85  | 
essence; all conceptual transformations occur in the translation  | 
|
86  | 
step.  | 
|
87  | 
||
88  | 
  \end{itemize}
 | 
|
89  | 
||
90  | 
\noindent From these steps, only the last two are carried out  | 
|
91  | 
outside the logic; by keeping this layer as thin as possible, the  | 
|
92  | 
amount of code to trust is kept to a minimum.  | 
|
| 59377 | 93  | 
\<close>  | 
| 28419 | 94  | 
|
95  | 
||
| 59377 | 96  | 
subsection \<open>The pre- and postprocessor \label{sec:preproc}\<close>
 | 
| 28419 | 97  | 
|
| 59377 | 98  | 
text \<open>  | 
| 38437 | 99  | 
Before selected function theorems are turned into abstract code, a  | 
100  | 
chain of definitional transformation steps is carried out:  | 
|
101  | 
  \emph{preprocessing}.  The preprocessor consists of two
 | 
|
102  | 
  components: a \emph{simpset} and \emph{function transformers}.
 | 
|
| 28419 | 103  | 
|
| 
54721
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
104  | 
The preprocessor simpset has a disparate brother, the  | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
105  | 
  \emph{postprocessor simpset}.
 | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
106  | 
In the theory-to-code scenario depicted in the picture  | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
107  | 
above, it plays no role. But if generated code is used  | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
108  | 
  to evaluate expressions (cf.~\secref{sec:evaluation}), the
 | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
109  | 
postprocessor simpset is applied to the resulting expression before this  | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
110  | 
is turned back.  | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
111  | 
|
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
112  | 
  The pre- and postprocessor \emph{simpsets} can apply the
 | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
113  | 
full generality of the Isabelle  | 
| 38437 | 114  | 
simplifier. Due to the interpretation of theorems as code  | 
| 32000 | 115  | 
equations, rewrites are applied to the right hand side and the  | 
116  | 
arguments of the left hand side of an equation, but never to the  | 
|
| 
54721
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
117  | 
constant heading the left hand side.  | 
| 28213 | 118  | 
|
| 61781 | 119  | 
Pre- and postprocessor can be setup to transfer between  | 
| 
54721
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
120  | 
expressions suitable for logical reasoning and expressions  | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
121  | 
suitable for execution. As example, take list membership; logically  | 
| 69597 | 122  | 
it is expressed as \<^term>\<open>x \<in> set xs\<close>. But for execution  | 
| 
54721
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
123  | 
the intermediate set is not desirable. Hence the following  | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
124  | 
specification:  | 
| 59377 | 125  | 
\<close>  | 
| 28419 | 126  | 
|
| 
54721
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
127  | 
definition %quote member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool"  | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
128  | 
where  | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
129  | 
[code_abbrev]: "member xs x \<longleftrightarrow> x \<in> set xs"  | 
| 28419 | 130  | 
|
| 59377 | 131  | 
text \<open>  | 
| 
54721
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
132  | 
  \noindent The \emph{@{attribute code_abbrev} attribute} declares
 | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
133  | 
its theorem a rewrite rule for the postprocessor and the symmetric  | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
134  | 
of its theorem as rewrite rule for the preprocessor. Together,  | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
135  | 
  this has the effect that expressions @{thm (rhs) member_def [no_vars]}
 | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
136  | 
  are replaced by @{thm (lhs) member_def [no_vars]} in generated code, but
 | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
137  | 
  are turned back into @{thm (rhs) member_def [no_vars]} if generated
 | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
138  | 
code is used for evaluation.  | 
| 38437 | 139  | 
|
| 
54721
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
140  | 
Rewrite rules for pre- or postprocessor may be  | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
141  | 
  declared independently using \emph{@{attribute code_unfold}}
 | 
| 
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
142  | 
  or \emph{@{attribute code_post}} respectively.
 | 
| 38437 | 143  | 
|
| 
54721
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
144  | 
  \emph{Function transformers} provide a very general
 | 
| 38437 | 145  | 
interface, transforming a list of function theorems to another list  | 
146  | 
of function theorems, provided that neither the heading constant nor  | 
|
| 69597 | 147  | 
its type change. The \<^term>\<open>0::nat\<close> / \<^const>\<open>Suc\<close> pattern  | 
| 69505 | 148  | 
  used in theory \<open>Code_Abstract_Nat\<close> (see \secref{abstract_nat})
 | 
| 
51171
 
e8b2d90da499
corrected and clarified Code_Binary_Nat vs. Code_Target_Nat
 
haftmann 
parents: 
51143 
diff
changeset
 | 
149  | 
uses this interface.  | 
| 28419 | 150  | 
|
| 
54721
 
22b888402278
be more explicit about pre- and postprocessor, particularly code_abbrev
 
haftmann 
parents: 
53015 
diff
changeset
 | 
151  | 
\noindent The current setup of the pre- and postprocessor may be inspected  | 
| 38505 | 152  | 
  using the @{command_def print_codeproc} command.  @{command_def
 | 
153  | 
  code_thms} (see \secref{sec:equations}) provides a convenient
 | 
|
154  | 
mechanism to inspect the impact of a preprocessor setup on code  | 
|
| 
68254
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
155  | 
  equations.  Attribute @{attribute code_preproc_trace} allows
 | 
| 
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
156  | 
for low-level tracing:  | 
| 59377 | 157  | 
\<close>  | 
| 28419 | 158  | 
|
| 
68254
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
159  | 
declare %quote [[code_preproc_trace]]  | 
| 
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
160  | 
|
| 
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
161  | 
declare %quote [[code_preproc_trace only: distinct remdups]]  | 
| 
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
162  | 
|
| 
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
163  | 
declare %quote [[code_preproc_trace off]]  | 
| 
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
164  | 
|
| 38437 | 165  | 
|
| 59377 | 166  | 
subsection \<open>Understanding code equations \label{sec:equations}\<close>
 | 
| 28419 | 167  | 
|
| 59377 | 168  | 
text \<open>  | 
| 38437 | 169  | 
  As told in \secref{sec:principle}, the notion of code equations is
 | 
170  | 
vital to code generation. Indeed most problems which occur in  | 
|
171  | 
practice can be resolved by an inspection of the underlying code  | 
|
172  | 
equations.  | 
|
| 28419 | 173  | 
|
| 38437 | 174  | 
It is possible to exchange the default code equations for constants  | 
175  | 
by explicitly proving alternative ones:  | 
|
| 59377 | 176  | 
\<close>  | 
| 28419 | 177  | 
|
| 38437 | 178  | 
lemma %quote [code]:  | 
179  | 
"dequeue (AQueue xs []) =  | 
|
180  | 
(if xs = [] then (None, AQueue [] [])  | 
|
181  | 
else dequeue (AQueue [] (rev xs)))"  | 
|
182  | 
"dequeue (AQueue xs (y # ys)) =  | 
|
183  | 
(Some y, AQueue xs ys)"  | 
|
184  | 
by (cases xs, simp_all) (cases "rev xs", simp_all)  | 
|
| 28213 | 185  | 
|
| 59377 | 186  | 
text \<open>  | 
| 69505 | 187  | 
\noindent The annotation \<open>[code]\<close> is an \<open>attribute\<close>  | 
| 38437 | 188  | 
which states that the given theorems should be considered as code  | 
| 69505 | 189  | 
equations for a \<open>fun\<close> statement -- the corresponding constant  | 
| 38437 | 190  | 
is determined syntactically. The resulting code:  | 
| 59377 | 191  | 
\<close>  | 
| 29794 | 192  | 
|
| 
69660
 
2bc2a8599369
canonical operation to typeset generated code makes dedicated environment obsolete
 
haftmann 
parents: 
69597 
diff
changeset
 | 
193  | 
text %quote \<open>  | 
| 
69697
 
4d95261fab5a
more conventional syntax for code_stmts antiquotation
 
haftmann 
parents: 
69660 
diff
changeset
 | 
194  | 
  @{code_stmts dequeue constant: dequeue (Haskell)}
 | 
| 59377 | 195  | 
\<close>  | 
| 29794 | 196  | 
|
| 59377 | 197  | 
text \<open>  | 
| 69597 | 198  | 
\noindent You may note that the equality test \<^term>\<open>xs = []\<close> has  | 
199  | 
been replaced by the predicate \<^term>\<open>List.null xs\<close>. This is due  | 
|
| 38437 | 200  | 
  to the default setup of the \qn{preprocessor}.
 | 
201  | 
||
202  | 
This possibility to select arbitrary code equations is the key  | 
|
203  | 
technique for program and datatype refinement (see  | 
|
| 39677 | 204  | 
  \secref{sec:refinement}).
 | 
| 38437 | 205  | 
|
206  | 
Due to the preprocessor, there is the distinction of raw code  | 
|
207  | 
equations (before preprocessing) and code equations (after  | 
|
208  | 
preprocessing).  | 
|
209  | 
||
| 38505 | 210  | 
  The first can be listed (among other data) using the @{command_def
 | 
211  | 
print_codesetup} command.  | 
|
| 38437 | 212  | 
|
213  | 
The code equations after preprocessing are already are blueprint of  | 
|
214  | 
  the generated program and can be inspected using the @{command
 | 
|
215  | 
code_thms} command:  | 
|
| 59377 | 216  | 
\<close>  | 
| 29794 | 217  | 
|
| 38437 | 218  | 
code_thms %quote dequeue  | 
| 28419 | 219  | 
|
| 59377 | 220  | 
text \<open>  | 
| 69597 | 221  | 
  \noindent This prints a table with the code equations for \<^const>\<open>dequeue\<close>, including \emph{all} code equations those equations depend
 | 
| 38437 | 222  | 
on recursively. These dependencies themselves can be visualized using  | 
| 38505 | 223  | 
  the @{command_def code_deps} command.
 | 
| 59377 | 224  | 
\<close>  | 
| 28419 | 225  | 
|
| 28213 | 226  | 
|
| 59377 | 227  | 
subsection \<open>Equality\<close>  | 
| 28213 | 228  | 
|
| 59377 | 229  | 
text \<open>  | 
| 38437 | 230  | 
Implementation of equality deserves some attention. Here an example  | 
231  | 
function involving polymorphic equality:  | 
|
| 59377 | 232  | 
\<close>  | 
| 28419 | 233  | 
|
| 28564 | 234  | 
primrec %quote collect_duplicates :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where  | 
| 28447 | 235  | 
"collect_duplicates xs ys [] = xs"  | 
| 38437 | 236  | 
| "collect_duplicates xs ys (z#zs) = (if z \<in> set xs  | 
237  | 
then if z \<in> set ys  | 
|
238  | 
then collect_duplicates xs ys zs  | 
|
239  | 
else collect_duplicates xs (z#ys) zs  | 
|
240  | 
else collect_duplicates (z#xs) (z#ys) zs)"  | 
|
| 28419 | 241  | 
|
| 59377 | 242  | 
text \<open>  | 
| 
37612
 
48fed6598be9
adapted to reorganization of auxiliary list operations; split off predicate compiler into separate theory
 
haftmann 
parents: 
37427 
diff
changeset
 | 
243  | 
\noindent During preprocessing, the membership test is rewritten,  | 
| 69597 | 244  | 
resulting in \<^const>\<open>List.member\<close>, which itself performs an explicit  | 
| 69505 | 245  | 
equality check, as can be seen in the corresponding \<open>SML\<close> code:  | 
| 59377 | 246  | 
\<close>  | 
| 28419 | 247  | 
|
| 
69660
 
2bc2a8599369
canonical operation to typeset generated code makes dedicated environment obsolete
 
haftmann 
parents: 
69597 
diff
changeset
 | 
248  | 
text %quote \<open>  | 
| 39683 | 249  | 
  @{code_stmts collect_duplicates (SML)}
 | 
| 59377 | 250  | 
\<close>  | 
| 28419 | 251  | 
|
| 59377 | 252  | 
text \<open>  | 
| 28419 | 253  | 
\noindent Obviously, polymorphic equality is implemented the Haskell  | 
| 38437 | 254  | 
way using a type class. How is this achieved? HOL introduces an  | 
| 69597 | 255  | 
  explicit class \<^class>\<open>equal\<close> with a corresponding operation \<^const>\<open>HOL.equal\<close> such that @{thm equal [no_vars]}.  The preprocessing
 | 
256  | 
framework does the rest by propagating the \<^class>\<open>equal\<close> constraints  | 
|
| 38437 | 257  | 
through all dependent code equations. For datatypes, instances of  | 
| 69597 | 258  | 
\<^class>\<open>equal\<close> are implicitly derived when possible. For other types,  | 
| 69505 | 259  | 
you may instantiate \<open>equal\<close> manually like any other type class.  | 
| 59377 | 260  | 
\<close>  | 
| 28419 | 261  | 
|
262  | 
||
| 59377 | 263  | 
subsection \<open>Explicit partiality \label{sec:partiality}\<close>
 | 
| 28462 | 264  | 
|
| 59377 | 265  | 
text \<open>  | 
| 28462 | 266  | 
Partiality usually enters the game by partial patterns, as  | 
267  | 
in the following example, again for amortised queues:  | 
|
| 59377 | 268  | 
\<close>  | 
| 28462 | 269  | 
|
| 29798 | 270  | 
definition %quote strict_dequeue :: "'a queue \<Rightarrow> 'a \<times> 'a queue" where  | 
271  | 
"strict_dequeue q = (case dequeue q  | 
|
272  | 
of (Some x, q') \<Rightarrow> (x, q'))"  | 
|
273  | 
||
274  | 
lemma %quote strict_dequeue_AQueue [code]:  | 
|
275  | 
"strict_dequeue (AQueue xs (y # ys)) = (y, AQueue xs ys)"  | 
|
276  | 
"strict_dequeue (AQueue xs []) =  | 
|
277  | 
(case rev xs of y # ys \<Rightarrow> (y, AQueue [] ys))"  | 
|
| 38437 | 278  | 
by (simp_all add: strict_dequeue_def) (cases xs, simp_all split: list.split)  | 
| 28462 | 279  | 
|
| 59377 | 280  | 
text \<open>  | 
| 28462 | 281  | 
\noindent In the corresponding code, there is no equation  | 
| 69597 | 282  | 
for the pattern \<^term>\<open>AQueue [] []\<close>:  | 
| 59377 | 283  | 
\<close>  | 
| 28462 | 284  | 
|
| 
69660
 
2bc2a8599369
canonical operation to typeset generated code makes dedicated environment obsolete
 
haftmann 
parents: 
69597 
diff
changeset
 | 
285  | 
text %quote \<open>  | 
| 
69697
 
4d95261fab5a
more conventional syntax for code_stmts antiquotation
 
haftmann 
parents: 
69660 
diff
changeset
 | 
286  | 
  @{code_stmts strict_dequeue constant: strict_dequeue (Haskell)}
 | 
| 59377 | 287  | 
\<close>  | 
| 28462 | 288  | 
|
| 59377 | 289  | 
text \<open>  | 
| 
68254
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
290  | 
\noindent In some cases it is desirable to state this  | 
| 28462 | 291  | 
  pseudo-\qt{partiality} more explicitly, e.g.~as follows:
 | 
| 59377 | 292  | 
\<close>  | 
| 28462 | 293  | 
|
| 28564 | 294  | 
axiomatization %quote empty_queue :: 'a  | 
| 28462 | 295  | 
|
| 29798 | 296  | 
definition %quote strict_dequeue' :: "'a queue \<Rightarrow> 'a \<times> 'a queue" where  | 
| 
68254
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
297  | 
"strict_dequeue' q = (case dequeue q of (Some x, q') \<Rightarrow> (x, q')  | 
| 
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
298  | 
| _ \<Rightarrow> empty_queue)"  | 
| 28213 | 299  | 
|
| 29798 | 300  | 
lemma %quote strict_dequeue'_AQueue [code]:  | 
301  | 
"strict_dequeue' (AQueue xs []) = (if xs = [] then empty_queue  | 
|
302  | 
else strict_dequeue' (AQueue [] (rev xs)))"  | 
|
303  | 
"strict_dequeue' (AQueue xs (y # ys)) =  | 
|
304  | 
(y, AQueue xs ys)"  | 
|
| 38437 | 305  | 
by (simp_all add: strict_dequeue'_def split: list.splits)  | 
| 28462 | 306  | 
|
| 59377 | 307  | 
text \<open>  | 
| 69597 | 308  | 
Observe that on the right hand side of the definition of \<^const>\<open>strict_dequeue'\<close>, the unspecified constant \<^const>\<open>empty_queue\<close> occurs.  | 
309  | 
An attempt to generate code for \<^const>\<open>strict_dequeue'\<close> would  | 
|
310  | 
make the code generator complain that \<^const>\<open>empty_queue\<close> has  | 
|
| 
68254
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
311  | 
no associated code equations. In most situations unimplemented  | 
| 
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
312  | 
constants indeed indicated a broken program; however such  | 
| 
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
313  | 
constants can also be thought of as function definitions which always fail,  | 
| 29798 | 314  | 
since there is never a successful pattern match on the left hand  | 
315  | 
side. In order to categorise a constant into that category  | 
|
| 
54890
 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 
haftmann 
parents: 
54721 
diff
changeset
 | 
316  | 
  explicitly, use the @{attribute code} attribute with
 | 
| 69505 | 317  | 
\<open>abort\<close>:  | 
| 59377 | 318  | 
\<close>  | 
| 28462 | 319  | 
|
| 
54890
 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 
haftmann 
parents: 
54721 
diff
changeset
 | 
320  | 
declare %quote [[code abort: empty_queue]]  | 
| 28462 | 321  | 
|
| 59377 | 322  | 
text \<open>  | 
| 28462 | 323  | 
\noindent Then the code generator will just insert an error or  | 
324  | 
exception at the appropriate position:  | 
|
| 59377 | 325  | 
\<close>  | 
| 28462 | 326  | 
|
| 
69660
 
2bc2a8599369
canonical operation to typeset generated code makes dedicated environment obsolete
 
haftmann 
parents: 
69597 
diff
changeset
 | 
327  | 
text %quote \<open>  | 
| 
69697
 
4d95261fab5a
more conventional syntax for code_stmts antiquotation
 
haftmann 
parents: 
69660 
diff
changeset
 | 
328  | 
  @{code_stmts strict_dequeue' constant: empty_queue strict_dequeue' (Haskell)}
 | 
| 59377 | 329  | 
\<close>  | 
| 28462 | 330  | 
|
| 59377 | 331  | 
text \<open>  | 
| 38437 | 332  | 
\noindent This feature however is rarely needed in practice. Note  | 
| 
68254
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
333  | 
that the HOL default setup already includes  | 
| 
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
334  | 
\<close>  | 
| 
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
335  | 
|
| 
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
336  | 
declare %quote [[code abort: undefined]]  | 
| 
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
337  | 
|
| 
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
338  | 
text \<open>  | 
| 69597 | 339  | 
\noindent -- hence \<^const>\<open>undefined\<close> can always be used in such  | 
| 
68254
 
3a7f257dcac7
more complete and more correct documentation on code generation
 
haftmann 
parents: 
61781 
diff
changeset
 | 
340  | 
situations.  | 
| 59377 | 341  | 
\<close>  | 
| 38437 | 342  | 
|
343  | 
||
| 59377 | 344  | 
subsection \<open>If something goes utterly wrong \label{sec:utterly_wrong}\<close>
 | 
| 38437 | 345  | 
|
| 59377 | 346  | 
text \<open>  | 
| 38437 | 347  | 
Under certain circumstances, the code generator fails to produce  | 
| 38440 | 348  | 
code entirely. To debug these, the following hints may prove  | 
349  | 
helpful:  | 
|
| 38437 | 350  | 
|
351  | 
  \begin{description}
 | 
|
352  | 
||
| 38440 | 353  | 
    \ditem{\emph{Check with a different target language}.}  Sometimes
 | 
354  | 
the situation gets more clear if you switch to another target  | 
|
355  | 
language; the code generated there might give some hints what  | 
|
356  | 
prevents the code generator to produce code for the desired  | 
|
357  | 
language.  | 
|
| 38437 | 358  | 
|
| 38440 | 359  | 
    \ditem{\emph{Inspect code equations}.}  Code equations are the central
 | 
| 43410 | 360  | 
carrier of code generation. Most problems occurring while generating  | 
| 38440 | 361  | 
code can be traced to single equations which are printed as part of  | 
362  | 
the error message. A closer inspection of those may offer the key  | 
|
363  | 
      for solving issues (cf.~\secref{sec:equations}).
 | 
|
| 38437 | 364  | 
|
| 38440 | 365  | 
    \ditem{\emph{Inspect preprocessor setup}.}  The preprocessor might
 | 
366  | 
transform code equations unexpectedly; to understand an  | 
|
367  | 
      inspection of its setup is necessary (cf.~\secref{sec:preproc}).
 | 
|
| 38437 | 368  | 
|
| 38440 | 369  | 
    \ditem{\emph{Generate exceptions}.}  If the code generator
 | 
370  | 
complains about missing code equations, in can be helpful to  | 
|
371  | 
implement the offending constants as exceptions  | 
|
372  | 
      (cf.~\secref{sec:partiality}); this allows at least for a formal
 | 
|
373  | 
generation of code, whose inspection may then give clues what is  | 
|
374  | 
wrong.  | 
|
| 38437 | 375  | 
|
| 38440 | 376  | 
    \ditem{\emph{Remove offending code equations}.}  If code
 | 
377  | 
generation is prevented by just a single equation, this can be  | 
|
378  | 
      removed (cf.~\secref{sec:equations}) to allow formal code
 | 
|
379  | 
generation, whose result in turn can be used to trace the  | 
|
380  | 
problem. The most prominent case here are mismatches in type  | 
|
381  | 
      class signatures (\qt{wellsortedness error}).
 | 
|
| 38437 | 382  | 
|
383  | 
  \end{description}
 | 
|
| 59377 | 384  | 
\<close>  | 
| 28213 | 385  | 
|
386  | 
end  |