# HG changeset patch # User haftmann # Date 1261553493 -3600 # Node ID 156a060d5d682044e0c071c1878f4b8e6d9cc3ba # Parent 3a7937841585a2f2d98683f24ee5981356877cb3# Parent 70210e9a8b4a97ad736d0b0f02ac291430c0e95d merged diff -r 3a7937841585 -r 156a060d5d68 NEWS --- a/NEWS Mon Dec 21 16:50:28 2009 +0000 +++ b/NEWS Wed Dec 23 08:31:33 2009 +0100 @@ -4,6 +4,12 @@ New in this Isabelle version ---------------------------- +*** Pure *** + +* Code generator: details of internal data cache have no impact on +the user space functionality any longer. + + *** HOL *** * Reorganized theory Sum_Type.thy; Inl and Inr now have @@ -13,7 +19,7 @@ * Complete_Lattice.thy: lemmas top_def and bot_def have been replaced by the more convenient lemmas Inf_empty and Sup_empty. Dropped lemmas -Inf_insert_simp adn Sup_insert_simp, which are subsumed by Inf_insert +Inf_insert_simp and Sup_insert_simp, which are subsumed by Inf_insert and Sup_insert. Lemmas Inf_UNIV and Sup_UNIV replace former Inf_Univ and Sup_Univ. Lemmas inf_top_right and sup_bot_right subsume inf_top and sup_bot respectively. INCOMPATIBILITY. diff -r 3a7937841585 -r 156a060d5d68 doc-src/Codegen/Thy/document/Codegen.tex --- a/doc-src/Codegen/Thy/document/Codegen.tex Mon Dec 21 16:50:28 2009 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1690 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Codegen}% -% -\isadelimtheory -\isanewline -\isanewline -% -\endisadelimtheory -% -\isatagtheory -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isadelimML -% -\endisadelimML -% -\isatagML -% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -% -\isamarkupchapter{Code generation from Isabelle theories% -} -\isamarkuptrue% -% -\isamarkupsection{Introduction% -} -\isamarkuptrue% -% -\isamarkupsubsection{Motivation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Executing formal specifications as programs is a well-established - topic in the theorem proving community. With increasing - application of theorem proving systems in the area of - software development and verification, its relevance manifests - for running test cases and rapid prototyping. In logical - calculi like constructive type theory, - a notion of executability is implicit due to the nature - of the calculus. In contrast, specifications in Isabelle - can be highly non-executable. In order to bridge - the gap between logic and executable specifications, - an explicit non-trivial transformation has to be applied: - code generation. - - This tutorial introduces a generic code generator for the - Isabelle system \cite{isa-tutorial}. - Generic in the sense that the - \qn{target language} for which code shall ultimately be - generated is not fixed but may be an arbitrary state-of-the-art - functional programming language (currently, the implementation - supports SML \cite{SML}, OCaml \cite{OCaml} and Haskell - \cite{haskell-revised-report}). - We aim to provide a - versatile environment - suitable for software development and verification, - structuring the process - of code generation into a small set of orthogonal principles - while achieving a big coverage of application areas - with maximum flexibility. - - Conceptually the code generator framework is part - of Isabelle's \isa{Pure} meta logic; the object logic - \isa{HOL} which is an extension of \isa{Pure} - already comes with a reasonable framework setup and thus provides - a good working horse for raising code-generation-driven - applications. So, we assume some familiarity and experience - with the ingredients of the \isa{HOL} \emph{Main} theory - (see also \cite{isa-tutorial}).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Overview% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The code generator aims to be usable with no further ado - in most cases while allowing for detailed customization. - This manifests in the structure of this tutorial: - we start with a generic example \secref{sec:example} - and introduce code generation concepts \secref{sec:concept}. - Section - \secref{sec:basics} explains how to use the framework naively, - presuming a reasonable default setup. Then, section - \secref{sec:advanced} deals with advanced topics, - introducing further aspects of the code generator framework - in a motivation-driven manner. Last, section \secref{sec:ml} - introduces the framework's internal programming interfaces. - - \begin{warn} - Ultimately, the code generator which this tutorial deals with - is supposed to replace the already established code generator - by Stefan Berghofer \cite{Berghofer-Nipkow:2002}. - So, for the moment, there are two distinct code generators - in Isabelle. - Also note that while the framework itself is - object-logic independent, only \isa{HOL} provides a reasonable - framework setup. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{An example: a simple theory of search trees \label{sec:example}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -When writing executable specifications using \isa{HOL}, - it is convenient to use - three existing packages: the datatype package for defining - datatypes, the function package for (recursive) functions, - and the class package for overloaded definitions. - - We develope a small theory of search trees; trees are represented - as a datatype with key type \isa{{\isacharprime}a} and value type \isa{{\isacharprime}b}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\isamarkupfalse% -\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree\ {\isacharequal}\ Leaf\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}linorder{\isachardoublequoteclose}\ {\isacharprime}b\isanewline -\ \ {\isacharbar}\ Branch\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharprime}a{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent Note that we have constrained the type of keys - to the class of total orders, \isa{linorder}. - - We define \isa{find} and \isa{update} functions:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ find\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isasymColon}linorder{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}find\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\ it\ {\isacharequal}\ {\isacharparenleft}if\ it\ {\isacharequal}\ key\ then\ Some\ val\ else\ None{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}find\ {\isacharparenleft}Branch\ t{\isadigit{1}}\ key\ t{\isadigit{2}}{\isacharparenright}\ it\ {\isacharequal}\ {\isacharparenleft}if\ it\ {\isasymle}\ key\ then\ find\ t{\isadigit{1}}\ it\ else\ find\ t{\isadigit{2}}\ it{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{fun}\isamarkupfalse% -\isanewline -\ \ update\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}linorder\ {\isasymtimes}\ {\isacharprime}b\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}\isanewline -\ \ \ \ if\ it\ {\isacharequal}\ key\ then\ Leaf\ key\ entry\isanewline -\ \ \ \ \ \ else\ if\ it\ {\isasymle}\ key\isanewline -\ \ \ \ \ \ then\ Branch\ {\isacharparenleft}Leaf\ it\ entry{\isacharparenright}\ it\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\isanewline -\ \ \ \ \ \ else\ Branch\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\ it\ {\isacharparenleft}Leaf\ it\ entry{\isacharparenright}\isanewline -\ \ \ {\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ {\isacharparenleft}Branch\ t{\isadigit{1}}\ key\ t{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}\isanewline -\ \ \ \ if\ it\ {\isasymle}\ key\isanewline -\ \ \ \ \ \ then\ {\isacharparenleft}Branch\ {\isacharparenleft}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ t{\isadigit{1}}{\isacharparenright}\ key\ t{\isadigit{2}}{\isacharparenright}\isanewline -\ \ \ \ \ \ else\ {\isacharparenleft}Branch\ t{\isadigit{1}}\ key\ {\isacharparenleft}update\ {\isacharparenleft}it{\isacharcomma}\ entry{\isacharparenright}\ t{\isadigit{2}}{\isacharparenright}{\isacharparenright}\isanewline -\ \ \ {\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent For testing purpose, we define a small example - using natural numbers \isa{nat} (which are a \isa{linorder}) - as keys and list of nats as values:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ example\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}nat{\isacharcomma}\ nat\ list{\isacharparenright}\ searchtree{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}example\ {\isacharequal}\ update\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharcomma}\ {\isacharbrackleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharbrackright}{\isacharparenright}\ {\isacharparenleft}update\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isacharcomma}\ {\isacharbrackleft}Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isacharbrackright}{\isacharparenright}\isanewline -\ \ \ \ {\isacharparenleft}update\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ {\isacharbrackleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharbrackright}{\isacharparenright}\ {\isacharparenleft}Leaf\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent Then we generate code% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ example\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}tree{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent which looks like: - \lstsml{Thy/examples/tree.ML}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Code generation concepts and process \label{sec:concept}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{figure}[h] - \centering - \includegraphics[width=0.7\textwidth]{codegen_process} - \caption{code generator -- processing overview} - \label{fig:process} - \end{figure} - - The code generator employs a notion of executability - for three foundational executable ingredients known - from functional programming: - \emph{defining equations}, \emph{datatypes}, and - \emph{type classes}. A defining equation as a first approximation - is a theorem of the form \isa{f\ t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n\ {\isasymequiv}\ t} - (an equation headed by a constant \isa{f} with arguments - \isa{t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n} and right hand side \isa{t}). - Code generation aims to turn defining equations - into a functional program by running through - a process (see figure \ref{fig:process}): - - \begin{itemize} - - \item Out of the vast collection of theorems proven in a - \qn{theory}, a reasonable subset modeling - defining equations is \qn{selected}. - - \item On those selected theorems, certain - transformations are carried out - (\qn{preprocessing}). Their purpose is to turn theorems - representing non- or badly executable - specifications into equivalent but executable counterparts. - The result is a structured collection of \qn{code theorems}. - - \item These \qn{code theorems} then are \qn{translated} - into an Haskell-like intermediate - language. - - \item Finally, out of the intermediate language the final - code in the desired \qn{target language} is \qn{serialized}. - - \end{itemize} - - From these steps, only the two last are carried out - outside the logic; by keeping this layer as - thin as possible, the amount of code to trust is - kept to a minimum.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Basics \label{sec:basics}% -} -\isamarkuptrue% -% -\isamarkupsubsection{Invoking the code generator% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Thanks to a reasonable setup of the \isa{HOL} theories, in - most cases code generation proceeds without further ado:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ fac\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ \ \ {\isachardoublequoteopen}fac\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}fac\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ Suc\ n\ {\isacharasterisk}\ fac\ n{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent This executable specification is now turned to SML code:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ fac\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}fac{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent The \isa{{\isasymEXPORTCODE}} command takes a space-separated list of - constants together with \qn{serialization directives} - These start with a \qn{target language} - identifier, followed by a file specification - where to write the generated code to. - - Internally, the defining equations for all selected - constants are taken, including any transitively required - constants, datatypes and classes, resulting in the following - code: - - \lstsml{Thy/examples/fac.ML} - - The code generator will complain when a required - ingredient does not provide a executable counterpart, - e.g.~generating code - for constants not yielding - a defining equation (e.g.~the Hilbert choice - operation \isa{SOME}):% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimML -% -\endisadelimML -% -\isatagML -% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ pick{\isacharunderscore}some\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}pick{\isacharunderscore}some\ xs\ {\isacharequal}\ {\isacharparenleft}SOME\ x{\isachardot}\ x\ {\isasymin}\ set\ xs{\isacharparenright}{\isachardoublequoteclose}% -\isadelimML -% -\endisadelimML -% -\isatagML -% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -\isanewline -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ pick{\isacharunderscore}some\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}fail{\isacharunderscore}const{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent will fail.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Theorem selection% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The list of all defining equations in a theory may be inspected - using the \isa{{\isasymPRINTCODESETUP}} command:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{print{\isacharunderscore}codesetup}\isamarkupfalse% -% -\begin{isamarkuptext}% -\noindent which displays a table of constant with corresponding - defining equations (the additional stuff displayed - shall not bother us for the moment). - - The typical \isa{HOL} tools are already set up in a way that - function definitions introduced by \isa{{\isasymDEFINITION}}, - \isa{{\isasymPRIMREC}}, \isa{{\isasymFUN}}, - \isa{{\isasymFUNCTION}}, \isa{{\isasymCONSTDEFS}}, - \isa{{\isasymRECDEF}} are implicitly propagated - to this defining equation table. Specific theorems may be - selected using an attribute: \emph{code func}. As example, - a weight selector function:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ pick\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}nat\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ list\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}pick\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ n\ {\isacharequal}\ {\isacharparenleft}let\ {\isacharparenleft}k{\isacharcomma}\ v{\isacharparenright}\ {\isacharequal}\ x\ in\isanewline -\ \ \ \ if\ n\ {\isacharless}\ k\ then\ v\ else\ pick\ xs\ {\isacharparenleft}n\ {\isacharminus}\ k{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent We want to eliminate the explicit destruction - of \isa{x} to \isa{{\isacharparenleft}k{\isacharcomma}\ v{\isacharparenright}}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}pick\ {\isacharparenleft}{\isacharparenleft}k{\isacharcomma}\ v{\isacharparenright}{\isacharhash}xs{\isacharparenright}\ n\ {\isacharequal}\ {\isacharparenleft}if\ n\ {\isacharless}\ k\ then\ v\ else\ pick\ xs\ {\isacharparenleft}n\ {\isacharminus}\ k{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ simp% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ pick\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}pick{\isadigit{1}}{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent This theorem now is used for generating code: - - \lstsml{Thy/examples/pick1.ML} - - \noindent The policy is that \emph{default equations} stemming from - \isa{{\isasymDEFINITION}}, - \isa{{\isasymPRIMREC}}, \isa{{\isasymFUN}}, - \isa{{\isasymFUNCTION}}, \isa{{\isasymCONSTDEFS}}, - \isa{{\isasymRECDEF}} statements are discarded as soon as an - equation is explicitly selected by means of \emph{code func}. - Further applications of \emph{code func} add theorems incrementally, - but syntactic redundancies are implicitly dropped. For example, - using a modified version of the \isa{fac} function - as defining equation, the then redundant (since - syntactically subsumed) original defining equations - are dropped. - - \begin{warn} - The attributes \emph{code} and \emph{code del} - associated with the existing code generator also apply to - the new one: \emph{code} implies \emph{code func}, - and \emph{code del} implies \emph{code func del}. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Type classes% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Type classes enter the game via the Isar class package. - For a short introduction how to use it, see \cite{isabelle-classes}; - here we just illustrate its impact on code generation. - - In a target language, type classes may be represented - natively (as in the case of Haskell). For languages - like SML, they are implemented using \emph{dictionaries}. - Our following example specifies a class \qt{null}, - assigning to each of its inhabitants a \qt{null} value:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{class}\isamarkupfalse% -\ null\ {\isacharequal}\ type\ {\isacharplus}\isanewline -\ \ \isakeyword{fixes}\ null\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\isanewline -\isanewline -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ head\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}null\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}head\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ null{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}head\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ x{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent We provide some instances for our \isa{null}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{instantiation}\isamarkupfalse% -\ option\ \isakeyword{and}\ list\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}type{\isacharparenright}\ null\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ {\isachardoublequoteopen}null\ {\isacharequal}\ None{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ {\isachardoublequoteopen}null\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\begin{isamarkuptext}% -\noindent Constructing a dummy example:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ {\isachardoublequoteopen}dummy\ {\isacharequal}\ head\ {\isacharbrackleft}Some\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ None{\isacharbrackright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -Type classes offer a suitable occasion to introduce - the Haskell serializer. Its usage is almost the same - as SML, but, in accordance with conventions - some Haskell systems enforce, each module ends - up in a single file. The module hierarchy is reflected in - the file system, with root directory given as file specification.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ dummy\ \isakeyword{in}\ Haskell\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lsthaskell{Thy/examples/Codegen.hs} - \noindent (we have left out all other modules). - - \medskip - - The whole code in SML with explicit dictionary passing:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ dummy\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}class{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/class.ML} - - \medskip - - \noindent or in OCaml:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ dummy\ \isakeyword{in}\ OCaml\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}class{\isachardot}ocaml{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/class.ocaml} - - \medskip The explicit association of constants - to classes can be inspected using the \isa{{\isasymPRINTCLASSES}} - command.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Recipes and advanced topics \label{sec:advanced}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In this tutorial, we do not attempt to give an exhaustive - description of the code generator framework; instead, - we cast a light on advanced topics by introducing - them together with practically motivated examples. Concerning - further reading, see - - \begin{itemize} - - \item the Isabelle/Isar Reference Manual \cite{isabelle-isar-ref} - for exhaustive syntax diagrams. - \item or \cite{Haftmann-Nipkow:2007:codegen} which deals with foundational issues - of the code generator framework. - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Library theories \label{sec:library}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The \isa{HOL} \isa{Main} theory already provides a code - generator setup - which should be suitable for most applications. Common extensions - and modifications are available by certain theories of the \isa{HOL} - library; beside being useful in applications, they may serve - as a tutorial for customizing the code generator setup. - - \begin{description} - - \item[\isa{Code{\isacharunderscore}Integer}] represents \isa{HOL} integers by big - integer literals in target languages. - \item[\isa{Code{\isacharunderscore}Char}] represents \isa{HOL} characters by - character literals in target languages. - \item[\isa{Code{\isacharunderscore}Char{\isacharunderscore}chr}] like \isa{Code{\isacharunderscore}Char}, - but also offers treatment of character codes; includes - \isa{Code{\isacharunderscore}Integer}. - \item[\isa{Efficient{\isacharunderscore}Nat}] \label{eff_nat} implements natural numbers by integers, - which in general will result in higher efficency; pattern - matching with \isa{{\isadigit{0}}} / \isa{Suc} - is eliminated; includes \isa{Code{\isacharunderscore}Integer}. - \item[\isa{Code{\isacharunderscore}Index}] provides an additional datatype - \isa{index} which is mapped to target-language built-in integers. - Useful for code setups which involve e.g. indexing of - target-language arrays. - \item[\isa{Code{\isacharunderscore}Message}] provides an additional datatype - \isa{message{\isacharunderscore}string} which is isomorphic to strings; - \isa{message{\isacharunderscore}string}s are mapped to target-language strings. - Useful for code setups which involve e.g. printing (error) messages. - - \end{description} - - \begin{warn} - When importing any of these theories, they should form the last - items in an import list. Since these theories adapt the - code generator setup in a non-conservative fashion, - strange effects may occur otherwise. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Preprocessing% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Before selected function theorems are turned into abstract - code, a chain of definitional transformation steps is carried - out: \emph{preprocessing}. In essence, the preprocessor - consists of two components: a \emph{simpset} and \emph{function transformers}. - - The \emph{simpset} allows to employ the full generality of the Isabelle - simplifier. Due to the interpretation of theorems - as defining equations, rewrites are applied to the right - hand side and the arguments of the left hand side of an - equation, but never to the constant heading the left hand side. - An important special case are \emph{inline theorems} which may be - declared an undeclared using the - \emph{code inline} or \emph{code inline del} attribute respectively. - Some common applications:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{itemize} -% -\begin{isamarkuptext}% -\item replacing non-executable constructs by executable ones:% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline -\ \ \ \ {\isachardoublequoteopen}x\ {\isasymin}\ set\ xs\ {\isasymlongleftrightarrow}\ x\ mem\ xs{\isachardoublequoteclose}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\item eliminating superfluous constants:% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline -\ \ \ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isachardoublequoteclose}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ simp% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\item replacing executable but inconvenient constructs:% -\end{isamarkuptext}% -\isamarkuptrue% -\ \ \isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ inline{\isacharbrackright}{\isacharcolon}\isanewline -\ \ \ \ {\isachardoublequoteopen}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongleftrightarrow}\ List{\isachardot}null\ xs{\isachardoublequoteclose}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\end{itemize} -% -\begin{isamarkuptext}% -\emph{Function transformers} provide a very general interface, - transforming a list of function theorems to another - list of function theorems, provided that neither the heading - constant nor its type change. The \isa{{\isadigit{0}}} / \isa{Suc} - pattern elimination implemented in - theory \isa{Efficient{\isacharunderscore}Nat} (see \secref{eff_nat}) uses this - interface. - - \noindent The current setup of the preprocessor may be inspected using - the \isa{{\isasymPRINTCODESETUP}} command. - - \begin{warn} - The attribute \emph{code unfold} - associated with the existing code generator also applies to - the new one: \emph{code unfold} implies \emph{code inline}. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Concerning operational equality% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Surely you have already noticed how equality is treated - by the code generator:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ collect{\isacharunderscore}duplicates\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ \ \ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ xs{\isachardoublequoteclose}\isanewline -\ \ {\isacharbar}\ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharparenleft}z{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ z\ {\isasymin}\ set\ xs\isanewline -\ \ \ \ \ \ then\ if\ z\ {\isasymin}\ set\ ys\isanewline -\ \ \ \ \ \ \ \ then\ collect{\isacharunderscore}duplicates\ xs\ ys\ zs\isanewline -\ \ \ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ xs\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs\isanewline -\ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ {\isacharparenleft}z{\isacharhash}xs{\isacharparenright}\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -The membership test during preprocessing is rewritten, - resulting in \isa{op\ mem}, which itself - performs an explicit equality check.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ collect{\isacharunderscore}duplicates\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}collect{\isacharunderscore}duplicates{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/collect_duplicates.ML}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -Obviously, polymorphic equality is implemented the Haskell - way using a type class. How is this achieved? HOL introduces - an explicit class \isa{eq} with a corresponding operation - \isa{eq{\isacharunderscore}class{\isachardot}eq} such that \isa{eq{\isacharunderscore}class{\isachardot}eq\ x\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharequal}\ y{\isacharparenright}}. - The preprocessing framework does the rest. - For datatypes, instances of \isa{eq} are implicitly derived - when possible. For other types, you may instantiate \isa{eq} - manually like any other type class. - - Though this \isa{eq} class is designed to get rarely in - the way, a subtlety - enters the stage when definitions of overloaded constants - are dependent on operational equality. For example, let - us define a lexicographic ordering on tuples:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{instantiation}\isamarkupfalse% -\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}ord{\isacharcomma}\ ord{\isacharparenright}\ ord\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ {\isacharbrackleft}code\ func\ del{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}p{\isadigit{1}}\ {\isacharless}\ p{\isadigit{2}}\ {\isasymlongleftrightarrow}\ {\isacharparenleft}let\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{1}}{\isacharsemicolon}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{2}}\ in\isanewline -\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ {\isacharbrackleft}code\ func\ del{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}p{\isadigit{1}}\ {\isasymle}\ p{\isadigit{2}}\ {\isasymlongleftrightarrow}\ {\isacharparenleft}let\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{1}}{\isacharsemicolon}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{2}}\ in\isanewline -\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ ord{\isacharunderscore}prod\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}ord{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}ord{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{unfolding}\isamarkupfalse% -\ less{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% -\ simp{\isacharunderscore}all% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -Then code generation will fail. Why? The definition - of \isa{op\ {\isasymle}} depends on equality on both arguments, - which are polymorphic and impose an additional \isa{eq} - class constraint, thus violating the type discipline - for class operations. - - The solution is to add \isa{eq} explicitly to the first sort arguments in the - code theorems:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ ord{\isacharunderscore}prod{\isacharunderscore}code\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}ord{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}ord{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{unfolding}\isamarkupfalse% -\ ord{\isacharunderscore}prod\ \isacommand{by}\isamarkupfalse% -\ rule{\isacharplus}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent Then code generation succeeds:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}eq{\isacharcomma}\ ord{\isacharbraceright}\ {\isasymtimes}\ {\isacharprime}b{\isasymColon}ord\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}lexicographic{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/lexicographic.ML}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -In general, code theorems for overloaded constants may have more - restrictive sort constraints than the underlying instance relation - between class and type constructor as long as the whole system of - constraints is coregular; code theorems violating coregularity - are rejected immediately. Consequently, it might be necessary - to delete disturbing theorems in the code theorem table, - as we have done here with the original definitions \isa{less{\isacharunderscore}prod{\isacharunderscore}def} - and \isa{less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def}. - - In some cases, the automatically derived defining equations - for equality on a particular type may not be appropriate. - As example, watch the following datatype representing - monomorphic parametric types (where type constructors - are referred to by natural numbers):% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\isamarkupfalse% -\ monotype\ {\isacharequal}\ Mono\ nat\ {\isachardoublequoteopen}monotype\ list{\isachardoublequoteclose}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -Then code generation for SML would fail with a message - that the generated code conains illegal mutual dependencies: - the theorem \isa{Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}\ {\isacharequal}\ Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}\ {\isasymequiv}\ tyco{\isadigit{1}}\ {\isacharequal}\ tyco{\isadigit{2}}\ {\isasymand}\ typargs{\isadigit{1}}\ {\isacharequal}\ typargs{\isadigit{2}}} already requires the - instance \isa{monotype\ {\isasymColon}\ eq}, which itself requires - \isa{Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}\ {\isacharequal}\ Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}\ {\isasymequiv}\ tyco{\isadigit{1}}\ {\isacharequal}\ tyco{\isadigit{2}}\ {\isasymand}\ typargs{\isadigit{1}}\ {\isacharequal}\ typargs{\isadigit{2}}}; Haskell has no problem with mutually - recursive \isa{instance} and \isa{function} definitions, - but the SML serializer does not support this. - - In such cases, you have to provide you own equality equations - involving auxiliary constants. In our case, - \isa{list{\isacharunderscore}all{\isadigit{2}}} can do the job:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ monotype{\isacharunderscore}eq{\isacharunderscore}list{\isacharunderscore}all{\isadigit{2}}\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}Mono\ tyco{\isadigit{1}}\ typargs{\isadigit{1}}\ {\isacharequal}\ Mono\ tyco{\isadigit{2}}\ typargs{\isadigit{2}}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ \ tyco{\isadigit{1}}\ {\isacharequal}\ tyco{\isadigit{2}}\ {\isasymand}\ list{\isacharunderscore}all{\isadigit{2}}\ {\isacharparenleft}op\ {\isacharequal}{\isacharparenright}\ typargs{\isadigit{1}}\ typargs{\isadigit{2}}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ list{\isacharunderscore}all{\isadigit{2}}{\isacharunderscore}eq\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -does not depend on instance \isa{monotype\ {\isasymColon}\ eq}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isacharequal}\ {\isacharcolon}{\isacharcolon}\ monotype\ {\isasymRightarrow}\ monotype\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}monotype{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/monotype.ML}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Programs as sets of theorems% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -As told in \secref{sec:concept}, code generation is based - on a structured collection of code theorems. - For explorative purpose, this collection - may be inspected using the \isa{{\isasymCODETHMS}} command:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{code{\isacharunderscore}thms}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ mod\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent prints a table with \emph{all} defining equations - for \isa{op\ mod}, including - \emph{all} defining equations those equations depend - on recursivly. \isa{{\isasymCODETHMS}} provides a convenient - mechanism to inspect the impact of a preprocessor setup - on defining equations. - - Similarly, the \isa{{\isasymCODEDEPS}} command shows a graph - visualizing dependencies between defining equations.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Constructor sets for datatypes% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Conceptually, any datatype is spanned by a set of - \emph{constructors} of type \isa{{\isasymtau}\ {\isacharequal}\ {\isasymdots}\ {\isasymRightarrow}\ {\isasymkappa}\ {\isasymalpha}\isactrlisub {\isadigit{1}}\ {\isasymdots}\ {\isasymalpha}\isactrlisub n} - where \isa{{\isacharbraceleft}{\isasymalpha}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymalpha}\isactrlisub n{\isacharbraceright}} is excactly the set of \emph{all} - type variables in \isa{{\isasymtau}}. The HOL datatype package - by default registers any new datatype in the table - of datatypes, which may be inspected using - the \isa{{\isasymPRINTCODESETUP}} command. - - In some cases, it may be convenient to alter or - extend this table; as an example, we will develope an alternative - representation of natural numbers as binary digits, whose - size does increase logarithmically with its value, not linear - \footnote{Indeed, the \isa{Efficient{\isacharunderscore}Nat} theory (see \ref{eff_nat}) - does something similar}. First, the digit representation:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{definition}\isamarkupfalse% -\ Dig{\isadigit{0}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ n\ {\isacharequal}\ {\isadigit{2}}\ {\isacharasterisk}\ n{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ Dig{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ n\ {\isacharequal}\ Suc\ {\isacharparenleft}{\isadigit{2}}\ {\isacharasterisk}\ n{\isacharparenright}{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\noindent We will use these two ">digits"< to represent natural numbers - in binary digits, e.g.:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ {\isadigit{4}}{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}{\isadigit{4}}{\isadigit{2}}\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}Dig{\isadigit{1}}\ {\isacharparenleft}Dig{\isadigit{0}}\ {\isacharparenleft}Dig{\isadigit{1}}\ {\isacharparenleft}Dig{\isadigit{0}}\ {\isadigit{1}}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp\ add{\isacharcolon}\ Dig{\isadigit{0}}{\isacharunderscore}def\ Dig{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent Of course we also have to provide proper code equations for - the operations, e.g. \isa{op\ {\isacharplus}}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ plus{\isacharunderscore}Dig\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}{\isadigit{0}}\ {\isacharplus}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}m\ {\isacharplus}\ {\isadigit{0}}\ {\isacharequal}\ m{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharplus}\ Dig{\isadigit{0}}\ n\ {\isacharequal}\ Dig{\isadigit{1}}\ n{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ m\ {\isacharplus}\ {\isadigit{1}}\ {\isacharequal}\ Dig{\isadigit{1}}\ m{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharplus}\ Dig{\isadigit{1}}\ n\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ m\ {\isacharplus}\ {\isadigit{1}}\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}m\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ m\ {\isacharplus}\ Dig{\isadigit{0}}\ n\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}m\ {\isacharplus}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{0}}\ m\ {\isacharplus}\ Dig{\isadigit{1}}\ n\ {\isacharequal}\ Dig{\isadigit{1}}\ {\isacharparenleft}m\ {\isacharplus}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ m\ {\isacharplus}\ Dig{\isadigit{0}}\ n\ {\isacharequal}\ Dig{\isadigit{1}}\ {\isacharparenleft}m\ {\isacharplus}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}Dig{\isadigit{1}}\ m\ {\isacharplus}\ Dig{\isadigit{1}}\ n\ {\isacharequal}\ Dig{\isadigit{0}}\ {\isacharparenleft}m\ {\isacharplus}\ n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ Dig{\isadigit{0}}{\isacharunderscore}def\ Dig{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent We then instruct the code generator to view \isa{{\isadigit{0}}}, - \isa{{\isadigit{1}}}, \isa{Dig{\isadigit{0}}} and \isa{Dig{\isadigit{1}}} as - datatype constructors:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{code{\isacharunderscore}datatype}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isadigit{0}}{\isasymColon}nat{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isadigit{1}}{\isasymColon}nat{\isachardoublequoteclose}\ Dig{\isadigit{0}}\ Dig{\isadigit{1}}% -\begin{isamarkuptext}% -\noindent For the former constructor \isa{Suc}, we provide a code - equation and remove some parts of the default code generator setup - which are an obstacle here:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ Suc{\isacharunderscore}Dig\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}Suc\ n\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ simp% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{declare}\isamarkupfalse% -\ One{\isacharunderscore}nat{\isacharunderscore}def\ {\isacharbrackleft}code\ inline\ del{\isacharbrackright}\isanewline -\isacommand{declare}\isamarkupfalse% -\ add{\isacharunderscore}Suc{\isacharunderscore}shift\ {\isacharbrackleft}code\ func\ del{\isacharbrackright}% -\begin{isamarkuptext}% -\noindent This yields the following code:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isacharplus}\ {\isasymColon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}nat{\isacharunderscore}binary{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/nat_binary.ML}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\medskip - - From this example, it can be easily glimpsed that using own constructor sets - is a little delicate since it changes the set of valid patterns for values - of that type. Without going into much detail, here some practical hints: - - \begin{itemize} - \item When changing the constuctor set for datatypes, take care to - provide an alternative for the \isa{case} combinator (e.g. by replacing - it using the preprocessor). - \item Values in the target language need not to be normalized -- different - values in the target language may represent the same value in the - logic (e.g. \isa{Dig{\isadigit{1}}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}}). - \item Usually, a good methodology to deal with the subleties of pattern - matching is to see the type as an abstract type: provide a set - of operations which operate on the concrete representation of the type, - and derive further operations by combinations of these primitive ones, - without relying on a particular representation. - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isataginvisible -\isacommand{code{\isacharunderscore}datatype}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isadigit{0}}{\isacharcolon}{\isacharcolon}nat{\isachardoublequoteclose}\ Suc\isanewline -\isacommand{declare}\isamarkupfalse% -\ plus{\isacharunderscore}Dig\ {\isacharbrackleft}code\ func\ del{\isacharbrackright}\isanewline -\isacommand{declare}\isamarkupfalse% -\ One{\isacharunderscore}nat{\isacharunderscore}def\ {\isacharbrackleft}code\ inline{\isacharbrackright}\isanewline -\isacommand{declare}\isamarkupfalse% -\ add{\isacharunderscore}Suc{\isacharunderscore}shift\ {\isacharbrackleft}code\ func{\isacharbrackright}\ \isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isadigit{0}}\ {\isacharplus}\ n\ {\isacharequal}\ {\isacharparenleft}n\ {\isasymColon}\ nat{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% -\ simp% -\endisataginvisible -{\isafoldinvisible}% -% -\isadeliminvisible -% -\endisadeliminvisible -% -\isamarkupsubsection{Customizing serialization% -} -\isamarkuptrue% -% -\isamarkupsubsubsection{Basics% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Consider the following function and its corresponding - SML code:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\isanewline -\ \ in{\isacharunderscore}interval\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymtimes}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isakeyword{where}\isanewline -\ \ {\isachardoublequoteopen}in{\isacharunderscore}interval\ {\isacharparenleft}k{\isacharcomma}\ l{\isacharparenright}\ n\ {\isasymlongleftrightarrow}\ k\ {\isasymle}\ n\ {\isasymand}\ n\ {\isasymle}\ l{\isachardoublequoteclose}% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ in{\isacharunderscore}interval\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}bool{\isacharunderscore}literal{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/bool_literal.ML} - - \noindent Though this is correct code, it is a little bit unsatisfactory: - boolean values and operators are materialized as distinguished - entities with have nothing to do with the SML-builtin notion - of \qt{bool}. This results in less readable code; - additionally, eager evaluation may cause programs to - loop or break which would perfectly terminate when - the existing SML \qt{bool} would be used. To map - the HOL \qt{bool} on SML \qt{bool}, we may use - \qn{custom serializations}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}type}\isamarkupfalse% -\ bool\isanewline -\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}bool{\isachardoublequoteclose}{\isacharparenright}\isanewline -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ True\ \isakeyword{and}\ False\ \isakeyword{and}\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline -\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}true{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}false{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}{\isacharunderscore}\ andalso\ {\isacharunderscore}{\isachardoublequoteclose}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -% -\begin{isamarkuptext}% -The \isa{{\isasymCODETYPE}} commad takes a type constructor - as arguments together with a list of custom serializations. - Each custom serialization starts with a target language - identifier followed by an expression, which during - code serialization is inserted whenever the type constructor - would occur. For constants, \isa{{\isasymCODECONST}} implements - the corresponding mechanism. Each ``\verb|_|'' in - a serialization expression is treated as a placeholder - for the type constructor's (the constant's) arguments.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ in{\isacharunderscore}interval\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}bool{\isacharunderscore}mlbool{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/bool_mlbool.ML} - - \noindent This still is not perfect: the parentheses - around the \qt{andalso} expression are superfluous. - Though the serializer - by no means attempts to imitate the rich Isabelle syntax - framework, it provides some common idioms, notably - associative infixes with precedences which may be used here:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isasymand}{\isachardoublequoteclose}\isanewline -\ \ {\isacharparenleft}SML\ \isakeyword{infixl}\ {\isadigit{1}}\ {\isachardoublequoteopen}andalso{\isachardoublequoteclose}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -\isanewline -\isanewline -\isacommand{export{\isacharunderscore}code}\isamarkupfalse% -\ in{\isacharunderscore}interval\ \ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}bool{\isacharunderscore}infix{\isachardot}ML{\isachardoublequoteclose}% -\begin{isamarkuptext}% -\lstsml{Thy/examples/bool_infix.ML} - - \medskip - - Next, we try to map HOL pairs to SML pairs, using the - infix ``\verb|*|'' type constructor and parentheses:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}type}\isamarkupfalse% -\ {\isacharasterisk}\isanewline -\ \ {\isacharparenleft}SML\ \isakeyword{infix}\ {\isadigit{2}}\ {\isachardoublequoteopen}{\isacharasterisk}{\isachardoublequoteclose}{\isacharparenright}\isanewline -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ Pair\isanewline -\ \ {\isacharparenleft}SML\ {\isachardoublequoteopen}{\isacharbang}{\isacharparenleft}{\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharcomma}{\isacharslash}\ {\isacharparenleft}{\isacharunderscore}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -% -\begin{isamarkuptext}% -The initial bang ``\verb|!|'' tells the serializer to never put - parentheses around the whole expression (they are already present), - while the parentheses around argument place holders - tell not to put parentheses around the arguments. - The slash ``\verb|/|'' (followed by arbitrary white space) - inserts a space which may be used as a break if necessary - during pretty printing. - - These examples give a glimpse what mechanisms - custom serializations provide; however their usage - requires careful thinking in order not to introduce - inconsistencies -- or, in other words: - custom serializations are completely axiomatic. - - A further noteworthy details is that any special - character in a custom serialization may be quoted - using ``\verb|'|''; thus, in - ``\verb|fn '_ => _|'' the first - ``\verb|_|'' is a proper underscore while the - second ``\verb|_|'' is a placeholder. - - The HOL theories provide further - examples for custom serializations.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Haskell serialization% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -For convenience, the default - HOL setup for Haskell maps the \isa{eq} class to - its counterpart in Haskell, giving custom serializations - for the class (\isa{{\isasymCODECLASS}}) and its operation:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}class}\isamarkupfalse% -\ eq\isanewline -\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Eq{\isachardoublequoteclose}\ \isakeyword{where}\ {\isachardoublequoteopen}op\ {\isacharequal}{\isachardoublequoteclose}\ {\isasymequiv}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharequal}{\isacharequal}{\isacharparenright}{\isachardoublequoteclose}{\isacharparenright}\isanewline -\isanewline -\isacommand{code{\isacharunderscore}const}\isamarkupfalse% -\ {\isachardoublequoteopen}op\ {\isacharequal}{\isachardoublequoteclose}\isanewline -\ \ {\isacharparenleft}Haskell\ \isakeyword{infixl}\ {\isadigit{4}}\ {\isachardoublequoteopen}{\isacharequal}{\isacharequal}{\isachardoublequoteclose}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -% -\begin{isamarkuptext}% -A problem now occurs whenever a type which - is an instance of \isa{eq} in HOL is mapped - on a Haskell-builtin type which is also an instance - of Haskell \isa{Eq}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{typedecl}\isamarkupfalse% -\ bar\isanewline -\isanewline -\isacommand{instantiation}\isamarkupfalse% -\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ {\isachardoublequoteopen}eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}x{\isasymColon}bar{\isacharparenright}\ y\ {\isasymlongleftrightarrow}\ x\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ default\ {\isacharparenleft}simp\ add{\isacharcolon}\ eq{\isacharunderscore}bar{\isacharunderscore}def{\isacharparenright}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -\isanewline -% -\isadelimtt -\isanewline -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}type}\isamarkupfalse% -\ bar\isanewline -\ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Integer{\isachardoublequoteclose}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -% -\begin{isamarkuptext}% -The code generator would produce - an additional instance, which of course is rejected. - To suppress this additional instance, use - \isa{{\isasymCODEINSTANCE}}:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtt -% -\endisadelimtt -% -\isatagtt -\isacommand{code{\isacharunderscore}instance}\isamarkupfalse% -\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline -\ \ {\isacharparenleft}Haskell\ {\isacharminus}{\isacharparenright}% -\endisatagtt -{\isafoldtt}% -% -\isadelimtt -% -\endisadelimtt -% -\isamarkupsubsubsection{Pretty printing% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The serializer provides ML interfaces to set up - pretty serializations for expressions like lists, numerals - and characters; these are - monolithic stubs and should only be used with the - theories introduced in \secref{sec:library}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Cyclic module dependencies% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Sometimes the awkward situation occurs that dependencies - between definitions introduce cyclic dependencies - between modules, which in the Haskell world leaves - you to the mercy of the Haskell implementation you are using, - while for SML code generation is not possible. - - A solution is to declare module names explicitly. - Let use assume the three cyclically dependent - modules are named \emph{A}, \emph{B} and \emph{C}. - Then, by stating% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{code{\isacharunderscore}modulename}\isamarkupfalse% -\ SML\isanewline -\ \ A\ ABC\isanewline -\ \ B\ ABC\isanewline -\ \ C\ ABC% -\begin{isamarkuptext}% -we explicitly map all those modules on \emph{ABC}, - resulting in an ad-hoc merge of this three modules - at serialization time.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Incremental code generation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Code generation is \emph{incremental}: theorems - and abstract intermediate code are cached and extended on demand. - The cache may be partially or fully dropped if the underlying - executable content of the theory changes. - Implementation of caching is supposed to transparently - hid away the details from the user. Anyway, caching - reaches the surface by using a slightly more general form - of the \isa{{\isasymCODETHMS}}, \isa{{\isasymCODEDEPS}} - and \isa{{\isasymEXPORTCODE}} commands: the list of constants - may be omitted. Then, all constants with code theorems - in the current cache are referred to.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{ML interfaces \label{sec:ml}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Since the code generator framework not only aims to provide - a nice Isar interface but also to form a base for - code-generation-based applications, here a short - description of the most important ML interfaces.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Executable theory content: \isa{Code}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -This Pure module implements the core notions of - executable content of a theory.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Managing executable content% -} -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{Code.add\_func}\verb|Code.add_func: thm -> theory -> theory| \\ - \indexml{Code.del\_func}\verb|Code.del_func: thm -> theory -> theory| \\ - \indexml{Code.add\_funcl}\verb|Code.add_funcl: string * thm list Susp.T -> theory -> theory| \\ - \indexml{Code.map\_pre}\verb|Code.map_pre: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory| \\ - \indexml{Code.map\_post}\verb|Code.map_post: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory| \\ - \indexml{Code.add\_functrans}\verb|Code.add_functrans: string * (theory -> thm list -> thm list option)|\isasep\isanewline% -\verb| -> theory -> theory| \\ - \indexml{Code.del\_functrans}\verb|Code.del_functrans: string -> theory -> theory| \\ - \indexml{Code.add\_datatype}\verb|Code.add_datatype: (string * typ) list -> theory -> theory| \\ - \indexml{Code.get\_datatype}\verb|Code.get_datatype: theory -> string|\isasep\isanewline% -\verb| -> (string * sort) list * (string * typ list) list| \\ - \indexml{Code.get\_datatype\_of\_constr}\verb|Code.get_datatype_of_constr: theory -> string -> string option| - \end{mldecls} - - \begin{description} - - \item \verb|Code.add_func|~\isa{thm}~\isa{thy} adds function - theorem \isa{thm} to executable content. - - \item \verb|Code.del_func|~\isa{thm}~\isa{thy} removes function - theorem \isa{thm} from executable content, if present. - - \item \verb|Code.add_funcl|~\isa{{\isacharparenleft}const{\isacharcomma}\ lthms{\isacharparenright}}~\isa{thy} adds - suspended defining equations \isa{lthms} for constant - \isa{const} to executable content. - - \item \verb|Code.map_pre|~\isa{f}~\isa{thy} changes - the preprocessor simpset. - - \item \verb|Code.add_functrans|~\isa{{\isacharparenleft}name{\isacharcomma}\ f{\isacharparenright}}~\isa{thy} adds - function transformer \isa{f} (named \isa{name}) to executable content; - \isa{f} is a transformer of the defining equations belonging - to a certain function definition, depending on the - current theory context. Returning \isa{NONE} indicates that no - transformation took place; otherwise, the whole process will be iterated - with the new defining equations. - - \item \verb|Code.del_functrans|~\isa{name}~\isa{thy} removes - function transformer named \isa{name} from executable content. - - \item \verb|Code.add_datatype|~\isa{cs}~\isa{thy} adds - a datatype to executable content, with generation - set \isa{cs}. - - \item \verb|Code.get_datatype_of_constr|~\isa{thy}~\isa{const} - returns type constructor corresponding to - constructor \isa{const}; returns \isa{NONE} - if \isa{const} is no constructor. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Auxiliary% -} -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexml{Code\_Unit.read\_const}\verb|Code.read_const: theory -> string -> string| \\ - \indexml{Code\_Unit.head\_func}\verb|Code.head_func: thm -> string * ((string * sort) list * typ)| \\ - \indexml{Code\_Unit.rewrite\_func}\verb|Code.rewrite_func: MetaSimplifier.simpset -> thm -> thm| \\ - \end{mldecls} - - \begin{description} - - \item \verb|Code.read_const|~\isa{thy}~\isa{s} - reads a constant as a concrete term expression \isa{s}. - - \item \verb|Code.head_func|~\isa{thm} - extracts the constant and its type from a defining equation \isa{thm}. - - \item \verb|Code.rewrite_func|~\isa{ss}~\isa{thm} - rewrites a defining equation \isa{thm} with a simpset \isa{ss}; - only arguments and right hand side are rewritten, - not the head of the defining equation. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupsubsection{Implementing code generator applications% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Implementing code generator applications on top - of the framework set out so far usually not only - involves using those primitive interfaces - but also storing code-dependent data and various - other things. - - \begin{warn} - Some interfaces discussed here have not reached - a final state yet. - Changes likely to occur in future. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Data depending on the theory's executable content% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Due to incrementality of code generation, changes in the - theory's executable content have to be propagated in a - certain fashion. Additionally, such changes may occur - not only during theory extension but also during theory - merge, which is a little bit nasty from an implementation - point of view. The framework provides a solution - to this technical challenge by providing a functorial - data slot \verb|CodeDataFun|; on instantiation - of this functor, the following types and operations - are required: - - \medskip - \begin{tabular}{l} - \isa{type\ T} \\ - \isa{val\ empty{\isacharcolon}\ T} \\ - \isa{val\ merge{\isacharcolon}\ Pretty{\isachardot}pp\ {\isasymrightarrow}\ T\ {\isacharasterisk}\ T\ {\isasymrightarrow}\ T} \\ - \isa{val\ purge{\isacharcolon}\ theory\ option\ {\isasymrightarrow}\ CodeUnit{\isachardot}const\ list\ option\ {\isasymrightarrow}\ T\ {\isasymrightarrow}\ T} - \end{tabular} - - \begin{description} - - \item \isa{T} the type of data to store. - - \item \isa{empty} initial (empty) data. - - \item \isa{merge} merging two data slots. - - \item \isa{purge}~\isa{thy}~\isa{consts} propagates changes in executable content; - if possible, the current theory context is handed over - as argument \isa{thy} (if there is no current theory context (e.g.~during - theory merge, \verb|NONE|); \isa{consts} indicates the kind - of change: \verb|NONE| stands for a fundamental change - which invalidates any existing code, \isa{SOME\ consts} - hints that executable content for constants \isa{consts} - has changed. - - \end{description} - - An instance of \verb|CodeDataFun| provides the following - interface: - - \medskip - \begin{tabular}{l} - \isa{get{\isacharcolon}\ theory\ {\isasymrightarrow}\ T} \\ - \isa{change{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ T{\isacharparenright}\ {\isasymrightarrow}\ T} \\ - \isa{change{\isacharunderscore}yield{\isacharcolon}\ theory\ {\isasymrightarrow}\ {\isacharparenleft}T\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T{\isacharparenright}\ {\isasymrightarrow}\ {\isacharprime}a\ {\isacharasterisk}\ T} - \end{tabular} - - \begin{description} - - \item \isa{get} retrieval of the current data. - - \item \isa{change} update of current data (cached!) - by giving a continuation. - - \item \isa{change{\isacharunderscore}yield} update with side result. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\emph{Happy proving, happy hacking!}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3a7937841585 -r 156a060d5d68 doc-src/IsarRef/Thy/HOL_Specific.thy --- a/doc-src/IsarRef/Thy/HOL_Specific.thy Mon Dec 21 16:50:28 2009 +0000 +++ b/doc-src/IsarRef/Thy/HOL_Specific.thy Wed Dec 23 08:31:33 2009 +0100 @@ -1139,7 +1139,7 @@ \end{matharray} \begin{rail} - 'export\_code' ( constexpr + ) ? \\ + 'export\_code' ( constexpr + ) \\ ( ( 'in' target ( 'module\_name' string ) ? \\ ( 'file' ( string | '-' ) ) ? ( '(' args ')' ) ?) + ) ? ; @@ -1216,10 +1216,8 @@ \item @{command (HOL) "export_code"} is the canonical interface for generating and serializing code: for a given list of constants, code - is generated for the specified target languages. Abstract code is - cached incrementally. If no constant is given, the currently cached - code is serialized. If no serialization instruction is given, only - abstract code is cached. + is generated for the specified target languages. If no serialization + instruction is given, only abstract code is generated internally. Constants may be specified by giving them literally, referring to all executable contants within a certain theory by giving @{text @@ -1239,20 +1237,21 @@ code internally in the context of the current ML session. Serializers take an optional list of arguments in parentheses. For - \emph{Haskell} a module name prefix may be given using the ``@{text + \emph{SML} and \emph{OCaml}, ``@{text no_signatures}`` omits + explicit module signatures. + + For \emph{Haskell} a module name prefix may be given using the ``@{text "root:"}'' argument; ``@{text string_classes}'' adds a ``@{verbatim "deriving (Read, Show)"}'' clause to each appropriate datatype declaration. \item @{command (HOL) "code_thms"} prints a list of theorems representing the corresponding program containing all given - constants; if no constants are given, the currently cached code - theorems are printed. + constants. \item @{command (HOL) "code_deps"} visualizes dependencies of theorems representing the corresponding program containing all given - constants; if no constants are given, the currently cached code - theorems are visualized. + constants. \item @{command (HOL) "code_datatype"} specifies a constructor set for a logical type. diff -r 3a7937841585 -r 156a060d5d68 doc-src/IsarRef/Thy/document/HOL_Specific.tex --- a/doc-src/IsarRef/Thy/document/HOL_Specific.tex Mon Dec 21 16:50:28 2009 +0000 +++ b/doc-src/IsarRef/Thy/document/HOL_Specific.tex Wed Dec 23 08:31:33 2009 +0100 @@ -1150,7 +1150,7 @@ \end{matharray} \begin{rail} - 'export\_code' ( constexpr + ) ? \\ + 'export\_code' ( constexpr + ) \\ ( ( 'in' target ( 'module\_name' string ) ? \\ ( 'file' ( string | '-' ) ) ? ( '(' args ')' ) ?) + ) ? ; @@ -1227,10 +1227,8 @@ \item \hyperlink{command.HOL.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} is the canonical interface for generating and serializing code: for a given list of constants, code - is generated for the specified target languages. Abstract code is - cached incrementally. If no constant is given, the currently cached - code is serialized. If no serialization instruction is given, only - abstract code is cached. + is generated for the specified target languages. If no serialization + instruction is given, only abstract code is generated internally. Constants may be specified by giving them literally, referring to all executable contants within a certain theory by giving \isa{{\isachardoublequote}name{\isachardot}{\isacharasterisk}{\isachardoublequote}}, or referring to \emph{all} executable constants currently @@ -1249,18 +1247,19 @@ code internally in the context of the current ML session. Serializers take an optional list of arguments in parentheses. For - \emph{Haskell} a module name prefix may be given using the ``\isa{{\isachardoublequote}root{\isacharcolon}{\isachardoublequote}}'' argument; ``\isa{string{\isacharunderscore}classes}'' adds a ``\verb|deriving (Read, Show)|'' clause to each appropriate datatype + \emph{SML} and \emph{OCaml}, ``\isa{no{\isacharunderscore}signatures}`` omits + explicit module signatures. + + For \emph{Haskell} a module name prefix may be given using the ``\isa{{\isachardoublequote}root{\isacharcolon}{\isachardoublequote}}'' argument; ``\isa{string{\isacharunderscore}classes}'' adds a ``\verb|deriving (Read, Show)|'' clause to each appropriate datatype declaration. \item \hyperlink{command.HOL.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}} prints a list of theorems representing the corresponding program containing all given - constants; if no constants are given, the currently cached code - theorems are printed. + constants. \item \hyperlink{command.HOL.code-deps}{\mbox{\isa{\isacommand{code{\isacharunderscore}deps}}}} visualizes dependencies of theorems representing the corresponding program containing all given - constants; if no constants are given, the currently cached code - theorems are visualized. + constants. \item \hyperlink{command.HOL.code-datatype}{\mbox{\isa{\isacommand{code{\isacharunderscore}datatype}}}} specifies a constructor set for a logical type. diff -r 3a7937841585 -r 156a060d5d68 src/Pure/Isar/code.ML --- a/src/Pure/Isar/code.ML Mon Dec 21 16:50:28 2009 +0000 +++ b/src/Pure/Isar/code.ML Wed Dec 23 08:31:33 2009 +0100 @@ -1,8 +1,9 @@ (* Title: Pure/Isar/code.ML Author: Florian Haftmann, TU Muenchen -Abstract executable code of theory. Management of data dependent on -executable code. Cache assumes non-concurrent processing of a single theory. +Abstract executable ingredients of theory. Management of data +dependent on executable ingredients as synchronized cache; purged +on any change of underlying executable ingredients. *) signature CODE = @@ -70,13 +71,11 @@ sig type T val empty: T - val purge: theory -> string list -> T -> T end; signature CODE_DATA = sig type T - val get: theory -> T val change: theory -> (T -> T) -> T val change_yield: theory -> (T -> 'a * T) -> 'a * T end; @@ -84,10 +83,7 @@ signature PRIVATE_CODE = sig include CODE - val declare_data: Object.T -> (theory -> string list -> Object.T -> Object.T) - -> serial - val get_data: serial * ('a -> Object.T) * (Object.T -> 'a) - -> theory -> 'a + val declare_data: Object.T -> serial val change_data: serial * ('a -> Object.T) * (Object.T -> 'a) -> theory -> ('a -> 'a) -> 'a val change_yield_data: serial * ('a -> Object.T) * (Object.T -> 'a) @@ -211,13 +207,9 @@ local -type kind = { - empty: Object.T, - purge: theory -> string list -> Object.T -> Object.T -}; +type kind = { empty: Object.T }; val kinds = Unsynchronized.ref (Datatab.empty: kind Datatab.table); -val kind_keys = Unsynchronized.ref ([]: serial list); fun invoke f k = case Datatab.lookup (! kinds) k of SOME kind => f kind @@ -225,20 +217,15 @@ in -fun declare_data empty purge = +fun declare_data empty = let val k = serial (); - val kind = {empty = empty, purge = purge}; - val _ = Unsynchronized.change kinds (Datatab.update (k, kind)); - val _ = Unsynchronized.change kind_keys (cons k); + val kind = { empty = empty }; + val _ = CRITICAL (fn () => Unsynchronized.change kinds (Datatab.update (k, kind))); in k end; fun invoke_init k = invoke (fn kind => #empty kind) k; -fun invoke_purge_all thy cs = - fold (fn k => Datatab.map_entry k - (invoke (fn kind => #purge kind thy cs) k)) (! kind_keys); - end; (*local*) @@ -247,26 +234,27 @@ local type data = Object.T Datatab.table; -val empty_data = Datatab.empty : data; +fun create_data data = Synchronized.var "code data" data; +fun empty_data () = create_data Datatab.empty; structure Code_Data = TheoryDataFun ( - type T = spec * data Unsynchronized.ref; + type T = spec * data Synchronized.var; val empty = (make_spec (false, (((Symtab.empty, Symtab.empty), Symtab.empty), - (Symtab.empty, (Symtab.empty, Symtab.empty)))), Unsynchronized.ref empty_data); - fun copy (spec, data) = (spec, Unsynchronized.ref (! data)); + (Symtab.empty, (Symtab.empty, Symtab.empty)))), empty_data ()); + fun copy (spec, data) = (spec, create_data (Synchronized.value data)); val extend = copy; fun merge _ ((spec1, _), (spec2, _)) = - (merge_spec (spec1, spec2), Unsynchronized.ref empty_data); + (merge_spec (spec1, spec2), empty_data ()); ); fun thy_data f thy = f ((snd o Code_Data.get) thy); -fun get_ensure_init kind data_ref = - case Datatab.lookup (! data_ref) kind +fun get_ensure_init kind data = + case Datatab.lookup (Synchronized.value data) kind of SOME x => x | NONE => let val y = invoke_init kind - in (Unsynchronized.change data_ref (Datatab.update (kind, y)); y) end; + in (Synchronized.change data (Datatab.update (kind, y)); y) end; in @@ -274,19 +262,12 @@ val the_exec = fst o Code_Data.get; -fun complete_class_params thy cs = - fold (fn c => case AxClass.inst_of_param thy c - of NONE => insert (op =) c - | SOME (c', _) => insert (op =) c' #> insert (op =) c) cs []; +fun map_exec_purge f = + Code_Data.map (fn (exec, data) => (f exec, empty_data ())); -fun map_exec_purge touched f thy = - Code_Data.map (fn (exec, data) => (f exec, Unsynchronized.ref (case touched - of SOME cs => invoke_purge_all thy (complete_class_params thy cs) (! data) - | NONE => empty_data))) thy; +val purge_data = (Code_Data.map o apsnd) (fn _ => empty_data ()); -val purge_data = (Code_Data.map o apsnd) (fn _ => Unsynchronized.ref empty_data); - -fun change_eqns delete c f = (map_exec_purge (SOME [c]) o map_eqns +fun change_eqns delete c f = (map_exec_purge o map_eqns o (if delete then Symtab.map_entry c else Symtab.map_default (c, ((false, (true, [])), []))) o apfst) (fn (_, eqns) => (true, f eqns)); @@ -323,15 +304,13 @@ (* access to data dependent on abstract executable code *) -fun get_data (kind, _, dest) = thy_data (get_ensure_init kind #> dest); - fun change_data (kind, mk, dest) = let fun chnge data_ref f = let val data = get_ensure_init kind data_ref; val data' = f (dest data); - in (Unsynchronized.change data_ref (Datatab.update (kind, mk data')); data') end; + in (Synchronized.change data_ref (Datatab.update (kind, mk data')); data') end; in thy_data chnge end; fun change_yield_data (kind, mk, dest) = @@ -340,7 +319,7 @@ let val data = get_ensure_init kind data_ref; val (x, data') = f (dest data); - in (x, (Unsynchronized.change data_ref (Datatab.update (kind, mk data')); data')) end; + in (x, (Synchronized.change data_ref (Datatab.update (kind, mk data')); data')) end; in thy_data chnge end; end; (*local*) @@ -707,7 +686,7 @@ fun add_type tyco thy = case Symtab.lookup ((snd o #types o Type.rep_tsig o Sign.tsig_of) thy) tyco of SOME (Type.Abbreviation (vs, _, _)) => - (map_exec_purge NONE o map_signatures o apfst) + (map_exec_purge o map_signatures o apfst) (Symtab.update (tyco, length vs)) thy | _ => error ("No such type abbreviation: " ^ quote tyco); @@ -723,7 +702,7 @@ error ("Illegal constant signature: " ^ Syntax.string_of_typ_global thy ty); in thy - |> (map_exec_purge NONE o map_signatures o apsnd) (Symtab.update (c, ty)) + |> (map_exec_purge o map_signatures o apsnd) (Symtab.update (c, ty)) end; val add_signature = gen_add_signature (K I) cert_signature; @@ -747,7 +726,7 @@ in thy |> fold (del_eqns o fst) cs - |> map_exec_purge NONE + |> map_exec_purge ((map_dtyps o Symtab.map_default (tyco, [])) (cons (serial (), vs_cos)) #> (map_cases o apfst) drop_outdated_cases) |> Type_Interpretation.data (tyco, serial ()) @@ -838,29 +817,27 @@ of [] => () | cs => error ("Non-constructor(s) in case certificate: " ^ commas (map quote cs)); val entry = (1 + Int.max (1, length case_pats), (k, case_pats)) - in (map_exec_purge (SOME [c]) o map_cases o apfst) (Symtab.update (c, entry)) thy end; + in (map_exec_purge o map_cases o apfst) (Symtab.update (c, entry)) thy end; fun add_undefined c thy = - (map_exec_purge (SOME [c]) o map_cases o apsnd) (Symtab.update (c, ())) thy; + (map_exec_purge o map_cases o apsnd) (Symtab.update (c, ())) thy; end; (*struct*) (** type-safe interfaces for data dependent on executable code **) -functor Code_Data_Fun(Data: CODE_DATA_ARGS): CODE_DATA = +functor Code_Data(Data: CODE_DATA_ARGS): CODE_DATA = struct type T = Data.T; exception Data of T; fun dest (Data x) = x -val kind = Code.declare_data (Data Data.empty) - (fn thy => fn cs => fn Data x => Data (Data.purge thy cs x)); +val kind = Code.declare_data (Data Data.empty); val data_op = (kind, Data, dest); -val get = Code.get_data data_op; val change = Code.change_data data_op; fun change_yield thy = Code.change_yield_data data_op thy; diff -r 3a7937841585 -r 156a060d5d68 src/Tools/Code/code_preproc.ML --- a/src/Tools/Code/code_preproc.ML Mon Dec 21 16:50:28 2009 +0000 +++ b/src/Tools/Code/code_preproc.ML Wed Dec 23 08:31:33 2009 +0100 @@ -445,22 +445,10 @@ (** store for preprocessed arities and code equations **) -structure Wellsorted = Code_Data_Fun +structure Wellsorted = Code_Data ( type T = ((string * class) * sort list) list * code_graph; val empty = ([], Graph.empty); - fun purge thy cs (arities, eqngr) = - let - val del_cs = ((Graph.all_preds eqngr - o filter (can (Graph.get_node eqngr))) cs); - val del_arities = del_cs - |> map_filter (AxClass.inst_of_param thy) - |> maps (fn (c, tyco) => - (map (rpair tyco) o Sign.complete_sort thy o the_list - o AxClass.class_of_param thy) c); - val arities' = fold (AList.delete (op =)) del_arities arities; - val eqngr' = Graph.del_nodes del_cs eqngr; - in (arities', eqngr') end; ); diff -r 3a7937841585 -r 156a060d5d68 src/Tools/Code/code_target.ML --- a/src/Tools/Code/code_target.ML Mon Dec 21 16:50:28 2009 +0000 +++ b/src/Tools/Code/code_target.ML Wed Dec 23 08:31:33 2009 +0100 @@ -356,15 +356,9 @@ (fn (c, name) => if member (op =) names4 name then SOME c else NONE) (cs2 ~~ names3); in fold (insert (op =)) cs5 cs1 end; -fun cached_program thy = - let - val (naming, program) = Code_Thingol.cached_program thy; - in (transitivly_non_empty_funs thy naming program, (naming, program)) end - fun export_code thy cs seris = let - val (cs', (naming, program)) = if null cs then cached_program thy - else Code_Thingol.consts_program thy cs; + val (cs', (naming, program)) = Code_Thingol.consts_program thy cs; fun mk_seri_dest dest = case dest of NONE => compile | SOME "-" => export @@ -514,7 +508,7 @@ val (inK, module_nameK, fileK) = ("in", "module_name", "file"); val code_exprP = - (Scan.repeat P.term_group + (Scan.repeat1 P.term_group -- Scan.repeat (P.$$$ inK |-- P.name -- Scan.option (P.$$$ module_nameK |-- P.name) -- Scan.option (P.$$$ fileK |-- P.name) diff -r 3a7937841585 -r 156a060d5d68 src/Tools/Code/code_thingol.ML --- a/src/Tools/Code/code_thingol.ML Mon Dec 21 16:50:28 2009 +0000 +++ b/src/Tools/Code/code_thingol.ML Wed Dec 23 08:31:33 2009 +0100 @@ -90,7 +90,6 @@ val canonize_thms: theory -> thm list -> thm list val read_const_exprs: theory -> string list -> string list * string list val consts_program: theory -> string list -> string list * (naming * program) - val cached_program: theory -> naming * program val eval_conv: theory -> (naming -> program -> ((string * sort) list * typscheme) * iterm -> string list -> cterm -> thm) -> cterm -> thm @@ -843,22 +842,12 @@ (* store *) -structure Program = Code_Data_Fun +structure Program = Code_Data ( type T = naming * program; val empty = (empty_naming, Graph.empty); - fun purge thy cs (naming, program) = - let - val names_delete = cs - |> map_filter (lookup_const naming) - |> filter (can (Graph.get_node program)) - |> Graph.all_preds program; - val program' = Graph.del_nodes names_delete program; - in (naming, program') end; ); -val cached_program = Program.get; - fun invoke_generation thy (algebra, eqngr) f name = Program.change_yield thy (fn naming_program => (NONE, naming_program) |> f thy algebra eqngr name @@ -943,10 +932,10 @@ fun code_depgr thy consts = let val (_, eqngr) = Code_Preproc.obtain thy consts []; - val select = Graph.all_succs eqngr consts; + val all_consts = Graph.all_succs eqngr consts; in eqngr - |> not (null consts) ? Graph.subgraph (member (op =) select) + |> Graph.subgraph (member (op =) all_consts) |> Graph.map_nodes ((apsnd o map o apfst) (AxClass.overload thy)) end; @@ -983,13 +972,13 @@ val _ = OuterSyntax.improper_command "code_thms" "print system of code equations for code" OuterKeyword.diag - (Scan.repeat P.term_group + (Scan.repeat1 P.term_group >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of))); val _ = OuterSyntax.improper_command "code_deps" "visualize dependencies of code equations for code" OuterKeyword.diag - (Scan.repeat P.term_group + (Scan.repeat1 P.term_group >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of))); diff -r 3a7937841585 -r 156a060d5d68 src/Tools/nbe.ML --- a/src/Tools/nbe.ML Mon Dec 21 16:50:28 2009 +0000 +++ b/src/Tools/nbe.ML Wed Dec 23 08:31:33 2009 +0100 @@ -505,18 +505,10 @@ (* function store *) -structure Nbe_Functions = Code_Data_Fun +structure Nbe_Functions = Code_Data ( type T = Code_Thingol.naming * ((Univ option * int) Graph.T * (int * string Inttab.table)); val empty = (Code_Thingol.empty_naming, (Graph.empty, (0, Inttab.empty))); - fun purge thy cs (naming, (gr, (maxidx, idx_tab))) = - let - val names_delete = cs - |> map_filter (Code_Thingol.lookup_const naming) - |> filter (can (Graph.get_node gr)) - |> Graph.all_preds gr; - val gr' = Graph.del_nodes names_delete gr; - in (naming, (gr', (maxidx, idx_tab))) end; ); (* compilation, evaluation and reification *)