--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Codegen/Adaptation.thy Mon Aug 27 23:00:38 2012 +0200
@@ -0,0 +1,347 @@
+theory Adaptation
+imports Setup
+begin
+
+setup %invisible {* Code_Target.extend_target ("\<SML>", ("SML", K I))
+ #> Code_Target.extend_target ("\<SMLdummy>", ("Haskell", K I)) *}
+
+section {* Adaptation to target languages \label{sec:adaptation} *}
+
+subsection {* Adapting code generation *}
+
+text {*
+ The aspects of code generation introduced so far have two aspects
+ in common:
+
+ \begin{itemize}
+
+ \item They act uniformly, without reference to a specific target
+ language.
+
+ \item They are \emph{safe} in the sense that as long as you trust
+ the code generator meta theory and implementation, you cannot
+ produce programs that yield results which are not derivable in
+ the logic.
+
+ \end{itemize}
+
+ \noindent In this section we will introduce means to \emph{adapt}
+ the serialiser to a specific target language, i.e.~to print program
+ fragments in a way which accommodates \qt{already existing}
+ ingredients of a target language environment, for three reasons:
+
+ \begin{itemize}
+ \item improving readability and aesthetics of generated code
+ \item gaining efficiency
+ \item interface with language parts which have no direct counterpart
+ in @{text "HOL"} (say, imperative data structures)
+ \end{itemize}
+
+ \noindent Generally, you should avoid using those features yourself
+ \emph{at any cost}:
+
+ \begin{itemize}
+
+ \item The safe configuration methods act uniformly on every target
+ language, whereas for adaptation you have to treat each target
+ language separately.
+
+ \item Application is extremely tedious since there is no
+ abstraction which would allow for a static check, making it easy
+ to produce garbage.
+
+ \item Subtle errors can be introduced unconsciously.
+
+ \end{itemize}
+
+ \noindent However, even if you ought refrain from setting up
+ adaptation yourself, already the @{text "HOL"} comes with some
+ reasonable default adaptations (say, using target language list
+ syntax). There also some common adaptation cases which you can
+ setup by importing particular library theories. In order to
+ understand these, we provide some clues here; these however are not
+ supposed to replace a careful study of the sources.
+*}
+
+
+subsection {* The adaptation principle *}
+
+text {*
+ Figure \ref{fig:adaptation} illustrates what \qt{adaptation} is
+ conceptually supposed to be:
+
+ \begin{figure}[here]
+ \includegraphics{adaptation}
+ \caption{The adaptation principle}
+ \label{fig:adaptation}
+ \end{figure}
+
+ \noindent In the tame view, code generation acts as broker between
+ @{text logic}, @{text "intermediate language"} and @{text "target
+ language"} by means of @{text translation} and @{text
+ serialisation}; for the latter, the serialiser has to observe the
+ structure of the @{text language} itself plus some @{text reserved}
+ keywords which have to be avoided for generated code. However, if
+ you consider @{text adaptation} mechanisms, the code generated by
+ the serializer is just the tip of the iceberg:
+
+ \begin{itemize}
+
+ \item @{text serialisation} can be \emph{parametrised} such that
+ logical entities are mapped to target-specific ones
+ (e.g. target-specific list syntax, see also
+ \secref{sec:adaptation_mechanisms})
+
+ \item Such parametrisations can involve references to a
+ target-specific standard @{text library} (e.g. using the @{text
+ Haskell} @{verbatim Maybe} type instead of the @{text HOL}
+ @{type "option"} type); if such are used, the corresponding
+ identifiers (in our example, @{verbatim Maybe}, @{verbatim
+ Nothing} and @{verbatim Just}) also have to be considered @{text
+ reserved}.
+
+ \item Even more, the user can enrich the library of the
+ target-language by providing code snippets (\qt{@{text
+ "includes"}}) which are prepended to any generated code (see
+ \secref{sec:include}); this typically also involves further
+ @{text reserved} identifiers.
+
+ \end{itemize}
+
+ \noindent As figure \ref{fig:adaptation} illustrates, all these
+ adaptation mechanisms have to act consistently; it is at the
+ discretion of the user to take care for this.
+*}
+
+subsection {* Common adaptation patterns *}
+
+text {*
+ The @{theory HOL} @{theory 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 @{text HOL} library; beside being useful in
+ applications, they may serve as a tutorial for customising the code
+ generator setup (see below \secref{sec:adaptation_mechanisms}).
+
+ \begin{description}
+
+ \item[@{text "Code_Integer"}] represents @{text HOL} integers by
+ big integer literals in target languages.
+
+ \item[@{text "Code_Char"}] represents @{text HOL} characters by
+ character literals in target languages.
+
+ \item[@{text "Code_Char_chr"}] like @{text "Code_Char"}, but
+ also offers treatment of character codes; includes @{text
+ "Code_Char"}.
+
+ \item[@{text "Efficient_Nat"}] \label{eff_nat} implements
+ natural numbers by integers, which in general will result in
+ higher efficiency; pattern matching with @{term "0\<Colon>nat"} /
+ @{const "Suc"} is eliminated; includes @{text "Code_Integer"}
+ and @{text "Code_Numeral"}.
+
+ \item[@{theory "Code_Numeral"}] provides an additional datatype
+ @{typ index} which is mapped to target-language built-in
+ integers. Useful for code setups which involve e.g.~indexing
+ of target-language arrays. Part of @{text "HOL-Main"}.
+
+ \item[@{theory "String"}] provides an additional datatype @{typ
+ String.literal} which is isomorphic to strings; @{typ
+ String.literal}s are mapped to target-language strings. Useful
+ for code setups which involve e.g.~printing (error) messages.
+ Part of @{text "HOL-Main"}.
+
+ \end{description}
+
+ \begin{warn}
+ When importing any of those theories which are not part of
+ @{text "HOL-Main"}, 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}
+*}
+
+
+subsection {* Parametrising serialisation \label{sec:adaptation_mechanisms} *}
+
+text {*
+ Consider the following function and its corresponding SML code:
+*}
+
+primrec %quote in_interval :: "nat \<times> nat \<Rightarrow> nat \<Rightarrow> bool" where
+ "in_interval (k, l) n \<longleftrightarrow> k \<le> n \<and> n \<le> l"
+(*<*)
+code_type %invisible bool
+ (SML)
+code_const %invisible True and False and "op \<and>" and Not
+ (SML and and and)
+(*>*)
+text %quotetypewriter {*
+ @{code_stmts in_interval (SML)}
+*}
+
+text {*
+ \noindent Though this is correct code, it is a little bit
+ unsatisfactory: boolean values and operators are materialised as
+ distinguished entities with have nothing to do with the SML-built-in
+ 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 @{verbatim
+ "bool"} would be used. To map the HOL @{typ bool} on SML @{verbatim
+ "bool"}, we may use \qn{custom serialisations}:
+*}
+
+code_type %quotett bool
+ (SML "bool")
+code_const %quotett True and False and "op \<and>"
+ (SML "true" and "false" and "_ andalso _")
+
+text {*
+ \noindent The @{command_def code_type} command takes a type constructor
+ as arguments together with a list of custom serialisations. Each
+ custom serialisation starts with a target language identifier
+ followed by an expression, which during code serialisation is
+ inserted whenever the type constructor would occur. For constants,
+ @{command_def code_const} implements the corresponding mechanism. Each
+ ``@{verbatim "_"}'' in a serialisation expression is treated as a
+ placeholder for the type constructor's (the constant's) arguments.
+*}
+
+text %quotetypewriter {*
+ @{code_stmts in_interval (SML)}
+*}
+
+text {*
+ \noindent This still is not perfect: the parentheses around the
+ \qt{andalso} expression are superfluous. Though the serialiser 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:
+*}
+
+code_const %quotett "op \<and>"
+ (SML infixl 1 "andalso")
+
+text %quotetypewriter {*
+ @{code_stmts in_interval (SML)}
+*}
+
+text {*
+ \noindent The attentive reader may ask how we assert that no
+ generated code will accidentally overwrite. For this reason the
+ serialiser has an internal table of identifiers which have to be
+ avoided to be used for new declarations. Initially, this table
+ typically contains the keywords of the target language. It can be
+ extended manually, thus avoiding accidental overwrites, using the
+ @{command_def "code_reserved"} command:
+*}
+
+code_reserved %quote "\<SMLdummy>" bool true false andalso
+
+text {*
+ \noindent Next, we try to map HOL pairs to SML pairs, using the
+ infix ``@{verbatim "*"}'' type constructor and parentheses:
+*}
+(*<*)
+code_type %invisible prod
+ (SML)
+code_const %invisible Pair
+ (SML)
+(*>*)
+code_type %quotett prod
+ (SML infix 2 "*")
+code_const %quotett Pair
+ (SML "!((_),/ (_))")
+
+text {*
+ \noindent The initial bang ``@{verbatim "!"}'' tells the serialiser
+ never to 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
+ ``@{verbatim "/"}'' (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 serialisations
+ provide; however their usage requires careful thinking in order not
+ to introduce inconsistencies -- or, in other words: custom
+ serialisations are completely axiomatic.
+
+ A further noteworthy detail is that any special character in a
+ custom serialisation may be quoted using ``@{verbatim "'"}''; thus,
+ in ``@{verbatim "fn '_ => _"}'' the first ``@{verbatim "_"}'' is a
+ proper underscore while the second ``@{verbatim "_"}'' is a
+ placeholder.
+*}
+
+
+subsection {* @{text Haskell} serialisation *}
+
+text {*
+ For convenience, the default @{text HOL} setup for @{text Haskell}
+ maps the @{class equal} class to its counterpart in @{text Haskell},
+ giving custom serialisations for the class @{class equal} (by command
+ @{command_def code_class}) and its operation @{const [source] HOL.equal}
+*}
+
+code_class %quotett equal
+ (Haskell "Eq")
+
+code_const %quotett "HOL.equal"
+ (Haskell infixl 4 "==")
+
+text {*
+ \noindent A problem now occurs whenever a type which is an instance
+ of @{class equal} in @{text HOL} is mapped on a @{text
+ Haskell}-built-in type which is also an instance of @{text Haskell}
+ @{text Eq}:
+*}
+
+typedecl %quote bar
+
+instantiation %quote bar :: equal
+begin
+
+definition %quote "HOL.equal (x\<Colon>bar) y \<longleftrightarrow> x = y"
+
+instance %quote by default (simp add: equal_bar_def)
+
+end %quote (*<*)
+
+(*>*) code_type %quotett bar
+ (Haskell "Integer")
+
+text {*
+ \noindent The code generator would produce an additional instance,
+ which of course is rejected by the @{text Haskell} compiler. To
+ suppress this additional instance, use @{command_def "code_instance"}:
+*}
+
+code_instance %quotett bar :: equal
+ (Haskell -)
+
+
+subsection {* Enhancing the target language context \label{sec:include} *}
+
+text {*
+ In rare cases it is necessary to \emph{enrich} the context of a
+ target language; this is accomplished using the @{command_def
+ "code_include"} command:
+*}
+
+code_include %quotett Haskell "Errno"
+{*errno i = error ("Error number: " ++ show i)*}
+
+code_reserved %quotett Haskell Errno
+
+text {*
+ \noindent Such named @{text include}s are then prepended to every
+ generated code. Inspect such code in order to find out how
+ @{command "code_include"} behaves with respect to a particular
+ target language.
+*}
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Codegen/Evaluation.thy Mon Aug 27 23:00:38 2012 +0200
@@ -0,0 +1,287 @@
+theory Evaluation
+imports Setup
+begin
+
+section {* Evaluation \label{sec:evaluation} *}
+
+text {*
+ Recalling \secref{sec:principle}, code generation turns a system of
+ equations into a program with the \emph{same} equational semantics.
+ As a consequence, this program can be used as a \emph{rewrite
+ engine} for terms: rewriting a term @{term "t"} using a program to a
+ term @{term "t'"} yields the theorems @{prop "t \<equiv> t'"}. This
+ application of code generation in the following is referred to as
+ \emph{evaluation}.
+*}
+
+
+subsection {* Evaluation techniques *}
+
+text {*
+ The existing infrastructure provides a rich palette of evaluation
+ techniques, each comprising different aspects:
+
+ \begin{description}
+
+ \item[Expressiveness.] Depending on how good symbolic computation
+ is supported, the class of terms which can be evaluated may be
+ bigger or smaller.
+
+ \item[Efficiency.] The more machine-near the technique, the
+ faster it is.
+
+ \item[Trustability.] Techniques which a huge (and also probably
+ more configurable infrastructure) are more fragile and less
+ trustable.
+
+ \end{description}
+*}
+
+
+subsubsection {* The simplifier (@{text simp}) *}
+
+text {*
+ The simplest way for evaluation is just using the simplifier with
+ the original code equations of the underlying program. This gives
+ fully symbolic evaluation and highest trustablity, with the usual
+ performance of the simplifier. Note that for operations on abstract
+ datatypes (cf.~\secref{sec:invariant}), the original theorems as
+ given by the users are used, not the modified ones.
+*}
+
+
+subsubsection {* Normalization by evaluation (@{text nbe}) *}
+
+text {*
+ Normalization by evaluation \cite{Aehlig-Haftmann-Nipkow:2008:nbe}
+ provides a comparably fast partially symbolic evaluation which
+ permits also normalization of functions and uninterpreted symbols;
+ the stack of code to be trusted is considerable.
+*}
+
+
+subsubsection {* Evaluation in ML (@{text code}) *}
+
+text {*
+ Highest performance can be achieved by evaluation in ML, at the cost
+ of being restricted to ground results and a layered stack of code to
+ be trusted, including code generator configurations by the user.
+
+ Evaluation is carried out in a target language \emph{Eval} which
+ inherits from \emph{SML} but for convenience uses parts of the
+ Isabelle runtime environment. The soundness of computation carried
+ out there depends crucially on the correctness of the code
+ generator setup; this is one of the reasons why you should not use
+ adaptation (see \secref{sec:adaptation}) frivolously.
+*}
+
+
+subsection {* Aspects of evaluation *}
+
+text {*
+ Each of the techniques can be combined with different aspects. The
+ most important distinction is between dynamic and static evaluation.
+ Dynamic evaluation takes the code generator configuration \qt{as it
+ is} at the point where evaluation is issued. Best example is the
+ @{command_def value} command which allows ad-hoc evaluation of
+ terms:
+*}
+
+value %quote "42 / (12 :: rat)"
+
+text {*
+ \noindent By default @{command value} tries all available evaluation
+ techniques and prints the result of the first succeeding one. A particular
+ technique may be specified in square brackets, e.g.
+*}
+
+value %quote [nbe] "42 / (12 :: rat)"
+
+text {*
+ To employ dynamic evaluation in the document generation, there is also
+ a @{text value} antiquotation. By default, it also tries all available evaluation
+ techniques and prints the result of the first succeeding one, unless a particular
+ technique is specified in square brackets.
+
+ Static evaluation freezes the code generator configuration at a
+ certain point and uses this context whenever evaluation is issued
+ later on. This is particularly appropriate for proof procedures
+ which use evaluation, since then the behaviour of evaluation is not
+ changed or even compromised later on by actions of the user.
+
+ As a technical complication, terms after evaluation in ML must be
+ turned into Isabelle's internal term representation again. Since
+ this is also configurable, it is never fully trusted. For this
+ reason, evaluation in ML comes with further aspects:
+
+ \begin{description}
+
+ \item[Plain evaluation.] A term is normalized using the provided
+ term reconstruction from ML to Isabelle; for applications which
+ do not need to be fully trusted.
+
+ \item[Property conversion.] Evaluates propositions; since these
+ are monomorphic, the term reconstruction is fixed once and for all
+ and therefore trustable.
+
+ \item[Conversion.] Evaluates an arbitrary term @{term "t"} first
+ by plain evaluation and certifies the result @{term "t'"} by
+ checking the equation @{term "t \<equiv> t'"} using property
+ conversion.
+
+ \end{description}
+
+ \noindent The picture is further complicated by the roles of
+ exceptions. Here three cases have to be distinguished:
+
+ \begin{itemize}
+
+ \item Evaluation of @{term t} terminates with a result @{term
+ "t'"}.
+
+ \item Evaluation of @{term t} terminates which en exception
+ indicating a pattern match failure or a non-implemented
+ function. As sketched in \secref{sec:partiality}, this can be
+ interpreted as partiality.
+
+ \item Evaluation raises any other kind of exception.
+
+ \end{itemize}
+
+ \noindent For conversions, the first case yields the equation @{term
+ "t = t'"}, the second defaults to reflexivity @{term "t = t"}.
+ Exceptions of the third kind are propagated to the user.
+
+ By default return values of plain evaluation are optional, yielding
+ @{text "SOME t'"} in the first case, @{text "NONE"} in the
+ second, and propagating the exception in the third case. A strict
+ variant of plain evaluation either yields @{text "t'"} or propagates
+ any exception, a liberal variant caputures any exception in a result
+ of type @{text "Exn.result"}.
+
+ For property conversion (which coincides with conversion except for
+ evaluation in ML), methods are provided which solve a given goal by
+ evaluation.
+*}
+
+
+subsection {* Schematic overview *}
+
+text {*
+ \newcommand{\ttsize}{\fontsize{5.8pt}{8pt}\selectfont}
+ \fontsize{9pt}{12pt}\selectfont
+ \begin{tabular}{ll||c|c|c}
+ & & @{text simp} & @{text nbe} & @{text code} \tabularnewline \hline \hline
+ \multirow{5}{1ex}{\rotatebox{90}{dynamic}}
+ & interactive evaluation
+ & @{command value} @{text "[simp]"} & @{command value} @{text "[nbe]"} & @{command value} @{text "[code]"}
+ \tabularnewline
+ & plain evaluation & & & \ttsize@{ML "Code_Evaluation.dynamic_value"} \tabularnewline \cline{2-5}
+ & evaluation method & @{method code_simp} & @{method normalization} & @{method eval} \tabularnewline
+ & property conversion & & & \ttsize@{ML "Code_Runtime.dynamic_holds_conv"} \tabularnewline \cline{2-5}
+ & conversion & \ttsize@{ML "Code_Simp.dynamic_conv"} & \ttsize@{ML "Nbe.dynamic_conv"}
+ & \ttsize@{ML "Code_Evaluation.dynamic_conv"} \tabularnewline \hline \hline
+ \multirow{3}{1ex}{\rotatebox{90}{static}}
+ & plain evaluation & & & \ttsize@{ML "Code_Evaluation.static_value"} \tabularnewline \cline{2-5}
+ & property conversion & &
+ & \ttsize@{ML "Code_Runtime.static_holds_conv"} \tabularnewline \cline{2-5}
+ & conversion & \ttsize@{ML "Code_Simp.static_conv"}
+ & \ttsize@{ML "Nbe.static_conv"}
+ & \ttsize@{ML "Code_Evaluation.static_conv"}
+ \end{tabular}
+*}
+
+
+subsection {* Intimate connection between logic and system runtime *}
+
+text {*
+ The toolbox of static evaluation conversions forms a reasonable base
+ to interweave generated code and system tools. However in some
+ situations more direct interaction is desirable.
+*}
+
+
+subsubsection {* Static embedding of generated code into system runtime -- the @{text code} antiquotation *}
+
+text {*
+ The @{text code} antiquotation allows to include constants from
+ generated code directly into ML system code, as in the following toy
+ example:
+*}
+
+datatype %quote form = T | F | And form form | Or form form (*<*)
+
+(*>*) ML %quotett {*
+ fun eval_form @{code T} = true
+ | eval_form @{code F} = false
+ | eval_form (@{code And} (p, q)) =
+ eval_form p andalso eval_form q
+ | eval_form (@{code Or} (p, q)) =
+ eval_form p orelse eval_form q;
+*}
+
+text {*
+ \noindent @{text code} takes as argument the name of a constant;
+ after the whole ML is read, the necessary code is generated
+ transparently and the corresponding constant names are inserted.
+ This technique also allows to use pattern matching on constructors
+ stemming from compiled datatypes. Note that the @{text code}
+ antiquotation may not refer to constants which carry adaptations;
+ here you have to refer to the corresponding adapted code directly.
+
+ For a less simplistic example, theory @{text Approximation} in
+ the @{text Decision_Procs} session is a good reference.
+*}
+
+
+subsubsection {* Static embedding of generated code into system runtime -- @{text code_reflect} *}
+
+text {*
+ The @{text code} antiquoation is lightweight, but the generated code
+ is only accessible while the ML section is processed. Sometimes this
+ is not appropriate, especially if the generated code contains datatype
+ declarations which are shared with other parts of the system. In these
+ cases, @{command_def code_reflect} can be used:
+*}
+
+code_reflect %quote Sum_Type
+ datatypes sum = Inl | Inr
+ functions "Sum_Type.Projl" "Sum_Type.Projr"
+
+text {*
+ \noindent @{command_def code_reflect} takes a structure name and
+ references to datatypes and functions; for these code is compiled
+ into the named ML structure and the \emph{Eval} target is modified
+ in a way that future code generation will reference these
+ precompiled versions of the given datatypes and functions. This
+ also allows to refer to the referenced datatypes and functions from
+ arbitrary ML code as well.
+
+ A typical example for @{command code_reflect} can be found in the
+ @{theory Predicate} theory.
+*}
+
+
+subsubsection {* Separate compilation -- @{text code_reflect} *}
+
+text {*
+ For technical reasons it is sometimes necessary to separate
+ generation and compilation of code which is supposed to be used in
+ the system runtime. For this @{command code_reflect} with an
+ optional @{text "file"} argument can be used:
+*}
+
+code_reflect %quote Rat
+ datatypes rat = Frct
+ functions Fract
+ "(plus :: rat \<Rightarrow> rat \<Rightarrow> rat)" "(minus :: rat \<Rightarrow> rat \<Rightarrow> rat)"
+ "(times :: rat \<Rightarrow> rat \<Rightarrow> rat)" "(divide :: rat \<Rightarrow> rat \<Rightarrow> rat)"
+ file "examples/rat.ML"
+
+text {*
+ \noindent This merely generates the referenced code to the given
+ file which can be included into the system runtime later on.
+*}
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Codegen/Foundations.thy Mon Aug 27 23:00:38 2012 +0200
@@ -0,0 +1,347 @@
+theory Foundations
+imports Introduction
+begin
+
+section {* Code generation foundations \label{sec:foundations} *}
+
+subsection {* Code generator architecture \label{sec:architecture} *}
+
+text {*
+ The code generator is actually a framework consisting of different
+ components which can be customised individually.
+
+ Conceptually all components operate on Isabelle's logic framework
+ @{theory Pure}. Practically, the object logic @{theory HOL}
+ provides the necessary facilities to make use of the code generator,
+ mainly since it is an extension of @{theory Pure}.
+
+ The constellation of the different components is visualized in the
+ following picture.
+
+ \begin{figure}[h]
+ \includegraphics{architecture}
+ \caption{Code generator architecture}
+ \label{fig:arch}
+ \end{figure}
+
+ \noindent Central to code generation is the notion of \emph{code
+ equations}. A code equation as a first approximation is a theorem
+ of the form @{text "f t\<^isub>1 t\<^isub>2 \<dots> t\<^isub>n \<equiv> t"} (an equation headed by a
+ constant @{text f} with arguments @{text "t\<^isub>1 t\<^isub>2 \<dots> t\<^isub>n"} and right
+ hand side @{text t}).
+
+ \begin{itemize}
+
+ \item Starting point of code generation is a collection of (raw)
+ code equations in a theory. It is not relevant where they stem
+ from, but typically they were either produced by specification
+ tools or proved explicitly by the user.
+
+ \item These raw code equations can be subjected to theorem
+ transformations. This \qn{preprocessor} (see
+ \secref{sec:preproc}) can apply the full expressiveness of
+ ML-based theorem transformations to code generation. The result
+ of preprocessing is a structured collection of code equations.
+
+ \item These code equations are \qn{translated} to a program in an
+ abstract intermediate language. Think of it as a kind of
+ \qt{Mini-Haskell} with four \qn{statements}: @{text data} (for
+ datatypes), @{text fun} (stemming from code equations), also
+ @{text class} and @{text inst} (for type classes).
+
+ \item Finally, the abstract program is \qn{serialised} into
+ concrete source code of a target language. This step only
+ produces concrete syntax but does not change the program in
+ essence; all conceptual transformations occur in the translation
+ step.
+
+ \end{itemize}
+
+ \noindent From these steps, only the last two 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.
+*}
+
+
+subsection {* The preprocessor \label{sec:preproc} *}
+
+text {*
+ Before selected function theorems are turned into abstract code, a
+ chain of definitional transformation steps is carried out:
+ \emph{preprocessing}. The preprocessor consists of two
+ components: a \emph{simpset} and \emph{function transformers}.
+
+ The \emph{simpset} can apply the full generality of the Isabelle
+ simplifier. Due to the interpretation of theorems as code
+ 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{unfold theorems}, which may be declared and removed using the
+ @{attribute code_unfold} or \emph{@{attribute code_unfold} del}
+ attribute, respectively.
+
+ Some common applications:
+*}
+
+text_raw {*
+ \begin{itemize}
+*}
+
+text {*
+ \item replacing non-executable constructs by executable ones:
+*}
+
+lemma %quote [code_unfold]:
+ "x \<in> set xs \<longleftrightarrow> List.member xs x" by (fact in_set_member)
+
+text {*
+ \item replacing executable but inconvenient constructs:
+*}
+
+lemma %quote [code_unfold]:
+ "xs = [] \<longleftrightarrow> List.null xs" by (fact eq_Nil_null)
+
+text {*
+ \item eliminating disturbing expressions:
+*}
+
+lemma %quote [code_unfold]:
+ "1 = Suc 0" by (fact One_nat_def)
+
+text_raw {*
+ \end{itemize}
+*}
+
+text {*
+ \noindent \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 @{term "0\<Colon>nat"} / @{const Suc} pattern
+ elimination implemented in theory @{text Efficient_Nat} (see
+ \secref{eff_nat}) uses this interface.
+
+ \noindent The current setup of the preprocessor may be inspected
+ using the @{command_def print_codeproc} command. @{command_def
+ code_thms} (see \secref{sec:equations}) provides a convenient
+ mechanism to inspect the impact of a preprocessor setup on code
+ equations.
+*}
+
+
+subsection {* Understanding code equations \label{sec:equations} *}
+
+text {*
+ As told in \secref{sec:principle}, the notion of code equations is
+ vital to code generation. Indeed most problems which occur in
+ practice can be resolved by an inspection of the underlying code
+ equations.
+
+ It is possible to exchange the default code equations for constants
+ by explicitly proving alternative ones:
+*}
+
+lemma %quote [code]:
+ "dequeue (AQueue xs []) =
+ (if xs = [] then (None, AQueue [] [])
+ else dequeue (AQueue [] (rev xs)))"
+ "dequeue (AQueue xs (y # ys)) =
+ (Some y, AQueue xs ys)"
+ by (cases xs, simp_all) (cases "rev xs", simp_all)
+
+text {*
+ \noindent The annotation @{text "[code]"} is an @{text attribute}
+ which states that the given theorems should be considered as code
+ equations for a @{text fun} statement -- the corresponding constant
+ is determined syntactically. The resulting code:
+*}
+
+text %quotetypewriter {*
+ @{code_stmts dequeue (consts) dequeue (Haskell)}
+*}
+
+text {*
+ \noindent You may note that the equality test @{term "xs = []"} has
+ been replaced by the predicate @{term "List.null xs"}. This is due
+ to the default setup of the \qn{preprocessor}.
+
+ This possibility to select arbitrary code equations is the key
+ technique for program and datatype refinement (see
+ \secref{sec:refinement}).
+
+ Due to the preprocessor, there is the distinction of raw code
+ equations (before preprocessing) and code equations (after
+ preprocessing).
+
+ The first can be listed (among other data) using the @{command_def
+ print_codesetup} command.
+
+ The code equations after preprocessing are already are blueprint of
+ the generated program and can be inspected using the @{command
+ code_thms} command:
+*}
+
+code_thms %quote dequeue
+
+text {*
+ \noindent This prints a table with the code equations for @{const
+ dequeue}, including \emph{all} code equations those equations depend
+ on recursively. These dependencies themselves can be visualized using
+ the @{command_def code_deps} command.
+*}
+
+
+subsection {* Equality *}
+
+text {*
+ Implementation of equality deserves some attention. Here an example
+ function involving polymorphic equality:
+*}
+
+primrec %quote collect_duplicates :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+ "collect_duplicates xs ys [] = xs"
+| "collect_duplicates xs ys (z#zs) = (if z \<in> set xs
+ then if z \<in> set ys
+ then collect_duplicates xs ys zs
+ else collect_duplicates xs (z#ys) zs
+ else collect_duplicates (z#xs) (z#ys) zs)"
+
+text {*
+ \noindent During preprocessing, the membership test is rewritten,
+ resulting in @{const List.member}, which itself performs an explicit
+ equality check, as can be seen in the corresponding @{text SML} code:
+*}
+
+text %quotetypewriter {*
+ @{code_stmts collect_duplicates (SML)}
+*}
+
+text {*
+ \noindent Obviously, polymorphic equality is implemented the Haskell
+ way using a type class. How is this achieved? HOL introduces an
+ explicit class @{class equal} with a corresponding operation @{const
+ HOL.equal} such that @{thm equal [no_vars]}. The preprocessing
+ framework does the rest by propagating the @{class equal} constraints
+ through all dependent code equations. For datatypes, instances of
+ @{class equal} are implicitly derived when possible. For other types,
+ you may instantiate @{text equal} manually like any other type class.
+*}
+
+
+subsection {* Explicit partiality \label{sec:partiality} *}
+
+text {*
+ Partiality usually enters the game by partial patterns, as
+ in the following example, again for amortised queues:
+*}
+
+definition %quote strict_dequeue :: "'a queue \<Rightarrow> 'a \<times> 'a queue" where
+ "strict_dequeue q = (case dequeue q
+ of (Some x, q') \<Rightarrow> (x, q'))"
+
+lemma %quote strict_dequeue_AQueue [code]:
+ "strict_dequeue (AQueue xs (y # ys)) = (y, AQueue xs ys)"
+ "strict_dequeue (AQueue xs []) =
+ (case rev xs of y # ys \<Rightarrow> (y, AQueue [] ys))"
+ by (simp_all add: strict_dequeue_def) (cases xs, simp_all split: list.split)
+
+text {*
+ \noindent In the corresponding code, there is no equation
+ for the pattern @{term "AQueue [] []"}:
+*}
+
+text %quotetypewriter {*
+ @{code_stmts strict_dequeue (consts) strict_dequeue (Haskell)}
+*}
+
+text {*
+ \noindent In some cases it is desirable to have this
+ pseudo-\qt{partiality} more explicitly, e.g.~as follows:
+*}
+
+axiomatization %quote empty_queue :: 'a
+
+definition %quote strict_dequeue' :: "'a queue \<Rightarrow> 'a \<times> 'a queue" where
+ "strict_dequeue' q = (case dequeue q of (Some x, q') \<Rightarrow> (x, q') | _ \<Rightarrow> empty_queue)"
+
+lemma %quote strict_dequeue'_AQueue [code]:
+ "strict_dequeue' (AQueue xs []) = (if xs = [] then empty_queue
+ else strict_dequeue' (AQueue [] (rev xs)))"
+ "strict_dequeue' (AQueue xs (y # ys)) =
+ (y, AQueue xs ys)"
+ by (simp_all add: strict_dequeue'_def split: list.splits)
+
+text {*
+ Observe that on the right hand side of the definition of @{const
+ "strict_dequeue'"}, the unspecified constant @{const empty_queue} occurs.
+
+ Normally, if constants without any code equations occur in a
+ program, the code generator complains (since in most cases this is
+ indeed an error). But such constants can also be thought
+ of as function definitions which always fail,
+ since there is never a successful pattern match on the left hand
+ side. In order to categorise a constant into that category
+ explicitly, use @{command_def "code_abort"}:
+*}
+
+code_abort %quote empty_queue
+
+text {*
+ \noindent Then the code generator will just insert an error or
+ exception at the appropriate position:
+*}
+
+text %quotetypewriter {*
+ @{code_stmts strict_dequeue' (consts) empty_queue strict_dequeue' (Haskell)}
+*}
+
+text {*
+ \noindent This feature however is rarely needed in practice. Note
+ also that the HOL default setup already declares @{const undefined}
+ as @{command "code_abort"}, which is most likely to be used in such
+ situations.
+*}
+
+
+subsection {* If something goes utterly wrong \label{sec:utterly_wrong} *}
+
+text {*
+ Under certain circumstances, the code generator fails to produce
+ code entirely. To debug these, the following hints may prove
+ helpful:
+
+ \begin{description}
+
+ \ditem{\emph{Check with a different target language}.} Sometimes
+ the situation gets more clear if you switch to another target
+ language; the code generated there might give some hints what
+ prevents the code generator to produce code for the desired
+ language.
+
+ \ditem{\emph{Inspect code equations}.} Code equations are the central
+ carrier of code generation. Most problems occurring while generating
+ code can be traced to single equations which are printed as part of
+ the error message. A closer inspection of those may offer the key
+ for solving issues (cf.~\secref{sec:equations}).
+
+ \ditem{\emph{Inspect preprocessor setup}.} The preprocessor might
+ transform code equations unexpectedly; to understand an
+ inspection of its setup is necessary (cf.~\secref{sec:preproc}).
+
+ \ditem{\emph{Generate exceptions}.} If the code generator
+ complains about missing code equations, in can be helpful to
+ implement the offending constants as exceptions
+ (cf.~\secref{sec:partiality}); this allows at least for a formal
+ generation of code, whose inspection may then give clues what is
+ wrong.
+
+ \ditem{\emph{Remove offending code equations}.} If code
+ generation is prevented by just a single equation, this can be
+ removed (cf.~\secref{sec:equations}) to allow formal code
+ generation, whose result in turn can be used to trace the
+ problem. The most prominent case here are mismatches in type
+ class signatures (\qt{wellsortedness error}).
+
+ \end{description}
+*}
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Codegen/Further.thy Mon Aug 27 23:00:38 2012 +0200
@@ -0,0 +1,351 @@
+theory Further
+imports Setup
+begin
+
+section {* Further issues \label{sec:further} *}
+
+subsection {* Specialities of the @{text Scala} target language \label{sec:scala} *}
+
+text {*
+ @{text Scala} deviates from languages of the ML family in a couple
+ of aspects; those which affect code generation mainly have to do with
+ @{text Scala}'s type system:
+
+ \begin{itemize}
+
+ \item @{text Scala} prefers tupled syntax over curried syntax.
+
+ \item @{text Scala} sacrifices Hindely-Milner type inference for a
+ much more rich type system with subtyping etc. For this reason
+ type arguments sometimes have to be given explicitly in square
+ brackets (mimicking System F syntax).
+
+ \item In contrast to @{text Haskell} where most specialities of
+ the type system are implemented using \emph{type classes},
+ @{text Scala} provides a sophisticated system of \emph{implicit
+ arguments}.
+
+ \end{itemize}
+
+ \noindent Concerning currying, the @{text Scala} serializer counts
+ arguments in code equations to determine how many arguments
+ shall be tupled; remaining arguments and abstractions in terms
+ rather than function definitions are always curried.
+
+ The second aspect affects user-defined adaptations with @{command
+ code_const}. For regular terms, the @{text Scala} serializer prints
+ all type arguments explicitly. For user-defined term adaptations
+ this is only possible for adaptations which take no arguments: here
+ the type arguments are just appended. Otherwise they are ignored;
+ hence user-defined adaptations for polymorphic constants have to be
+ designed very carefully to avoid ambiguity.
+
+ Isabelle's type classes are mapped onto @{text Scala} implicits; in
+ cases with diamonds in the subclass hierarchy this can lead to
+ ambiguities in the generated code:
+*}
+
+class %quote class1 =
+ fixes foo :: "'a \<Rightarrow> 'a"
+
+class %quote class2 = class1
+
+class %quote class3 = class1
+
+text {*
+ \noindent Here both @{class class2} and @{class class3} inherit from @{class class1},
+ forming the upper part of a diamond.
+*}
+
+definition %quote bar :: "'a :: {class2, class3} \<Rightarrow> 'a" where
+ "bar = foo"
+
+text {*
+ \noindent This yields the following code:
+*}
+
+text %quotetypewriter {*
+ @{code_stmts bar (Scala)}
+*}
+
+text {*
+ \noindent This code is rejected by the @{text Scala} compiler: in
+ the definition of @{text bar}, it is not clear from where to derive
+ the implicit argument for @{text foo}.
+
+ The solution to the problem is to close the diamond by a further
+ class with inherits from both @{class class2} and @{class class3}:
+*}
+
+class %quote class4 = class2 + class3
+
+text {*
+ \noindent Then the offending code equation can be restricted to
+ @{class class4}:
+*}
+
+lemma %quote [code]:
+ "(bar :: 'a::class4 \<Rightarrow> 'a) = foo"
+ by (simp only: bar_def)
+
+text {*
+ \noindent with the following code:
+*}
+
+text %quotetypewriter {*
+ @{code_stmts bar (Scala)}
+*}
+
+text {*
+ \noindent which exposes no ambiguity.
+
+ Since the preprocessor (cf.~\secref{sec:preproc}) propagates sort
+ constraints through a system of code equations, it is usually not
+ very difficult to identify the set of code equations which actually
+ needs more restricted sort constraints.
+*}
+
+subsection {* Modules namespace *}
+
+text {*
+ When invoking the @{command export_code} command it is possible to
+ leave out the @{keyword "module_name"} part; then code is
+ distributed over different modules, where the module name space
+ roughly is induced by the Isabelle theory name space.
+
+ Then sometimes the awkward situation occurs that dependencies
+ between definitions introduce cyclic dependencies between modules,
+ which in the @{text Haskell} world leaves you to the mercy of the
+ @{text Haskell} implementation you are using, while for @{text
+ SML}/@{text OCaml} 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
+*}
+
+code_modulename %quote SML
+ A ABC
+ B ABC
+ C ABC
+
+text {*
+ \noindent we explicitly map all those modules on \emph{ABC},
+ resulting in an ad-hoc merge of this three modules at serialisation
+ time.
+*}
+
+subsection {* Locales and interpretation *}
+
+text {*
+ A technical issue comes to surface when generating code from
+ specifications stemming from locale interpretation.
+
+ Let us assume a locale specifying a power operation on arbitrary
+ types:
+*}
+
+locale %quote power =
+ fixes power :: "'a \<Rightarrow> 'b \<Rightarrow> 'b"
+ assumes power_commute: "power x \<circ> power y = power y \<circ> power x"
+begin
+
+text {*
+ \noindent Inside that locale we can lift @{text power} to exponent
+ lists by means of specification relative to that locale:
+*}
+
+primrec %quote powers :: "'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
+ "powers [] = id"
+| "powers (x # xs) = power x \<circ> powers xs"
+
+lemma %quote powers_append:
+ "powers (xs @ ys) = powers xs \<circ> powers ys"
+ by (induct xs) simp_all
+
+lemma %quote powers_power:
+ "powers xs \<circ> power x = power x \<circ> powers xs"
+ by (induct xs)
+ (simp_all del: o_apply id_apply add: o_assoc [symmetric],
+ simp del: o_apply add: o_assoc power_commute)
+
+lemma %quote powers_rev:
+ "powers (rev xs) = powers xs"
+ by (induct xs) (simp_all add: powers_append powers_power)
+
+end %quote
+
+text {*
+ After an interpretation of this locale (say, @{command_def
+ interpretation} @{text "fun_power:"} @{term [source] "power (\<lambda>n (f
+ :: 'a \<Rightarrow> 'a). f ^^ n)"}), one would expect to have a constant @{text
+ "fun_power.powers :: nat list \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"} for which code
+ can be generated. But this not the case: internally, the term
+ @{text "fun_power.powers"} is an abbreviation for the foundational
+ term @{term [source] "power.powers (\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n)"}
+ (see \cite{isabelle-locale} for the details behind).
+
+ Fortunately, with minor effort the desired behaviour can be
+ achieved. First, a dedicated definition of the constant on which
+ the local @{text "powers"} after interpretation is supposed to be
+ mapped on:
+*}
+
+definition %quote funpows :: "nat list \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
+ [code del]: "funpows = power.powers (\<lambda>n f. f ^^ n)"
+
+text {*
+ \noindent In general, the pattern is @{text "c = t"} where @{text c}
+ is the name of the future constant and @{text t} the foundational
+ term corresponding to the local constant after interpretation.
+
+ The interpretation itself is enriched with an equation @{text "t = c"}:
+*}
+
+interpretation %quote fun_power: power "\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n" where
+ "power.powers (\<lambda>n f. f ^^ n) = funpows"
+ by unfold_locales
+ (simp_all add: fun_eq_iff funpow_mult mult_commute funpows_def)
+
+text {*
+ \noindent This additional equation is trivially proved by the
+ definition itself.
+
+ After this setup procedure, code generation can continue as usual:
+*}
+
+text %quotetypewriter {*
+ @{code_stmts funpows (consts) Nat.funpow funpows (Haskell)}
+*}
+
+
+subsection {* Imperative data structures *}
+
+text {*
+ If you consider imperative data structures as inevitable for a
+ specific application, you should consider \emph{Imperative
+ Functional Programming with Isabelle/HOL}
+ \cite{bulwahn-et-al:2008:imperative}; the framework described there
+ is available in session @{text Imperative_HOL}, together with a
+ short primer document.
+*}
+
+
+subsection {* ML system interfaces \label{sec:ml} *}
+
+text {*
+ 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 fundamental ML
+ interfaces.
+*}
+
+subsubsection {* Managing executable content *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML Code.read_const: "theory -> string -> string"} \\
+ @{index_ML Code.add_eqn: "thm -> theory -> theory"} \\
+ @{index_ML Code.del_eqn: "thm -> theory -> theory"} \\
+ @{index_ML Code_Preproc.map_pre: "(simpset -> simpset) -> theory -> theory"} \\
+ @{index_ML Code_Preproc.map_post: "(simpset -> simpset) -> theory -> theory"} \\
+ @{index_ML Code_Preproc.add_functrans: "
+ string * (theory -> (thm * bool) list -> (thm * bool) list option)
+ -> theory -> theory"} \\
+ @{index_ML Code_Preproc.del_functrans: "string -> theory -> theory"} \\
+ @{index_ML Code.add_datatype: "(string * typ) list -> theory -> theory"} \\
+ @{index_ML Code.get_type: "theory -> string
+ -> ((string * sort) list * (string * ((string * sort) list * typ list)) list) * bool"} \\
+ @{index_ML Code.get_type_of_constr_or_abstr: "theory -> string -> (string * bool) option"}
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML Code.read_const}~@{text thy}~@{text s}
+ reads a constant as a concrete term expression @{text s}.
+
+ \item @{ML Code.add_eqn}~@{text "thm"}~@{text "thy"} adds function
+ theorem @{text "thm"} to executable content.
+
+ \item @{ML Code.del_eqn}~@{text "thm"}~@{text "thy"} removes function
+ theorem @{text "thm"} from executable content, if present.
+
+ \item @{ML Code_Preproc.map_pre}~@{text "f"}~@{text "thy"} changes
+ the preprocessor simpset.
+
+ \item @{ML Code_Preproc.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds
+ function transformer @{text f} (named @{text name}) to executable content;
+ @{text f} is a transformer of the code equations belonging
+ to a certain function definition, depending on the
+ current theory context. Returning @{text NONE} indicates that no
+ transformation took place; otherwise, the whole process will be iterated
+ with the new code equations.
+
+ \item @{ML Code_Preproc.del_functrans}~@{text "name"}~@{text "thy"} removes
+ function transformer named @{text name} from executable content.
+
+ \item @{ML Code.add_datatype}~@{text cs}~@{text thy} adds
+ a datatype to executable content, with generation
+ set @{text cs}.
+
+ \item @{ML Code.get_type_of_constr_or_abstr}~@{text "thy"}~@{text "const"}
+ returns type constructor corresponding to
+ constructor @{text const}; returns @{text NONE}
+ if @{text const} is no constructor.
+
+ \end{description}
+*}
+
+
+subsubsection {* Data depending on the theory's executable content *}
+
+text {*
+ 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.
+
+ 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 @{ML_functor Code_Data}; on instantiation of this functor, the
+ following types and operations are required:
+
+ \medskip
+ \begin{tabular}{l}
+ @{text "type T"} \\
+ @{text "val empty: T"} \\
+ \end{tabular}
+
+ \begin{description}
+
+ \item @{text T} the type of data to store.
+
+ \item @{text empty} initial (empty) data.
+
+ \end{description}
+
+ \noindent An instance of @{ML_functor Code_Data} provides the
+ following interface:
+
+ \medskip
+ \begin{tabular}{l}
+ @{text "change: theory \<rightarrow> (T \<rightarrow> T) \<rightarrow> T"} \\
+ @{text "change_yield: theory \<rightarrow> (T \<rightarrow> 'a * T) \<rightarrow> 'a * T"}
+ \end{tabular}
+
+ \begin{description}
+
+ \item @{text change} update of current data (cached!) by giving a
+ continuation.
+
+ \item @{text change_yield} update with side result.
+
+ \end{description}
+*}
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Codegen/Inductive_Predicate.thy Mon Aug 27 23:00:38 2012 +0200
@@ -0,0 +1,275 @@
+theory Inductive_Predicate
+imports Setup
+begin
+
+(*<*)
+hide_const %invisible append
+
+inductive %invisible append where
+ "append [] ys ys"
+| "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"
+
+lemma %invisible append: "append xs ys zs = (xs @ ys = zs)"
+ by (induct xs arbitrary: ys zs) (auto elim: append.cases intro: append.intros)
+
+lemmas lexordp_def =
+ lexordp_def [unfolded lexord_def mem_Collect_eq split]
+(*>*)
+
+section {* Inductive Predicates \label{sec:inductive} *}
+
+text {*
+ The @{text "predicate compiler"} is an extension of the code generator
+ which turns inductive specifications into equational ones, from
+ which in turn executable code can be generated. The mechanisms of
+ this compiler are described in detail in
+ \cite{Berghofer-Bulwahn-Haftmann:2009:TPHOL}.
+
+ Consider the simple predicate @{const append} given by these two
+ introduction rules:
+*}
+
+text %quote {*
+ @{thm append.intros(1)[of ys]} \\
+ @{thm append.intros(2)[of xs ys zs x]}
+*}
+
+text {*
+ \noindent To invoke the compiler, simply use @{command_def "code_pred"}:
+*}
+
+code_pred %quote append .
+
+text {*
+ \noindent The @{command "code_pred"} command takes the name of the
+ inductive predicate and then you put a period to discharge a trivial
+ correctness proof. The compiler infers possible modes for the
+ predicate and produces the derived code equations. Modes annotate
+ which (parts of the) arguments are to be taken as input, and which
+ output. Modes are similar to types, but use the notation @{text "i"}
+ for input and @{text "o"} for output.
+
+ For @{term "append"}, the compiler can infer the following modes:
+ \begin{itemize}
+ \item @{text "i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool"}
+ \item @{text "i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool"}
+ \item @{text "o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool"}
+ \end{itemize}
+ You can compute sets of predicates using @{command_def "values"}:
+*}
+
+values %quote "{zs. append [(1::nat),2,3] [4,5] zs}"
+
+text {* \noindent outputs @{text "{[1, 2, 3, 4, 5]}"}, and *}
+
+values %quote "{(xs, ys). append xs ys [(2::nat),3]}"
+
+text {* \noindent outputs @{text "{([], [2, 3]), ([2], [3]), ([2, 3], [])}"}. *}
+
+text {*
+ \noindent If you are only interested in the first elements of the
+ set comprehension (with respect to a depth-first search on the
+ introduction rules), you can pass an argument to @{command "values"}
+ to specify the number of elements you want:
+*}
+
+values %quote 1 "{(xs, ys). append xs ys [(1::nat), 2, 3, 4]}"
+values %quote 3 "{(xs, ys). append xs ys [(1::nat), 2, 3, 4]}"
+
+text {*
+ \noindent The @{command "values"} command can only compute set
+ comprehensions for which a mode has been inferred.
+
+ The code equations for a predicate are made available as theorems with
+ the suffix @{text "equation"}, and can be inspected with:
+*}
+
+thm %quote append.equation
+
+text {*
+ \noindent More advanced options are described in the following subsections.
+*}
+
+subsection {* Alternative names for functions *}
+
+text {*
+ By default, the functions generated from a predicate are named after
+ the predicate with the mode mangled into the name (e.g., @{text
+ "append_i_i_o"}). You can specify your own names as follows:
+*}
+
+code_pred %quote (modes: i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool as concat,
+ o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool as split,
+ i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool as suffix) append .
+
+subsection {* Alternative introduction rules *}
+
+text {*
+ Sometimes the introduction rules of an predicate are not executable
+ because they contain non-executable constants or specific modes
+ could not be inferred. It is also possible that the introduction
+ rules yield a function that loops forever due to the execution in a
+ depth-first search manner. Therefore, you can declare alternative
+ introduction rules for predicates with the attribute @{attribute
+ "code_pred_intro"}. For example, the transitive closure is defined
+ by:
+*}
+
+text %quote {*
+ @{lemma [source] "r a b \<Longrightarrow> tranclp r a b" by (fact tranclp.intros(1))}\\
+ @{lemma [source] "tranclp r a b \<Longrightarrow> r b c \<Longrightarrow> tranclp r a c" by (fact tranclp.intros(2))}
+*}
+
+text {*
+ \noindent These rules do not suit well for executing the transitive
+ closure with the mode @{text "(i \<Rightarrow> o \<Rightarrow> bool) \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool"}, as
+ the second rule will cause an infinite loop in the recursive call.
+ This can be avoided using the following alternative rules which are
+ declared to the predicate compiler by the attribute @{attribute
+ "code_pred_intro"}:
+*}
+
+lemma %quote [code_pred_intro]:
+ "r a b \<Longrightarrow> tranclp r a b"
+ "r a b \<Longrightarrow> tranclp r b c \<Longrightarrow> tranclp r a c"
+by auto
+
+text {*
+ \noindent After declaring all alternative rules for the transitive
+ closure, you invoke @{command "code_pred"} as usual. As you have
+ declared alternative rules for the predicate, you are urged to prove
+ that these introduction rules are complete, i.e., that you can
+ derive an elimination rule for the alternative rules:
+*}
+
+code_pred %quote tranclp
+proof -
+ case tranclp
+ from this converse_tranclpE [OF tranclp.prems] show thesis by metis
+qed
+
+text {*
+ \noindent Alternative rules can also be used for constants that have
+ not been defined inductively. For example, the lexicographic order
+ which is defined as:
+*}
+
+text %quote {*
+ @{thm [display] lexordp_def [of r]}
+*}
+
+text {*
+ \noindent To make it executable, you can derive the following two
+ rules and prove the elimination rule:
+*}
+
+lemma %quote [code_pred_intro]:
+ "append xs (a # v) ys \<Longrightarrow> lexordp r xs ys"
+(*<*)unfolding lexordp_def by (auto simp add: append)(*>*)
+
+lemma %quote [code_pred_intro]:
+ "append u (a # v) xs \<Longrightarrow> append u (b # w) ys \<Longrightarrow> r a b
+ \<Longrightarrow> lexordp r xs ys"
+(*<*)unfolding lexordp_def append apply simp
+apply (rule disjI2) by auto(*>*)
+
+code_pred %quote lexordp
+(*<*)proof -
+ fix r xs ys
+ assume lexord: "lexordp r xs ys"
+ assume 1: "\<And>r' xs' ys' a v. r = r' \<Longrightarrow> xs = xs' \<Longrightarrow> ys = ys'
+ \<Longrightarrow> append xs' (a # v) ys' \<Longrightarrow> thesis"
+ assume 2: "\<And>r' xs' ys' u a v b w. r = r' \<Longrightarrow> xs = xs' \<Longrightarrow> ys = ys'
+ \<Longrightarrow> append u (a # v) xs' \<Longrightarrow> append u (b # w) ys' \<Longrightarrow> r' a b \<Longrightarrow> thesis"
+ {
+ assume "\<exists>a v. ys = xs @ a # v"
+ from this 1 have thesis
+ by (fastforce simp add: append)
+ } moreover
+ {
+ assume "\<exists>u a b v w. r a b \<and> xs = u @ a # v \<and> ys = u @ b # w"
+ from this 2 have thesis by (fastforce simp add: append)
+ } moreover
+ note lexord
+ ultimately show thesis
+ unfolding lexordp_def
+ by fastforce
+qed(*>*)
+
+
+subsection {* Options for values *}
+
+text {*
+ In the presence of higher-order predicates, multiple modes for some
+ predicate could be inferred that are not disambiguated by the
+ pattern of the set comprehension. To disambiguate the modes for the
+ arguments of a predicate, you can state the modes explicitly in the
+ @{command "values"} command. Consider the simple predicate @{term
+ "succ"}:
+*}
+
+inductive %quote succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
+ "succ 0 (Suc 0)"
+| "succ x y \<Longrightarrow> succ (Suc x) (Suc y)"
+
+code_pred %quote succ .
+
+text {*
+ \noindent For this, the predicate compiler can infer modes @{text "o
+ \<Rightarrow> o \<Rightarrow> bool"}, @{text "i \<Rightarrow> o \<Rightarrow> bool"}, @{text "o \<Rightarrow> i \<Rightarrow> bool"} and
+ @{text "i \<Rightarrow> i \<Rightarrow> bool"}. The invocation of @{command "values"}
+ @{text "{n. tranclp succ 10 n}"} loops, as multiple modes for the
+ predicate @{text "succ"} are possible and here the first mode @{text
+ "o \<Rightarrow> o \<Rightarrow> bool"} is chosen. To choose another mode for the argument,
+ you can declare the mode for the argument between the @{command
+ "values"} and the number of elements.
+*}
+
+values %quote [mode: i \<Rightarrow> o \<Rightarrow> bool] 1 "{n. tranclp succ 10 n}" (*FIMXE does not terminate for n\<ge>1*)
+values %quote [mode: o \<Rightarrow> i \<Rightarrow> bool] 1 "{n. tranclp succ n 10}"
+
+
+subsection {* Embedding into functional code within Isabelle/HOL *}
+
+text {*
+ To embed the computation of an inductive predicate into functions
+ that are defined in Isabelle/HOL, you have a number of options:
+
+ \begin{itemize}
+
+ \item You want to use the first-order predicate with the mode
+ where all arguments are input. Then you can use the predicate directly, e.g.
+
+ \begin{quote}
+ @{text "valid_suffix ys zs = "} \\
+ @{text "(if append [Suc 0, 2] ys zs then Some ys else None)"}
+ \end{quote}
+
+ \item If you know that the execution returns only one value (it is
+ deterministic), then you can use the combinator @{term
+ "Predicate.the"}, e.g., a functional concatenation of lists is
+ defined with
+
+ \begin{quote}
+ @{term "functional_concat xs ys = Predicate.the (append_i_i_o xs ys)"}
+ \end{quote}
+
+ Note that if the evaluation does not return a unique value, it
+ raises a run-time error @{term "not_unique"}.
+
+ \end{itemize}
+*}
+
+
+subsection {* Further Examples *}
+
+text {*
+ Further examples for compiling inductive predicates can be found in
+ the @{text "HOL/ex/Predicate_Compile_ex.thy"} theory file. There are
+ also some examples in the Archive of Formal Proofs, notably in the
+ @{text "POPLmark-deBruijn"} and the @{text "FeatherweightJava"}
+ sessions.
+*}
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Codegen/Introduction.thy Mon Aug 27 23:00:38 2012 +0200
@@ -0,0 +1,242 @@
+theory Introduction
+imports Setup
+begin
+
+section {* Introduction *}
+
+text {*
+ This tutorial introduces the code generator facilities of @{text
+ "Isabelle/HOL"}. It allows to turn (a certain class of) HOL
+ specifications into corresponding executable code in the programming
+ languages @{text SML} \cite{SML}, @{text OCaml} \cite{OCaml},
+ @{text Haskell} \cite{haskell-revised-report} and @{text Scala}
+ \cite{scala-overview-tech-report}.
+
+ To profit from this tutorial, some familiarity and experience with
+ @{theory HOL} \cite{isa-tutorial} and its basic theories is assumed.
+*}
+
+
+subsection {* Code generation principle: shallow embedding \label{sec:principle} *}
+
+text {*
+ The key concept for understanding Isabelle's code generation is
+ \emph{shallow embedding}: logical entities like constants, types and
+ classes are identified with corresponding entities in the target
+ language. In particular, the carrier of a generated program's
+ semantics are \emph{equational theorems} from the logic. If we view
+ a generated program as an implementation of a higher-order rewrite
+ system, then every rewrite step performed by the program can be
+ simulated in the logic, which guarantees partial correctness
+ \cite{Haftmann-Nipkow:2010:code}.
+*}
+
+
+subsection {* A quick start with the Isabelle/HOL toolbox \label{sec:queue_example} *}
+
+text {*
+ In a HOL theory, the @{command_def datatype} and @{command_def
+ definition}/@{command_def primrec}/@{command_def fun} declarations
+ form the core of a functional programming language. By default
+ equational theorems stemming from those are used for generated code,
+ therefore \qt{naive} code generation can proceed without further
+ ado.
+
+ For example, here a simple \qt{implementation} of amortised queues:
+*}
+
+datatype %quote 'a queue = AQueue "'a list" "'a list"
+
+definition %quote empty :: "'a queue" where
+ "empty = AQueue [] []"
+
+primrec %quote enqueue :: "'a \<Rightarrow> 'a queue \<Rightarrow> 'a queue" where
+ "enqueue x (AQueue xs ys) = AQueue (x # xs) ys"
+
+fun %quote dequeue :: "'a queue \<Rightarrow> 'a option \<times> 'a queue" where
+ "dequeue (AQueue [] []) = (None, AQueue [] [])"
+ | "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)"
+ | "dequeue (AQueue xs []) =
+ (case rev xs of y # ys \<Rightarrow> (Some y, AQueue [] ys))" (*<*)
+
+lemma %invisible dequeue_nonempty_Nil [simp]:
+ "xs \<noteq> [] \<Longrightarrow> dequeue (AQueue xs []) = (case rev xs of y # ys \<Rightarrow> (Some y, AQueue [] ys))"
+ by (cases xs) (simp_all split: list.splits) (*>*)
+
+text {* \noindent Then we can generate code e.g.~for @{text SML} as follows: *}
+
+export_code %quote empty dequeue enqueue in SML
+ module_name Example file "examples/example.ML"
+
+text {* \noindent resulting in the following code: *}
+
+text %quotetypewriter {*
+ @{code_stmts empty enqueue dequeue (SML)}
+*}
+
+text {*
+ \noindent The @{command_def export_code} command takes a
+ space-separated list of constants for which code shall be generated;
+ anything else needed for those is added implicitly. Then follows a
+ target language identifier and a freely chosen module name. A file
+ name denotes the destination to store the generated code. Note that
+ the semantics of the destination depends on the target language: for
+ @{text SML}, @{text OCaml} and @{text Scala} it denotes a \emph{file},
+ for @{text Haskell} it denotes a \emph{directory} where a file named as the
+ module name (with extension @{text ".hs"}) is written:
+*}
+
+export_code %quote empty dequeue enqueue in Haskell
+ module_name Example file "examples/"
+
+text {*
+ \noindent This is the corresponding code:
+*}
+
+text %quotetypewriter {*
+ @{code_stmts empty enqueue dequeue (Haskell)}
+*}
+
+text {*
+ \noindent For more details about @{command export_code} see
+ \secref{sec:further}.
+*}
+
+
+subsection {* Type classes *}
+
+text {*
+ Code can also be generated from type classes in a Haskell-like
+ manner. For illustration here an example from abstract algebra:
+*}
+
+class %quote semigroup =
+ fixes mult :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<otimes>" 70)
+ assumes assoc: "(x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"
+
+class %quote monoid = semigroup +
+ fixes neutral :: 'a ("\<one>")
+ assumes neutl: "\<one> \<otimes> x = x"
+ and neutr: "x \<otimes> \<one> = x"
+
+instantiation %quote nat :: monoid
+begin
+
+primrec %quote mult_nat where
+ "0 \<otimes> n = (0\<Colon>nat)"
+ | "Suc m \<otimes> n = n + m \<otimes> n"
+
+definition %quote neutral_nat where
+ "\<one> = Suc 0"
+
+lemma %quote add_mult_distrib:
+ fixes n m q :: nat
+ shows "(n + m) \<otimes> q = n \<otimes> q + m \<otimes> q"
+ by (induct n) simp_all
+
+instance %quote proof
+ fix m n q :: nat
+ show "m \<otimes> n \<otimes> q = m \<otimes> (n \<otimes> q)"
+ by (induct m) (simp_all add: add_mult_distrib)
+ show "\<one> \<otimes> n = n"
+ by (simp add: neutral_nat_def)
+ show "m \<otimes> \<one> = m"
+ by (induct m) (simp_all add: neutral_nat_def)
+qed
+
+end %quote
+
+text {*
+ \noindent We define the natural operation of the natural numbers
+ on monoids:
+*}
+
+primrec %quote (in monoid) pow :: "nat \<Rightarrow> 'a \<Rightarrow> 'a" where
+ "pow 0 a = \<one>"
+ | "pow (Suc n) a = a \<otimes> pow n a"
+
+text {*
+ \noindent This we use to define the discrete exponentiation
+ function:
+*}
+
+definition %quote bexp :: "nat \<Rightarrow> nat" where
+ "bexp n = pow n (Suc (Suc 0))"
+
+text {*
+ \noindent The corresponding code in Haskell uses that language's
+ native classes:
+*}
+
+text %quotetypewriter {*
+ @{code_stmts bexp (Haskell)}
+*}
+
+text {*
+ \noindent This is a convenient place to show how explicit dictionary
+ construction manifests in generated code -- the same example in
+ @{text SML}:
+*}
+
+text %quotetypewriter {*
+ @{code_stmts bexp (SML)}
+*}
+
+text {*
+ \noindent Note the parameters with trailing underscore (@{verbatim
+ "A_"}), which are the dictionary parameters.
+*}
+
+
+subsection {* How to continue from here *}
+
+text {*
+ What you have seen so far should be already enough in a lot of
+ cases. If you are content with this, you can quit reading here.
+
+ Anyway, to understand situations where problems occur or to increase
+ the scope of code generation beyond default, it is necessary to gain
+ some understanding how the code generator actually works:
+
+ \begin{itemize}
+
+ \item The foundations of the code generator are described in
+ \secref{sec:foundations}.
+
+ \item In particular \secref{sec:utterly_wrong} gives hints how to
+ debug situations where code generation does not succeed as
+ expected.
+
+ \item The scope and quality of generated code can be increased
+ dramatically by applying refinement techniques, which are
+ introduced in \secref{sec:refinement}.
+
+ \item Inductive predicates can be turned executable using an
+ extension of the code generator \secref{sec:inductive}.
+
+ \item If you want to utilize code generation to obtain fast
+ evaluators e.g.~for decision procedures, have a look at
+ \secref{sec:evaluation}.
+
+ \item You may want to skim over the more technical sections
+ \secref{sec:adaptation} and \secref{sec:further}.
+
+ \item The target language Scala \cite{scala-overview-tech-report}
+ comes with some specialities discussed in \secref{sec:scala}.
+
+ \item For exhaustive syntax diagrams etc. you should visit the
+ Isabelle/Isar Reference Manual \cite{isabelle-isar-ref}.
+
+ \end{itemize}
+
+ \bigskip
+
+ \begin{center}\fbox{\fbox{\begin{minipage}{8cm}
+
+ \begin{center}\textit{Happy proving, happy hacking!}\end{center}
+
+ \end{minipage}}}\end{center}
+*}
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Codegen/Refinement.thy Mon Aug 27 23:00:38 2012 +0200
@@ -0,0 +1,274 @@
+theory Refinement
+imports Setup
+begin
+
+section {* Program and datatype refinement \label{sec:refinement} *}
+
+text {*
+ Code generation by shallow embedding (cf.~\secref{sec:principle})
+ allows to choose code equations and datatype constructors freely,
+ given that some very basic syntactic properties are met; this
+ flexibility opens up mechanisms for refinement which allow to extend
+ the scope and quality of generated code dramatically.
+*}
+
+
+subsection {* Program refinement *}
+
+text {*
+ Program refinement works by choosing appropriate code equations
+ explicitly (cf.~\secref{sec:equations}); as example, we use Fibonacci
+ numbers:
+*}
+
+fun %quote fib :: "nat \<Rightarrow> nat" where
+ "fib 0 = 0"
+ | "fib (Suc 0) = Suc 0"
+ | "fib (Suc (Suc n)) = fib n + fib (Suc n)"
+
+text {*
+ \noindent The runtime of the corresponding code grows exponential due
+ to two recursive calls:
+*}
+
+text %quotetypewriter {*
+ @{code_stmts fib (consts) fib (Haskell)}
+*}
+
+text {*
+ \noindent A more efficient implementation would use dynamic
+ programming, e.g.~sharing of common intermediate results between
+ recursive calls. This idea is expressed by an auxiliary operation
+ which computes a Fibonacci number and its successor simultaneously:
+*}
+
+definition %quote fib_step :: "nat \<Rightarrow> nat \<times> nat" where
+ "fib_step n = (fib (Suc n), fib n)"
+
+text {*
+ \noindent This operation can be implemented by recursion using
+ dynamic programming:
+*}
+
+lemma %quote [code]:
+ "fib_step 0 = (Suc 0, 0)"
+ "fib_step (Suc n) = (let (m, q) = fib_step n in (m + q, m))"
+ by (simp_all add: fib_step_def)
+
+text {*
+ \noindent What remains is to implement @{const fib} by @{const
+ fib_step} as follows:
+*}
+
+lemma %quote [code]:
+ "fib 0 = 0"
+ "fib (Suc n) = fst (fib_step n)"
+ by (simp_all add: fib_step_def)
+
+text {*
+ \noindent The resulting code shows only linear growth of runtime:
+*}
+
+text %quotetypewriter {*
+ @{code_stmts fib (consts) fib fib_step (Haskell)}
+*}
+
+
+subsection {* Datatype refinement *}
+
+text {*
+ Selecting specific code equations \emph{and} datatype constructors
+ leads to datatype refinement. As an example, we will develop an
+ alternative representation of the queue example given in
+ \secref{sec:queue_example}. The amortised representation is
+ convenient for generating code but exposes its \qt{implementation}
+ details, which may be cumbersome when proving theorems about it.
+ Therefore, here is a simple, straightforward representation of
+ queues:
+*}
+
+datatype %quote 'a queue = Queue "'a list"
+
+definition %quote empty :: "'a queue" where
+ "empty = Queue []"
+
+primrec %quote enqueue :: "'a \<Rightarrow> 'a queue \<Rightarrow> 'a queue" where
+ "enqueue x (Queue xs) = Queue (xs @ [x])"
+
+fun %quote dequeue :: "'a queue \<Rightarrow> 'a option \<times> 'a queue" where
+ "dequeue (Queue []) = (None, Queue [])"
+ | "dequeue (Queue (x # xs)) = (Some x, Queue xs)"
+
+text {*
+ \noindent This we can use directly for proving; for executing,
+ we provide an alternative characterisation:
+*}
+
+definition %quote AQueue :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a queue" where
+ "AQueue xs ys = Queue (ys @ rev xs)"
+
+code_datatype %quote AQueue
+
+text {*
+ \noindent Here we define a \qt{constructor} @{const "AQueue"} which
+ is defined in terms of @{text "Queue"} and interprets its arguments
+ according to what the \emph{content} of an amortised queue is supposed
+ to be.
+
+ The prerequisite for datatype constructors is only syntactical: a
+ constructor must be of type @{text "\<tau> = \<dots> \<Rightarrow> \<kappa> \<alpha>\<^isub>1 \<dots> \<alpha>\<^isub>n"} where @{text
+ "{\<alpha>\<^isub>1, \<dots>, \<alpha>\<^isub>n}"} is exactly the set of \emph{all} type variables in
+ @{text "\<tau>"}; then @{text "\<kappa>"} is its corresponding datatype. The
+ HOL datatype package by default registers any new datatype with its
+ constructors, but this may be changed using @{command_def
+ code_datatype}; the currently chosen constructors can be inspected
+ using the @{command print_codesetup} command.
+
+ Equipped with this, we are able to prove the following equations
+ for our primitive queue operations which \qt{implement} the simple
+ queues in an amortised fashion:
+*}
+
+lemma %quote empty_AQueue [code]:
+ "empty = AQueue [] []"
+ by (simp add: AQueue_def empty_def)
+
+lemma %quote enqueue_AQueue [code]:
+ "enqueue x (AQueue xs ys) = AQueue (x # xs) ys"
+ by (simp add: AQueue_def)
+
+lemma %quote dequeue_AQueue [code]:
+ "dequeue (AQueue xs []) =
+ (if xs = [] then (None, AQueue [] [])
+ else dequeue (AQueue [] (rev xs)))"
+ "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)"
+ by (simp_all add: AQueue_def)
+
+text {*
+ \noindent It is good style, although no absolute requirement, to
+ provide code equations for the original artefacts of the implemented
+ type, if possible; in our case, these are the datatype constructor
+ @{const Queue} and the case combinator @{const queue_case}:
+*}
+
+lemma %quote Queue_AQueue [code]:
+ "Queue = AQueue []"
+ by (simp add: AQueue_def fun_eq_iff)
+
+lemma %quote queue_case_AQueue [code]:
+ "queue_case f (AQueue xs ys) = f (ys @ rev xs)"
+ by (simp add: AQueue_def)
+
+text {*
+ \noindent The resulting code looks as expected:
+*}
+
+text %quotetypewriter {*
+ @{code_stmts empty enqueue dequeue Queue queue_case (SML)}
+*}
+
+text {*
+ The same techniques can also be applied to types which are not
+ specified as datatypes, e.g.~type @{typ int} is originally specified
+ as quotient type by means of @{command_def typedef}, but for code
+ generation constants allowing construction of binary numeral values
+ are used as constructors for @{typ int}.
+
+ This approach however fails if the representation of a type demands
+ invariants; this issue is discussed in the next section.
+*}
+
+
+subsection {* Datatype refinement involving invariants \label{sec:invariant} *}
+
+text {*
+ Datatype representation involving invariants require a dedicated
+ setup for the type and its primitive operations. As a running
+ example, we implement a type @{text "'a dlist"} of list consisting
+ of distinct elements.
+
+ The first step is to decide on which representation the abstract
+ type (in our example @{text "'a dlist"}) should be implemented.
+ Here we choose @{text "'a list"}. Then a conversion from the concrete
+ type to the abstract type must be specified, here:
+*}
+
+text %quote {*
+ @{term_type Dlist}
+*}
+
+text {*
+ \noindent Next follows the specification of a suitable \emph{projection},
+ i.e.~a conversion from abstract to concrete type:
+*}
+
+text %quote {*
+ @{term_type list_of_dlist}
+*}
+
+text {*
+ \noindent This projection must be specified such that the following
+ \emph{abstract datatype certificate} can be proven:
+*}
+
+lemma %quote [code abstype]:
+ "Dlist (list_of_dlist dxs) = dxs"
+ by (fact Dlist_list_of_dlist)
+
+text {*
+ \noindent Note that so far the invariant on representations
+ (@{term_type distinct}) has never been mentioned explicitly:
+ the invariant is only referred to implicitly: all values in
+ set @{term "{xs. list_of_dlist (Dlist xs) = xs}"} are invariant,
+ and in our example this is exactly @{term "{xs. distinct xs}"}.
+
+ The primitive operations on @{typ "'a dlist"} are specified
+ indirectly using the projection @{const list_of_dlist}. For
+ the empty @{text "dlist"}, @{const Dlist.empty}, we finally want
+ the code equation
+*}
+
+text %quote {*
+ @{term "Dlist.empty = Dlist []"}
+*}
+
+text {*
+ \noindent This we have to prove indirectly as follows:
+*}
+
+lemma %quote [code abstract]:
+ "list_of_dlist Dlist.empty = []"
+ by (fact list_of_dlist_empty)
+
+text {*
+ \noindent This equation logically encodes both the desired code
+ equation and that the expression @{const Dlist} is applied to obeys
+ the implicit invariant. Equations for insertion and removal are
+ similar:
+*}
+
+lemma %quote [code abstract]:
+ "list_of_dlist (Dlist.insert x dxs) = List.insert x (list_of_dlist dxs)"
+ by (fact list_of_dlist_insert)
+
+lemma %quote [code abstract]:
+ "list_of_dlist (Dlist.remove x dxs) = remove1 x (list_of_dlist dxs)"
+ by (fact list_of_dlist_remove)
+
+text {*
+ \noindent Then the corresponding code is as follows:
+*}
+
+text %quotetypewriter {*
+ @{code_stmts Dlist.empty Dlist.insert Dlist.remove list_of_dlist (Haskell)}
+*} (*(types) dlist (consts) dempty dinsert dremove list_of List.member insert remove *)
+
+text {*
+ Typical data structures implemented by representations involving
+ invariants are available in the library, theory @{theory Mapping}
+ specifies key-value-mappings (type @{typ "('a, 'b) mapping"});
+ these can be implemented by red-black-trees (theory @{theory RBT}).
+*}
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Codegen/Setup.thy Mon Aug 27 23:00:38 2012 +0200
@@ -0,0 +1,37 @@
+theory Setup
+imports
+ Complex_Main
+ "~~/src/HOL/Library/Dlist"
+ "~~/src/HOL/Library/RBT"
+ "~~/src/HOL/Library/Mapping"
+begin
+
+(* FIXME avoid writing into source directory *)
+ML {*
+ Isabelle_System.mkdirs (Path.append (Thy_Load.master_directory @{theory}) (Path.basic "examples"))
+*}
+
+ML_file "../antiquote_setup.ML"
+ML_file "../more_antiquote.ML"
+
+setup {*
+ Antiquote_Setup.setup #>
+ More_Antiquote.setup #>
+let
+ val typ = Simple_Syntax.read_typ;
+in
+ Sign.del_modesyntax_i (Symbol.xsymbolsN, false)
+ [("_constrain", typ "logic => type => logic", Mixfix ("_\<Colon>_", [4, 0], 3)),
+ ("_constrain", typ "prop' => type => prop'", Mixfix ("_\<Colon>_", [4, 0], 3))] #>
+ Sign.add_modesyntax_i (Symbol.xsymbolsN, false)
+ [("_constrain", typ "logic => type => logic", Mixfix ("_ \<Colon> _", [4, 0], 3)),
+ ("_constrain", typ "prop' => type => prop'", Mixfix ("_ \<Colon> _", [4, 0], 3))]
+end
+*}
+
+setup {* Code_Target.set_default_code_width 74 *}
+
+declare [[names_unique = false]]
+
+end
+
--- a/doc-src/Codegen/Thy/Adaptation.thy Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,347 +0,0 @@
-theory Adaptation
-imports Setup
-begin
-
-setup %invisible {* Code_Target.extend_target ("\<SML>", ("SML", K I))
- #> Code_Target.extend_target ("\<SMLdummy>", ("Haskell", K I)) *}
-
-section {* Adaptation to target languages \label{sec:adaptation} *}
-
-subsection {* Adapting code generation *}
-
-text {*
- The aspects of code generation introduced so far have two aspects
- in common:
-
- \begin{itemize}
-
- \item They act uniformly, without reference to a specific target
- language.
-
- \item They are \emph{safe} in the sense that as long as you trust
- the code generator meta theory and implementation, you cannot
- produce programs that yield results which are not derivable in
- the logic.
-
- \end{itemize}
-
- \noindent In this section we will introduce means to \emph{adapt}
- the serialiser to a specific target language, i.e.~to print program
- fragments in a way which accommodates \qt{already existing}
- ingredients of a target language environment, for three reasons:
-
- \begin{itemize}
- \item improving readability and aesthetics of generated code
- \item gaining efficiency
- \item interface with language parts which have no direct counterpart
- in @{text "HOL"} (say, imperative data structures)
- \end{itemize}
-
- \noindent Generally, you should avoid using those features yourself
- \emph{at any cost}:
-
- \begin{itemize}
-
- \item The safe configuration methods act uniformly on every target
- language, whereas for adaptation you have to treat each target
- language separately.
-
- \item Application is extremely tedious since there is no
- abstraction which would allow for a static check, making it easy
- to produce garbage.
-
- \item Subtle errors can be introduced unconsciously.
-
- \end{itemize}
-
- \noindent However, even if you ought refrain from setting up
- adaptation yourself, already the @{text "HOL"} comes with some
- reasonable default adaptations (say, using target language list
- syntax). There also some common adaptation cases which you can
- setup by importing particular library theories. In order to
- understand these, we provide some clues here; these however are not
- supposed to replace a careful study of the sources.
-*}
-
-
-subsection {* The adaptation principle *}
-
-text {*
- Figure \ref{fig:adaptation} illustrates what \qt{adaptation} is
- conceptually supposed to be:
-
- \begin{figure}[here]
- \includegraphics{adaptation}
- \caption{The adaptation principle}
- \label{fig:adaptation}
- \end{figure}
-
- \noindent In the tame view, code generation acts as broker between
- @{text logic}, @{text "intermediate language"} and @{text "target
- language"} by means of @{text translation} and @{text
- serialisation}; for the latter, the serialiser has to observe the
- structure of the @{text language} itself plus some @{text reserved}
- keywords which have to be avoided for generated code. However, if
- you consider @{text adaptation} mechanisms, the code generated by
- the serializer is just the tip of the iceberg:
-
- \begin{itemize}
-
- \item @{text serialisation} can be \emph{parametrised} such that
- logical entities are mapped to target-specific ones
- (e.g. target-specific list syntax, see also
- \secref{sec:adaptation_mechanisms})
-
- \item Such parametrisations can involve references to a
- target-specific standard @{text library} (e.g. using the @{text
- Haskell} @{verbatim Maybe} type instead of the @{text HOL}
- @{type "option"} type); if such are used, the corresponding
- identifiers (in our example, @{verbatim Maybe}, @{verbatim
- Nothing} and @{verbatim Just}) also have to be considered @{text
- reserved}.
-
- \item Even more, the user can enrich the library of the
- target-language by providing code snippets (\qt{@{text
- "includes"}}) which are prepended to any generated code (see
- \secref{sec:include}); this typically also involves further
- @{text reserved} identifiers.
-
- \end{itemize}
-
- \noindent As figure \ref{fig:adaptation} illustrates, all these
- adaptation mechanisms have to act consistently; it is at the
- discretion of the user to take care for this.
-*}
-
-subsection {* Common adaptation patterns *}
-
-text {*
- The @{theory HOL} @{theory 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 @{text HOL} library; beside being useful in
- applications, they may serve as a tutorial for customising the code
- generator setup (see below \secref{sec:adaptation_mechanisms}).
-
- \begin{description}
-
- \item[@{text "Code_Integer"}] represents @{text HOL} integers by
- big integer literals in target languages.
-
- \item[@{text "Code_Char"}] represents @{text HOL} characters by
- character literals in target languages.
-
- \item[@{text "Code_Char_chr"}] like @{text "Code_Char"}, but
- also offers treatment of character codes; includes @{text
- "Code_Char"}.
-
- \item[@{text "Efficient_Nat"}] \label{eff_nat} implements
- natural numbers by integers, which in general will result in
- higher efficiency; pattern matching with @{term "0\<Colon>nat"} /
- @{const "Suc"} is eliminated; includes @{text "Code_Integer"}
- and @{text "Code_Numeral"}.
-
- \item[@{theory "Code_Numeral"}] provides an additional datatype
- @{typ index} which is mapped to target-language built-in
- integers. Useful for code setups which involve e.g.~indexing
- of target-language arrays. Part of @{text "HOL-Main"}.
-
- \item[@{theory "String"}] provides an additional datatype @{typ
- String.literal} which is isomorphic to strings; @{typ
- String.literal}s are mapped to target-language strings. Useful
- for code setups which involve e.g.~printing (error) messages.
- Part of @{text "HOL-Main"}.
-
- \end{description}
-
- \begin{warn}
- When importing any of those theories which are not part of
- @{text "HOL-Main"}, 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}
-*}
-
-
-subsection {* Parametrising serialisation \label{sec:adaptation_mechanisms} *}
-
-text {*
- Consider the following function and its corresponding SML code:
-*}
-
-primrec %quote in_interval :: "nat \<times> nat \<Rightarrow> nat \<Rightarrow> bool" where
- "in_interval (k, l) n \<longleftrightarrow> k \<le> n \<and> n \<le> l"
-(*<*)
-code_type %invisible bool
- (SML)
-code_const %invisible True and False and "op \<and>" and Not
- (SML and and and)
-(*>*)
-text %quotetypewriter {*
- @{code_stmts in_interval (SML)}
-*}
-
-text {*
- \noindent Though this is correct code, it is a little bit
- unsatisfactory: boolean values and operators are materialised as
- distinguished entities with have nothing to do with the SML-built-in
- 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 @{verbatim
- "bool"} would be used. To map the HOL @{typ bool} on SML @{verbatim
- "bool"}, we may use \qn{custom serialisations}:
-*}
-
-code_type %quotett bool
- (SML "bool")
-code_const %quotett True and False and "op \<and>"
- (SML "true" and "false" and "_ andalso _")
-
-text {*
- \noindent The @{command_def code_type} command takes a type constructor
- as arguments together with a list of custom serialisations. Each
- custom serialisation starts with a target language identifier
- followed by an expression, which during code serialisation is
- inserted whenever the type constructor would occur. For constants,
- @{command_def code_const} implements the corresponding mechanism. Each
- ``@{verbatim "_"}'' in a serialisation expression is treated as a
- placeholder for the type constructor's (the constant's) arguments.
-*}
-
-text %quotetypewriter {*
- @{code_stmts in_interval (SML)}
-*}
-
-text {*
- \noindent This still is not perfect: the parentheses around the
- \qt{andalso} expression are superfluous. Though the serialiser 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:
-*}
-
-code_const %quotett "op \<and>"
- (SML infixl 1 "andalso")
-
-text %quotetypewriter {*
- @{code_stmts in_interval (SML)}
-*}
-
-text {*
- \noindent The attentive reader may ask how we assert that no
- generated code will accidentally overwrite. For this reason the
- serialiser has an internal table of identifiers which have to be
- avoided to be used for new declarations. Initially, this table
- typically contains the keywords of the target language. It can be
- extended manually, thus avoiding accidental overwrites, using the
- @{command_def "code_reserved"} command:
-*}
-
-code_reserved %quote "\<SMLdummy>" bool true false andalso
-
-text {*
- \noindent Next, we try to map HOL pairs to SML pairs, using the
- infix ``@{verbatim "*"}'' type constructor and parentheses:
-*}
-(*<*)
-code_type %invisible prod
- (SML)
-code_const %invisible Pair
- (SML)
-(*>*)
-code_type %quotett prod
- (SML infix 2 "*")
-code_const %quotett Pair
- (SML "!((_),/ (_))")
-
-text {*
- \noindent The initial bang ``@{verbatim "!"}'' tells the serialiser
- never to 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
- ``@{verbatim "/"}'' (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 serialisations
- provide; however their usage requires careful thinking in order not
- to introduce inconsistencies -- or, in other words: custom
- serialisations are completely axiomatic.
-
- A further noteworthy detail is that any special character in a
- custom serialisation may be quoted using ``@{verbatim "'"}''; thus,
- in ``@{verbatim "fn '_ => _"}'' the first ``@{verbatim "_"}'' is a
- proper underscore while the second ``@{verbatim "_"}'' is a
- placeholder.
-*}
-
-
-subsection {* @{text Haskell} serialisation *}
-
-text {*
- For convenience, the default @{text HOL} setup for @{text Haskell}
- maps the @{class equal} class to its counterpart in @{text Haskell},
- giving custom serialisations for the class @{class equal} (by command
- @{command_def code_class}) and its operation @{const [source] HOL.equal}
-*}
-
-code_class %quotett equal
- (Haskell "Eq")
-
-code_const %quotett "HOL.equal"
- (Haskell infixl 4 "==")
-
-text {*
- \noindent A problem now occurs whenever a type which is an instance
- of @{class equal} in @{text HOL} is mapped on a @{text
- Haskell}-built-in type which is also an instance of @{text Haskell}
- @{text Eq}:
-*}
-
-typedecl %quote bar
-
-instantiation %quote bar :: equal
-begin
-
-definition %quote "HOL.equal (x\<Colon>bar) y \<longleftrightarrow> x = y"
-
-instance %quote by default (simp add: equal_bar_def)
-
-end %quote (*<*)
-
-(*>*) code_type %quotett bar
- (Haskell "Integer")
-
-text {*
- \noindent The code generator would produce an additional instance,
- which of course is rejected by the @{text Haskell} compiler. To
- suppress this additional instance, use @{command_def "code_instance"}:
-*}
-
-code_instance %quotett bar :: equal
- (Haskell -)
-
-
-subsection {* Enhancing the target language context \label{sec:include} *}
-
-text {*
- In rare cases it is necessary to \emph{enrich} the context of a
- target language; this is accomplished using the @{command_def
- "code_include"} command:
-*}
-
-code_include %quotett Haskell "Errno"
-{*errno i = error ("Error number: " ++ show i)*}
-
-code_reserved %quotett Haskell Errno
-
-text {*
- \noindent Such named @{text include}s are then prepended to every
- generated code. Inspect such code in order to find out how
- @{command "code_include"} behaves with respect to a particular
- target language.
-*}
-
-end
-
--- a/doc-src/Codegen/Thy/Evaluation.thy Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,287 +0,0 @@
-theory Evaluation
-imports Setup
-begin
-
-section {* Evaluation \label{sec:evaluation} *}
-
-text {*
- Recalling \secref{sec:principle}, code generation turns a system of
- equations into a program with the \emph{same} equational semantics.
- As a consequence, this program can be used as a \emph{rewrite
- engine} for terms: rewriting a term @{term "t"} using a program to a
- term @{term "t'"} yields the theorems @{prop "t \<equiv> t'"}. This
- application of code generation in the following is referred to as
- \emph{evaluation}.
-*}
-
-
-subsection {* Evaluation techniques *}
-
-text {*
- The existing infrastructure provides a rich palette of evaluation
- techniques, each comprising different aspects:
-
- \begin{description}
-
- \item[Expressiveness.] Depending on how good symbolic computation
- is supported, the class of terms which can be evaluated may be
- bigger or smaller.
-
- \item[Efficiency.] The more machine-near the technique, the
- faster it is.
-
- \item[Trustability.] Techniques which a huge (and also probably
- more configurable infrastructure) are more fragile and less
- trustable.
-
- \end{description}
-*}
-
-
-subsubsection {* The simplifier (@{text simp}) *}
-
-text {*
- The simplest way for evaluation is just using the simplifier with
- the original code equations of the underlying program. This gives
- fully symbolic evaluation and highest trustablity, with the usual
- performance of the simplifier. Note that for operations on abstract
- datatypes (cf.~\secref{sec:invariant}), the original theorems as
- given by the users are used, not the modified ones.
-*}
-
-
-subsubsection {* Normalization by evaluation (@{text nbe}) *}
-
-text {*
- Normalization by evaluation \cite{Aehlig-Haftmann-Nipkow:2008:nbe}
- provides a comparably fast partially symbolic evaluation which
- permits also normalization of functions and uninterpreted symbols;
- the stack of code to be trusted is considerable.
-*}
-
-
-subsubsection {* Evaluation in ML (@{text code}) *}
-
-text {*
- Highest performance can be achieved by evaluation in ML, at the cost
- of being restricted to ground results and a layered stack of code to
- be trusted, including code generator configurations by the user.
-
- Evaluation is carried out in a target language \emph{Eval} which
- inherits from \emph{SML} but for convenience uses parts of the
- Isabelle runtime environment. The soundness of computation carried
- out there depends crucially on the correctness of the code
- generator setup; this is one of the reasons why you should not use
- adaptation (see \secref{sec:adaptation}) frivolously.
-*}
-
-
-subsection {* Aspects of evaluation *}
-
-text {*
- Each of the techniques can be combined with different aspects. The
- most important distinction is between dynamic and static evaluation.
- Dynamic evaluation takes the code generator configuration \qt{as it
- is} at the point where evaluation is issued. Best example is the
- @{command_def value} command which allows ad-hoc evaluation of
- terms:
-*}
-
-value %quote "42 / (12 :: rat)"
-
-text {*
- \noindent By default @{command value} tries all available evaluation
- techniques and prints the result of the first succeeding one. A particular
- technique may be specified in square brackets, e.g.
-*}
-
-value %quote [nbe] "42 / (12 :: rat)"
-
-text {*
- To employ dynamic evaluation in the document generation, there is also
- a @{text value} antiquotation. By default, it also tries all available evaluation
- techniques and prints the result of the first succeeding one, unless a particular
- technique is specified in square brackets.
-
- Static evaluation freezes the code generator configuration at a
- certain point and uses this context whenever evaluation is issued
- later on. This is particularly appropriate for proof procedures
- which use evaluation, since then the behaviour of evaluation is not
- changed or even compromised later on by actions of the user.
-
- As a technical complication, terms after evaluation in ML must be
- turned into Isabelle's internal term representation again. Since
- this is also configurable, it is never fully trusted. For this
- reason, evaluation in ML comes with further aspects:
-
- \begin{description}
-
- \item[Plain evaluation.] A term is normalized using the provided
- term reconstruction from ML to Isabelle; for applications which
- do not need to be fully trusted.
-
- \item[Property conversion.] Evaluates propositions; since these
- are monomorphic, the term reconstruction is fixed once and for all
- and therefore trustable.
-
- \item[Conversion.] Evaluates an arbitrary term @{term "t"} first
- by plain evaluation and certifies the result @{term "t'"} by
- checking the equation @{term "t \<equiv> t'"} using property
- conversion.
-
- \end{description}
-
- \noindent The picture is further complicated by the roles of
- exceptions. Here three cases have to be distinguished:
-
- \begin{itemize}
-
- \item Evaluation of @{term t} terminates with a result @{term
- "t'"}.
-
- \item Evaluation of @{term t} terminates which en exception
- indicating a pattern match failure or a non-implemented
- function. As sketched in \secref{sec:partiality}, this can be
- interpreted as partiality.
-
- \item Evaluation raises any other kind of exception.
-
- \end{itemize}
-
- \noindent For conversions, the first case yields the equation @{term
- "t = t'"}, the second defaults to reflexivity @{term "t = t"}.
- Exceptions of the third kind are propagated to the user.
-
- By default return values of plain evaluation are optional, yielding
- @{text "SOME t'"} in the first case, @{text "NONE"} in the
- second, and propagating the exception in the third case. A strict
- variant of plain evaluation either yields @{text "t'"} or propagates
- any exception, a liberal variant caputures any exception in a result
- of type @{text "Exn.result"}.
-
- For property conversion (which coincides with conversion except for
- evaluation in ML), methods are provided which solve a given goal by
- evaluation.
-*}
-
-
-subsection {* Schematic overview *}
-
-text {*
- \newcommand{\ttsize}{\fontsize{5.8pt}{8pt}\selectfont}
- \fontsize{9pt}{12pt}\selectfont
- \begin{tabular}{ll||c|c|c}
- & & @{text simp} & @{text nbe} & @{text code} \tabularnewline \hline \hline
- \multirow{5}{1ex}{\rotatebox{90}{dynamic}}
- & interactive evaluation
- & @{command value} @{text "[simp]"} & @{command value} @{text "[nbe]"} & @{command value} @{text "[code]"}
- \tabularnewline
- & plain evaluation & & & \ttsize@{ML "Code_Evaluation.dynamic_value"} \tabularnewline \cline{2-5}
- & evaluation method & @{method code_simp} & @{method normalization} & @{method eval} \tabularnewline
- & property conversion & & & \ttsize@{ML "Code_Runtime.dynamic_holds_conv"} \tabularnewline \cline{2-5}
- & conversion & \ttsize@{ML "Code_Simp.dynamic_conv"} & \ttsize@{ML "Nbe.dynamic_conv"}
- & \ttsize@{ML "Code_Evaluation.dynamic_conv"} \tabularnewline \hline \hline
- \multirow{3}{1ex}{\rotatebox{90}{static}}
- & plain evaluation & & & \ttsize@{ML "Code_Evaluation.static_value"} \tabularnewline \cline{2-5}
- & property conversion & &
- & \ttsize@{ML "Code_Runtime.static_holds_conv"} \tabularnewline \cline{2-5}
- & conversion & \ttsize@{ML "Code_Simp.static_conv"}
- & \ttsize@{ML "Nbe.static_conv"}
- & \ttsize@{ML "Code_Evaluation.static_conv"}
- \end{tabular}
-*}
-
-
-subsection {* Intimate connection between logic and system runtime *}
-
-text {*
- The toolbox of static evaluation conversions forms a reasonable base
- to interweave generated code and system tools. However in some
- situations more direct interaction is desirable.
-*}
-
-
-subsubsection {* Static embedding of generated code into system runtime -- the @{text code} antiquotation *}
-
-text {*
- The @{text code} antiquotation allows to include constants from
- generated code directly into ML system code, as in the following toy
- example:
-*}
-
-datatype %quote form = T | F | And form form | Or form form (*<*)
-
-(*>*) ML %quotett {*
- fun eval_form @{code T} = true
- | eval_form @{code F} = false
- | eval_form (@{code And} (p, q)) =
- eval_form p andalso eval_form q
- | eval_form (@{code Or} (p, q)) =
- eval_form p orelse eval_form q;
-*}
-
-text {*
- \noindent @{text code} takes as argument the name of a constant;
- after the whole ML is read, the necessary code is generated
- transparently and the corresponding constant names are inserted.
- This technique also allows to use pattern matching on constructors
- stemming from compiled datatypes. Note that the @{text code}
- antiquotation may not refer to constants which carry adaptations;
- here you have to refer to the corresponding adapted code directly.
-
- For a less simplistic example, theory @{text Approximation} in
- the @{text Decision_Procs} session is a good reference.
-*}
-
-
-subsubsection {* Static embedding of generated code into system runtime -- @{text code_reflect} *}
-
-text {*
- The @{text code} antiquoation is lightweight, but the generated code
- is only accessible while the ML section is processed. Sometimes this
- is not appropriate, especially if the generated code contains datatype
- declarations which are shared with other parts of the system. In these
- cases, @{command_def code_reflect} can be used:
-*}
-
-code_reflect %quote Sum_Type
- datatypes sum = Inl | Inr
- functions "Sum_Type.Projl" "Sum_Type.Projr"
-
-text {*
- \noindent @{command_def code_reflect} takes a structure name and
- references to datatypes and functions; for these code is compiled
- into the named ML structure and the \emph{Eval} target is modified
- in a way that future code generation will reference these
- precompiled versions of the given datatypes and functions. This
- also allows to refer to the referenced datatypes and functions from
- arbitrary ML code as well.
-
- A typical example for @{command code_reflect} can be found in the
- @{theory Predicate} theory.
-*}
-
-
-subsubsection {* Separate compilation -- @{text code_reflect} *}
-
-text {*
- For technical reasons it is sometimes necessary to separate
- generation and compilation of code which is supposed to be used in
- the system runtime. For this @{command code_reflect} with an
- optional @{text "file"} argument can be used:
-*}
-
-code_reflect %quote Rat
- datatypes rat = Frct
- functions Fract
- "(plus :: rat \<Rightarrow> rat \<Rightarrow> rat)" "(minus :: rat \<Rightarrow> rat \<Rightarrow> rat)"
- "(times :: rat \<Rightarrow> rat \<Rightarrow> rat)" "(divide :: rat \<Rightarrow> rat \<Rightarrow> rat)"
- file "examples/rat.ML"
-
-text {*
- \noindent This merely generates the referenced code to the given
- file which can be included into the system runtime later on.
-*}
-
-end
-
--- a/doc-src/Codegen/Thy/Foundations.thy Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,347 +0,0 @@
-theory Foundations
-imports Introduction
-begin
-
-section {* Code generation foundations \label{sec:foundations} *}
-
-subsection {* Code generator architecture \label{sec:architecture} *}
-
-text {*
- The code generator is actually a framework consisting of different
- components which can be customised individually.
-
- Conceptually all components operate on Isabelle's logic framework
- @{theory Pure}. Practically, the object logic @{theory HOL}
- provides the necessary facilities to make use of the code generator,
- mainly since it is an extension of @{theory Pure}.
-
- The constellation of the different components is visualized in the
- following picture.
-
- \begin{figure}[h]
- \includegraphics{architecture}
- \caption{Code generator architecture}
- \label{fig:arch}
- \end{figure}
-
- \noindent Central to code generation is the notion of \emph{code
- equations}. A code equation as a first approximation is a theorem
- of the form @{text "f t\<^isub>1 t\<^isub>2 \<dots> t\<^isub>n \<equiv> t"} (an equation headed by a
- constant @{text f} with arguments @{text "t\<^isub>1 t\<^isub>2 \<dots> t\<^isub>n"} and right
- hand side @{text t}).
-
- \begin{itemize}
-
- \item Starting point of code generation is a collection of (raw)
- code equations in a theory. It is not relevant where they stem
- from, but typically they were either produced by specification
- tools or proved explicitly by the user.
-
- \item These raw code equations can be subjected to theorem
- transformations. This \qn{preprocessor} (see
- \secref{sec:preproc}) can apply the full expressiveness of
- ML-based theorem transformations to code generation. The result
- of preprocessing is a structured collection of code equations.
-
- \item These code equations are \qn{translated} to a program in an
- abstract intermediate language. Think of it as a kind of
- \qt{Mini-Haskell} with four \qn{statements}: @{text data} (for
- datatypes), @{text fun} (stemming from code equations), also
- @{text class} and @{text inst} (for type classes).
-
- \item Finally, the abstract program is \qn{serialised} into
- concrete source code of a target language. This step only
- produces concrete syntax but does not change the program in
- essence; all conceptual transformations occur in the translation
- step.
-
- \end{itemize}
-
- \noindent From these steps, only the last two 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.
-*}
-
-
-subsection {* The preprocessor \label{sec:preproc} *}
-
-text {*
- Before selected function theorems are turned into abstract code, a
- chain of definitional transformation steps is carried out:
- \emph{preprocessing}. The preprocessor consists of two
- components: a \emph{simpset} and \emph{function transformers}.
-
- The \emph{simpset} can apply the full generality of the Isabelle
- simplifier. Due to the interpretation of theorems as code
- 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{unfold theorems}, which may be declared and removed using the
- @{attribute code_unfold} or \emph{@{attribute code_unfold} del}
- attribute, respectively.
-
- Some common applications:
-*}
-
-text_raw {*
- \begin{itemize}
-*}
-
-text {*
- \item replacing non-executable constructs by executable ones:
-*}
-
-lemma %quote [code_unfold]:
- "x \<in> set xs \<longleftrightarrow> List.member xs x" by (fact in_set_member)
-
-text {*
- \item replacing executable but inconvenient constructs:
-*}
-
-lemma %quote [code_unfold]:
- "xs = [] \<longleftrightarrow> List.null xs" by (fact eq_Nil_null)
-
-text {*
- \item eliminating disturbing expressions:
-*}
-
-lemma %quote [code_unfold]:
- "1 = Suc 0" by (fact One_nat_def)
-
-text_raw {*
- \end{itemize}
-*}
-
-text {*
- \noindent \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 @{term "0\<Colon>nat"} / @{const Suc} pattern
- elimination implemented in theory @{text Efficient_Nat} (see
- \secref{eff_nat}) uses this interface.
-
- \noindent The current setup of the preprocessor may be inspected
- using the @{command_def print_codeproc} command. @{command_def
- code_thms} (see \secref{sec:equations}) provides a convenient
- mechanism to inspect the impact of a preprocessor setup on code
- equations.
-*}
-
-
-subsection {* Understanding code equations \label{sec:equations} *}
-
-text {*
- As told in \secref{sec:principle}, the notion of code equations is
- vital to code generation. Indeed most problems which occur in
- practice can be resolved by an inspection of the underlying code
- equations.
-
- It is possible to exchange the default code equations for constants
- by explicitly proving alternative ones:
-*}
-
-lemma %quote [code]:
- "dequeue (AQueue xs []) =
- (if xs = [] then (None, AQueue [] [])
- else dequeue (AQueue [] (rev xs)))"
- "dequeue (AQueue xs (y # ys)) =
- (Some y, AQueue xs ys)"
- by (cases xs, simp_all) (cases "rev xs", simp_all)
-
-text {*
- \noindent The annotation @{text "[code]"} is an @{text attribute}
- which states that the given theorems should be considered as code
- equations for a @{text fun} statement -- the corresponding constant
- is determined syntactically. The resulting code:
-*}
-
-text %quotetypewriter {*
- @{code_stmts dequeue (consts) dequeue (Haskell)}
-*}
-
-text {*
- \noindent You may note that the equality test @{term "xs = []"} has
- been replaced by the predicate @{term "List.null xs"}. This is due
- to the default setup of the \qn{preprocessor}.
-
- This possibility to select arbitrary code equations is the key
- technique for program and datatype refinement (see
- \secref{sec:refinement}).
-
- Due to the preprocessor, there is the distinction of raw code
- equations (before preprocessing) and code equations (after
- preprocessing).
-
- The first can be listed (among other data) using the @{command_def
- print_codesetup} command.
-
- The code equations after preprocessing are already are blueprint of
- the generated program and can be inspected using the @{command
- code_thms} command:
-*}
-
-code_thms %quote dequeue
-
-text {*
- \noindent This prints a table with the code equations for @{const
- dequeue}, including \emph{all} code equations those equations depend
- on recursively. These dependencies themselves can be visualized using
- the @{command_def code_deps} command.
-*}
-
-
-subsection {* Equality *}
-
-text {*
- Implementation of equality deserves some attention. Here an example
- function involving polymorphic equality:
-*}
-
-primrec %quote collect_duplicates :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
- "collect_duplicates xs ys [] = xs"
-| "collect_duplicates xs ys (z#zs) = (if z \<in> set xs
- then if z \<in> set ys
- then collect_duplicates xs ys zs
- else collect_duplicates xs (z#ys) zs
- else collect_duplicates (z#xs) (z#ys) zs)"
-
-text {*
- \noindent During preprocessing, the membership test is rewritten,
- resulting in @{const List.member}, which itself performs an explicit
- equality check, as can be seen in the corresponding @{text SML} code:
-*}
-
-text %quotetypewriter {*
- @{code_stmts collect_duplicates (SML)}
-*}
-
-text {*
- \noindent Obviously, polymorphic equality is implemented the Haskell
- way using a type class. How is this achieved? HOL introduces an
- explicit class @{class equal} with a corresponding operation @{const
- HOL.equal} such that @{thm equal [no_vars]}. The preprocessing
- framework does the rest by propagating the @{class equal} constraints
- through all dependent code equations. For datatypes, instances of
- @{class equal} are implicitly derived when possible. For other types,
- you may instantiate @{text equal} manually like any other type class.
-*}
-
-
-subsection {* Explicit partiality \label{sec:partiality} *}
-
-text {*
- Partiality usually enters the game by partial patterns, as
- in the following example, again for amortised queues:
-*}
-
-definition %quote strict_dequeue :: "'a queue \<Rightarrow> 'a \<times> 'a queue" where
- "strict_dequeue q = (case dequeue q
- of (Some x, q') \<Rightarrow> (x, q'))"
-
-lemma %quote strict_dequeue_AQueue [code]:
- "strict_dequeue (AQueue xs (y # ys)) = (y, AQueue xs ys)"
- "strict_dequeue (AQueue xs []) =
- (case rev xs of y # ys \<Rightarrow> (y, AQueue [] ys))"
- by (simp_all add: strict_dequeue_def) (cases xs, simp_all split: list.split)
-
-text {*
- \noindent In the corresponding code, there is no equation
- for the pattern @{term "AQueue [] []"}:
-*}
-
-text %quotetypewriter {*
- @{code_stmts strict_dequeue (consts) strict_dequeue (Haskell)}
-*}
-
-text {*
- \noindent In some cases it is desirable to have this
- pseudo-\qt{partiality} more explicitly, e.g.~as follows:
-*}
-
-axiomatization %quote empty_queue :: 'a
-
-definition %quote strict_dequeue' :: "'a queue \<Rightarrow> 'a \<times> 'a queue" where
- "strict_dequeue' q = (case dequeue q of (Some x, q') \<Rightarrow> (x, q') | _ \<Rightarrow> empty_queue)"
-
-lemma %quote strict_dequeue'_AQueue [code]:
- "strict_dequeue' (AQueue xs []) = (if xs = [] then empty_queue
- else strict_dequeue' (AQueue [] (rev xs)))"
- "strict_dequeue' (AQueue xs (y # ys)) =
- (y, AQueue xs ys)"
- by (simp_all add: strict_dequeue'_def split: list.splits)
-
-text {*
- Observe that on the right hand side of the definition of @{const
- "strict_dequeue'"}, the unspecified constant @{const empty_queue} occurs.
-
- Normally, if constants without any code equations occur in a
- program, the code generator complains (since in most cases this is
- indeed an error). But such constants can also be thought
- of as function definitions which always fail,
- since there is never a successful pattern match on the left hand
- side. In order to categorise a constant into that category
- explicitly, use @{command_def "code_abort"}:
-*}
-
-code_abort %quote empty_queue
-
-text {*
- \noindent Then the code generator will just insert an error or
- exception at the appropriate position:
-*}
-
-text %quotetypewriter {*
- @{code_stmts strict_dequeue' (consts) empty_queue strict_dequeue' (Haskell)}
-*}
-
-text {*
- \noindent This feature however is rarely needed in practice. Note
- also that the HOL default setup already declares @{const undefined}
- as @{command "code_abort"}, which is most likely to be used in such
- situations.
-*}
-
-
-subsection {* If something goes utterly wrong \label{sec:utterly_wrong} *}
-
-text {*
- Under certain circumstances, the code generator fails to produce
- code entirely. To debug these, the following hints may prove
- helpful:
-
- \begin{description}
-
- \ditem{\emph{Check with a different target language}.} Sometimes
- the situation gets more clear if you switch to another target
- language; the code generated there might give some hints what
- prevents the code generator to produce code for the desired
- language.
-
- \ditem{\emph{Inspect code equations}.} Code equations are the central
- carrier of code generation. Most problems occurring while generating
- code can be traced to single equations which are printed as part of
- the error message. A closer inspection of those may offer the key
- for solving issues (cf.~\secref{sec:equations}).
-
- \ditem{\emph{Inspect preprocessor setup}.} The preprocessor might
- transform code equations unexpectedly; to understand an
- inspection of its setup is necessary (cf.~\secref{sec:preproc}).
-
- \ditem{\emph{Generate exceptions}.} If the code generator
- complains about missing code equations, in can be helpful to
- implement the offending constants as exceptions
- (cf.~\secref{sec:partiality}); this allows at least for a formal
- generation of code, whose inspection may then give clues what is
- wrong.
-
- \ditem{\emph{Remove offending code equations}.} If code
- generation is prevented by just a single equation, this can be
- removed (cf.~\secref{sec:equations}) to allow formal code
- generation, whose result in turn can be used to trace the
- problem. The most prominent case here are mismatches in type
- class signatures (\qt{wellsortedness error}).
-
- \end{description}
-*}
-
-end
--- a/doc-src/Codegen/Thy/Further.thy Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,351 +0,0 @@
-theory Further
-imports Setup
-begin
-
-section {* Further issues \label{sec:further} *}
-
-subsection {* Specialities of the @{text Scala} target language \label{sec:scala} *}
-
-text {*
- @{text Scala} deviates from languages of the ML family in a couple
- of aspects; those which affect code generation mainly have to do with
- @{text Scala}'s type system:
-
- \begin{itemize}
-
- \item @{text Scala} prefers tupled syntax over curried syntax.
-
- \item @{text Scala} sacrifices Hindely-Milner type inference for a
- much more rich type system with subtyping etc. For this reason
- type arguments sometimes have to be given explicitly in square
- brackets (mimicking System F syntax).
-
- \item In contrast to @{text Haskell} where most specialities of
- the type system are implemented using \emph{type classes},
- @{text Scala} provides a sophisticated system of \emph{implicit
- arguments}.
-
- \end{itemize}
-
- \noindent Concerning currying, the @{text Scala} serializer counts
- arguments in code equations to determine how many arguments
- shall be tupled; remaining arguments and abstractions in terms
- rather than function definitions are always curried.
-
- The second aspect affects user-defined adaptations with @{command
- code_const}. For regular terms, the @{text Scala} serializer prints
- all type arguments explicitly. For user-defined term adaptations
- this is only possible for adaptations which take no arguments: here
- the type arguments are just appended. Otherwise they are ignored;
- hence user-defined adaptations for polymorphic constants have to be
- designed very carefully to avoid ambiguity.
-
- Isabelle's type classes are mapped onto @{text Scala} implicits; in
- cases with diamonds in the subclass hierarchy this can lead to
- ambiguities in the generated code:
-*}
-
-class %quote class1 =
- fixes foo :: "'a \<Rightarrow> 'a"
-
-class %quote class2 = class1
-
-class %quote class3 = class1
-
-text {*
- \noindent Here both @{class class2} and @{class class3} inherit from @{class class1},
- forming the upper part of a diamond.
-*}
-
-definition %quote bar :: "'a :: {class2, class3} \<Rightarrow> 'a" where
- "bar = foo"
-
-text {*
- \noindent This yields the following code:
-*}
-
-text %quotetypewriter {*
- @{code_stmts bar (Scala)}
-*}
-
-text {*
- \noindent This code is rejected by the @{text Scala} compiler: in
- the definition of @{text bar}, it is not clear from where to derive
- the implicit argument for @{text foo}.
-
- The solution to the problem is to close the diamond by a further
- class with inherits from both @{class class2} and @{class class3}:
-*}
-
-class %quote class4 = class2 + class3
-
-text {*
- \noindent Then the offending code equation can be restricted to
- @{class class4}:
-*}
-
-lemma %quote [code]:
- "(bar :: 'a::class4 \<Rightarrow> 'a) = foo"
- by (simp only: bar_def)
-
-text {*
- \noindent with the following code:
-*}
-
-text %quotetypewriter {*
- @{code_stmts bar (Scala)}
-*}
-
-text {*
- \noindent which exposes no ambiguity.
-
- Since the preprocessor (cf.~\secref{sec:preproc}) propagates sort
- constraints through a system of code equations, it is usually not
- very difficult to identify the set of code equations which actually
- needs more restricted sort constraints.
-*}
-
-subsection {* Modules namespace *}
-
-text {*
- When invoking the @{command export_code} command it is possible to
- leave out the @{keyword "module_name"} part; then code is
- distributed over different modules, where the module name space
- roughly is induced by the Isabelle theory name space.
-
- Then sometimes the awkward situation occurs that dependencies
- between definitions introduce cyclic dependencies between modules,
- which in the @{text Haskell} world leaves you to the mercy of the
- @{text Haskell} implementation you are using, while for @{text
- SML}/@{text OCaml} 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
-*}
-
-code_modulename %quote SML
- A ABC
- B ABC
- C ABC
-
-text {*
- \noindent we explicitly map all those modules on \emph{ABC},
- resulting in an ad-hoc merge of this three modules at serialisation
- time.
-*}
-
-subsection {* Locales and interpretation *}
-
-text {*
- A technical issue comes to surface when generating code from
- specifications stemming from locale interpretation.
-
- Let us assume a locale specifying a power operation on arbitrary
- types:
-*}
-
-locale %quote power =
- fixes power :: "'a \<Rightarrow> 'b \<Rightarrow> 'b"
- assumes power_commute: "power x \<circ> power y = power y \<circ> power x"
-begin
-
-text {*
- \noindent Inside that locale we can lift @{text power} to exponent
- lists by means of specification relative to that locale:
-*}
-
-primrec %quote powers :: "'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
- "powers [] = id"
-| "powers (x # xs) = power x \<circ> powers xs"
-
-lemma %quote powers_append:
- "powers (xs @ ys) = powers xs \<circ> powers ys"
- by (induct xs) simp_all
-
-lemma %quote powers_power:
- "powers xs \<circ> power x = power x \<circ> powers xs"
- by (induct xs)
- (simp_all del: o_apply id_apply add: o_assoc [symmetric],
- simp del: o_apply add: o_assoc power_commute)
-
-lemma %quote powers_rev:
- "powers (rev xs) = powers xs"
- by (induct xs) (simp_all add: powers_append powers_power)
-
-end %quote
-
-text {*
- After an interpretation of this locale (say, @{command_def
- interpretation} @{text "fun_power:"} @{term [source] "power (\<lambda>n (f
- :: 'a \<Rightarrow> 'a). f ^^ n)"}), one would expect to have a constant @{text
- "fun_power.powers :: nat list \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"} for which code
- can be generated. But this not the case: internally, the term
- @{text "fun_power.powers"} is an abbreviation for the foundational
- term @{term [source] "power.powers (\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n)"}
- (see \cite{isabelle-locale} for the details behind).
-
- Fortunately, with minor effort the desired behaviour can be
- achieved. First, a dedicated definition of the constant on which
- the local @{text "powers"} after interpretation is supposed to be
- mapped on:
-*}
-
-definition %quote funpows :: "nat list \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
- [code del]: "funpows = power.powers (\<lambda>n f. f ^^ n)"
-
-text {*
- \noindent In general, the pattern is @{text "c = t"} where @{text c}
- is the name of the future constant and @{text t} the foundational
- term corresponding to the local constant after interpretation.
-
- The interpretation itself is enriched with an equation @{text "t = c"}:
-*}
-
-interpretation %quote fun_power: power "\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n" where
- "power.powers (\<lambda>n f. f ^^ n) = funpows"
- by unfold_locales
- (simp_all add: fun_eq_iff funpow_mult mult_commute funpows_def)
-
-text {*
- \noindent This additional equation is trivially proved by the
- definition itself.
-
- After this setup procedure, code generation can continue as usual:
-*}
-
-text %quotetypewriter {*
- @{code_stmts funpows (consts) Nat.funpow funpows (Haskell)}
-*}
-
-
-subsection {* Imperative data structures *}
-
-text {*
- If you consider imperative data structures as inevitable for a
- specific application, you should consider \emph{Imperative
- Functional Programming with Isabelle/HOL}
- \cite{bulwahn-et-al:2008:imperative}; the framework described there
- is available in session @{text Imperative_HOL}, together with a
- short primer document.
-*}
-
-
-subsection {* ML system interfaces \label{sec:ml} *}
-
-text {*
- 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 fundamental ML
- interfaces.
-*}
-
-subsubsection {* Managing executable content *}
-
-text %mlref {*
- \begin{mldecls}
- @{index_ML Code.read_const: "theory -> string -> string"} \\
- @{index_ML Code.add_eqn: "thm -> theory -> theory"} \\
- @{index_ML Code.del_eqn: "thm -> theory -> theory"} \\
- @{index_ML Code_Preproc.map_pre: "(simpset -> simpset) -> theory -> theory"} \\
- @{index_ML Code_Preproc.map_post: "(simpset -> simpset) -> theory -> theory"} \\
- @{index_ML Code_Preproc.add_functrans: "
- string * (theory -> (thm * bool) list -> (thm * bool) list option)
- -> theory -> theory"} \\
- @{index_ML Code_Preproc.del_functrans: "string -> theory -> theory"} \\
- @{index_ML Code.add_datatype: "(string * typ) list -> theory -> theory"} \\
- @{index_ML Code.get_type: "theory -> string
- -> ((string * sort) list * (string * ((string * sort) list * typ list)) list) * bool"} \\
- @{index_ML Code.get_type_of_constr_or_abstr: "theory -> string -> (string * bool) option"}
- \end{mldecls}
-
- \begin{description}
-
- \item @{ML Code.read_const}~@{text thy}~@{text s}
- reads a constant as a concrete term expression @{text s}.
-
- \item @{ML Code.add_eqn}~@{text "thm"}~@{text "thy"} adds function
- theorem @{text "thm"} to executable content.
-
- \item @{ML Code.del_eqn}~@{text "thm"}~@{text "thy"} removes function
- theorem @{text "thm"} from executable content, if present.
-
- \item @{ML Code_Preproc.map_pre}~@{text "f"}~@{text "thy"} changes
- the preprocessor simpset.
-
- \item @{ML Code_Preproc.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds
- function transformer @{text f} (named @{text name}) to executable content;
- @{text f} is a transformer of the code equations belonging
- to a certain function definition, depending on the
- current theory context. Returning @{text NONE} indicates that no
- transformation took place; otherwise, the whole process will be iterated
- with the new code equations.
-
- \item @{ML Code_Preproc.del_functrans}~@{text "name"}~@{text "thy"} removes
- function transformer named @{text name} from executable content.
-
- \item @{ML Code.add_datatype}~@{text cs}~@{text thy} adds
- a datatype to executable content, with generation
- set @{text cs}.
-
- \item @{ML Code.get_type_of_constr_or_abstr}~@{text "thy"}~@{text "const"}
- returns type constructor corresponding to
- constructor @{text const}; returns @{text NONE}
- if @{text const} is no constructor.
-
- \end{description}
-*}
-
-
-subsubsection {* Data depending on the theory's executable content *}
-
-text {*
- 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.
-
- 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 @{ML_functor Code_Data}; on instantiation of this functor, the
- following types and operations are required:
-
- \medskip
- \begin{tabular}{l}
- @{text "type T"} \\
- @{text "val empty: T"} \\
- \end{tabular}
-
- \begin{description}
-
- \item @{text T} the type of data to store.
-
- \item @{text empty} initial (empty) data.
-
- \end{description}
-
- \noindent An instance of @{ML_functor Code_Data} provides the
- following interface:
-
- \medskip
- \begin{tabular}{l}
- @{text "change: theory \<rightarrow> (T \<rightarrow> T) \<rightarrow> T"} \\
- @{text "change_yield: theory \<rightarrow> (T \<rightarrow> 'a * T) \<rightarrow> 'a * T"}
- \end{tabular}
-
- \begin{description}
-
- \item @{text change} update of current data (cached!) by giving a
- continuation.
-
- \item @{text change_yield} update with side result.
-
- \end{description}
-*}
-
-end
-
--- a/doc-src/Codegen/Thy/Inductive_Predicate.thy Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,275 +0,0 @@
-theory Inductive_Predicate
-imports Setup
-begin
-
-(*<*)
-hide_const %invisible append
-
-inductive %invisible append where
- "append [] ys ys"
-| "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"
-
-lemma %invisible append: "append xs ys zs = (xs @ ys = zs)"
- by (induct xs arbitrary: ys zs) (auto elim: append.cases intro: append.intros)
-
-lemmas lexordp_def =
- lexordp_def [unfolded lexord_def mem_Collect_eq split]
-(*>*)
-
-section {* Inductive Predicates \label{sec:inductive} *}
-
-text {*
- The @{text "predicate compiler"} is an extension of the code generator
- which turns inductive specifications into equational ones, from
- which in turn executable code can be generated. The mechanisms of
- this compiler are described in detail in
- \cite{Berghofer-Bulwahn-Haftmann:2009:TPHOL}.
-
- Consider the simple predicate @{const append} given by these two
- introduction rules:
-*}
-
-text %quote {*
- @{thm append.intros(1)[of ys]} \\
- @{thm append.intros(2)[of xs ys zs x]}
-*}
-
-text {*
- \noindent To invoke the compiler, simply use @{command_def "code_pred"}:
-*}
-
-code_pred %quote append .
-
-text {*
- \noindent The @{command "code_pred"} command takes the name of the
- inductive predicate and then you put a period to discharge a trivial
- correctness proof. The compiler infers possible modes for the
- predicate and produces the derived code equations. Modes annotate
- which (parts of the) arguments are to be taken as input, and which
- output. Modes are similar to types, but use the notation @{text "i"}
- for input and @{text "o"} for output.
-
- For @{term "append"}, the compiler can infer the following modes:
- \begin{itemize}
- \item @{text "i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool"}
- \item @{text "i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool"}
- \item @{text "o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool"}
- \end{itemize}
- You can compute sets of predicates using @{command_def "values"}:
-*}
-
-values %quote "{zs. append [(1::nat),2,3] [4,5] zs}"
-
-text {* \noindent outputs @{text "{[1, 2, 3, 4, 5]}"}, and *}
-
-values %quote "{(xs, ys). append xs ys [(2::nat),3]}"
-
-text {* \noindent outputs @{text "{([], [2, 3]), ([2], [3]), ([2, 3], [])}"}. *}
-
-text {*
- \noindent If you are only interested in the first elements of the
- set comprehension (with respect to a depth-first search on the
- introduction rules), you can pass an argument to @{command "values"}
- to specify the number of elements you want:
-*}
-
-values %quote 1 "{(xs, ys). append xs ys [(1::nat), 2, 3, 4]}"
-values %quote 3 "{(xs, ys). append xs ys [(1::nat), 2, 3, 4]}"
-
-text {*
- \noindent The @{command "values"} command can only compute set
- comprehensions for which a mode has been inferred.
-
- The code equations for a predicate are made available as theorems with
- the suffix @{text "equation"}, and can be inspected with:
-*}
-
-thm %quote append.equation
-
-text {*
- \noindent More advanced options are described in the following subsections.
-*}
-
-subsection {* Alternative names for functions *}
-
-text {*
- By default, the functions generated from a predicate are named after
- the predicate with the mode mangled into the name (e.g., @{text
- "append_i_i_o"}). You can specify your own names as follows:
-*}
-
-code_pred %quote (modes: i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool as concat,
- o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool as split,
- i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool as suffix) append .
-
-subsection {* Alternative introduction rules *}
-
-text {*
- Sometimes the introduction rules of an predicate are not executable
- because they contain non-executable constants or specific modes
- could not be inferred. It is also possible that the introduction
- rules yield a function that loops forever due to the execution in a
- depth-first search manner. Therefore, you can declare alternative
- introduction rules for predicates with the attribute @{attribute
- "code_pred_intro"}. For example, the transitive closure is defined
- by:
-*}
-
-text %quote {*
- @{lemma [source] "r a b \<Longrightarrow> tranclp r a b" by (fact tranclp.intros(1))}\\
- @{lemma [source] "tranclp r a b \<Longrightarrow> r b c \<Longrightarrow> tranclp r a c" by (fact tranclp.intros(2))}
-*}
-
-text {*
- \noindent These rules do not suit well for executing the transitive
- closure with the mode @{text "(i \<Rightarrow> o \<Rightarrow> bool) \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool"}, as
- the second rule will cause an infinite loop in the recursive call.
- This can be avoided using the following alternative rules which are
- declared to the predicate compiler by the attribute @{attribute
- "code_pred_intro"}:
-*}
-
-lemma %quote [code_pred_intro]:
- "r a b \<Longrightarrow> tranclp r a b"
- "r a b \<Longrightarrow> tranclp r b c \<Longrightarrow> tranclp r a c"
-by auto
-
-text {*
- \noindent After declaring all alternative rules for the transitive
- closure, you invoke @{command "code_pred"} as usual. As you have
- declared alternative rules for the predicate, you are urged to prove
- that these introduction rules are complete, i.e., that you can
- derive an elimination rule for the alternative rules:
-*}
-
-code_pred %quote tranclp
-proof -
- case tranclp
- from this converse_tranclpE [OF tranclp.prems] show thesis by metis
-qed
-
-text {*
- \noindent Alternative rules can also be used for constants that have
- not been defined inductively. For example, the lexicographic order
- which is defined as:
-*}
-
-text %quote {*
- @{thm [display] lexordp_def [of r]}
-*}
-
-text {*
- \noindent To make it executable, you can derive the following two
- rules and prove the elimination rule:
-*}
-
-lemma %quote [code_pred_intro]:
- "append xs (a # v) ys \<Longrightarrow> lexordp r xs ys"
-(*<*)unfolding lexordp_def by (auto simp add: append)(*>*)
-
-lemma %quote [code_pred_intro]:
- "append u (a # v) xs \<Longrightarrow> append u (b # w) ys \<Longrightarrow> r a b
- \<Longrightarrow> lexordp r xs ys"
-(*<*)unfolding lexordp_def append apply simp
-apply (rule disjI2) by auto(*>*)
-
-code_pred %quote lexordp
-(*<*)proof -
- fix r xs ys
- assume lexord: "lexordp r xs ys"
- assume 1: "\<And>r' xs' ys' a v. r = r' \<Longrightarrow> xs = xs' \<Longrightarrow> ys = ys'
- \<Longrightarrow> append xs' (a # v) ys' \<Longrightarrow> thesis"
- assume 2: "\<And>r' xs' ys' u a v b w. r = r' \<Longrightarrow> xs = xs' \<Longrightarrow> ys = ys'
- \<Longrightarrow> append u (a # v) xs' \<Longrightarrow> append u (b # w) ys' \<Longrightarrow> r' a b \<Longrightarrow> thesis"
- {
- assume "\<exists>a v. ys = xs @ a # v"
- from this 1 have thesis
- by (fastforce simp add: append)
- } moreover
- {
- assume "\<exists>u a b v w. r a b \<and> xs = u @ a # v \<and> ys = u @ b # w"
- from this 2 have thesis by (fastforce simp add: append)
- } moreover
- note lexord
- ultimately show thesis
- unfolding lexordp_def
- by fastforce
-qed(*>*)
-
-
-subsection {* Options for values *}
-
-text {*
- In the presence of higher-order predicates, multiple modes for some
- predicate could be inferred that are not disambiguated by the
- pattern of the set comprehension. To disambiguate the modes for the
- arguments of a predicate, you can state the modes explicitly in the
- @{command "values"} command. Consider the simple predicate @{term
- "succ"}:
-*}
-
-inductive %quote succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
- "succ 0 (Suc 0)"
-| "succ x y \<Longrightarrow> succ (Suc x) (Suc y)"
-
-code_pred %quote succ .
-
-text {*
- \noindent For this, the predicate compiler can infer modes @{text "o
- \<Rightarrow> o \<Rightarrow> bool"}, @{text "i \<Rightarrow> o \<Rightarrow> bool"}, @{text "o \<Rightarrow> i \<Rightarrow> bool"} and
- @{text "i \<Rightarrow> i \<Rightarrow> bool"}. The invocation of @{command "values"}
- @{text "{n. tranclp succ 10 n}"} loops, as multiple modes for the
- predicate @{text "succ"} are possible and here the first mode @{text
- "o \<Rightarrow> o \<Rightarrow> bool"} is chosen. To choose another mode for the argument,
- you can declare the mode for the argument between the @{command
- "values"} and the number of elements.
-*}
-
-values %quote [mode: i \<Rightarrow> o \<Rightarrow> bool] 1 "{n. tranclp succ 10 n}" (*FIMXE does not terminate for n\<ge>1*)
-values %quote [mode: o \<Rightarrow> i \<Rightarrow> bool] 1 "{n. tranclp succ n 10}"
-
-
-subsection {* Embedding into functional code within Isabelle/HOL *}
-
-text {*
- To embed the computation of an inductive predicate into functions
- that are defined in Isabelle/HOL, you have a number of options:
-
- \begin{itemize}
-
- \item You want to use the first-order predicate with the mode
- where all arguments are input. Then you can use the predicate directly, e.g.
-
- \begin{quote}
- @{text "valid_suffix ys zs = "} \\
- @{text "(if append [Suc 0, 2] ys zs then Some ys else None)"}
- \end{quote}
-
- \item If you know that the execution returns only one value (it is
- deterministic), then you can use the combinator @{term
- "Predicate.the"}, e.g., a functional concatenation of lists is
- defined with
-
- \begin{quote}
- @{term "functional_concat xs ys = Predicate.the (append_i_i_o xs ys)"}
- \end{quote}
-
- Note that if the evaluation does not return a unique value, it
- raises a run-time error @{term "not_unique"}.
-
- \end{itemize}
-*}
-
-
-subsection {* Further Examples *}
-
-text {*
- Further examples for compiling inductive predicates can be found in
- the @{text "HOL/ex/Predicate_Compile_ex.thy"} theory file. There are
- also some examples in the Archive of Formal Proofs, notably in the
- @{text "POPLmark-deBruijn"} and the @{text "FeatherweightJava"}
- sessions.
-*}
-
-end
-
--- a/doc-src/Codegen/Thy/Introduction.thy Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,242 +0,0 @@
-theory Introduction
-imports Setup
-begin
-
-section {* Introduction *}
-
-text {*
- This tutorial introduces the code generator facilities of @{text
- "Isabelle/HOL"}. It allows to turn (a certain class of) HOL
- specifications into corresponding executable code in the programming
- languages @{text SML} \cite{SML}, @{text OCaml} \cite{OCaml},
- @{text Haskell} \cite{haskell-revised-report} and @{text Scala}
- \cite{scala-overview-tech-report}.
-
- To profit from this tutorial, some familiarity and experience with
- @{theory HOL} \cite{isa-tutorial} and its basic theories is assumed.
-*}
-
-
-subsection {* Code generation principle: shallow embedding \label{sec:principle} *}
-
-text {*
- The key concept for understanding Isabelle's code generation is
- \emph{shallow embedding}: logical entities like constants, types and
- classes are identified with corresponding entities in the target
- language. In particular, the carrier of a generated program's
- semantics are \emph{equational theorems} from the logic. If we view
- a generated program as an implementation of a higher-order rewrite
- system, then every rewrite step performed by the program can be
- simulated in the logic, which guarantees partial correctness
- \cite{Haftmann-Nipkow:2010:code}.
-*}
-
-
-subsection {* A quick start with the Isabelle/HOL toolbox \label{sec:queue_example} *}
-
-text {*
- In a HOL theory, the @{command_def datatype} and @{command_def
- definition}/@{command_def primrec}/@{command_def fun} declarations
- form the core of a functional programming language. By default
- equational theorems stemming from those are used for generated code,
- therefore \qt{naive} code generation can proceed without further
- ado.
-
- For example, here a simple \qt{implementation} of amortised queues:
-*}
-
-datatype %quote 'a queue = AQueue "'a list" "'a list"
-
-definition %quote empty :: "'a queue" where
- "empty = AQueue [] []"
-
-primrec %quote enqueue :: "'a \<Rightarrow> 'a queue \<Rightarrow> 'a queue" where
- "enqueue x (AQueue xs ys) = AQueue (x # xs) ys"
-
-fun %quote dequeue :: "'a queue \<Rightarrow> 'a option \<times> 'a queue" where
- "dequeue (AQueue [] []) = (None, AQueue [] [])"
- | "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)"
- | "dequeue (AQueue xs []) =
- (case rev xs of y # ys \<Rightarrow> (Some y, AQueue [] ys))" (*<*)
-
-lemma %invisible dequeue_nonempty_Nil [simp]:
- "xs \<noteq> [] \<Longrightarrow> dequeue (AQueue xs []) = (case rev xs of y # ys \<Rightarrow> (Some y, AQueue [] ys))"
- by (cases xs) (simp_all split: list.splits) (*>*)
-
-text {* \noindent Then we can generate code e.g.~for @{text SML} as follows: *}
-
-export_code %quote empty dequeue enqueue in SML
- module_name Example file "examples/example.ML"
-
-text {* \noindent resulting in the following code: *}
-
-text %quotetypewriter {*
- @{code_stmts empty enqueue dequeue (SML)}
-*}
-
-text {*
- \noindent The @{command_def export_code} command takes a
- space-separated list of constants for which code shall be generated;
- anything else needed for those is added implicitly. Then follows a
- target language identifier and a freely chosen module name. A file
- name denotes the destination to store the generated code. Note that
- the semantics of the destination depends on the target language: for
- @{text SML}, @{text OCaml} and @{text Scala} it denotes a \emph{file},
- for @{text Haskell} it denotes a \emph{directory} where a file named as the
- module name (with extension @{text ".hs"}) is written:
-*}
-
-export_code %quote empty dequeue enqueue in Haskell
- module_name Example file "examples/"
-
-text {*
- \noindent This is the corresponding code:
-*}
-
-text %quotetypewriter {*
- @{code_stmts empty enqueue dequeue (Haskell)}
-*}
-
-text {*
- \noindent For more details about @{command export_code} see
- \secref{sec:further}.
-*}
-
-
-subsection {* Type classes *}
-
-text {*
- Code can also be generated from type classes in a Haskell-like
- manner. For illustration here an example from abstract algebra:
-*}
-
-class %quote semigroup =
- fixes mult :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<otimes>" 70)
- assumes assoc: "(x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"
-
-class %quote monoid = semigroup +
- fixes neutral :: 'a ("\<one>")
- assumes neutl: "\<one> \<otimes> x = x"
- and neutr: "x \<otimes> \<one> = x"
-
-instantiation %quote nat :: monoid
-begin
-
-primrec %quote mult_nat where
- "0 \<otimes> n = (0\<Colon>nat)"
- | "Suc m \<otimes> n = n + m \<otimes> n"
-
-definition %quote neutral_nat where
- "\<one> = Suc 0"
-
-lemma %quote add_mult_distrib:
- fixes n m q :: nat
- shows "(n + m) \<otimes> q = n \<otimes> q + m \<otimes> q"
- by (induct n) simp_all
-
-instance %quote proof
- fix m n q :: nat
- show "m \<otimes> n \<otimes> q = m \<otimes> (n \<otimes> q)"
- by (induct m) (simp_all add: add_mult_distrib)
- show "\<one> \<otimes> n = n"
- by (simp add: neutral_nat_def)
- show "m \<otimes> \<one> = m"
- by (induct m) (simp_all add: neutral_nat_def)
-qed
-
-end %quote
-
-text {*
- \noindent We define the natural operation of the natural numbers
- on monoids:
-*}
-
-primrec %quote (in monoid) pow :: "nat \<Rightarrow> 'a \<Rightarrow> 'a" where
- "pow 0 a = \<one>"
- | "pow (Suc n) a = a \<otimes> pow n a"
-
-text {*
- \noindent This we use to define the discrete exponentiation
- function:
-*}
-
-definition %quote bexp :: "nat \<Rightarrow> nat" where
- "bexp n = pow n (Suc (Suc 0))"
-
-text {*
- \noindent The corresponding code in Haskell uses that language's
- native classes:
-*}
-
-text %quotetypewriter {*
- @{code_stmts bexp (Haskell)}
-*}
-
-text {*
- \noindent This is a convenient place to show how explicit dictionary
- construction manifests in generated code -- the same example in
- @{text SML}:
-*}
-
-text %quotetypewriter {*
- @{code_stmts bexp (SML)}
-*}
-
-text {*
- \noindent Note the parameters with trailing underscore (@{verbatim
- "A_"}), which are the dictionary parameters.
-*}
-
-
-subsection {* How to continue from here *}
-
-text {*
- What you have seen so far should be already enough in a lot of
- cases. If you are content with this, you can quit reading here.
-
- Anyway, to understand situations where problems occur or to increase
- the scope of code generation beyond default, it is necessary to gain
- some understanding how the code generator actually works:
-
- \begin{itemize}
-
- \item The foundations of the code generator are described in
- \secref{sec:foundations}.
-
- \item In particular \secref{sec:utterly_wrong} gives hints how to
- debug situations where code generation does not succeed as
- expected.
-
- \item The scope and quality of generated code can be increased
- dramatically by applying refinement techniques, which are
- introduced in \secref{sec:refinement}.
-
- \item Inductive predicates can be turned executable using an
- extension of the code generator \secref{sec:inductive}.
-
- \item If you want to utilize code generation to obtain fast
- evaluators e.g.~for decision procedures, have a look at
- \secref{sec:evaluation}.
-
- \item You may want to skim over the more technical sections
- \secref{sec:adaptation} and \secref{sec:further}.
-
- \item The target language Scala \cite{scala-overview-tech-report}
- comes with some specialities discussed in \secref{sec:scala}.
-
- \item For exhaustive syntax diagrams etc. you should visit the
- Isabelle/Isar Reference Manual \cite{isabelle-isar-ref}.
-
- \end{itemize}
-
- \bigskip
-
- \begin{center}\fbox{\fbox{\begin{minipage}{8cm}
-
- \begin{center}\textit{Happy proving, happy hacking!}\end{center}
-
- \end{minipage}}}\end{center}
-*}
-
-end
-
--- a/doc-src/Codegen/Thy/Refinement.thy Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,274 +0,0 @@
-theory Refinement
-imports Setup
-begin
-
-section {* Program and datatype refinement \label{sec:refinement} *}
-
-text {*
- Code generation by shallow embedding (cf.~\secref{sec:principle})
- allows to choose code equations and datatype constructors freely,
- given that some very basic syntactic properties are met; this
- flexibility opens up mechanisms for refinement which allow to extend
- the scope and quality of generated code dramatically.
-*}
-
-
-subsection {* Program refinement *}
-
-text {*
- Program refinement works by choosing appropriate code equations
- explicitly (cf.~\secref{sec:equations}); as example, we use Fibonacci
- numbers:
-*}
-
-fun %quote fib :: "nat \<Rightarrow> nat" where
- "fib 0 = 0"
- | "fib (Suc 0) = Suc 0"
- | "fib (Suc (Suc n)) = fib n + fib (Suc n)"
-
-text {*
- \noindent The runtime of the corresponding code grows exponential due
- to two recursive calls:
-*}
-
-text %quotetypewriter {*
- @{code_stmts fib (consts) fib (Haskell)}
-*}
-
-text {*
- \noindent A more efficient implementation would use dynamic
- programming, e.g.~sharing of common intermediate results between
- recursive calls. This idea is expressed by an auxiliary operation
- which computes a Fibonacci number and its successor simultaneously:
-*}
-
-definition %quote fib_step :: "nat \<Rightarrow> nat \<times> nat" where
- "fib_step n = (fib (Suc n), fib n)"
-
-text {*
- \noindent This operation can be implemented by recursion using
- dynamic programming:
-*}
-
-lemma %quote [code]:
- "fib_step 0 = (Suc 0, 0)"
- "fib_step (Suc n) = (let (m, q) = fib_step n in (m + q, m))"
- by (simp_all add: fib_step_def)
-
-text {*
- \noindent What remains is to implement @{const fib} by @{const
- fib_step} as follows:
-*}
-
-lemma %quote [code]:
- "fib 0 = 0"
- "fib (Suc n) = fst (fib_step n)"
- by (simp_all add: fib_step_def)
-
-text {*
- \noindent The resulting code shows only linear growth of runtime:
-*}
-
-text %quotetypewriter {*
- @{code_stmts fib (consts) fib fib_step (Haskell)}
-*}
-
-
-subsection {* Datatype refinement *}
-
-text {*
- Selecting specific code equations \emph{and} datatype constructors
- leads to datatype refinement. As an example, we will develop an
- alternative representation of the queue example given in
- \secref{sec:queue_example}. The amortised representation is
- convenient for generating code but exposes its \qt{implementation}
- details, which may be cumbersome when proving theorems about it.
- Therefore, here is a simple, straightforward representation of
- queues:
-*}
-
-datatype %quote 'a queue = Queue "'a list"
-
-definition %quote empty :: "'a queue" where
- "empty = Queue []"
-
-primrec %quote enqueue :: "'a \<Rightarrow> 'a queue \<Rightarrow> 'a queue" where
- "enqueue x (Queue xs) = Queue (xs @ [x])"
-
-fun %quote dequeue :: "'a queue \<Rightarrow> 'a option \<times> 'a queue" where
- "dequeue (Queue []) = (None, Queue [])"
- | "dequeue (Queue (x # xs)) = (Some x, Queue xs)"
-
-text {*
- \noindent This we can use directly for proving; for executing,
- we provide an alternative characterisation:
-*}
-
-definition %quote AQueue :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a queue" where
- "AQueue xs ys = Queue (ys @ rev xs)"
-
-code_datatype %quote AQueue
-
-text {*
- \noindent Here we define a \qt{constructor} @{const "AQueue"} which
- is defined in terms of @{text "Queue"} and interprets its arguments
- according to what the \emph{content} of an amortised queue is supposed
- to be.
-
- The prerequisite for datatype constructors is only syntactical: a
- constructor must be of type @{text "\<tau> = \<dots> \<Rightarrow> \<kappa> \<alpha>\<^isub>1 \<dots> \<alpha>\<^isub>n"} where @{text
- "{\<alpha>\<^isub>1, \<dots>, \<alpha>\<^isub>n}"} is exactly the set of \emph{all} type variables in
- @{text "\<tau>"}; then @{text "\<kappa>"} is its corresponding datatype. The
- HOL datatype package by default registers any new datatype with its
- constructors, but this may be changed using @{command_def
- code_datatype}; the currently chosen constructors can be inspected
- using the @{command print_codesetup} command.
-
- Equipped with this, we are able to prove the following equations
- for our primitive queue operations which \qt{implement} the simple
- queues in an amortised fashion:
-*}
-
-lemma %quote empty_AQueue [code]:
- "empty = AQueue [] []"
- by (simp add: AQueue_def empty_def)
-
-lemma %quote enqueue_AQueue [code]:
- "enqueue x (AQueue xs ys) = AQueue (x # xs) ys"
- by (simp add: AQueue_def)
-
-lemma %quote dequeue_AQueue [code]:
- "dequeue (AQueue xs []) =
- (if xs = [] then (None, AQueue [] [])
- else dequeue (AQueue [] (rev xs)))"
- "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)"
- by (simp_all add: AQueue_def)
-
-text {*
- \noindent It is good style, although no absolute requirement, to
- provide code equations for the original artefacts of the implemented
- type, if possible; in our case, these are the datatype constructor
- @{const Queue} and the case combinator @{const queue_case}:
-*}
-
-lemma %quote Queue_AQueue [code]:
- "Queue = AQueue []"
- by (simp add: AQueue_def fun_eq_iff)
-
-lemma %quote queue_case_AQueue [code]:
- "queue_case f (AQueue xs ys) = f (ys @ rev xs)"
- by (simp add: AQueue_def)
-
-text {*
- \noindent The resulting code looks as expected:
-*}
-
-text %quotetypewriter {*
- @{code_stmts empty enqueue dequeue Queue queue_case (SML)}
-*}
-
-text {*
- The same techniques can also be applied to types which are not
- specified as datatypes, e.g.~type @{typ int} is originally specified
- as quotient type by means of @{command_def typedef}, but for code
- generation constants allowing construction of binary numeral values
- are used as constructors for @{typ int}.
-
- This approach however fails if the representation of a type demands
- invariants; this issue is discussed in the next section.
-*}
-
-
-subsection {* Datatype refinement involving invariants \label{sec:invariant} *}
-
-text {*
- Datatype representation involving invariants require a dedicated
- setup for the type and its primitive operations. As a running
- example, we implement a type @{text "'a dlist"} of list consisting
- of distinct elements.
-
- The first step is to decide on which representation the abstract
- type (in our example @{text "'a dlist"}) should be implemented.
- Here we choose @{text "'a list"}. Then a conversion from the concrete
- type to the abstract type must be specified, here:
-*}
-
-text %quote {*
- @{term_type Dlist}
-*}
-
-text {*
- \noindent Next follows the specification of a suitable \emph{projection},
- i.e.~a conversion from abstract to concrete type:
-*}
-
-text %quote {*
- @{term_type list_of_dlist}
-*}
-
-text {*
- \noindent This projection must be specified such that the following
- \emph{abstract datatype certificate} can be proven:
-*}
-
-lemma %quote [code abstype]:
- "Dlist (list_of_dlist dxs) = dxs"
- by (fact Dlist_list_of_dlist)
-
-text {*
- \noindent Note that so far the invariant on representations
- (@{term_type distinct}) has never been mentioned explicitly:
- the invariant is only referred to implicitly: all values in
- set @{term "{xs. list_of_dlist (Dlist xs) = xs}"} are invariant,
- and in our example this is exactly @{term "{xs. distinct xs}"}.
-
- The primitive operations on @{typ "'a dlist"} are specified
- indirectly using the projection @{const list_of_dlist}. For
- the empty @{text "dlist"}, @{const Dlist.empty}, we finally want
- the code equation
-*}
-
-text %quote {*
- @{term "Dlist.empty = Dlist []"}
-*}
-
-text {*
- \noindent This we have to prove indirectly as follows:
-*}
-
-lemma %quote [code abstract]:
- "list_of_dlist Dlist.empty = []"
- by (fact list_of_dlist_empty)
-
-text {*
- \noindent This equation logically encodes both the desired code
- equation and that the expression @{const Dlist} is applied to obeys
- the implicit invariant. Equations for insertion and removal are
- similar:
-*}
-
-lemma %quote [code abstract]:
- "list_of_dlist (Dlist.insert x dxs) = List.insert x (list_of_dlist dxs)"
- by (fact list_of_dlist_insert)
-
-lemma %quote [code abstract]:
- "list_of_dlist (Dlist.remove x dxs) = remove1 x (list_of_dlist dxs)"
- by (fact list_of_dlist_remove)
-
-text {*
- \noindent Then the corresponding code is as follows:
-*}
-
-text %quotetypewriter {*
- @{code_stmts Dlist.empty Dlist.insert Dlist.remove list_of_dlist (Haskell)}
-*} (*(types) dlist (consts) dempty dinsert dremove list_of List.member insert remove *)
-
-text {*
- Typical data structures implemented by representations involving
- invariants are available in the library, theory @{theory Mapping}
- specifies key-value-mappings (type @{typ "('a, 'b) mapping"});
- these can be implemented by red-black-trees (theory @{theory RBT}).
-*}
-
-end
-
--- a/doc-src/Codegen/Thy/Setup.thy Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,32 +0,0 @@
-theory Setup
-imports
- Complex_Main
- "~~/src/HOL/Library/Dlist"
- "~~/src/HOL/Library/RBT"
- "~~/src/HOL/Library/Mapping"
-begin
-
-ML_file "../../antiquote_setup.ML"
-ML_file "../../more_antiquote.ML"
-
-setup {*
- Antiquote_Setup.setup #>
- More_Antiquote.setup #>
-let
- val typ = Simple_Syntax.read_typ;
-in
- Sign.del_modesyntax_i (Symbol.xsymbolsN, false)
- [("_constrain", typ "logic => type => logic", Mixfix ("_\<Colon>_", [4, 0], 3)),
- ("_constrain", typ "prop' => type => prop'", Mixfix ("_\<Colon>_", [4, 0], 3))] #>
- Sign.add_modesyntax_i (Symbol.xsymbolsN, false)
- [("_constrain", typ "logic => type => logic", Mixfix ("_ \<Colon> _", [4, 0], 3)),
- ("_constrain", typ "prop' => type => prop'", Mixfix ("_ \<Colon> _", [4, 0], 3))]
-end
-*}
-
-setup {* Code_Target.set_default_code_width 74 *}
-
-declare [[names_unique = false]]
-
-end
-
--- a/doc-src/Codegen/Thy/document/Adaptation.tex Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,651 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{Adaptation}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-\isacommand{theory}\isamarkupfalse%
-\ Adaptation\isanewline
-\isakeyword{imports}\ Setup\isanewline
-\isakeyword{begin}%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-\isanewline
-%
-\endisadelimtheory
-%
-\isadeliminvisible
-\isanewline
-%
-\endisadeliminvisible
-%
-\isataginvisible
-\isacommand{setup}\isamarkupfalse%
-\ {\isaliteral{7B2A}{\isacharverbatimopen}}\ Code{\isaliteral{5F}{\isacharunderscore}}Target{\isaliteral{2E}{\isachardot}}extend{\isaliteral{5F}{\isacharunderscore}}target\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C534D4C3E}{\isasymSML}}{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequote}}SML{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2C}{\isacharcomma}}\ K\ I{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ {\isaliteral{23}{\isacharhash}}{\isaliteral{3E}{\isachargreater}}\ Code{\isaliteral{5F}{\isacharunderscore}}Target{\isaliteral{2E}{\isachardot}}extend{\isaliteral{5F}{\isacharunderscore}}target\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C534D4C64756D6D793E}{\isasymSMLdummy}}{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequote}}Haskell{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2C}{\isacharcomma}}\ K\ I{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2A7D}{\isacharverbatimclose}}%
-\endisataginvisible
-{\isafoldinvisible}%
-%
-\isadeliminvisible
-%
-\endisadeliminvisible
-%
-\isamarkupsection{Adaptation to target languages \label{sec:adaptation}%
-}
-\isamarkuptrue%
-%
-\isamarkupsubsection{Adapting code generation%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-The aspects of code generation introduced so far have two aspects
- in common:
-
- \begin{itemize}
-
- \item They act uniformly, without reference to a specific target
- language.
-
- \item They are \emph{safe} in the sense that as long as you trust
- the code generator meta theory and implementation, you cannot
- produce programs that yield results which are not derivable in
- the logic.
-
- \end{itemize}
-
- \noindent In this section we will introduce means to \emph{adapt}
- the serialiser to a specific target language, i.e.~to print program
- fragments in a way which accommodates \qt{already existing}
- ingredients of a target language environment, for three reasons:
-
- \begin{itemize}
- \item improving readability and aesthetics of generated code
- \item gaining efficiency
- \item interface with language parts which have no direct counterpart
- in \isa{HOL} (say, imperative data structures)
- \end{itemize}
-
- \noindent Generally, you should avoid using those features yourself
- \emph{at any cost}:
-
- \begin{itemize}
-
- \item The safe configuration methods act uniformly on every target
- language, whereas for adaptation you have to treat each target
- language separately.
-
- \item Application is extremely tedious since there is no
- abstraction which would allow for a static check, making it easy
- to produce garbage.
-
- \item Subtle errors can be introduced unconsciously.
-
- \end{itemize}
-
- \noindent However, even if you ought refrain from setting up
- adaptation yourself, already the \isa{HOL} comes with some
- reasonable default adaptations (say, using target language list
- syntax). There also some common adaptation cases which you can
- setup by importing particular library theories. In order to
- understand these, we provide some clues here; these however are not
- supposed to replace a careful study of the sources.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{The adaptation principle%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Figure \ref{fig:adaptation} illustrates what \qt{adaptation} is
- conceptually supposed to be:
-
- \begin{figure}[here]
- \includegraphics{adaptation}
- \caption{The adaptation principle}
- \label{fig:adaptation}
- \end{figure}
-
- \noindent In the tame view, code generation acts as broker between
- \isa{logic}, \isa{intermediate\ language} and \isa{target\ language} by means of \isa{translation} and \isa{serialisation}; for the latter, the serialiser has to observe the
- structure of the \isa{language} itself plus some \isa{reserved}
- keywords which have to be avoided for generated code. However, if
- you consider \isa{adaptation} mechanisms, the code generated by
- the serializer is just the tip of the iceberg:
-
- \begin{itemize}
-
- \item \isa{serialisation} can be \emph{parametrised} such that
- logical entities are mapped to target-specific ones
- (e.g. target-specific list syntax, see also
- \secref{sec:adaptation_mechanisms})
-
- \item Such parametrisations can involve references to a
- target-specific standard \isa{library} (e.g. using the \isa{Haskell} \verb|Maybe| type instead of the \isa{HOL}
- \isa{option} type); if such are used, the corresponding
- identifiers (in our example, \verb|Maybe|, \verb|Nothing| and \verb|Just|) also have to be considered \isa{reserved}.
-
- \item Even more, the user can enrich the library of the
- target-language by providing code snippets (\qt{\isa{includes}}) which are prepended to any generated code (see
- \secref{sec:include}); this typically also involves further
- \isa{reserved} identifiers.
-
- \end{itemize}
-
- \noindent As figure \ref{fig:adaptation} illustrates, all these
- adaptation mechanisms have to act consistently; it is at the
- discretion of the user to take care for this.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Common adaptation patterns%
-}
-\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 customising the code
- generator setup (see below \secref{sec:adaptation_mechanisms}).
-
- \begin{description}
-
- \item[\isa{Code{\isaliteral{5F}{\isacharunderscore}}Integer}] represents \isa{HOL} integers by
- big integer literals in target languages.
-
- \item[\isa{Code{\isaliteral{5F}{\isacharunderscore}}Char}] represents \isa{HOL} characters by
- character literals in target languages.
-
- \item[\isa{Code{\isaliteral{5F}{\isacharunderscore}}Char{\isaliteral{5F}{\isacharunderscore}}chr}] like \isa{Code{\isaliteral{5F}{\isacharunderscore}}Char}, but
- also offers treatment of character codes; includes \isa{Code{\isaliteral{5F}{\isacharunderscore}}Char}.
-
- \item[\isa{Efficient{\isaliteral{5F}{\isacharunderscore}}Nat}] \label{eff_nat} implements
- natural numbers by integers, which in general will result in
- higher efficiency; pattern matching with \isa{{\isadigit{0}}} /
- \isa{Suc} is eliminated; includes \isa{Code{\isaliteral{5F}{\isacharunderscore}}Integer}
- and \isa{Code{\isaliteral{5F}{\isacharunderscore}}Numeral}.
-
- \item[\isa{Code{\isaliteral{5F}{\isacharunderscore}}Numeral}] 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. Part of \isa{HOL{\isaliteral{2D}{\isacharminus}}Main}.
-
- \item[\isa{String}] provides an additional datatype \isa{String{\isaliteral{2E}{\isachardot}}literal} which is isomorphic to strings; \isa{String{\isaliteral{2E}{\isachardot}}literal}s are mapped to target-language strings. Useful
- for code setups which involve e.g.~printing (error) messages.
- Part of \isa{HOL{\isaliteral{2D}{\isacharminus}}Main}.
-
- \end{description}
-
- \begin{warn}
- When importing any of those theories which are not part of
- \isa{HOL{\isaliteral{2D}{\isacharminus}}Main}, 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{Parametrising serialisation \label{sec:adaptation_mechanisms}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Consider the following function and its corresponding SML code:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{primrec}\isamarkupfalse%
-\ in{\isaliteral{5F}{\isacharunderscore}}interval\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}nat\ {\isaliteral{5C3C74696D65733E}{\isasymtimes}}\ nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}in{\isaliteral{5F}{\isacharunderscore}}interval\ {\isaliteral{28}{\isacharparenleft}}k{\isaliteral{2C}{\isacharcomma}}\ l{\isaliteral{29}{\isacharparenright}}\ n\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ k\ {\isaliteral{5C3C6C653E}{\isasymle}}\ n\ {\isaliteral{5C3C616E643E}{\isasymand}}\ n\ {\isaliteral{5C3C6C653E}{\isasymle}}\ l{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isadeliminvisible
-%
-\endisadeliminvisible
-%
-\isataginvisible
-%
-\endisataginvisible
-{\isafoldinvisible}%
-%
-\isadeliminvisible
-%
-\endisadeliminvisible
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-structure\ Example\ {\isaliteral{3A}{\isacharcolon}}\ sig\isanewline
-\ \ datatype\ nat\ {\isaliteral{3D}{\isacharequal}}\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{7C}{\isacharbar}}\ Suc\ of\ nat\isanewline
-\ \ datatype\ boola\ {\isaliteral{3D}{\isacharequal}}\ True\ {\isaliteral{7C}{\isacharbar}}\ False\isanewline
-\ \ val\ conj\ {\isaliteral{3A}{\isacharcolon}}\ boola\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ boola\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ boola\isanewline
-\ \ val\ less{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3A}{\isacharcolon}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ boola\isanewline
-\ \ val\ less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3A}{\isacharcolon}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ boola\isanewline
-\ \ val\ in{\isaliteral{5F}{\isacharunderscore}}interval\ {\isaliteral{3A}{\isacharcolon}}\ nat\ {\isaliteral{2A}{\isacharasterisk}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ boola\isanewline
-end\ {\isaliteral{3D}{\isacharequal}}\ struct\isanewline
-\isanewline
-datatype\ nat\ {\isaliteral{3D}{\isacharequal}}\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{7C}{\isacharbar}}\ Suc\ of\ nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-datatype\ boola\ {\isaliteral{3D}{\isacharequal}}\ True\ {\isaliteral{7C}{\isacharbar}}\ False{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ conj\ p\ True\ {\isaliteral{3D}{\isacharequal}}\ p\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ conj\ p\ False\ {\isaliteral{3D}{\isacharequal}}\ False\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ conj\ True\ p\ {\isaliteral{3D}{\isacharequal}}\ p\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ conj\ False\ p\ {\isaliteral{3D}{\isacharequal}}\ False{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ less{\isaliteral{5F}{\isacharunderscore}}nat\ m\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ m\ n\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ less{\isaliteral{5F}{\isacharunderscore}}nat\ n\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3D}{\isacharequal}}\ False\isanewline
-and\ less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{28}{\isacharparenleft}}Suc\ m{\isaliteral{29}{\isacharparenright}}\ n\ {\isaliteral{3D}{\isacharequal}}\ less{\isaliteral{5F}{\isacharunderscore}}nat\ m\ n\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ n\ {\isaliteral{3D}{\isacharequal}}\ True{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ in{\isaliteral{5F}{\isacharunderscore}}interval\ {\isaliteral{28}{\isacharparenleft}}k{\isaliteral{2C}{\isacharcomma}}\ l{\isaliteral{29}{\isacharparenright}}\ n\ {\isaliteral{3D}{\isacharequal}}\ conj\ {\isaliteral{28}{\isacharparenleft}}less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ k\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ n\ l{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-end{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}struct\ Example{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\begin{isamarkuptext}%
-\noindent Though this is correct code, it is a little bit
- unsatisfactory: boolean values and operators are materialised as
- distinguished entities with have nothing to do with the SML-built-in
- 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 \verb|bool| would be used. To map the HOL \isa{bool} on SML \verb|bool|, we may use \qn{custom serialisations}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotett
-%
-\endisadelimquotett
-%
-\isatagquotett
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}type}\isamarkupfalse%
-\ bool\isanewline
-\ \ {\isaliteral{28}{\isacharparenleft}}SML\ {\isaliteral{22}{\isachardoublequoteopen}}bool{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}const}\isamarkupfalse%
-\ True\ \isakeyword{and}\ False\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}op\ {\isaliteral{5C3C616E643E}{\isasymand}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ {\isaliteral{28}{\isacharparenleft}}SML\ {\isaliteral{22}{\isachardoublequoteopen}}true{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}false{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5F}{\isacharunderscore}}\ andalso\ {\isaliteral{5F}{\isacharunderscore}}{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}%
-\endisatagquotett
-{\isafoldquotett}%
-%
-\isadelimquotett
-%
-\endisadelimquotett
-%
-\begin{isamarkuptext}%
-\noindent The \indexdef{}{command}{code\_type}\hypertarget{command.code-type}{\hyperlink{command.code-type}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}type}}}}} command takes a type constructor
- as arguments together with a list of custom serialisations. Each
- custom serialisation starts with a target language identifier
- followed by an expression, which during code serialisation is
- inserted whenever the type constructor would occur. For constants,
- \indexdef{}{command}{code\_const}\hypertarget{command.code-const}{\hyperlink{command.code-const}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}const}}}}} implements the corresponding mechanism. Each
- ``\verb|_|'' in a serialisation expression is treated as a
- placeholder for the type constructor's (the constant's) arguments.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-structure\ Example\ {\isaliteral{3A}{\isacharcolon}}\ sig\isanewline
-\ \ datatype\ nat\ {\isaliteral{3D}{\isacharequal}}\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{7C}{\isacharbar}}\ Suc\ of\ nat\isanewline
-\ \ val\ less{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3A}{\isacharcolon}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ bool\isanewline
-\ \ val\ less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3A}{\isacharcolon}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ bool\isanewline
-\ \ val\ in{\isaliteral{5F}{\isacharunderscore}}interval\ {\isaliteral{3A}{\isacharcolon}}\ nat\ {\isaliteral{2A}{\isacharasterisk}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ bool\isanewline
-end\ {\isaliteral{3D}{\isacharequal}}\ struct\isanewline
-\isanewline
-datatype\ nat\ {\isaliteral{3D}{\isacharequal}}\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{7C}{\isacharbar}}\ Suc\ of\ nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ less{\isaliteral{5F}{\isacharunderscore}}nat\ m\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ m\ n\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ less{\isaliteral{5F}{\isacharunderscore}}nat\ n\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3D}{\isacharequal}}\ false\isanewline
-and\ less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{28}{\isacharparenleft}}Suc\ m{\isaliteral{29}{\isacharparenright}}\ n\ {\isaliteral{3D}{\isacharequal}}\ less{\isaliteral{5F}{\isacharunderscore}}nat\ m\ n\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ n\ {\isaliteral{3D}{\isacharequal}}\ true{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ in{\isaliteral{5F}{\isacharunderscore}}interval\ {\isaliteral{28}{\isacharparenleft}}k{\isaliteral{2C}{\isacharcomma}}\ l{\isaliteral{29}{\isacharparenright}}\ n\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ k\ n{\isaliteral{29}{\isacharparenright}}\ andalso\ {\isaliteral{28}{\isacharparenleft}}less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ n\ l{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-end{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}struct\ Example{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\begin{isamarkuptext}%
-\noindent This still is not perfect: the parentheses around the
- \qt{andalso} expression are superfluous. Though the serialiser 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%
-%
-\isadelimquotett
-%
-\endisadelimquotett
-%
-\isatagquotett
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}const}\isamarkupfalse%
-\ {\isaliteral{22}{\isachardoublequoteopen}}op\ {\isaliteral{5C3C616E643E}{\isasymand}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ {\isaliteral{28}{\isacharparenleft}}SML\ \isakeyword{infixl}\ {\isadigit{1}}\ {\isaliteral{22}{\isachardoublequoteopen}}andalso{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}%
-\endisatagquotett
-{\isafoldquotett}%
-%
-\isadelimquotett
-%
-\endisadelimquotett
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-structure\ Example\ {\isaliteral{3A}{\isacharcolon}}\ sig\isanewline
-\ \ datatype\ nat\ {\isaliteral{3D}{\isacharequal}}\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{7C}{\isacharbar}}\ Suc\ of\ nat\isanewline
-\ \ val\ less{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3A}{\isacharcolon}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ bool\isanewline
-\ \ val\ less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3A}{\isacharcolon}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ bool\isanewline
-\ \ val\ in{\isaliteral{5F}{\isacharunderscore}}interval\ {\isaliteral{3A}{\isacharcolon}}\ nat\ {\isaliteral{2A}{\isacharasterisk}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ bool\isanewline
-end\ {\isaliteral{3D}{\isacharequal}}\ struct\isanewline
-\isanewline
-datatype\ nat\ {\isaliteral{3D}{\isacharequal}}\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{7C}{\isacharbar}}\ Suc\ of\ nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ less{\isaliteral{5F}{\isacharunderscore}}nat\ m\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ m\ n\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ less{\isaliteral{5F}{\isacharunderscore}}nat\ n\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3D}{\isacharequal}}\ false\isanewline
-and\ less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{28}{\isacharparenleft}}Suc\ m{\isaliteral{29}{\isacharparenright}}\ n\ {\isaliteral{3D}{\isacharequal}}\ less{\isaliteral{5F}{\isacharunderscore}}nat\ m\ n\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ n\ {\isaliteral{3D}{\isacharequal}}\ true{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ in{\isaliteral{5F}{\isacharunderscore}}interval\ {\isaliteral{28}{\isacharparenleft}}k{\isaliteral{2C}{\isacharcomma}}\ l{\isaliteral{29}{\isacharparenright}}\ n\ {\isaliteral{3D}{\isacharequal}}\ less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ k\ n\ andalso\ less{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}nat\ n\ l{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-end{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}struct\ Example{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\begin{isamarkuptext}%
-\noindent The attentive reader may ask how we assert that no
- generated code will accidentally overwrite. For this reason the
- serialiser has an internal table of identifiers which have to be
- avoided to be used for new declarations. Initially, this table
- typically contains the keywords of the target language. It can be
- extended manually, thus avoiding accidental overwrites, using the
- \indexdef{}{command}{code\_reserved}\hypertarget{command.code-reserved}{\hyperlink{command.code-reserved}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reserved}}}}} command:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reserved}\isamarkupfalse%
-\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C534D4C64756D6D793E}{\isasymSMLdummy}}{\isaliteral{22}{\isachardoublequoteclose}}\ bool\ true\ false\ andalso%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent Next, we try to map HOL pairs to SML pairs, using the
- infix ``\verb|*|'' type constructor and parentheses:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadeliminvisible
-%
-\endisadeliminvisible
-%
-\isataginvisible
-%
-\endisataginvisible
-{\isafoldinvisible}%
-%
-\isadeliminvisible
-%
-\endisadeliminvisible
-%
-\isadelimquotett
-%
-\endisadelimquotett
-%
-\isatagquotett
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}type}\isamarkupfalse%
-\ prod\isanewline
-\ \ {\isaliteral{28}{\isacharparenleft}}SML\ \isakeyword{infix}\ {\isadigit{2}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}const}\isamarkupfalse%
-\ Pair\isanewline
-\ \ {\isaliteral{28}{\isacharparenleft}}SML\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{21}{\isacharbang}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}{\isaliteral{2F}{\isacharslash}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}%
-\endisatagquotett
-{\isafoldquotett}%
-%
-\isadelimquotett
-%
-\endisadelimquotett
-%
-\begin{isamarkuptext}%
-\noindent The initial bang ``\verb|!|'' tells the serialiser
- never to 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 serialisations
- provide; however their usage requires careful thinking in order not
- to introduce inconsistencies -- or, in other words: custom
- serialisations are completely axiomatic.
-
- A further noteworthy detail is that any special character in a
- custom serialisation may be quoted using ``\verb|'|''; thus,
- in ``\verb|fn '_ => _|'' the first ``\verb|_|'' is a
- proper underscore while the second ``\verb|_|'' is a
- placeholder.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{\isa{Haskell} serialisation%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-For convenience, the default \isa{HOL} setup for \isa{Haskell}
- maps the \isa{equal} class to its counterpart in \isa{Haskell},
- giving custom serialisations for the class \isa{equal} (by command
- \indexdef{}{command}{code\_class}\hypertarget{command.code-class}{\hyperlink{command.code-class}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}class}}}}}) and its operation \isa{HOL{\isaliteral{2E}{\isachardot}}equal}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotett
-%
-\endisadelimquotett
-%
-\isatagquotett
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}class}\isamarkupfalse%
-\ equal\isanewline
-\ \ {\isaliteral{28}{\isacharparenleft}}Haskell\ {\isaliteral{22}{\isachardoublequoteopen}}Eq{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\isanewline
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}const}\isamarkupfalse%
-\ {\isaliteral{22}{\isachardoublequoteopen}}HOL{\isaliteral{2E}{\isachardot}}equal{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ {\isaliteral{28}{\isacharparenleft}}Haskell\ \isakeyword{infixl}\ {\isadigit{4}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}%
-\endisatagquotett
-{\isafoldquotett}%
-%
-\isadelimquotett
-%
-\endisadelimquotett
-%
-\begin{isamarkuptext}%
-\noindent A problem now occurs whenever a type which is an instance
- of \isa{equal} in \isa{HOL} is mapped on a \isa{Haskell}-built-in type which is also an instance of \isa{Haskell}
- \isa{Eq}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{typedecl}\isamarkupfalse%
-\ bar\isanewline
-\isanewline
-\isacommand{instantiation}\isamarkupfalse%
-\ bar\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ equal\isanewline
-\isakeyword{begin}\isanewline
-\isanewline
-\isacommand{definition}\isamarkupfalse%
-\ {\isaliteral{22}{\isachardoublequoteopen}}HOL{\isaliteral{2E}{\isachardot}}equal\ {\isaliteral{28}{\isacharparenleft}}x{\isaliteral{5C3C436F6C6F6E3E}{\isasymColon}}bar{\isaliteral{29}{\isacharparenright}}\ y\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{instance}\isamarkupfalse%
-\ \isacommand{by}\isamarkupfalse%
-\ default\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ equal{\isaliteral{5F}{\isacharunderscore}}bar{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\isanewline
-\isanewline
-\isacommand{end}\isamarkupfalse%
-%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isadelimquotett
-\ %
-\endisadelimquotett
-%
-\isatagquotett
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}type}\isamarkupfalse%
-\ bar\isanewline
-\ \ {\isaliteral{28}{\isacharparenleft}}Haskell\ {\isaliteral{22}{\isachardoublequoteopen}}Integer{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}%
-\endisatagquotett
-{\isafoldquotett}%
-%
-\isadelimquotett
-%
-\endisadelimquotett
-%
-\begin{isamarkuptext}%
-\noindent The code generator would produce an additional instance,
- which of course is rejected by the \isa{Haskell} compiler. To
- suppress this additional instance, use \indexdef{}{command}{code\_instance}\hypertarget{command.code-instance}{\hyperlink{command.code-instance}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}instance}}}}}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotett
-%
-\endisadelimquotett
-%
-\isatagquotett
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}instance}\isamarkupfalse%
-\ bar\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ equal\isanewline
-\ \ {\isaliteral{28}{\isacharparenleft}}Haskell\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{29}{\isacharparenright}}%
-\endisatagquotett
-{\isafoldquotett}%
-%
-\isadelimquotett
-%
-\endisadelimquotett
-%
-\isamarkupsubsection{Enhancing the target language context \label{sec:include}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-In rare cases it is necessary to \emph{enrich} the context of a
- target language; this is accomplished using the \indexdef{}{command}{code\_include}\hypertarget{command.code-include}{\hyperlink{command.code-include}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}include}}}}} command:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotett
-%
-\endisadelimquotett
-%
-\isatagquotett
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}include}\isamarkupfalse%
-\ Haskell\ {\isaliteral{22}{\isachardoublequoteopen}}Errno{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-{\isaliteral{7B2A}{\isacharverbatimopen}}errno\ i\ {\isaliteral{3D}{\isacharequal}}\ error\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequote}}Error\ number{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequote}}\ {\isaliteral{2B}{\isacharplus}}{\isaliteral{2B}{\isacharplus}}\ show\ i{\isaliteral{29}{\isacharparenright}}{\isaliteral{2A7D}{\isacharverbatimclose}}\isanewline
-\isanewline
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reserved}\isamarkupfalse%
-\ Haskell\ Errno%
-\endisatagquotett
-{\isafoldquotett}%
-%
-\isadelimquotett
-%
-\endisadelimquotett
-%
-\begin{isamarkuptext}%
-\noindent Such named \isa{include}s are then prepended to every
- generated code. Inspect such code in order to find out how
- \hyperlink{command.code-include}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}include}}}} behaves with respect to a particular
- target language.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-\isacommand{end}\isamarkupfalse%
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-\isanewline
-\isanewline
-\end{isabellebody}%
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End:
--- a/doc-src/Codegen/Thy/document/Evaluation.tex Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,435 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{Evaluation}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-\isacommand{theory}\isamarkupfalse%
-\ Evaluation\isanewline
-\isakeyword{imports}\ Setup\isanewline
-\isakeyword{begin}%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isamarkupsection{Evaluation \label{sec:evaluation}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Recalling \secref{sec:principle}, code generation turns a system of
- equations into a program with the \emph{same} equational semantics.
- As a consequence, this program can be used as a \emph{rewrite
- engine} for terms: rewriting a term \isa{t} using a program to a
- term \isa{t{\isaliteral{27}{\isacharprime}}} yields the theorems \isa{t\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{27}{\isacharprime}}}. This
- application of code generation in the following is referred to as
- \emph{evaluation}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Evaluation techniques%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-The existing infrastructure provides a rich palette of evaluation
- techniques, each comprising different aspects:
-
- \begin{description}
-
- \item[Expressiveness.] Depending on how good symbolic computation
- is supported, the class of terms which can be evaluated may be
- bigger or smaller.
-
- \item[Efficiency.] The more machine-near the technique, the
- faster it is.
-
- \item[Trustability.] Techniques which a huge (and also probably
- more configurable infrastructure) are more fragile and less
- trustable.
-
- \end{description}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsubsection{The simplifier (\isa{simp})%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-The simplest way for evaluation is just using the simplifier with
- the original code equations of the underlying program. This gives
- fully symbolic evaluation and highest trustablity, with the usual
- performance of the simplifier. Note that for operations on abstract
- datatypes (cf.~\secref{sec:invariant}), the original theorems as
- given by the users are used, not the modified ones.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsubsection{Normalization by evaluation (\isa{nbe})%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Normalization by evaluation \cite{Aehlig-Haftmann-Nipkow:2008:nbe}
- provides a comparably fast partially symbolic evaluation which
- permits also normalization of functions and uninterpreted symbols;
- the stack of code to be trusted is considerable.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsubsection{Evaluation in ML (\isa{code})%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Highest performance can be achieved by evaluation in ML, at the cost
- of being restricted to ground results and a layered stack of code to
- be trusted, including code generator configurations by the user.
-
- Evaluation is carried out in a target language \emph{Eval} which
- inherits from \emph{SML} but for convenience uses parts of the
- Isabelle runtime environment. The soundness of computation carried
- out there depends crucially on the correctness of the code
- generator setup; this is one of the reasons why you should not use
- adaptation (see \secref{sec:adaptation}) frivolously.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Aspects of evaluation%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Each of the techniques can be combined with different aspects. The
- most important distinction is between dynamic and static evaluation.
- Dynamic evaluation takes the code generator configuration \qt{as it
- is} at the point where evaluation is issued. Best example is the
- \indexdef{}{command}{value}\hypertarget{command.value}{\hyperlink{command.value}{\mbox{\isa{\isacommand{value}}}}} command which allows ad-hoc evaluation of
- terms:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{value}\isamarkupfalse%
-\ {\isaliteral{22}{\isachardoublequoteopen}}{\isadigit{4}}{\isadigit{2}}\ {\isaliteral{2F}{\isacharslash}}\ {\isaliteral{28}{\isacharparenleft}}{\isadigit{1}}{\isadigit{2}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ rat{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent By default \hyperlink{command.value}{\mbox{\isa{\isacommand{value}}}} tries all available evaluation
- techniques and prints the result of the first succeeding one. A particular
- technique may be specified in square brackets, e.g.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{value}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}nbe{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isadigit{4}}{\isadigit{2}}\ {\isaliteral{2F}{\isacharslash}}\ {\isaliteral{28}{\isacharparenleft}}{\isadigit{1}}{\isadigit{2}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ rat{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-To employ dynamic evaluation in the document generation, there is also
- a \isa{value} antiquotation. By default, it also tries all available evaluation
- techniques and prints the result of the first succeeding one, unless a particular
- technique is specified in square brackets.
-
- Static evaluation freezes the code generator configuration at a
- certain point and uses this context whenever evaluation is issued
- later on. This is particularly appropriate for proof procedures
- which use evaluation, since then the behaviour of evaluation is not
- changed or even compromised later on by actions of the user.
-
- As a technical complication, terms after evaluation in ML must be
- turned into Isabelle's internal term representation again. Since
- this is also configurable, it is never fully trusted. For this
- reason, evaluation in ML comes with further aspects:
-
- \begin{description}
-
- \item[Plain evaluation.] A term is normalized using the provided
- term reconstruction from ML to Isabelle; for applications which
- do not need to be fully trusted.
-
- \item[Property conversion.] Evaluates propositions; since these
- are monomorphic, the term reconstruction is fixed once and for all
- and therefore trustable.
-
- \item[Conversion.] Evaluates an arbitrary term \isa{t} first
- by plain evaluation and certifies the result \isa{t{\isaliteral{27}{\isacharprime}}} by
- checking the equation \isa{t\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{27}{\isacharprime}}} using property
- conversion.
-
- \end{description}
-
- \noindent The picture is further complicated by the roles of
- exceptions. Here three cases have to be distinguished:
-
- \begin{itemize}
-
- \item Evaluation of \isa{t} terminates with a result \isa{t{\isaliteral{27}{\isacharprime}}}.
-
- \item Evaluation of \isa{t} terminates which en exception
- indicating a pattern match failure or a non-implemented
- function. As sketched in \secref{sec:partiality}, this can be
- interpreted as partiality.
-
- \item Evaluation raises any other kind of exception.
-
- \end{itemize}
-
- \noindent For conversions, the first case yields the equation \isa{t\ {\isaliteral{3D}{\isacharequal}}\ t{\isaliteral{27}{\isacharprime}}}, the second defaults to reflexivity \isa{t\ {\isaliteral{3D}{\isacharequal}}\ t}.
- Exceptions of the third kind are propagated to the user.
-
- By default return values of plain evaluation are optional, yielding
- \isa{SOME\ t{\isaliteral{27}{\isacharprime}}} in the first case, \isa{NONE} in the
- second, and propagating the exception in the third case. A strict
- variant of plain evaluation either yields \isa{t{\isaliteral{27}{\isacharprime}}} or propagates
- any exception, a liberal variant caputures any exception in a result
- of type \isa{Exn{\isaliteral{2E}{\isachardot}}result}.
-
- For property conversion (which coincides with conversion except for
- evaluation in ML), methods are provided which solve a given goal by
- evaluation.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Schematic overview%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-\newcommand{\ttsize}{\fontsize{5.8pt}{8pt}\selectfont}
- \fontsize{9pt}{12pt}\selectfont
- \begin{tabular}{ll||c|c|c}
- & & \isa{simp} & \isa{nbe} & \isa{code} \tabularnewline \hline \hline
- \multirow{5}{1ex}{\rotatebox{90}{dynamic}}
- & interactive evaluation
- & \hyperlink{command.value}{\mbox{\isa{\isacommand{value}}}} \isa{{\isaliteral{5B}{\isacharbrackleft}}simp{\isaliteral{5D}{\isacharbrackright}}} & \hyperlink{command.value}{\mbox{\isa{\isacommand{value}}}} \isa{{\isaliteral{5B}{\isacharbrackleft}}nbe{\isaliteral{5D}{\isacharbrackright}}} & \hyperlink{command.value}{\mbox{\isa{\isacommand{value}}}} \isa{{\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5D}{\isacharbrackright}}}
- \tabularnewline
- & plain evaluation & & & \ttsize\verb|Code_Evaluation.dynamic_value| \tabularnewline \cline{2-5}
- & evaluation method & \hyperlink{method.code-simp}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}simp}}} & \hyperlink{method.normalization}{\mbox{\isa{normalization}}} & \hyperlink{method.eval}{\mbox{\isa{eval}}} \tabularnewline
- & property conversion & & & \ttsize\verb|Code_Runtime.dynamic_holds_conv| \tabularnewline \cline{2-5}
- & conversion & \ttsize\verb|Code_Simp.dynamic_conv| & \ttsize\verb|Nbe.dynamic_conv|
- & \ttsize\verb|Code_Evaluation.dynamic_conv| \tabularnewline \hline \hline
- \multirow{3}{1ex}{\rotatebox{90}{static}}
- & plain evaluation & & & \ttsize\verb|Code_Evaluation.static_value| \tabularnewline \cline{2-5}
- & property conversion & &
- & \ttsize\verb|Code_Runtime.static_holds_conv| \tabularnewline \cline{2-5}
- & conversion & \ttsize\verb|Code_Simp.static_conv|
- & \ttsize\verb|Nbe.static_conv|
- & \ttsize\verb|Code_Evaluation.static_conv|
- \end{tabular}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Intimate connection between logic and system runtime%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-The toolbox of static evaluation conversions forms a reasonable base
- to interweave generated code and system tools. However in some
- situations more direct interaction is desirable.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsubsection{Static embedding of generated code into system runtime -- the \isa{code} antiquotation%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-The \isa{code} antiquotation allows to include constants from
- generated code directly into ML system code, as in the following toy
- example:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{datatype}\isamarkupfalse%
-\ form\ {\isaliteral{3D}{\isacharequal}}\ T\ {\isaliteral{7C}{\isacharbar}}\ F\ {\isaliteral{7C}{\isacharbar}}\ And\ form\ form\ {\isaliteral{7C}{\isacharbar}}\ Or\ form\ form\ %
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isadelimquotett
-\ %
-\endisadelimquotett
-%
-\isatagquotett
-\isacommand{ML}\isamarkupfalse%
-\ {\isaliteral{7B2A}{\isacharverbatimopen}}\isanewline
-\ \ fun\ eval{\isaliteral{5F}{\isacharunderscore}}form\ %
-\isaantiq
-code\ T{}%
-\endisaantiq
-\ {\isaliteral{3D}{\isacharequal}}\ true\isanewline
-\ \ \ \ {\isaliteral{7C}{\isacharbar}}\ eval{\isaliteral{5F}{\isacharunderscore}}form\ %
-\isaantiq
-code\ F{}%
-\endisaantiq
-\ {\isaliteral{3D}{\isacharequal}}\ false\isanewline
-\ \ \ \ {\isaliteral{7C}{\isacharbar}}\ eval{\isaliteral{5F}{\isacharunderscore}}form\ {\isaliteral{28}{\isacharparenleft}}%
-\isaantiq
-code\ And{}%
-\endisaantiq
-\ {\isaliteral{28}{\isacharparenleft}}p{\isaliteral{2C}{\isacharcomma}}\ q{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ \ \ \ \ \ \ eval{\isaliteral{5F}{\isacharunderscore}}form\ p\ andalso\ eval{\isaliteral{5F}{\isacharunderscore}}form\ q\isanewline
-\ \ \ \ {\isaliteral{7C}{\isacharbar}}\ eval{\isaliteral{5F}{\isacharunderscore}}form\ {\isaliteral{28}{\isacharparenleft}}%
-\isaantiq
-code\ Or{}%
-\endisaantiq
-\ {\isaliteral{28}{\isacharparenleft}}p{\isaliteral{2C}{\isacharcomma}}\ q{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ \ \ \ \ \ \ eval{\isaliteral{5F}{\isacharunderscore}}form\ p\ orelse\ eval{\isaliteral{5F}{\isacharunderscore}}form\ q{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-{\isaliteral{2A7D}{\isacharverbatimclose}}%
-\endisatagquotett
-{\isafoldquotett}%
-%
-\isadelimquotett
-%
-\endisadelimquotett
-%
-\begin{isamarkuptext}%
-\noindent \isa{code} takes as argument the name of a constant;
- after the whole ML is read, the necessary code is generated
- transparently and the corresponding constant names are inserted.
- This technique also allows to use pattern matching on constructors
- stemming from compiled datatypes. Note that the \isa{code}
- antiquotation may not refer to constants which carry adaptations;
- here you have to refer to the corresponding adapted code directly.
-
- For a less simplistic example, theory \isa{Approximation} in
- the \isa{Decision{\isaliteral{5F}{\isacharunderscore}}Procs} session is a good reference.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsubsection{Static embedding of generated code into system runtime -- \isa{code{\isaliteral{5F}{\isacharunderscore}}reflect}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-The \isa{code} antiquoation is lightweight, but the generated code
- is only accessible while the ML section is processed. Sometimes this
- is not appropriate, especially if the generated code contains datatype
- declarations which are shared with other parts of the system. In these
- cases, \indexdef{}{command}{code\_reflect}\hypertarget{command.code-reflect}{\hyperlink{command.code-reflect}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reflect}}}}} can be used:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reflect}\isamarkupfalse%
-\ Sum{\isaliteral{5F}{\isacharunderscore}}Type\isanewline
-\ \ \isakeyword{datatypes}\ sum\ {\isaliteral{3D}{\isacharequal}}\ Inl\ {\isaliteral{7C}{\isacharbar}}\ Inr\isanewline
-\ \ \isakeyword{functions}\ {\isaliteral{22}{\isachardoublequoteopen}}Sum{\isaliteral{5F}{\isacharunderscore}}Type{\isaliteral{2E}{\isachardot}}Projl{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{22}{\isachardoublequoteopen}}Sum{\isaliteral{5F}{\isacharunderscore}}Type{\isaliteral{2E}{\isachardot}}Projr{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent \indexdef{}{command}{code\_reflect}\hypertarget{command.code-reflect}{\hyperlink{command.code-reflect}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reflect}}}}} takes a structure name and
- references to datatypes and functions; for these code is compiled
- into the named ML structure and the \emph{Eval} target is modified
- in a way that future code generation will reference these
- precompiled versions of the given datatypes and functions. This
- also allows to refer to the referenced datatypes and functions from
- arbitrary ML code as well.
-
- A typical example for \hyperlink{command.code-reflect}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reflect}}}} can be found in the
- \isa{Predicate} theory.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsubsection{Separate compilation -- \isa{code{\isaliteral{5F}{\isacharunderscore}}reflect}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-For technical reasons it is sometimes necessary to separate
- generation and compilation of code which is supposed to be used in
- the system runtime. For this \hyperlink{command.code-reflect}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reflect}}}} with an
- optional \isa{file} argument can be used:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reflect}\isamarkupfalse%
-\ Rat\isanewline
-\ \ \isakeyword{datatypes}\ rat\ {\isaliteral{3D}{\isacharequal}}\ Frct\isanewline
-\ \ \isakeyword{functions}\ Fract\isanewline
-\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}plus\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ rat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ rat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ rat{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}minus\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ rat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ rat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ rat{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}times\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ rat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ rat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ rat{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}divide\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ rat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ rat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ rat{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isakeyword{file}\ {\isaliteral{22}{\isachardoublequoteopen}}examples{\isaliteral{2F}{\isacharslash}}rat{\isaliteral{2E}{\isachardot}}ML{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent This merely generates the referenced code to the given
- file which can be included into the system runtime later on.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-\isacommand{end}\isamarkupfalse%
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-\isanewline
-\isanewline
-\end{isabellebody}%
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End:
--- a/doc-src/Codegen/Thy/document/Foundations.tex Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,613 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{Foundations}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-\isacommand{theory}\isamarkupfalse%
-\ Foundations\isanewline
-\isakeyword{imports}\ Introduction\isanewline
-\isakeyword{begin}%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isamarkupsection{Code generation foundations \label{sec:foundations}%
-}
-\isamarkuptrue%
-%
-\isamarkupsubsection{Code generator architecture \label{sec:architecture}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-The code generator is actually a framework consisting of different
- components which can be customised individually.
-
- Conceptually all components operate on Isabelle's logic framework
- \isa{Pure}. Practically, the object logic \isa{HOL}
- provides the necessary facilities to make use of the code generator,
- mainly since it is an extension of \isa{Pure}.
-
- The constellation of the different components is visualized in the
- following picture.
-
- \begin{figure}[h]
- \includegraphics{architecture}
- \caption{Code generator architecture}
- \label{fig:arch}
- \end{figure}
-
- \noindent Central to code generation is the notion of \emph{code
- equations}. A code equation as a first approximation is a theorem
- of the form \isa{f\ t\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ t\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{2}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ t\isaliteral{5C3C5E697375623E}{}\isactrlisub n\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t} (an equation headed by a
- constant \isa{f} with arguments \isa{t\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ t\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{2}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ t\isaliteral{5C3C5E697375623E}{}\isactrlisub n} and right
- hand side \isa{t}).
-
- \begin{itemize}
-
- \item Starting point of code generation is a collection of (raw)
- code equations in a theory. It is not relevant where they stem
- from, but typically they were either produced by specification
- tools or proved explicitly by the user.
-
- \item These raw code equations can be subjected to theorem
- transformations. This \qn{preprocessor} (see
- \secref{sec:preproc}) can apply the full expressiveness of
- ML-based theorem transformations to code generation. The result
- of preprocessing is a structured collection of code equations.
-
- \item These code equations are \qn{translated} to a program in an
- abstract intermediate language. Think of it as a kind of
- \qt{Mini-Haskell} with four \qn{statements}: \isa{data} (for
- datatypes), \isa{fun} (stemming from code equations), also
- \isa{class} and \isa{inst} (for type classes).
-
- \item Finally, the abstract program is \qn{serialised} into
- concrete source code of a target language. This step only
- produces concrete syntax but does not change the program in
- essence; all conceptual transformations occur in the translation
- step.
-
- \end{itemize}
-
- \noindent From these steps, only the last two 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%
-%
-\isamarkupsubsection{The preprocessor \label{sec:preproc}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Before selected function theorems are turned into abstract code, a
- chain of definitional transformation steps is carried out:
- \emph{preprocessing}. The preprocessor consists of two
- components: a \emph{simpset} and \emph{function transformers}.
-
- The \emph{simpset} can apply the full generality of the Isabelle
- simplifier. Due to the interpretation of theorems as code
- 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{unfold theorems}, which may be declared and removed using the
- \hyperlink{attribute.code-unfold}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}unfold}}} or \emph{\hyperlink{attribute.code-unfold}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}unfold}}} del}
- attribute, respectively.
-
- Some common applications:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\begin{itemize}
-%
-\begin{isamarkuptext}%
-\item replacing non-executable constructs by executable ones:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{lemma}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5F}{\isacharunderscore}}unfold{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ set\ xs\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ List{\isaliteral{2E}{\isachardot}}member\ xs\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}fact\ in{\isaliteral{5F}{\isacharunderscore}}set{\isaliteral{5F}{\isacharunderscore}}member{\isaliteral{29}{\isacharparenright}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\item replacing executable but inconvenient constructs:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{lemma}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5F}{\isacharunderscore}}unfold{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}xs\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ List{\isaliteral{2E}{\isachardot}}null\ xs{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}fact\ eq{\isaliteral{5F}{\isacharunderscore}}Nil{\isaliteral{5F}{\isacharunderscore}}null{\isaliteral{29}{\isacharparenright}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\item eliminating disturbing expressions:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{lemma}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5F}{\isacharunderscore}}unfold{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isadigit{1}}\ {\isaliteral{3D}{\isacharequal}}\ Suc\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}fact\ One{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\end{itemize}
-%
-\begin{isamarkuptext}%
-\noindent \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{\isaliteral{5F}{\isacharunderscore}}Nat} (see
- \secref{eff_nat}) uses this interface.
-
- \noindent The current setup of the preprocessor may be inspected
- using the \indexdef{}{command}{print\_codeproc}\hypertarget{command.print-codeproc}{\hyperlink{command.print-codeproc}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}codeproc}}}}} command. \indexdef{}{command}{code\_thms}\hypertarget{command.code-thms}{\hyperlink{command.code-thms}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}thms}}}}} (see \secref{sec:equations}) provides a convenient
- mechanism to inspect the impact of a preprocessor setup on code
- equations.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Understanding code equations \label{sec:equations}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-As told in \secref{sec:principle}, the notion of code equations is
- vital to code generation. Indeed most problems which occur in
- practice can be resolved by an inspection of the underlying code
- equations.
-
- It is possible to exchange the default code equations for constants
- by explicitly proving alternative ones:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{lemma}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}if\ xs\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ then\ {\isaliteral{28}{\isacharparenleft}}None{\isaliteral{2C}{\isacharcomma}}\ AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ \ \ \ \ \ else\ dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{28}{\isacharparenleft}}rev\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{23}{\isacharhash}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}Some\ y{\isaliteral{2C}{\isacharcomma}}\ AQueue\ xs\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}cases\ xs{\isaliteral{2C}{\isacharcomma}}\ simp{\isaliteral{5F}{\isacharunderscore}}all{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}cases\ {\isaliteral{22}{\isachardoublequoteopen}}rev\ xs{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{2C}{\isacharcomma}}\ simp{\isaliteral{5F}{\isacharunderscore}}all{\isaliteral{29}{\isacharparenright}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent The annotation \isa{{\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5D}{\isacharbrackright}}} is an \isa{attribute}
- which states that the given theorems should be considered as code
- equations for a \isa{fun} statement -- the corresponding constant
- is determined syntactically. The resulting code:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-dequeue\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ forall\ a{\isaliteral{2E}{\isachardot}}\ Queue\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{28}{\isacharparenleft}}Maybe\ a{\isaliteral{2C}{\isacharcomma}}\ Queue\ a{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{3A}{\isacharcolon}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}Just\ y{\isaliteral{2C}{\isacharcomma}}\ AQueue\ xs\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ {\isaliteral{28}{\isacharparenleft}}if\ null\ xs\ then\ {\isaliteral{28}{\isacharparenleft}}Nothing{\isaliteral{2C}{\isacharcomma}}\ AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ \ \ else\ dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{28}{\isacharparenleft}}reverse\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\begin{isamarkuptext}%
-\noindent You may note that the equality test \isa{xs\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}} has
- been replaced by the predicate \isa{List{\isaliteral{2E}{\isachardot}}null\ xs}. This is due
- to the default setup of the \qn{preprocessor}.
-
- This possibility to select arbitrary code equations is the key
- technique for program and datatype refinement (see
- \secref{sec:refinement}).
-
- Due to the preprocessor, there is the distinction of raw code
- equations (before preprocessing) and code equations (after
- preprocessing).
-
- The first can be listed (among other data) using the \indexdef{}{command}{print\_codesetup}\hypertarget{command.print-codesetup}{\hyperlink{command.print-codesetup}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}codesetup}}}}} command.
-
- The code equations after preprocessing are already are blueprint of
- the generated program and can be inspected using the \hyperlink{command.code-thms}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}thms}}}} command:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}thms}\isamarkupfalse%
-\ dequeue%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent This prints a table with the code equations for \isa{dequeue}, including \emph{all} code equations those equations depend
- on recursively. These dependencies themselves can be visualized using
- the \indexdef{}{command}{code\_deps}\hypertarget{command.code-deps}{\hyperlink{command.code-deps}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}deps}}}}} command.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Equality%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Implementation of equality deserves some attention. Here an example
- function involving polymorphic equality:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{primrec}\isamarkupfalse%
-\ collect{\isaliteral{5F}{\isacharunderscore}}duplicates\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ list{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}collect{\isaliteral{5F}{\isacharunderscore}}duplicates\ xs\ ys\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ xs{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}collect{\isaliteral{5F}{\isacharunderscore}}duplicates\ xs\ ys\ {\isaliteral{28}{\isacharparenleft}}z{\isaliteral{23}{\isacharhash}}zs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}if\ z\ {\isaliteral{5C3C696E3E}{\isasymin}}\ set\ xs\isanewline
-\ \ \ \ then\ if\ z\ {\isaliteral{5C3C696E3E}{\isasymin}}\ set\ ys\isanewline
-\ \ \ \ \ \ then\ collect{\isaliteral{5F}{\isacharunderscore}}duplicates\ xs\ ys\ zs\isanewline
-\ \ \ \ \ \ else\ collect{\isaliteral{5F}{\isacharunderscore}}duplicates\ xs\ {\isaliteral{28}{\isacharparenleft}}z{\isaliteral{23}{\isacharhash}}ys{\isaliteral{29}{\isacharparenright}}\ zs\isanewline
-\ \ \ \ else\ collect{\isaliteral{5F}{\isacharunderscore}}duplicates\ {\isaliteral{28}{\isacharparenleft}}z{\isaliteral{23}{\isacharhash}}xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}z{\isaliteral{23}{\isacharhash}}ys{\isaliteral{29}{\isacharparenright}}\ zs{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent During preprocessing, the membership test is rewritten,
- resulting in \isa{List{\isaliteral{2E}{\isachardot}}member}, which itself performs an explicit
- equality check, as can be seen in the corresponding \isa{SML} code:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-structure\ Example\ {\isaliteral{3A}{\isacharcolon}}\ sig\isanewline
-\ \ type\ {\isaliteral{27}{\isacharprime}}a\ equal\isanewline
-\ \ val\ equal\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ equal\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ bool\isanewline
-\ \ val\ eq\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ equal\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ bool\isanewline
-\ \ val\ member\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ equal\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ bool\isanewline
-\ \ val\ collect{\isaliteral{5F}{\isacharunderscore}}duplicates\ {\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ \ \ {\isaliteral{27}{\isacharprime}}a\ equal\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ list\isanewline
-end\ {\isaliteral{3D}{\isacharequal}}\ struct\isanewline
-\isanewline
-type\ {\isaliteral{27}{\isacharprime}}a\ equal\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B}{\isacharbraceleft}}equal\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ bool{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-val\ equal\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{23}{\isacharhash}}equal\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ equal\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ bool{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ eq\ A{\isaliteral{5F}{\isacharunderscore}}\ a\ b\ {\isaliteral{3D}{\isacharequal}}\ equal\ A{\isaliteral{5F}{\isacharunderscore}}\ a\ b{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ member\ A{\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ y\ {\isaliteral{3D}{\isacharequal}}\ false\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ member\ A{\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ xs{\isaliteral{29}{\isacharparenright}}\ y\ {\isaliteral{3D}{\isacharequal}}\ eq\ A{\isaliteral{5F}{\isacharunderscore}}\ x\ y\ orelse\ member\ A{\isaliteral{5F}{\isacharunderscore}}\ xs\ y{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ collect{\isaliteral{5F}{\isacharunderscore}}duplicates\ A{\isaliteral{5F}{\isacharunderscore}}\ xs\ ys\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ xs\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ collect{\isaliteral{5F}{\isacharunderscore}}duplicates\ A{\isaliteral{5F}{\isacharunderscore}}\ xs\ ys\ {\isaliteral{28}{\isacharparenleft}}z\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ zs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ \ \ {\isaliteral{28}{\isacharparenleft}}if\ member\ A{\isaliteral{5F}{\isacharunderscore}}\ xs\ z\isanewline
-\ \ \ \ \ \ then\ {\isaliteral{28}{\isacharparenleft}}if\ member\ A{\isaliteral{5F}{\isacharunderscore}}\ ys\ z\ then\ collect{\isaliteral{5F}{\isacharunderscore}}duplicates\ A{\isaliteral{5F}{\isacharunderscore}}\ xs\ ys\ zs\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ else\ collect{\isaliteral{5F}{\isacharunderscore}}duplicates\ A{\isaliteral{5F}{\isacharunderscore}}\ xs\ {\isaliteral{28}{\isacharparenleft}}z\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ ys{\isaliteral{29}{\isacharparenright}}\ zs{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ \ \ \ \ else\ collect{\isaliteral{5F}{\isacharunderscore}}duplicates\ A{\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{28}{\isacharparenleft}}z\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}z\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ ys{\isaliteral{29}{\isacharparenright}}\ zs{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-end{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}struct\ Example{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\begin{isamarkuptext}%
-\noindent Obviously, polymorphic equality is implemented the Haskell
- way using a type class. How is this achieved? HOL introduces an
- explicit class \isa{equal} with a corresponding operation \isa{equal{\isaliteral{5F}{\isacharunderscore}}class{\isaliteral{2E}{\isachardot}}equal} such that \isa{equal{\isaliteral{5F}{\isacharunderscore}}class{\isaliteral{2E}{\isachardot}}equal\ {\isaliteral{3D}{\isacharequal}}\ op\ {\isaliteral{3D}{\isacharequal}}}. The preprocessing
- framework does the rest by propagating the \isa{equal} constraints
- through all dependent code equations. For datatypes, instances of
- \isa{equal} are implicitly derived when possible. For other types,
- you may instantiate \isa{equal} manually like any other type class.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Explicit partiality \label{sec:partiality}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Partiality usually enters the game by partial patterns, as
- in the following example, again for amortised queues:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{definition}\isamarkupfalse%
-\ strict{\isaliteral{5F}{\isacharunderscore}}dequeue\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C74696D65733E}{\isasymtimes}}\ {\isaliteral{27}{\isacharprime}}a\ queue{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}strict{\isaliteral{5F}{\isacharunderscore}}dequeue\ q\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}case\ dequeue\ q\isanewline
-\ \ \ \ of\ {\isaliteral{28}{\isacharparenleft}}Some\ x{\isaliteral{2C}{\isacharcomma}}\ q{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}x{\isaliteral{2C}{\isacharcomma}}\ q{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{lemma}\isamarkupfalse%
-\ strict{\isaliteral{5F}{\isacharunderscore}}dequeue{\isaliteral{5F}{\isacharunderscore}}AQueue\ {\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}strict{\isaliteral{5F}{\isacharunderscore}}dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{23}{\isacharhash}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}y{\isaliteral{2C}{\isacharcomma}}\ AQueue\ xs\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}strict{\isaliteral{5F}{\isacharunderscore}}dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ \ \ {\isaliteral{28}{\isacharparenleft}}case\ rev\ xs\ of\ y\ {\isaliteral{23}{\isacharhash}}\ ys\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}y{\isaliteral{2C}{\isacharcomma}}\ AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}simp{\isaliteral{5F}{\isacharunderscore}}all\ add{\isaliteral{3A}{\isacharcolon}}\ strict{\isaliteral{5F}{\isacharunderscore}}dequeue{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}cases\ xs{\isaliteral{2C}{\isacharcomma}}\ simp{\isaliteral{5F}{\isacharunderscore}}all\ split{\isaliteral{3A}{\isacharcolon}}\ list{\isaliteral{2E}{\isachardot}}split{\isaliteral{29}{\isacharparenright}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent In the corresponding code, there is no equation
- for the pattern \isa{AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-strict{\isaliteral{5F}{\isacharunderscore}}dequeue\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ forall\ a{\isaliteral{2E}{\isachardot}}\ Queue\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{28}{\isacharparenleft}}a{\isaliteral{2C}{\isacharcomma}}\ Queue\ a{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-strict{\isaliteral{5F}{\isacharunderscore}}dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ let\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-\ \ \ \ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{3A}{\isacharcolon}}\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ reverse\ xs{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\ \ {\isaliteral{7D}{\isacharbraceright}}\ in\ {\isaliteral{28}{\isacharparenleft}}y{\isaliteral{2C}{\isacharcomma}}\ AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-strict{\isaliteral{5F}{\isacharunderscore}}dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{3A}{\isacharcolon}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}y{\isaliteral{2C}{\isacharcomma}}\ AQueue\ xs\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\begin{isamarkuptext}%
-\noindent In some cases it is desirable to have this
- pseudo-\qt{partiality} more explicitly, e.g.~as follows:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{axiomatization}\isamarkupfalse%
-\ empty{\isaliteral{5F}{\isacharunderscore}}queue\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\isanewline
-\isanewline
-\isacommand{definition}\isamarkupfalse%
-\ strict{\isaliteral{5F}{\isacharunderscore}}dequeue{\isaliteral{27}{\isacharprime}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C74696D65733E}{\isasymtimes}}\ {\isaliteral{27}{\isacharprime}}a\ queue{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}strict{\isaliteral{5F}{\isacharunderscore}}dequeue{\isaliteral{27}{\isacharprime}}\ q\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}case\ dequeue\ q\ of\ {\isaliteral{28}{\isacharparenleft}}Some\ x{\isaliteral{2C}{\isacharcomma}}\ q{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}x{\isaliteral{2C}{\isacharcomma}}\ q{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ empty{\isaliteral{5F}{\isacharunderscore}}queue{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{lemma}\isamarkupfalse%
-\ strict{\isaliteral{5F}{\isacharunderscore}}dequeue{\isaliteral{27}{\isacharprime}}{\isaliteral{5F}{\isacharunderscore}}AQueue\ {\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}strict{\isaliteral{5F}{\isacharunderscore}}dequeue{\isaliteral{27}{\isacharprime}}\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}if\ xs\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ then\ empty{\isaliteral{5F}{\isacharunderscore}}queue\isanewline
-\ \ \ \ \ else\ strict{\isaliteral{5F}{\isacharunderscore}}dequeue{\isaliteral{27}{\isacharprime}}\ {\isaliteral{28}{\isacharparenleft}}AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{28}{\isacharparenleft}}rev\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}strict{\isaliteral{5F}{\isacharunderscore}}dequeue{\isaliteral{27}{\isacharprime}}\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{23}{\isacharhash}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}y{\isaliteral{2C}{\isacharcomma}}\ AQueue\ xs\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}simp{\isaliteral{5F}{\isacharunderscore}}all\ add{\isaliteral{3A}{\isacharcolon}}\ strict{\isaliteral{5F}{\isacharunderscore}}dequeue{\isaliteral{27}{\isacharprime}}{\isaliteral{5F}{\isacharunderscore}}def\ split{\isaliteral{3A}{\isacharcolon}}\ list{\isaliteral{2E}{\isachardot}}splits{\isaliteral{29}{\isacharparenright}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-Observe that on the right hand side of the definition of \isa{strict{\isaliteral{5F}{\isacharunderscore}}dequeue{\isaliteral{27}{\isacharprime}}}, the unspecified constant \isa{empty{\isaliteral{5F}{\isacharunderscore}}queue} occurs.
-
- Normally, if constants without any code equations occur in a
- program, the code generator complains (since in most cases this is
- indeed an error). But such constants can also be thought
- of as function definitions which always fail,
- since there is never a successful pattern match on the left hand
- side. In order to categorise a constant into that category
- explicitly, use \indexdef{}{command}{code\_abort}\hypertarget{command.code-abort}{\hyperlink{command.code-abort}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}abort}}}}}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}abort}\isamarkupfalse%
-\ empty{\isaliteral{5F}{\isacharunderscore}}queue%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent Then the code generator will just insert an error or
- exception at the appropriate position:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-empty{\isaliteral{5F}{\isacharunderscore}}queue\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ forall\ a{\isaliteral{2E}{\isachardot}}\ a{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-empty{\isaliteral{5F}{\isacharunderscore}}queue\ {\isaliteral{3D}{\isacharequal}}\ error\ {\isaliteral{22}{\isachardoublequote}}empty{\isaliteral{5F}{\isacharunderscore}}queue{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-strict{\isaliteral{5F}{\isacharunderscore}}dequeue\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ forall\ a{\isaliteral{2E}{\isachardot}}\ Queue\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{28}{\isacharparenleft}}a{\isaliteral{2C}{\isacharcomma}}\ Queue\ a{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-strict{\isaliteral{5F}{\isacharunderscore}}dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{3A}{\isacharcolon}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}y{\isaliteral{2C}{\isacharcomma}}\ AQueue\ xs\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-strict{\isaliteral{5F}{\isacharunderscore}}dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ {\isaliteral{28}{\isacharparenleft}}if\ null\ xs\ then\ empty{\isaliteral{5F}{\isacharunderscore}}queue\isanewline
-\ \ \ \ else\ strict{\isaliteral{5F}{\isacharunderscore}}dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{28}{\isacharparenleft}}reverse\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\begin{isamarkuptext}%
-\noindent This feature however is rarely needed in practice. Note
- also that the HOL default setup already declares \isa{undefined}
- as \hyperlink{command.code-abort}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}abort}}}}, which is most likely to be used in such
- situations.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{If something goes utterly wrong \label{sec:utterly_wrong}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Under certain circumstances, the code generator fails to produce
- code entirely. To debug these, the following hints may prove
- helpful:
-
- \begin{description}
-
- \ditem{\emph{Check with a different target language}.} Sometimes
- the situation gets more clear if you switch to another target
- language; the code generated there might give some hints what
- prevents the code generator to produce code for the desired
- language.
-
- \ditem{\emph{Inspect code equations}.} Code equations are the central
- carrier of code generation. Most problems occurring while generating
- code can be traced to single equations which are printed as part of
- the error message. A closer inspection of those may offer the key
- for solving issues (cf.~\secref{sec:equations}).
-
- \ditem{\emph{Inspect preprocessor setup}.} The preprocessor might
- transform code equations unexpectedly; to understand an
- inspection of its setup is necessary (cf.~\secref{sec:preproc}).
-
- \ditem{\emph{Generate exceptions}.} If the code generator
- complains about missing code equations, in can be helpful to
- implement the offending constants as exceptions
- (cf.~\secref{sec:partiality}); this allows at least for a formal
- generation of code, whose inspection may then give clues what is
- wrong.
-
- \ditem{\emph{Remove offending code equations}.} If code
- generation is prevented by just a single equation, this can be
- removed (cf.~\secref{sec:equations}) to allow formal code
- generation, whose result in turn can be used to trace the
- problem. The most prominent case here are mismatches in type
- class signatures (\qt{wellsortedness error}).
-
- \end{description}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-\isacommand{end}\isamarkupfalse%
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-\isanewline
-\end{isabellebody}%
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End:
--- a/doc-src/Codegen/Thy/document/Further.tex Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,624 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{Further}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-\isacommand{theory}\isamarkupfalse%
-\ Further\isanewline
-\isakeyword{imports}\ Setup\isanewline
-\isakeyword{begin}%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isamarkupsection{Further issues \label{sec:further}%
-}
-\isamarkuptrue%
-%
-\isamarkupsubsection{Specialities of the \isa{Scala} target language \label{sec:scala}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-\isa{Scala} deviates from languages of the ML family in a couple
- of aspects; those which affect code generation mainly have to do with
- \isa{Scala}'s type system:
-
- \begin{itemize}
-
- \item \isa{Scala} prefers tupled syntax over curried syntax.
-
- \item \isa{Scala} sacrifices Hindely-Milner type inference for a
- much more rich type system with subtyping etc. For this reason
- type arguments sometimes have to be given explicitly in square
- brackets (mimicking System F syntax).
-
- \item In contrast to \isa{Haskell} where most specialities of
- the type system are implemented using \emph{type classes},
- \isa{Scala} provides a sophisticated system of \emph{implicit
- arguments}.
-
- \end{itemize}
-
- \noindent Concerning currying, the \isa{Scala} serializer counts
- arguments in code equations to determine how many arguments
- shall be tupled; remaining arguments and abstractions in terms
- rather than function definitions are always curried.
-
- The second aspect affects user-defined adaptations with \hyperlink{command.code-const}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}const}}}}. For regular terms, the \isa{Scala} serializer prints
- all type arguments explicitly. For user-defined term adaptations
- this is only possible for adaptations which take no arguments: here
- the type arguments are just appended. Otherwise they are ignored;
- hence user-defined adaptations for polymorphic constants have to be
- designed very carefully to avoid ambiguity.
-
- Isabelle's type classes are mapped onto \isa{Scala} implicits; in
- cases with diamonds in the subclass hierarchy this can lead to
- ambiguities in the generated code:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{class}\isamarkupfalse%
-\ class{\isadigit{1}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ \isakeyword{fixes}\ foo\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{class}\isamarkupfalse%
-\ class{\isadigit{2}}\ {\isaliteral{3D}{\isacharequal}}\ class{\isadigit{1}}\isanewline
-\isanewline
-\isacommand{class}\isamarkupfalse%
-\ class{\isadigit{3}}\ {\isaliteral{3D}{\isacharequal}}\ class{\isadigit{1}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent Here both \isa{class{\isadigit{2}}} and \isa{class{\isadigit{3}}} inherit from \isa{class{\isadigit{1}}},
- forming the upper part of a diamond.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{definition}\isamarkupfalse%
-\ bar\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{7B}{\isacharbraceleft}}class{\isadigit{2}}{\isaliteral{2C}{\isacharcomma}}\ class{\isadigit{3}}{\isaliteral{7D}{\isacharbraceright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}bar\ {\isaliteral{3D}{\isacharequal}}\ foo{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent This yields the following code:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-object\ Example\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-\isanewline
-trait\ class{\isadigit{1}}{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-\ \ val\ {\isaliteral{60}{\isacharbackquote}}Example{\isaliteral{2E}{\isachardot}}foo{\isaliteral{60}{\isacharbackquote}}{\isaliteral{3A}{\isacharcolon}}\ A\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ A\isanewline
-{\isaliteral{7D}{\isacharbraceright}}\isanewline
-def\ foo{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{28}{\isacharparenleft}}a{\isaliteral{3A}{\isacharcolon}}\ A{\isaliteral{29}{\isacharparenright}}{\isaliteral{28}{\isacharparenleft}}implicit\ A{\isaliteral{3A}{\isacharcolon}}\ class{\isadigit{1}}{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{3A}{\isacharcolon}}\ A\ {\isaliteral{3D}{\isacharequal}}\ A{\isaliteral{2E}{\isachardot}}{\isaliteral{60}{\isacharbackquote}}Example{\isaliteral{2E}{\isachardot}}foo{\isaliteral{60}{\isacharbackquote}}{\isaliteral{28}{\isacharparenleft}}a{\isaliteral{29}{\isacharparenright}}\isanewline
-\isanewline
-trait\ class{\isadigit{2}}{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}\ extends\ class{\isadigit{1}}{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-{\isaliteral{7D}{\isacharbraceright}}\isanewline
-\isanewline
-trait\ class{\isadigit{3}}{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}\ extends\ class{\isadigit{1}}{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-{\isaliteral{7D}{\isacharbraceright}}\isanewline
-\isanewline
-def\ bar{\isaliteral{5B}{\isacharbrackleft}}A\ {\isaliteral{3A}{\isacharcolon}}\ class{\isadigit{2}}\ {\isaliteral{3A}{\isacharcolon}}\ class{\isadigit{3}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ A\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ A\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}a{\isaliteral{3A}{\isacharcolon}}\ A{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ foo{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{28}{\isacharparenleft}}a{\isaliteral{29}{\isacharparenright}}\isanewline
-\isanewline
-{\isaliteral{7D}{\isacharbraceright}}\ {\isaliteral{2F}{\isacharslash}}{\isaliteral{2A}{\isacharasterisk}}\ object\ Example\ {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{2F}{\isacharslash}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\begin{isamarkuptext}%
-\noindent This code is rejected by the \isa{Scala} compiler: in
- the definition of \isa{bar}, it is not clear from where to derive
- the implicit argument for \isa{foo}.
-
- The solution to the problem is to close the diamond by a further
- class with inherits from both \isa{class{\isadigit{2}}} and \isa{class{\isadigit{3}}}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{class}\isamarkupfalse%
-\ class{\isadigit{4}}\ {\isaliteral{3D}{\isacharequal}}\ class{\isadigit{2}}\ {\isaliteral{2B}{\isacharplus}}\ class{\isadigit{3}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent Then the offending code equation can be restricted to
- \isa{class{\isadigit{4}}}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{lemma}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}bar\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}class{\isadigit{4}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ foo{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}simp\ only{\isaliteral{3A}{\isacharcolon}}\ bar{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent with the following code:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-object\ Example\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-\isanewline
-trait\ class{\isadigit{1}}{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-\ \ val\ {\isaliteral{60}{\isacharbackquote}}Example{\isaliteral{2E}{\isachardot}}foo{\isaliteral{60}{\isacharbackquote}}{\isaliteral{3A}{\isacharcolon}}\ A\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ A\isanewline
-{\isaliteral{7D}{\isacharbraceright}}\isanewline
-def\ foo{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{28}{\isacharparenleft}}a{\isaliteral{3A}{\isacharcolon}}\ A{\isaliteral{29}{\isacharparenright}}{\isaliteral{28}{\isacharparenleft}}implicit\ A{\isaliteral{3A}{\isacharcolon}}\ class{\isadigit{1}}{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{3A}{\isacharcolon}}\ A\ {\isaliteral{3D}{\isacharequal}}\ A{\isaliteral{2E}{\isachardot}}{\isaliteral{60}{\isacharbackquote}}Example{\isaliteral{2E}{\isachardot}}foo{\isaliteral{60}{\isacharbackquote}}{\isaliteral{28}{\isacharparenleft}}a{\isaliteral{29}{\isacharparenright}}\isanewline
-\isanewline
-trait\ class{\isadigit{3}}{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}\ extends\ class{\isadigit{1}}{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-{\isaliteral{7D}{\isacharbraceright}}\isanewline
-\isanewline
-trait\ class{\isadigit{2}}{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}\ extends\ class{\isadigit{1}}{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-{\isaliteral{7D}{\isacharbraceright}}\isanewline
-\isanewline
-trait\ class{\isadigit{4}}{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}\ extends\ class{\isadigit{2}}{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}\ with\ class{\isadigit{3}}{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-{\isaliteral{7D}{\isacharbraceright}}\isanewline
-\isanewline
-def\ bar{\isaliteral{5B}{\isacharbrackleft}}A\ {\isaliteral{3A}{\isacharcolon}}\ class{\isadigit{4}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ A\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ A\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}a{\isaliteral{3A}{\isacharcolon}}\ A{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ foo{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{28}{\isacharparenleft}}a{\isaliteral{29}{\isacharparenright}}\isanewline
-\isanewline
-{\isaliteral{7D}{\isacharbraceright}}\ {\isaliteral{2F}{\isacharslash}}{\isaliteral{2A}{\isacharasterisk}}\ object\ Example\ {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{2F}{\isacharslash}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\begin{isamarkuptext}%
-\noindent which exposes no ambiguity.
-
- Since the preprocessor (cf.~\secref{sec:preproc}) propagates sort
- constraints through a system of code equations, it is usually not
- very difficult to identify the set of code equations which actually
- needs more restricted sort constraints.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Modules namespace%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-When invoking the \hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isaliteral{5F}{\isacharunderscore}}code}}}} command it is possible to
- leave out the \hyperlink{keyword.module-name}{\mbox{\isa{\isakeyword{module{\isaliteral{5F}{\isacharunderscore}}name}}}} part; then code is
- distributed over different modules, where the module name space
- roughly is induced by the Isabelle theory name space.
-
- Then sometimes the awkward situation occurs that dependencies
- between definitions introduce cyclic dependencies between modules,
- which in the \isa{Haskell} world leaves you to the mercy of the
- \isa{Haskell} implementation you are using, while for \isa{SML}/\isa{OCaml} 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%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}modulename}\isamarkupfalse%
-\ SML\isanewline
-\ \ A\ ABC\isanewline
-\ \ B\ ABC\isanewline
-\ \ C\ ABC%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent we explicitly map all those modules on \emph{ABC},
- resulting in an ad-hoc merge of this three modules at serialisation
- time.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Locales and interpretation%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-A technical issue comes to surface when generating code from
- specifications stemming from locale interpretation.
-
- Let us assume a locale specifying a power operation on arbitrary
- types:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{locale}\isamarkupfalse%
-\ power\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ \isakeyword{fixes}\ power\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isakeyword{assumes}\ power{\isaliteral{5F}{\isacharunderscore}}commute{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}power\ x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ power\ y\ {\isaliteral{3D}{\isacharequal}}\ power\ y\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ power\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{begin}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent Inside that locale we can lift \isa{power} to exponent
- lists by means of specification relative to that locale:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{primrec}\isamarkupfalse%
-\ powers\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}powers\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ id{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}powers\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{23}{\isacharhash}}\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ power\ x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ powers\ xs{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{lemma}\isamarkupfalse%
-\ powers{\isaliteral{5F}{\isacharunderscore}}append{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}powers\ {\isaliteral{28}{\isacharparenleft}}xs\ {\isaliteral{40}{\isacharat}}\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ powers\ xs\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ powers\ ys{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}induct\ xs{\isaliteral{29}{\isacharparenright}}\ simp{\isaliteral{5F}{\isacharunderscore}}all\isanewline
-\isanewline
-\isacommand{lemma}\isamarkupfalse%
-\ powers{\isaliteral{5F}{\isacharunderscore}}power{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}powers\ xs\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ power\ x\ {\isaliteral{3D}{\isacharequal}}\ power\ x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ powers\ xs{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}induct\ xs{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ \ \ {\isaliteral{28}{\isacharparenleft}}simp{\isaliteral{5F}{\isacharunderscore}}all\ del{\isaliteral{3A}{\isacharcolon}}\ o{\isaliteral{5F}{\isacharunderscore}}apply\ id{\isaliteral{5F}{\isacharunderscore}}apply\ add{\isaliteral{3A}{\isacharcolon}}\ o{\isaliteral{5F}{\isacharunderscore}}assoc\ {\isaliteral{5B}{\isacharbrackleft}}symmetric{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2C}{\isacharcomma}}\isanewline
-\ \ \ \ \ \ simp\ del{\isaliteral{3A}{\isacharcolon}}\ o{\isaliteral{5F}{\isacharunderscore}}apply\ add{\isaliteral{3A}{\isacharcolon}}\ o{\isaliteral{5F}{\isacharunderscore}}assoc\ power{\isaliteral{5F}{\isacharunderscore}}commute{\isaliteral{29}{\isacharparenright}}\isanewline
-\isanewline
-\isacommand{lemma}\isamarkupfalse%
-\ powers{\isaliteral{5F}{\isacharunderscore}}rev{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}powers\ {\isaliteral{28}{\isacharparenleft}}rev\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ powers\ xs{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}induct\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}simp{\isaliteral{5F}{\isacharunderscore}}all\ add{\isaliteral{3A}{\isacharcolon}}\ powers{\isaliteral{5F}{\isacharunderscore}}append\ powers{\isaliteral{5F}{\isacharunderscore}}power{\isaliteral{29}{\isacharparenright}}\isanewline
-\isanewline
-\isacommand{end}\isamarkupfalse%
-%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-After an interpretation of this locale (say, \indexdef{}{command}{interpretation}\hypertarget{command.interpretation}{\hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}}} \isa{fun{\isaliteral{5F}{\isacharunderscore}}power{\isaliteral{3A}{\isacharcolon}}} \isa{{\isaliteral{22}{\isachardoublequote}}power\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}n\ {\isaliteral{28}{\isacharparenleft}}f\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a{\isaliteral{29}{\isacharparenright}}{\isaliteral{2E}{\isachardot}}\ f\ {\isaliteral{5E}{\isacharcircum}}{\isaliteral{5E}{\isacharcircum}}\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}), one would expect to have a constant \isa{fun{\isaliteral{5F}{\isacharunderscore}}power{\isaliteral{2E}{\isachardot}}powers\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\ list\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a} for which code
- can be generated. But this not the case: internally, the term
- \isa{fun{\isaliteral{5F}{\isacharunderscore}}power{\isaliteral{2E}{\isachardot}}powers} is an abbreviation for the foundational
- term \isa{{\isaliteral{22}{\isachardoublequote}}power{\isaliteral{2E}{\isachardot}}powers\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}n\ {\isaliteral{28}{\isacharparenleft}}f\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a{\isaliteral{29}{\isacharparenright}}{\isaliteral{2E}{\isachardot}}\ f\ {\isaliteral{5E}{\isacharcircum}}{\isaliteral{5E}{\isacharcircum}}\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}
- (see \cite{isabelle-locale} for the details behind).
-
- Fortunately, with minor effort the desired behaviour can be
- achieved. First, a dedicated definition of the constant on which
- the local \isa{powers} after interpretation is supposed to be
- mapped on:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{definition}\isamarkupfalse%
-\ funpows\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}nat\ list\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ {\isaliteral{5B}{\isacharbrackleft}}code\ del{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}funpows\ {\isaliteral{3D}{\isacharequal}}\ power{\isaliteral{2E}{\isachardot}}powers\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}n\ f{\isaliteral{2E}{\isachardot}}\ f\ {\isaliteral{5E}{\isacharcircum}}{\isaliteral{5E}{\isacharcircum}}\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent In general, the pattern is \isa{c\ {\isaliteral{3D}{\isacharequal}}\ t} where \isa{c}
- is the name of the future constant and \isa{t} the foundational
- term corresponding to the local constant after interpretation.
-
- The interpretation itself is enriched with an equation \isa{t\ {\isaliteral{3D}{\isacharequal}}\ c}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{interpretation}\isamarkupfalse%
-\ fun{\isaliteral{5F}{\isacharunderscore}}power{\isaliteral{3A}{\isacharcolon}}\ power\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}n\ {\isaliteral{28}{\isacharparenleft}}f\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a{\isaliteral{29}{\isacharparenright}}{\isaliteral{2E}{\isachardot}}\ f\ {\isaliteral{5E}{\isacharcircum}}{\isaliteral{5E}{\isacharcircum}}\ n{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}power{\isaliteral{2E}{\isachardot}}powers\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}n\ f{\isaliteral{2E}{\isachardot}}\ f\ {\isaliteral{5E}{\isacharcircum}}{\isaliteral{5E}{\isacharcircum}}\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ funpows{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ unfold{\isaliteral{5F}{\isacharunderscore}}locales\isanewline
-\ \ \ \ {\isaliteral{28}{\isacharparenleft}}simp{\isaliteral{5F}{\isacharunderscore}}all\ add{\isaliteral{3A}{\isacharcolon}}\ fun{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}iff\ funpow{\isaliteral{5F}{\isacharunderscore}}mult\ mult{\isaliteral{5F}{\isacharunderscore}}commute\ funpows{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent This additional equation is trivially proved by the
- definition itself.
-
- After this setup procedure, code generation can continue as usual:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-funpow\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ forall\ a{\isaliteral{2E}{\isachardot}}\ Nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{28}{\isacharparenleft}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ a{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-funpow\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ f\ {\isaliteral{3D}{\isacharequal}}\ id{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-funpow\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ f\ {\isaliteral{3D}{\isacharequal}}\ f\ {\isaliteral{2E}{\isachardot}}\ funpow\ n\ f{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-funpows\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ forall\ a{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5B}{\isacharbrackleft}}Nat{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{28}{\isacharparenleft}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ a{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-funpows\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ id{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-funpows\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3A}{\isacharcolon}}\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ funpow\ x\ {\isaliteral{2E}{\isachardot}}\ funpows\ xs{\isaliteral{3B}{\isacharsemicolon}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isamarkupsubsection{Imperative data structures%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-If you consider imperative data structures as inevitable for a
- specific application, you should consider \emph{Imperative
- Functional Programming with Isabelle/HOL}
- \cite{bulwahn-et-al:2008:imperative}; the framework described there
- is available in session \isa{Imperative{\isaliteral{5F}{\isacharunderscore}}HOL}, together with a
- short primer document.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{ML system 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 fundamental ML
- interfaces.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsubsection{Managing executable content%
-}
-\isamarkuptrue%
-%
-\isadelimmlref
-%
-\endisadelimmlref
-%
-\isatagmlref
-%
-\begin{isamarkuptext}%
-\begin{mldecls}
- \indexdef{}{ML}{Code.read\_const}\verb|Code.read_const: theory -> string -> string| \\
- \indexdef{}{ML}{Code.add\_eqn}\verb|Code.add_eqn: thm -> theory -> theory| \\
- \indexdef{}{ML}{Code.del\_eqn}\verb|Code.del_eqn: thm -> theory -> theory| \\
- \indexdef{}{ML}{Code\_Preproc.map\_pre}\verb|Code_Preproc.map_pre: (simpset -> simpset) -> theory -> theory| \\
- \indexdef{}{ML}{Code\_Preproc.map\_post}\verb|Code_Preproc.map_post: (simpset -> simpset) -> theory -> theory| \\
- \indexdef{}{ML}{Code\_Preproc.add\_functrans}\verb|Code_Preproc.add_functrans: |\isasep\isanewline%
-\verb| string * (theory -> (thm * bool) list -> (thm * bool) list option)|\isasep\isanewline%
-\verb| -> theory -> theory| \\
- \indexdef{}{ML}{Code\_Preproc.del\_functrans}\verb|Code_Preproc.del_functrans: string -> theory -> theory| \\
- \indexdef{}{ML}{Code.add\_datatype}\verb|Code.add_datatype: (string * typ) list -> theory -> theory| \\
- \indexdef{}{ML}{Code.get\_type}\verb|Code.get_type: theory -> string|\isasep\isanewline%
-\verb| -> ((string * sort) list * (string * ((string * sort) list * typ list)) list) * bool| \\
- \indexdef{}{ML}{Code.get\_type\_of\_constr\_or\_abstr}\verb|Code.get_type_of_constr_or_abstr: theory -> string -> (string * bool) option|
- \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.add_eqn|~\isa{thm}~\isa{thy} adds function
- theorem \isa{thm} to executable content.
-
- \item \verb|Code.del_eqn|~\isa{thm}~\isa{thy} removes function
- theorem \isa{thm} from executable content, if present.
-
- \item \verb|Code_Preproc.map_pre|~\isa{f}~\isa{thy} changes
- the preprocessor simpset.
-
- \item \verb|Code_Preproc.add_functrans|~\isa{{\isaliteral{28}{\isacharparenleft}}name{\isaliteral{2C}{\isacharcomma}}\ f{\isaliteral{29}{\isacharparenright}}}~\isa{thy} adds
- function transformer \isa{f} (named \isa{name}) to executable content;
- \isa{f} is a transformer of the code 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 code equations.
-
- \item \verb|Code_Preproc.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_type_of_constr_or_abstr|~\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
-%
-\isamarkupsubsubsection{Data depending on the theory's executable content%
-}
-\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.
-
- 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|Code_Data|; on instantiation of this functor, the
- following types and operations are required:
-
- \medskip
- \begin{tabular}{l}
- \isa{type\ T} \\
- \isa{val\ empty{\isaliteral{3A}{\isacharcolon}}\ T} \\
- \end{tabular}
-
- \begin{description}
-
- \item \isa{T} the type of data to store.
-
- \item \isa{empty} initial (empty) data.
-
- \end{description}
-
- \noindent An instance of \verb|Code_Data| provides the
- following interface:
-
- \medskip
- \begin{tabular}{l}
- \isa{change{\isaliteral{3A}{\isacharcolon}}\ theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}T\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ T{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ T} \\
- \isa{change{\isaliteral{5F}{\isacharunderscore}}yield{\isaliteral{3A}{\isacharcolon}}\ theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}T\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2A}{\isacharasterisk}}\ T{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2A}{\isacharasterisk}}\ T}
- \end{tabular}
-
- \begin{description}
-
- \item \isa{change} update of current data (cached!) by giving a
- continuation.
-
- \item \isa{change{\isaliteral{5F}{\isacharunderscore}}yield} update with side result.
-
- \end{description}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-\isacommand{end}\isamarkupfalse%
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-\isanewline
-\isanewline
-\end{isabellebody}%
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End:
--- a/doc-src/Codegen/Thy/document/Inductive_Predicate.tex Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,499 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{Inductive{\isaliteral{5F}{\isacharunderscore}}Predicate}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-\isacommand{theory}\isamarkupfalse%
-\ Inductive{\isaliteral{5F}{\isacharunderscore}}Predicate\isanewline
-\isakeyword{imports}\ Setup\isanewline
-\isakeyword{begin}\isanewline
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isadeliminvisible
-%
-\endisadeliminvisible
-%
-\isataginvisible
-%
-\endisataginvisible
-{\isafoldinvisible}%
-%
-\isadeliminvisible
-%
-\endisadeliminvisible
-%
-\isamarkupsection{Inductive Predicates \label{sec:inductive}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-The \isa{predicate\ compiler} is an extension of the code generator
- which turns inductive specifications into equational ones, from
- which in turn executable code can be generated. The mechanisms of
- this compiler are described in detail in
- \cite{Berghofer-Bulwahn-Haftmann:2009:TPHOL}.
-
- Consider the simple predicate \isa{append} given by these two
- introduction rules:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-%
-\begin{isamarkuptext}%
-\isa{append\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ ys\ ys} \\
- \isa{append\ xs\ ys\ zs\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ append\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{23}{\isacharhash}}\ xs{\isaliteral{29}{\isacharparenright}}\ ys\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{23}{\isacharhash}}\ zs{\isaliteral{29}{\isacharparenright}}}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent To invoke the compiler, simply use \indexdef{}{command}{code\_pred}\hypertarget{command.code-pred}{\hyperlink{command.code-pred}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}pred}}}}}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}pred}\isamarkupfalse%
-\ append\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse%
-%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent The \hyperlink{command.code-pred}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}pred}}}} command takes the name of the
- inductive predicate and then you put a period to discharge a trivial
- correctness proof. The compiler infers possible modes for the
- predicate and produces the derived code equations. Modes annotate
- which (parts of the) arguments are to be taken as input, and which
- output. Modes are similar to types, but use the notation \isa{i}
- for input and \isa{o} for output.
-
- For \isa{append}, the compiler can infer the following modes:
- \begin{itemize}
- \item \isa{i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool}
- \item \isa{i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool}
- \item \isa{o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool}
- \end{itemize}
- You can compute sets of predicates using \indexdef{}{command}{values}\hypertarget{command.values}{\hyperlink{command.values}{\mbox{\isa{\isacommand{values}}}}}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{values}\isamarkupfalse%
-\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{7B}{\isacharbraceleft}}zs{\isaliteral{2E}{\isachardot}}\ append\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{1}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}nat{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}{\isadigit{2}}{\isaliteral{2C}{\isacharcomma}}{\isadigit{3}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isadigit{4}}{\isaliteral{2C}{\isacharcomma}}{\isadigit{5}}{\isaliteral{5D}{\isacharbrackright}}\ zs{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent outputs \isa{{\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{5B}{\isacharbrackleft}}{\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{2}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{3}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{4}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{5}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{7D}{\isacharbraceright}}}, and%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{values}\isamarkupfalse%
-\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{28}{\isacharparenleft}}xs{\isaliteral{2C}{\isacharcomma}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{2E}{\isachardot}}\ append\ xs\ ys\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{2}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}nat{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}{\isadigit{3}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent outputs \isa{{\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isadigit{2}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{3}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5B}{\isacharbrackleft}}{\isadigit{2}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isadigit{3}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5B}{\isacharbrackleft}}{\isadigit{2}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{3}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{7D}{\isacharbraceright}}}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-\noindent If you are only interested in the first elements of the
- set comprehension (with respect to a depth-first search on the
- introduction rules), you can pass an argument to \hyperlink{command.values}{\mbox{\isa{\isacommand{values}}}}
- to specify the number of elements you want:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{values}\isamarkupfalse%
-\ {\isadigit{1}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{28}{\isacharparenleft}}xs{\isaliteral{2C}{\isacharcomma}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{2E}{\isachardot}}\ append\ xs\ ys\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{1}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}nat{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{2}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{3}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{4}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isacommand{values}\isamarkupfalse%
-\ {\isadigit{3}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{28}{\isacharparenleft}}xs{\isaliteral{2C}{\isacharcomma}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{2E}{\isachardot}}\ append\ xs\ ys\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{1}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}nat{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{2}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{3}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{4}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent The \hyperlink{command.values}{\mbox{\isa{\isacommand{values}}}} command can only compute set
- comprehensions for which a mode has been inferred.
-
- The code equations for a predicate are made available as theorems with
- the suffix \isa{equation}, and can be inspected with:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{thm}\isamarkupfalse%
-\ append{\isaliteral{2E}{\isachardot}}equation%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent More advanced options are described in the following subsections.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Alternative names for functions%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-By default, the functions generated from a predicate are named after
- the predicate with the mode mangled into the name (e.g., \isa{append{\isaliteral{5F}{\isacharunderscore}}i{\isaliteral{5F}{\isacharunderscore}}i{\isaliteral{5F}{\isacharunderscore}}o}). You can specify your own names as follows:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}pred}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}modes{\isaliteral{3A}{\isacharcolon}}\ i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool\ as\ concat{\isaliteral{2C}{\isacharcomma}}\isanewline
-\ \ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool\ as\ split{\isaliteral{2C}{\isacharcomma}}\isanewline
-\ \ i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool\ as\ suffix{\isaliteral{29}{\isacharparenright}}\ append\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse%
-%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isamarkupsubsection{Alternative introduction rules%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Sometimes the introduction rules of an predicate are not executable
- because they contain non-executable constants or specific modes
- could not be inferred. It is also possible that the introduction
- rules yield a function that loops forever due to the execution in a
- depth-first search manner. Therefore, you can declare alternative
- introduction rules for predicates with the attribute \hyperlink{attribute.code-pred-intro}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}pred{\isaliteral{5F}{\isacharunderscore}}intro}}}. For example, the transitive closure is defined
- by:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-%
-\begin{isamarkuptext}%
-\isa{{\isaliteral{22}{\isachardoublequote}}r\ a\ b\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ tranclp\ r\ a\ b{\isaliteral{22}{\isachardoublequote}}}\\
- \isa{{\isaliteral{22}{\isachardoublequote}}tranclp\ r\ a\ b\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ r\ b\ c\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ tranclp\ r\ a\ c{\isaliteral{22}{\isachardoublequote}}}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent These rules do not suit well for executing the transitive
- closure with the mode \isa{{\isaliteral{28}{\isacharparenleft}}i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool}, as
- the second rule will cause an infinite loop in the recursive call.
- This can be avoided using the following alternative rules which are
- declared to the predicate compiler by the attribute \hyperlink{attribute.code-pred-intro}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}pred{\isaliteral{5F}{\isacharunderscore}}intro}}}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{lemma}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5F}{\isacharunderscore}}pred{\isaliteral{5F}{\isacharunderscore}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}r\ a\ b\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ tranclp\ r\ a\ b{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}r\ a\ b\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ tranclp\ r\ b\ c\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ tranclp\ r\ a\ c{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isacommand{by}\isamarkupfalse%
-\ auto%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent After declaring all alternative rules for the transitive
- closure, you invoke \hyperlink{command.code-pred}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}pred}}}} as usual. As you have
- declared alternative rules for the predicate, you are urged to prove
- that these introduction rules are complete, i.e., that you can
- derive an elimination rule for the alternative rules:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}pred}\isamarkupfalse%
-\ tranclp\isanewline
-\isacommand{proof}\isamarkupfalse%
-\ {\isaliteral{2D}{\isacharminus}}\isanewline
-\ \ \isacommand{case}\isamarkupfalse%
-\ tranclp\isanewline
-\ \ \isacommand{from}\isamarkupfalse%
-\ this\ converse{\isaliteral{5F}{\isacharunderscore}}tranclpE\ {\isaliteral{5B}{\isacharbrackleft}}OF\ tranclp{\isaliteral{2E}{\isachardot}}prems{\isaliteral{5D}{\isacharbrackright}}\ \isacommand{show}\isamarkupfalse%
-\ thesis\ \isacommand{by}\isamarkupfalse%
-\ metis\isanewline
-\isacommand{qed}\isamarkupfalse%
-%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent Alternative rules can also be used for constants that have
- not been defined inductively. For example, the lexicographic order
- which is defined as:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-%
-\begin{isamarkuptext}%
-\begin{isabelle}%
-lexordp\ r\ {\isaliteral{3F}{\isacharquery}}xs\ {\isaliteral{3F}{\isacharquery}}ys\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\isanewline
-{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}a\ v{\isaliteral{2E}{\isachardot}}\ {\isaliteral{3F}{\isacharquery}}ys\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{3F}{\isacharquery}}xs\ {\isaliteral{40}{\isacharat}}\ a\ {\isaliteral{23}{\isacharhash}}\ v\ {\isaliteral{5C3C6F723E}{\isasymor}}\isanewline
-\isaindent{{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}a\ v{\isaliteral{2E}{\isachardot}}\ }{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}u\ a\ b\ v\ w{\isaliteral{2E}{\isachardot}}\ r\ a\ b\ {\isaliteral{5C3C616E643E}{\isasymand}}\ {\isaliteral{3F}{\isacharquery}}xs\ {\isaliteral{3D}{\isacharequal}}\ u\ {\isaliteral{40}{\isacharat}}\ a\ {\isaliteral{23}{\isacharhash}}\ v\ {\isaliteral{5C3C616E643E}{\isasymand}}\ {\isaliteral{3F}{\isacharquery}}ys\ {\isaliteral{3D}{\isacharequal}}\ u\ {\isaliteral{40}{\isacharat}}\ b\ {\isaliteral{23}{\isacharhash}}\ w{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}%
-\end{isabelle}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent To make it executable, you can derive the following two
- rules and prove the elimination rule:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{lemma}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5F}{\isacharunderscore}}pred{\isaliteral{5F}{\isacharunderscore}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}append\ xs\ {\isaliteral{28}{\isacharparenleft}}a\ {\isaliteral{23}{\isacharhash}}\ v{\isaliteral{29}{\isacharparenright}}\ ys\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ lexordp\ r\ xs\ ys{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isacommand{lemma}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5F}{\isacharunderscore}}pred{\isaliteral{5F}{\isacharunderscore}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}append\ u\ {\isaliteral{28}{\isacharparenleft}}a\ {\isaliteral{23}{\isacharhash}}\ v{\isaliteral{29}{\isacharparenright}}\ xs\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ append\ u\ {\isaliteral{28}{\isacharparenleft}}b\ {\isaliteral{23}{\isacharhash}}\ w{\isaliteral{29}{\isacharparenright}}\ ys\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ r\ a\ b\isanewline
-\ \ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ lexordp\ r\ xs\ ys{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}pred}\isamarkupfalse%
-\ lexordp%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isamarkupsubsection{Options for values%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-In the presence of higher-order predicates, multiple modes for some
- predicate could be inferred that are not disambiguated by the
- pattern of the set comprehension. To disambiguate the modes for the
- arguments of a predicate, you can state the modes explicitly in the
- \hyperlink{command.values}{\mbox{\isa{\isacommand{values}}}} command. Consider the simple predicate \isa{succ}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{inductive}\isamarkupfalse%
-\ succ\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}succ\ {\isadigit{0}}\ {\isaliteral{28}{\isacharparenleft}}Suc\ {\isadigit{0}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}succ\ x\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ succ\ {\isaliteral{28}{\isacharparenleft}}Suc\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}Suc\ y{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}pred}\isamarkupfalse%
-\ succ\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse%
-%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent For this, the predicate compiler can infer modes \isa{o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool}, \isa{i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool}, \isa{o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool} and
- \isa{i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool}. The invocation of \hyperlink{command.values}{\mbox{\isa{\isacommand{values}}}}
- \isa{{\isaliteral{7B}{\isacharbraceleft}}n{\isaliteral{2E}{\isachardot}}\ tranclp\ succ\ {\isadigit{1}}{\isadigit{0}}\ n{\isaliteral{7D}{\isacharbraceright}}} loops, as multiple modes for the
- predicate \isa{succ} are possible and here the first mode \isa{o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool} is chosen. To choose another mode for the argument,
- you can declare the mode for the argument between the \hyperlink{command.values}{\mbox{\isa{\isacommand{values}}}} and the number of elements.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{values}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}mode{\isaliteral{3A}{\isacharcolon}}\ i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{5D}{\isacharbrackright}}\ {\isadigit{1}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{7B}{\isacharbraceleft}}n{\isaliteral{2E}{\isachardot}}\ tranclp\ succ\ {\isadigit{1}}{\isadigit{0}}\ n{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isanewline
-\isacommand{values}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}mode{\isaliteral{3A}{\isacharcolon}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{5D}{\isacharbrackright}}\ {\isadigit{1}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{7B}{\isacharbraceleft}}n{\isaliteral{2E}{\isachardot}}\ tranclp\ succ\ n\ {\isadigit{1}}{\isadigit{0}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isamarkupsubsection{Embedding into functional code within Isabelle/HOL%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-To embed the computation of an inductive predicate into functions
- that are defined in Isabelle/HOL, you have a number of options:
-
- \begin{itemize}
-
- \item You want to use the first-order predicate with the mode
- where all arguments are input. Then you can use the predicate directly, e.g.
-
- \begin{quote}
- \isa{valid{\isaliteral{5F}{\isacharunderscore}}suffix\ ys\ zs\ {\isaliteral{3D}{\isacharequal}}} \\
- \isa{{\isaliteral{28}{\isacharparenleft}}if\ append\ {\isaliteral{5B}{\isacharbrackleft}}Suc\ {\isadigit{0}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{2}}{\isaliteral{5D}{\isacharbrackright}}\ ys\ zs\ then\ Some\ ys\ else\ None{\isaliteral{29}{\isacharparenright}}}
- \end{quote}
-
- \item If you know that the execution returns only one value (it is
- deterministic), then you can use the combinator \isa{Predicate{\isaliteral{2E}{\isachardot}}the}, e.g., a functional concatenation of lists is
- defined with
-
- \begin{quote}
- \isa{functional{\isaliteral{5F}{\isacharunderscore}}concat\ xs\ ys\ {\isaliteral{3D}{\isacharequal}}\ Predicate{\isaliteral{2E}{\isachardot}}the\ {\isaliteral{28}{\isacharparenleft}}append{\isaliteral{5F}{\isacharunderscore}}i{\isaliteral{5F}{\isacharunderscore}}i{\isaliteral{5F}{\isacharunderscore}}o\ xs\ ys{\isaliteral{29}{\isacharparenright}}}
- \end{quote}
-
- Note that if the evaluation does not return a unique value, it
- raises a run-time error \isa{not{\isaliteral{5F}{\isacharunderscore}}unique}.
-
- \end{itemize}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Further Examples%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Further examples for compiling inductive predicates can be found in
- the \isa{HOL{\isaliteral{2F}{\isacharslash}}ex{\isaliteral{2F}{\isacharslash}}Predicate{\isaliteral{5F}{\isacharunderscore}}Compile{\isaliteral{5F}{\isacharunderscore}}ex{\isaliteral{2E}{\isachardot}}thy} theory file. There are
- also some examples in the Archive of Formal Proofs, notably in the
- \isa{POPLmark{\isaliteral{2D}{\isacharminus}}deBruijn} and the \isa{FeatherweightJava}
- sessions.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-\isacommand{end}\isamarkupfalse%
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-\isanewline
-\isanewline
-\end{isabellebody}%
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End:
--- a/doc-src/Codegen/Thy/document/Introduction.tex Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,597 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{Introduction}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-\isacommand{theory}\isamarkupfalse%
-\ Introduction\isanewline
-\isakeyword{imports}\ Setup\isanewline
-\isakeyword{begin}%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isamarkupsection{Introduction%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-This tutorial introduces the code generator facilities of \isa{Isabelle{\isaliteral{2F}{\isacharslash}}HOL}. It allows to turn (a certain class of) HOL
- specifications into corresponding executable code in the programming
- languages \isa{SML} \cite{SML}, \isa{OCaml} \cite{OCaml},
- \isa{Haskell} \cite{haskell-revised-report} and \isa{Scala}
- \cite{scala-overview-tech-report}.
-
- To profit from this tutorial, some familiarity and experience with
- \isa{HOL} \cite{isa-tutorial} and its basic theories is assumed.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Code generation principle: shallow embedding \label{sec:principle}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-The key concept for understanding Isabelle's code generation is
- \emph{shallow embedding}: logical entities like constants, types and
- classes are identified with corresponding entities in the target
- language. In particular, the carrier of a generated program's
- semantics are \emph{equational theorems} from the logic. If we view
- a generated program as an implementation of a higher-order rewrite
- system, then every rewrite step performed by the program can be
- simulated in the logic, which guarantees partial correctness
- \cite{Haftmann-Nipkow:2010:code}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{A quick start with the Isabelle/HOL toolbox \label{sec:queue_example}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-In a HOL theory, the \indexdef{}{command}{datatype}\hypertarget{command.datatype}{\hyperlink{command.datatype}{\mbox{\isa{\isacommand{datatype}}}}} and \indexdef{}{command}{definition}\hypertarget{command.definition}{\hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}}/\indexdef{}{command}{primrec}\hypertarget{command.primrec}{\hyperlink{command.primrec}{\mbox{\isa{\isacommand{primrec}}}}}/\indexdef{}{command}{fun}\hypertarget{command.fun}{\hyperlink{command.fun}{\mbox{\isa{\isacommand{fun}}}}} declarations
- form the core of a functional programming language. By default
- equational theorems stemming from those are used for generated code,
- therefore \qt{naive} code generation can proceed without further
- ado.
-
- For example, here a simple \qt{implementation} of amortised queues:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{datatype}\isamarkupfalse%
-\ {\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ list{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ list{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{definition}\isamarkupfalse%
-\ empty\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ queue{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}empty\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{primrec}\isamarkupfalse%
-\ enqueue\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ queue{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}enqueue\ x\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{23}{\isacharhash}}\ xs{\isaliteral{29}{\isacharparenright}}\ ys{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{fun}\isamarkupfalse%
-\ dequeue\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ option\ {\isaliteral{5C3C74696D65733E}{\isasymtimes}}\ {\isaliteral{27}{\isacharprime}}a\ queue{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}None{\isaliteral{2C}{\isacharcomma}}\ AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{23}{\isacharhash}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}Some\ y{\isaliteral{2C}{\isacharcomma}}\ AQueue\ xs\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}case\ rev\ xs\ of\ y\ {\isaliteral{23}{\isacharhash}}\ ys\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}Some\ y{\isaliteral{2C}{\isacharcomma}}\ AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ %
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isadeliminvisible
-%
-\endisadeliminvisible
-%
-\isataginvisible
-%
-\endisataginvisible
-{\isafoldinvisible}%
-%
-\isadeliminvisible
-%
-\endisadeliminvisible
-%
-\begin{isamarkuptext}%
-\noindent Then we can generate code e.g.~for \isa{SML} as follows:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{export{\isaliteral{5F}{\isacharunderscore}}code}\isamarkupfalse%
-\ empty\ dequeue\ enqueue\ \isakeyword{in}\ SML\isanewline
-\ \ \isakeyword{module{\isaliteral{5F}{\isacharunderscore}}name}\ Example\ \isakeyword{file}\ {\isaliteral{22}{\isachardoublequoteopen}}examples{\isaliteral{2F}{\isacharslash}}example{\isaliteral{2E}{\isachardot}}ML{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent resulting in the following code:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-structure\ Example\ {\isaliteral{3A}{\isacharcolon}}\ sig\isanewline
-\ \ val\ fold\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}b\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}b\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}b\isanewline
-\ \ val\ rev\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ list\isanewline
-\ \ datatype\ {\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ of\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{27}{\isacharprime}}a\ list\isanewline
-\ \ val\ empty\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ queue\isanewline
-\ \ val\ dequeue\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ option\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{27}{\isacharprime}}a\ queue\isanewline
-\ \ val\ enqueue\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ queue\isanewline
-end\ {\isaliteral{3D}{\isacharequal}}\ struct\isanewline
-\isanewline
-fun\ fold\ f\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ xs{\isaliteral{29}{\isacharparenright}}\ s\ {\isaliteral{3D}{\isacharequal}}\ fold\ f\ xs\ {\isaliteral{28}{\isacharparenleft}}f\ x\ s{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ fold\ f\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ s\ {\isaliteral{3D}{\isacharequal}}\ s{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ rev\ xs\ {\isaliteral{3D}{\isacharequal}}\ fold\ {\isaliteral{28}{\isacharparenleft}}fn\ a\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ fn\ b\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ a\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ b{\isaliteral{29}{\isacharparenright}}\ xs\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-datatype\ {\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ of\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{27}{\isacharprime}}a\ list{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-val\ empty\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}NONE{\isaliteral{2C}{\isacharcomma}}\ AQueue\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ {\isaliteral{28}{\isacharparenleft}}xs{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}SOME\ y{\isaliteral{2C}{\isacharcomma}}\ AQueue\ {\isaliteral{28}{\isacharparenleft}}xs{\isaliteral{2C}{\isacharcomma}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ {\isaliteral{28}{\isacharparenleft}}v\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ va{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ \ \ let\isanewline
-\ \ \ \ \ \ val\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ ys\ {\isaliteral{3D}{\isacharequal}}\ rev\ {\isaliteral{28}{\isacharparenleft}}v\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ va{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\ \ \ \ in\isanewline
-\ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}SOME\ y{\isaliteral{2C}{\isacharcomma}}\ AQueue\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2C}{\isacharcomma}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ \ \ end{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ enqueue\ x\ {\isaliteral{28}{\isacharparenleft}}AQueue\ {\isaliteral{28}{\isacharparenleft}}xs{\isaliteral{2C}{\isacharcomma}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ xs{\isaliteral{2C}{\isacharcomma}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-end{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}struct\ Example{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\begin{isamarkuptext}%
-\noindent The \indexdef{}{command}{export\_code}\hypertarget{command.export-code}{\hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isaliteral{5F}{\isacharunderscore}}code}}}}} command takes a
- space-separated list of constants for which code shall be generated;
- anything else needed for those is added implicitly. Then follows a
- target language identifier and a freely chosen module name. A file
- name denotes the destination to store the generated code. Note that
- the semantics of the destination depends on the target language: for
- \isa{SML}, \isa{OCaml} and \isa{Scala} it denotes a \emph{file},
- for \isa{Haskell} it denotes a \emph{directory} where a file named as the
- module name (with extension \isa{{\isaliteral{2E}{\isachardot}}hs}) is written:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{export{\isaliteral{5F}{\isacharunderscore}}code}\isamarkupfalse%
-\ empty\ dequeue\ enqueue\ \isakeyword{in}\ Haskell\isanewline
-\ \ \isakeyword{module{\isaliteral{5F}{\isacharunderscore}}name}\ Example\ \isakeyword{file}\ {\isaliteral{22}{\isachardoublequoteopen}}examples{\isaliteral{2F}{\isacharslash}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent This is the corresponding code:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-module\ Example\ where\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-\isanewline
-import\ Prelude\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2F}{\isacharslash}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3C}{\isacharless}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3C}{\isacharless}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2B}{\isacharplus}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2D}{\isacharminus}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2F}{\isacharslash}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\isanewline
-\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3C}{\isacharless}}{\isaliteral{3C}{\isacharless}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{26}{\isacharampersand}}{\isaliteral{26}{\isacharampersand}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5E}{\isacharcircum}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5E}{\isacharcircum}}{\isaliteral{5E}{\isacharcircum}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2E}{\isachardot}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{24}{\isachardollar}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{24}{\isachardollar}}{\isaliteral{21}{\isacharbang}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2B}{\isacharplus}}{\isaliteral{2B}{\isacharplus}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\isanewline
-\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{21}{\isacharbang}}{\isaliteral{21}{\isacharbang}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ Eq{\isaliteral{2C}{\isacharcomma}}\ error{\isaliteral{2C}{\isacharcomma}}\ id{\isaliteral{2C}{\isacharcomma}}\ return{\isaliteral{2C}{\isacharcomma}}\ not{\isaliteral{2C}{\isacharcomma}}\ fst{\isaliteral{2C}{\isacharcomma}}\ snd{\isaliteral{2C}{\isacharcomma}}\ map{\isaliteral{2C}{\isacharcomma}}\ filter{\isaliteral{2C}{\isacharcomma}}\ concat{\isaliteral{2C}{\isacharcomma}}\isanewline
-\ \ concatMap{\isaliteral{2C}{\isacharcomma}}\ reverse{\isaliteral{2C}{\isacharcomma}}\ zip{\isaliteral{2C}{\isacharcomma}}\ null{\isaliteral{2C}{\isacharcomma}}\ takeWhile{\isaliteral{2C}{\isacharcomma}}\ dropWhile{\isaliteral{2C}{\isacharcomma}}\ all{\isaliteral{2C}{\isacharcomma}}\ any{\isaliteral{2C}{\isacharcomma}}\ Integer{\isaliteral{2C}{\isacharcomma}}\isanewline
-\ \ negate{\isaliteral{2C}{\isacharcomma}}\ abs{\isaliteral{2C}{\isacharcomma}}\ divMod{\isaliteral{2C}{\isacharcomma}}\ String{\isaliteral{2C}{\isacharcomma}}\ Bool{\isaliteral{28}{\isacharparenleft}}True{\isaliteral{2C}{\isacharcomma}}\ False{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ Maybe{\isaliteral{28}{\isacharparenleft}}Nothing{\isaliteral{2C}{\isacharcomma}}\ Just{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-import\ qualified\ Prelude{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-data\ Queue\ a\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ {\isaliteral{5B}{\isacharbrackleft}}a{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5B}{\isacharbrackleft}}a{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-empty\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ forall\ a{\isaliteral{2E}{\isachardot}}\ Queue\ a{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-empty\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-dequeue\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ forall\ a{\isaliteral{2E}{\isachardot}}\ Queue\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{28}{\isacharparenleft}}Maybe\ a{\isaliteral{2C}{\isacharcomma}}\ Queue\ a{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}Nothing{\isaliteral{2C}{\isacharcomma}}\ AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{3A}{\isacharcolon}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}Just\ y{\isaliteral{2C}{\isacharcomma}}\ AQueue\ xs\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ {\isaliteral{28}{\isacharparenleft}}v\ {\isaliteral{3A}{\isacharcolon}}\ va{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ let\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-\ \ \ \ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{3A}{\isacharcolon}}\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ reverse\ {\isaliteral{28}{\isacharparenleft}}v\ {\isaliteral{3A}{\isacharcolon}}\ va{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\ \ {\isaliteral{7D}{\isacharbraceright}}\ in\ {\isaliteral{28}{\isacharparenleft}}Just\ y{\isaliteral{2C}{\isacharcomma}}\ AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-enqueue\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ forall\ a{\isaliteral{2E}{\isachardot}}\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ Queue\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ Queue\ a{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-enqueue\ x\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3A}{\isacharcolon}}\ xs{\isaliteral{29}{\isacharparenright}}\ ys{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-{\isaliteral{7D}{\isacharbraceright}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\begin{isamarkuptext}%
-\noindent For more details about \hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isaliteral{5F}{\isacharunderscore}}code}}}} see
- \secref{sec:further}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Type classes%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Code can also be generated from type classes in a Haskell-like
- manner. For illustration here an example from abstract algebra:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{class}\isamarkupfalse%
-\ semigroup\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ \isakeyword{fixes}\ mult\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infixl}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{7}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ \isakeyword{assumes}\ assoc{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ z\ {\isaliteral{3D}{\isacharequal}}\ x\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{class}\isamarkupfalse%
-\ monoid\ {\isaliteral{3D}{\isacharequal}}\ semigroup\ {\isaliteral{2B}{\isacharplus}}\isanewline
-\ \ \isakeyword{fixes}\ neutral\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6F6E653E}{\isasymone}}{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ \isakeyword{assumes}\ neutl{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6F6E653E}{\isasymone}}\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ x\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \ \ \isakeyword{and}\ neutr{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ {\isaliteral{5C3C6F6E653E}{\isasymone}}\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{instantiation}\isamarkupfalse%
-\ nat\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ monoid\isanewline
-\isakeyword{begin}\isanewline
-\isanewline
-\isacommand{primrec}\isamarkupfalse%
-\ mult{\isaliteral{5F}{\isacharunderscore}}nat\ \isakeyword{where}\isanewline
-\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isadigit{0}}\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ n\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isadigit{0}}{\isaliteral{5C3C436F6C6F6E3E}{\isasymColon}}nat{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}Suc\ m\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ n\ {\isaliteral{3D}{\isacharequal}}\ n\ {\isaliteral{2B}{\isacharplus}}\ m\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{definition}\isamarkupfalse%
-\ neutral{\isaliteral{5F}{\isacharunderscore}}nat\ \isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6F6E653E}{\isasymone}}\ {\isaliteral{3D}{\isacharequal}}\ Suc\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{lemma}\isamarkupfalse%
-\ add{\isaliteral{5F}{\isacharunderscore}}mult{\isaliteral{5F}{\isacharunderscore}}distrib{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ \isakeyword{fixes}\ n\ m\ q\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\isanewline
-\ \ \isakeyword{shows}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2B}{\isacharplus}}\ m{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ q\ {\isaliteral{3D}{\isacharequal}}\ n\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ q\ {\isaliteral{2B}{\isacharplus}}\ m\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ q{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}induct\ n{\isaliteral{29}{\isacharparenright}}\ simp{\isaliteral{5F}{\isacharunderscore}}all\isanewline
-\isanewline
-\isacommand{instance}\isamarkupfalse%
-\ \isacommand{proof}\isamarkupfalse%
-\isanewline
-\ \ \isacommand{fix}\isamarkupfalse%
-\ m\ n\ q\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isaliteral{22}{\isachardoublequoteopen}}m\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ n\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ q\ {\isaliteral{3D}{\isacharequal}}\ m\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ q{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}induct\ m{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}simp{\isaliteral{5F}{\isacharunderscore}}all\ add{\isaliteral{3A}{\isacharcolon}}\ add{\isaliteral{5F}{\isacharunderscore}}mult{\isaliteral{5F}{\isacharunderscore}}distrib{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6F6E653E}{\isasymone}}\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ n\ {\isaliteral{3D}{\isacharequal}}\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ neutral{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isaliteral{22}{\isachardoublequoteopen}}m\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ {\isaliteral{5C3C6F6E653E}{\isasymone}}\ {\isaliteral{3D}{\isacharequal}}\ m{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}induct\ m{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}simp{\isaliteral{5F}{\isacharunderscore}}all\ add{\isaliteral{3A}{\isacharcolon}}\ neutral{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\isanewline
-\isacommand{qed}\isamarkupfalse%
-\isanewline
-\isanewline
-\isacommand{end}\isamarkupfalse%
-%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent We define the natural operation of the natural numbers
- on monoids:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{primrec}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{in}\ monoid{\isaliteral{29}{\isacharparenright}}\ pow\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}pow\ {\isadigit{0}}\ a\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5C3C6F6E653E}{\isasymone}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}pow\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ a\ {\isaliteral{3D}{\isacharequal}}\ a\ {\isaliteral{5C3C6F74696D65733E}{\isasymotimes}}\ pow\ n\ a{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent This we use to define the discrete exponentiation
- function:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{definition}\isamarkupfalse%
-\ bexp\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}bexp\ n\ {\isaliteral{3D}{\isacharequal}}\ pow\ n\ {\isaliteral{28}{\isacharparenleft}}Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ {\isadigit{0}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent The corresponding code in Haskell uses that language's
- native classes:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-module\ Example\ where\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-\isanewline
-import\ Prelude\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2F}{\isacharslash}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3C}{\isacharless}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3C}{\isacharless}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2B}{\isacharplus}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2D}{\isacharminus}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2F}{\isacharslash}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\isanewline
-\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3C}{\isacharless}}{\isaliteral{3C}{\isacharless}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{26}{\isacharampersand}}{\isaliteral{26}{\isacharampersand}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5E}{\isacharcircum}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5E}{\isacharcircum}}{\isaliteral{5E}{\isacharcircum}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2E}{\isachardot}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{24}{\isachardollar}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{24}{\isachardollar}}{\isaliteral{21}{\isacharbang}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2B}{\isacharplus}}{\isaliteral{2B}{\isacharplus}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\isanewline
-\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{21}{\isacharbang}}{\isaliteral{21}{\isacharbang}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ Eq{\isaliteral{2C}{\isacharcomma}}\ error{\isaliteral{2C}{\isacharcomma}}\ id{\isaliteral{2C}{\isacharcomma}}\ return{\isaliteral{2C}{\isacharcomma}}\ not{\isaliteral{2C}{\isacharcomma}}\ fst{\isaliteral{2C}{\isacharcomma}}\ snd{\isaliteral{2C}{\isacharcomma}}\ map{\isaliteral{2C}{\isacharcomma}}\ filter{\isaliteral{2C}{\isacharcomma}}\ concat{\isaliteral{2C}{\isacharcomma}}\isanewline
-\ \ concatMap{\isaliteral{2C}{\isacharcomma}}\ reverse{\isaliteral{2C}{\isacharcomma}}\ zip{\isaliteral{2C}{\isacharcomma}}\ null{\isaliteral{2C}{\isacharcomma}}\ takeWhile{\isaliteral{2C}{\isacharcomma}}\ dropWhile{\isaliteral{2C}{\isacharcomma}}\ all{\isaliteral{2C}{\isacharcomma}}\ any{\isaliteral{2C}{\isacharcomma}}\ Integer{\isaliteral{2C}{\isacharcomma}}\isanewline
-\ \ negate{\isaliteral{2C}{\isacharcomma}}\ abs{\isaliteral{2C}{\isacharcomma}}\ divMod{\isaliteral{2C}{\isacharcomma}}\ String{\isaliteral{2C}{\isacharcomma}}\ Bool{\isaliteral{28}{\isacharparenleft}}True{\isaliteral{2C}{\isacharcomma}}\ False{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ Maybe{\isaliteral{28}{\isacharparenleft}}Nothing{\isaliteral{2C}{\isacharcomma}}\ Just{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-import\ qualified\ Prelude{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-data\ Nat\ {\isaliteral{3D}{\isacharequal}}\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{7C}{\isacharbar}}\ Suc\ Nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-plus{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ Nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ Nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ Nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-plus{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{28}{\isacharparenleft}}Suc\ m{\isaliteral{29}{\isacharparenright}}\ n\ {\isaliteral{3D}{\isacharequal}}\ plus{\isaliteral{5F}{\isacharunderscore}}nat\ m\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-plus{\isaliteral{5F}{\isacharunderscore}}nat\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ n\ {\isaliteral{3D}{\isacharequal}}\ n{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-class\ Semigroup\ a\ where\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-\ \ mult\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ a{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-class\ {\isaliteral{28}{\isacharparenleft}}Semigroup\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ Monoid\ a\ where\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-\ \ neutral\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ a{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-pow\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ forall\ a{\isaliteral{2E}{\isachardot}}\ {\isaliteral{28}{\isacharparenleft}}Monoid\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ Nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ a{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-pow\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ a\ {\isaliteral{3D}{\isacharequal}}\ neutral{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-pow\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ a\ {\isaliteral{3D}{\isacharequal}}\ mult\ a\ {\isaliteral{28}{\isacharparenleft}}pow\ n\ a{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-mult{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ Nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ Nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ Nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-mult{\isaliteral{5F}{\isacharunderscore}}nat\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ n\ {\isaliteral{3D}{\isacharequal}}\ Zero{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-mult{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{28}{\isacharparenleft}}Suc\ m{\isaliteral{29}{\isacharparenright}}\ n\ {\isaliteral{3D}{\isacharequal}}\ plus{\isaliteral{5F}{\isacharunderscore}}nat\ n\ {\isaliteral{28}{\isacharparenleft}}mult{\isaliteral{5F}{\isacharunderscore}}nat\ m\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-instance\ Semigroup\ Nat\ where\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-\ \ mult\ {\isaliteral{3D}{\isacharequal}}\ mult{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-neutral{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ Nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-neutral{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3D}{\isacharequal}}\ Suc\ Zero{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-instance\ Monoid\ Nat\ where\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-\ \ neutral\ {\isaliteral{3D}{\isacharequal}}\ neutral{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-bexp\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ Nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ Nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-bexp\ n\ {\isaliteral{3D}{\isacharequal}}\ pow\ n\ {\isaliteral{28}{\isacharparenleft}}Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ Zero{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-{\isaliteral{7D}{\isacharbraceright}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\begin{isamarkuptext}%
-\noindent This is a convenient place to show how explicit dictionary
- construction manifests in generated code -- the same example in
- \isa{SML}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-structure\ Example\ {\isaliteral{3A}{\isacharcolon}}\ sig\isanewline
-\ \ datatype\ nat\ {\isaliteral{3D}{\isacharequal}}\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{7C}{\isacharbar}}\ Suc\ of\ nat\isanewline
-\ \ val\ plus{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3A}{\isacharcolon}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ nat\isanewline
-\ \ type\ {\isaliteral{27}{\isacharprime}}a\ semigroup\isanewline
-\ \ val\ mult\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ semigroup\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\isanewline
-\ \ type\ {\isaliteral{27}{\isacharprime}}a\ monoid\isanewline
-\ \ val\ semigroup{\isaliteral{5F}{\isacharunderscore}}monoid\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ monoid\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ semigroup\isanewline
-\ \ val\ neutral\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ monoid\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\isanewline
-\ \ val\ pow\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ monoid\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\isanewline
-\ \ val\ mult{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3A}{\isacharcolon}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ nat\isanewline
-\ \ val\ semigroup{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3A}{\isacharcolon}}\ nat\ semigroup\isanewline
-\ \ val\ neutral{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3A}{\isacharcolon}}\ nat\isanewline
-\ \ val\ monoid{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3A}{\isacharcolon}}\ nat\ monoid\isanewline
-\ \ val\ bexp\ {\isaliteral{3A}{\isacharcolon}}\ nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ nat\isanewline
-end\ {\isaliteral{3D}{\isacharequal}}\ struct\isanewline
-\isanewline
-datatype\ nat\ {\isaliteral{3D}{\isacharequal}}\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{7C}{\isacharbar}}\ Suc\ of\ nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ plus{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{28}{\isacharparenleft}}Suc\ m{\isaliteral{29}{\isacharparenright}}\ n\ {\isaliteral{3D}{\isacharequal}}\ plus{\isaliteral{5F}{\isacharunderscore}}nat\ m\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ plus{\isaliteral{5F}{\isacharunderscore}}nat\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ n\ {\isaliteral{3D}{\isacharequal}}\ n{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-type\ {\isaliteral{27}{\isacharprime}}a\ semigroup\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B}{\isacharbraceleft}}mult\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-val\ mult\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{23}{\isacharhash}}mult\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ semigroup\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-type\ {\isaliteral{27}{\isacharprime}}a\ monoid\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B}{\isacharbraceleft}}semigroup{\isaliteral{5F}{\isacharunderscore}}monoid\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ semigroup{\isaliteral{2C}{\isacharcomma}}\ neutral\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-val\ semigroup{\isaliteral{5F}{\isacharunderscore}}monoid\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{23}{\isacharhash}}semigroup{\isaliteral{5F}{\isacharunderscore}}monoid\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ monoid\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ semigroup{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-val\ neutral\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{23}{\isacharhash}}neutral\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ monoid\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ pow\ A{\isaliteral{5F}{\isacharunderscore}}\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ a\ {\isaliteral{3D}{\isacharequal}}\ neutral\ A{\isaliteral{5F}{\isacharunderscore}}\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ pow\ A{\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ a\ {\isaliteral{3D}{\isacharequal}}\ mult\ {\isaliteral{28}{\isacharparenleft}}semigroup{\isaliteral{5F}{\isacharunderscore}}monoid\ A{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{29}{\isacharparenright}}\ a\ {\isaliteral{28}{\isacharparenleft}}pow\ A{\isaliteral{5F}{\isacharunderscore}}\ n\ a{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ mult{\isaliteral{5F}{\isacharunderscore}}nat\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ n\ {\isaliteral{3D}{\isacharequal}}\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ mult{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{28}{\isacharparenleft}}Suc\ m{\isaliteral{29}{\isacharparenright}}\ n\ {\isaliteral{3D}{\isacharequal}}\ plus{\isaliteral{5F}{\isacharunderscore}}nat\ n\ {\isaliteral{28}{\isacharparenleft}}mult{\isaliteral{5F}{\isacharunderscore}}nat\ m\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-val\ semigroup{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B}{\isacharbraceleft}}mult\ {\isaliteral{3D}{\isacharequal}}\ mult{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{7D}{\isacharbraceright}}\ {\isaliteral{3A}{\isacharcolon}}\ nat\ semigroup{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-val\ neutral{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3A}{\isacharcolon}}\ nat\ {\isaliteral{3D}{\isacharequal}}\ Suc\ Zero{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-val\ monoid{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B}{\isacharbraceleft}}semigroup{\isaliteral{5F}{\isacharunderscore}}monoid\ {\isaliteral{3D}{\isacharequal}}\ semigroup{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{2C}{\isacharcomma}}\ neutral\ {\isaliteral{3D}{\isacharequal}}\ neutral{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{7D}{\isacharbraceright}}\isanewline
-\ \ {\isaliteral{3A}{\isacharcolon}}\ nat\ monoid{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ bexp\ n\ {\isaliteral{3D}{\isacharequal}}\ pow\ monoid{\isaliteral{5F}{\isacharunderscore}}nat\ n\ {\isaliteral{28}{\isacharparenleft}}Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ Zero{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-end{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}struct\ Example{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\begin{isamarkuptext}%
-\noindent Note the parameters with trailing underscore (\verb|A_|), which are the dictionary parameters.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{How to continue from here%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-What you have seen so far should be already enough in a lot of
- cases. If you are content with this, you can quit reading here.
-
- Anyway, to understand situations where problems occur or to increase
- the scope of code generation beyond default, it is necessary to gain
- some understanding how the code generator actually works:
-
- \begin{itemize}
-
- \item The foundations of the code generator are described in
- \secref{sec:foundations}.
-
- \item In particular \secref{sec:utterly_wrong} gives hints how to
- debug situations where code generation does not succeed as
- expected.
-
- \item The scope and quality of generated code can be increased
- dramatically by applying refinement techniques, which are
- introduced in \secref{sec:refinement}.
-
- \item Inductive predicates can be turned executable using an
- extension of the code generator \secref{sec:inductive}.
-
- \item If you want to utilize code generation to obtain fast
- evaluators e.g.~for decision procedures, have a look at
- \secref{sec:evaluation}.
-
- \item You may want to skim over the more technical sections
- \secref{sec:adaptation} and \secref{sec:further}.
-
- \item The target language Scala \cite{scala-overview-tech-report}
- comes with some specialities discussed in \secref{sec:scala}.
-
- \item For exhaustive syntax diagrams etc. you should visit the
- Isabelle/Isar Reference Manual \cite{isabelle-isar-ref}.
-
- \end{itemize}
-
- \bigskip
-
- \begin{center}\fbox{\fbox{\begin{minipage}{8cm}
-
- \begin{center}\textit{Happy proving, happy hacking!}\end{center}
-
- \end{minipage}}}\end{center}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-\isacommand{end}\isamarkupfalse%
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-\isanewline
-\isanewline
-\end{isabellebody}%
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End:
--- a/doc-src/Codegen/Thy/document/Refinement.tex Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,660 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{Refinement}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-\isacommand{theory}\isamarkupfalse%
-\ Refinement\isanewline
-\isakeyword{imports}\ Setup\isanewline
-\isakeyword{begin}%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isamarkupsection{Program and datatype refinement \label{sec:refinement}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Code generation by shallow embedding (cf.~\secref{sec:principle})
- allows to choose code equations and datatype constructors freely,
- given that some very basic syntactic properties are met; this
- flexibility opens up mechanisms for refinement which allow to extend
- the scope and quality of generated code dramatically.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Program refinement%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Program refinement works by choosing appropriate code equations
- explicitly (cf.~\secref{sec:equations}); as example, we use Fibonacci
- numbers:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{fun}\isamarkupfalse%
-\ fib\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}fib\ {\isadigit{0}}\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}fib\ {\isaliteral{28}{\isacharparenleft}}Suc\ {\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Suc\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}fib\ {\isaliteral{28}{\isacharparenleft}}Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ fib\ n\ {\isaliteral{2B}{\isacharplus}}\ fib\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent The runtime of the corresponding code grows exponential due
- to two recursive calls:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-fib\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ Nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ Nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-fib\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3D}{\isacharequal}}\ Zero{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-fib\ {\isaliteral{28}{\isacharparenleft}}Suc\ Zero{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Suc\ Zero{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-fib\ {\isaliteral{28}{\isacharparenleft}}Suc\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ plus{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{28}{\isacharparenleft}}fib\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}fib\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\begin{isamarkuptext}%
-\noindent A more efficient implementation would use dynamic
- programming, e.g.~sharing of common intermediate results between
- recursive calls. This idea is expressed by an auxiliary operation
- which computes a Fibonacci number and its successor simultaneously:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{definition}\isamarkupfalse%
-\ fib{\isaliteral{5F}{\isacharunderscore}}step\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat\ {\isaliteral{5C3C74696D65733E}{\isasymtimes}}\ nat{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}fib{\isaliteral{5F}{\isacharunderscore}}step\ n\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}fib\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ fib\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent This operation can be implemented by recursion using
- dynamic programming:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{lemma}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}fib{\isaliteral{5F}{\isacharunderscore}}step\ {\isadigit{0}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}Suc\ {\isadigit{0}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{0}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}fib{\isaliteral{5F}{\isacharunderscore}}step\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}let\ {\isaliteral{28}{\isacharparenleft}}m{\isaliteral{2C}{\isacharcomma}}\ q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ fib{\isaliteral{5F}{\isacharunderscore}}step\ n\ in\ {\isaliteral{28}{\isacharparenleft}}m\ {\isaliteral{2B}{\isacharplus}}\ q{\isaliteral{2C}{\isacharcomma}}\ m{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}simp{\isaliteral{5F}{\isacharunderscore}}all\ add{\isaliteral{3A}{\isacharcolon}}\ fib{\isaliteral{5F}{\isacharunderscore}}step{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent What remains is to implement \isa{fib} by \isa{fib{\isaliteral{5F}{\isacharunderscore}}step} as follows:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{lemma}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}fib\ {\isadigit{0}}\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}fib\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ fst\ {\isaliteral{28}{\isacharparenleft}}fib{\isaliteral{5F}{\isacharunderscore}}step\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}simp{\isaliteral{5F}{\isacharunderscore}}all\ add{\isaliteral{3A}{\isacharcolon}}\ fib{\isaliteral{5F}{\isacharunderscore}}step{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent The resulting code shows only linear growth of runtime:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-fib{\isaliteral{5F}{\isacharunderscore}}step\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ Nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{28}{\isacharparenleft}}Nat{\isaliteral{2C}{\isacharcomma}}\ Nat{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-fib{\isaliteral{5F}{\isacharunderscore}}step\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ let\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}m{\isaliteral{2C}{\isacharcomma}}\ q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ fib{\isaliteral{5F}{\isacharunderscore}}step\ n{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{7D}{\isacharbraceright}}\ in\ {\isaliteral{28}{\isacharparenleft}}plus{\isaliteral{5F}{\isacharunderscore}}nat\ m\ q{\isaliteral{2C}{\isacharcomma}}\ m{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-fib{\isaliteral{5F}{\isacharunderscore}}step\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}Suc\ Zero{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{2C}{\isacharcomma}}\ Zero{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fib\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ Nat\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ Nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-fib\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ fst\ {\isaliteral{28}{\isacharparenleft}}fib{\isaliteral{5F}{\isacharunderscore}}step\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-fib\ Zero{\isaliteral{5F}{\isacharunderscore}}nat\ {\isaliteral{3D}{\isacharequal}}\ Zero{\isaliteral{5F}{\isacharunderscore}}nat{\isaliteral{3B}{\isacharsemicolon}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isamarkupsubsection{Datatype refinement%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Selecting specific code equations \emph{and} datatype constructors
- leads to datatype refinement. As an example, we will develop an
- alternative representation of the queue example given in
- \secref{sec:queue_example}. The amortised representation is
- convenient for generating code but exposes its \qt{implementation}
- details, which may be cumbersome when proving theorems about it.
- Therefore, here is a simple, straightforward representation of
- queues:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{datatype}\isamarkupfalse%
-\ {\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{3D}{\isacharequal}}\ Queue\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ list{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{definition}\isamarkupfalse%
-\ empty\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ queue{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}empty\ {\isaliteral{3D}{\isacharequal}}\ Queue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{primrec}\isamarkupfalse%
-\ enqueue\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ queue{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}enqueue\ x\ {\isaliteral{28}{\isacharparenleft}}Queue\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Queue\ {\isaliteral{28}{\isacharparenleft}}xs\ {\isaliteral{40}{\isacharat}}\ {\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{fun}\isamarkupfalse%
-\ dequeue\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ option\ {\isaliteral{5C3C74696D65733E}{\isasymtimes}}\ {\isaliteral{27}{\isacharprime}}a\ queue{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}dequeue\ {\isaliteral{28}{\isacharparenleft}}Queue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}None{\isaliteral{2C}{\isacharcomma}}\ Queue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}dequeue\ {\isaliteral{28}{\isacharparenleft}}Queue\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{23}{\isacharhash}}\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}Some\ x{\isaliteral{2C}{\isacharcomma}}\ Queue\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent This we can use directly for proving; for executing,
- we provide an alternative characterisation:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{definition}\isamarkupfalse%
-\ AQueue\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ queue{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}AQueue\ xs\ ys\ {\isaliteral{3D}{\isacharequal}}\ Queue\ {\isaliteral{28}{\isacharparenleft}}ys\ {\isaliteral{40}{\isacharat}}\ rev\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{code{\isaliteral{5F}{\isacharunderscore}}datatype}\isamarkupfalse%
-\ AQueue%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent Here we define a \qt{constructor} \isa{AQueue} which
- is defined in terms of \isa{Queue} and interprets its arguments
- according to what the \emph{content} of an amortised queue is supposed
- to be.
-
- The prerequisite for datatype constructors is only syntactical: a
- constructor must be of type \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6B617070613E}{\isasymkappa}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n} where \isa{{\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{7D}{\isacharbraceright}}} is exactly the set of \emph{all} type variables in
- \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}}; then \isa{{\isaliteral{5C3C6B617070613E}{\isasymkappa}}} is its corresponding datatype. The
- HOL datatype package by default registers any new datatype with its
- constructors, but this may be changed using \indexdef{}{command}{code\_datatype}\hypertarget{command.code-datatype}{\hyperlink{command.code-datatype}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}datatype}}}}}; the currently chosen constructors can be inspected
- using the \hyperlink{command.print-codesetup}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}codesetup}}}} command.
-
- Equipped with this, we are able to prove the following equations
- for our primitive queue operations which \qt{implement} the simple
- queues in an amortised fashion:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{lemma}\isamarkupfalse%
-\ empty{\isaliteral{5F}{\isacharunderscore}}AQueue\ {\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}empty\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ AQueue{\isaliteral{5F}{\isacharunderscore}}def\ empty{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\isanewline
-\isanewline
-\isacommand{lemma}\isamarkupfalse%
-\ enqueue{\isaliteral{5F}{\isacharunderscore}}AQueue\ {\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}enqueue\ x\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{23}{\isacharhash}}\ xs{\isaliteral{29}{\isacharparenright}}\ ys{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ AQueue{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}\isanewline
-\isanewline
-\isacommand{lemma}\isamarkupfalse%
-\ dequeue{\isaliteral{5F}{\isacharunderscore}}AQueue\ {\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ \ \ {\isaliteral{28}{\isacharparenleft}}if\ xs\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ then\ {\isaliteral{28}{\isacharparenleft}}None{\isaliteral{2C}{\isacharcomma}}\ AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ \ \ else\ dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{28}{\isacharparenleft}}rev\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{23}{\isacharhash}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}Some\ y{\isaliteral{2C}{\isacharcomma}}\ AQueue\ xs\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}simp{\isaliteral{5F}{\isacharunderscore}}all\ add{\isaliteral{3A}{\isacharcolon}}\ AQueue{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent It is good style, although no absolute requirement, to
- provide code equations for the original artefacts of the implemented
- type, if possible; in our case, these are the datatype constructor
- \isa{Queue} and the case combinator \isa{queue{\isaliteral{5F}{\isacharunderscore}}case}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{lemma}\isamarkupfalse%
-\ Queue{\isaliteral{5F}{\isacharunderscore}}AQueue\ {\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}Queue\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ AQueue{\isaliteral{5F}{\isacharunderscore}}def\ fun{\isaliteral{5F}{\isacharunderscore}}eq{\isaliteral{5F}{\isacharunderscore}}iff{\isaliteral{29}{\isacharparenright}}\isanewline
-\isanewline
-\isacommand{lemma}\isamarkupfalse%
-\ queue{\isaliteral{5F}{\isacharunderscore}}case{\isaliteral{5F}{\isacharunderscore}}AQueue\ {\isaliteral{5B}{\isacharbrackleft}}code{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}queue{\isaliteral{5F}{\isacharunderscore}}case\ f\ {\isaliteral{28}{\isacharparenleft}}AQueue\ xs\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ f\ {\isaliteral{28}{\isacharparenleft}}ys\ {\isaliteral{40}{\isacharat}}\ rev\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ AQueue{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent The resulting code looks as expected:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-structure\ Example\ {\isaliteral{3A}{\isacharcolon}}\ sig\isanewline
-\ \ val\ fold\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}b\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}b\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}b\isanewline
-\ \ val\ rev\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ list\isanewline
-\ \ val\ null\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ bool\isanewline
-\ \ datatype\ {\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ of\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{27}{\isacharprime}}a\ list\isanewline
-\ \ val\ empty\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ queue\isanewline
-\ \ val\ queue\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ queue\isanewline
-\ \ val\ dequeue\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ option\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{27}{\isacharprime}}a\ queue\isanewline
-\ \ val\ enqueue\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ queue\isanewline
-\ \ val\ queue{\isaliteral{5F}{\isacharunderscore}}case\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{27}{\isacharprime}}b\isanewline
-end\ {\isaliteral{3D}{\isacharequal}}\ struct\isanewline
-\isanewline
-fun\ fold\ f\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ xs{\isaliteral{29}{\isacharparenright}}\ s\ {\isaliteral{3D}{\isacharequal}}\ fold\ f\ xs\ {\isaliteral{28}{\isacharparenleft}}f\ x\ s{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ fold\ f\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ s\ {\isaliteral{3D}{\isacharequal}}\ s{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ rev\ xs\ {\isaliteral{3D}{\isacharequal}}\ fold\ {\isaliteral{28}{\isacharparenleft}}fn\ a\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ fn\ b\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ a\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ b{\isaliteral{29}{\isacharparenright}}\ xs\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ null\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ true\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ null\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ false{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-datatype\ {\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ of\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{27}{\isacharprime}}a\ list{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-val\ empty\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ queue\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ queue\ x\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2C}{\isacharcomma}}\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ {\isaliteral{28}{\isacharparenleft}}xs{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}SOME\ y{\isaliteral{2C}{\isacharcomma}}\ AQueue\ {\isaliteral{28}{\isacharparenleft}}xs{\isaliteral{2C}{\isacharcomma}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ {\isaliteral{7C}{\isacharbar}}\ dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ {\isaliteral{28}{\isacharparenleft}}xs{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ \ \ {\isaliteral{28}{\isacharparenleft}}if\ null\ xs\ then\ {\isaliteral{28}{\isacharparenleft}}NONE{\isaliteral{2C}{\isacharcomma}}\ AQueue\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ \ \ \ \ else\ dequeue\ {\isaliteral{28}{\isacharparenleft}}AQueue\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2C}{\isacharcomma}}\ rev\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ enqueue\ x\ {\isaliteral{28}{\isacharparenleft}}AQueue\ {\isaliteral{28}{\isacharparenleft}}xs{\isaliteral{2C}{\isacharcomma}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ AQueue\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ xs{\isaliteral{2C}{\isacharcomma}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-fun\ queue{\isaliteral{5F}{\isacharunderscore}}case\ f\ {\isaliteral{28}{\isacharparenleft}}AQueue\ {\isaliteral{28}{\isacharparenleft}}xs{\isaliteral{2C}{\isacharcomma}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ f\ {\isaliteral{28}{\isacharparenleft}}ys\ {\isaliteral{40}{\isacharat}}\ rev\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-end{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}struct\ Example{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\begin{isamarkuptext}%
-The same techniques can also be applied to types which are not
- specified as datatypes, e.g.~type \isa{int} is originally specified
- as quotient type by means of \indexdef{}{command}{typedef}\hypertarget{command.typedef}{\hyperlink{command.typedef}{\mbox{\isa{\isacommand{typedef}}}}}, but for code
- generation constants allowing construction of binary numeral values
- are used as constructors for \isa{int}.
-
- This approach however fails if the representation of a type demands
- invariants; this issue is discussed in the next section.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Datatype refinement involving invariants \label{sec:invariant}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Datatype representation involving invariants require a dedicated
- setup for the type and its primitive operations. As a running
- example, we implement a type \isa{{\isaliteral{27}{\isacharprime}}a\ dlist} of list consisting
- of distinct elements.
-
- The first step is to decide on which representation the abstract
- type (in our example \isa{{\isaliteral{27}{\isacharprime}}a\ dlist}) should be implemented.
- Here we choose \isa{{\isaliteral{27}{\isacharprime}}a\ list}. Then a conversion from the concrete
- type to the abstract type must be specified, here:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-%
-\begin{isamarkuptext}%
-\isa{Dlist\ {\isaliteral{5C3C436F6C6F6E3E}{\isasymColon}}\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ dlist}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent Next follows the specification of a suitable \emph{projection},
- i.e.~a conversion from abstract to concrete type:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-%
-\begin{isamarkuptext}%
-\isa{list{\isaliteral{5F}{\isacharunderscore}}of{\isaliteral{5F}{\isacharunderscore}}dlist\ {\isaliteral{5C3C436F6C6F6E3E}{\isasymColon}}\ {\isaliteral{27}{\isacharprime}}a\ dlist\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ list}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent This projection must be specified such that the following
- \emph{abstract datatype certificate} can be proven:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{lemma}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}code\ abstype{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}Dlist\ {\isaliteral{28}{\isacharparenleft}}list{\isaliteral{5F}{\isacharunderscore}}of{\isaliteral{5F}{\isacharunderscore}}dlist\ dxs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ dxs{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}fact\ Dlist{\isaliteral{5F}{\isacharunderscore}}list{\isaliteral{5F}{\isacharunderscore}}of{\isaliteral{5F}{\isacharunderscore}}dlist{\isaliteral{29}{\isacharparenright}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent Note that so far the invariant on representations
- (\isa{distinct\ {\isaliteral{5C3C436F6C6F6E3E}{\isasymColon}}\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool}) has never been mentioned explicitly:
- the invariant is only referred to implicitly: all values in
- set \isa{{\isaliteral{7B}{\isacharbraceleft}}xs{\isaliteral{2E}{\isachardot}}\ list{\isaliteral{5F}{\isacharunderscore}}of{\isaliteral{5F}{\isacharunderscore}}dlist\ {\isaliteral{28}{\isacharparenleft}}Dlist\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ xs{\isaliteral{7D}{\isacharbraceright}}} are invariant,
- and in our example this is exactly \isa{{\isaliteral{7B}{\isacharbraceleft}}xs{\isaliteral{2E}{\isachardot}}\ distinct\ xs{\isaliteral{7D}{\isacharbraceright}}}.
-
- The primitive operations on \isa{{\isaliteral{27}{\isacharprime}}a\ dlist} are specified
- indirectly using the projection \isa{list{\isaliteral{5F}{\isacharunderscore}}of{\isaliteral{5F}{\isacharunderscore}}dlist}. For
- the empty \isa{dlist}, \isa{Dlist{\isaliteral{2E}{\isachardot}}empty}, we finally want
- the code equation%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-%
-\begin{isamarkuptext}%
-\isa{Dlist{\isaliteral{2E}{\isachardot}}empty\ {\isaliteral{3D}{\isacharequal}}\ Dlist\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent This we have to prove indirectly as follows:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{lemma}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}code\ abstract{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}list{\isaliteral{5F}{\isacharunderscore}}of{\isaliteral{5F}{\isacharunderscore}}dlist\ Dlist{\isaliteral{2E}{\isachardot}}empty\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}fact\ list{\isaliteral{5F}{\isacharunderscore}}of{\isaliteral{5F}{\isacharunderscore}}dlist{\isaliteral{5F}{\isacharunderscore}}empty{\isaliteral{29}{\isacharparenright}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent This equation logically encodes both the desired code
- equation and that the expression \isa{Dlist} is applied to obeys
- the implicit invariant. Equations for insertion and removal are
- similar:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\isatagquote
-\isacommand{lemma}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}code\ abstract{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}list{\isaliteral{5F}{\isacharunderscore}}of{\isaliteral{5F}{\isacharunderscore}}dlist\ {\isaliteral{28}{\isacharparenleft}}Dlist{\isaliteral{2E}{\isachardot}}insert\ x\ dxs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ List{\isaliteral{2E}{\isachardot}}insert\ x\ {\isaliteral{28}{\isacharparenleft}}list{\isaliteral{5F}{\isacharunderscore}}of{\isaliteral{5F}{\isacharunderscore}}dlist\ dxs{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}fact\ list{\isaliteral{5F}{\isacharunderscore}}of{\isaliteral{5F}{\isacharunderscore}}dlist{\isaliteral{5F}{\isacharunderscore}}insert{\isaliteral{29}{\isacharparenright}}\isanewline
-\isanewline
-\isacommand{lemma}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}code\ abstract{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}list{\isaliteral{5F}{\isacharunderscore}}of{\isaliteral{5F}{\isacharunderscore}}dlist\ {\isaliteral{28}{\isacharparenleft}}Dlist{\isaliteral{2E}{\isachardot}}remove\ x\ dxs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ remove{\isadigit{1}}\ x\ {\isaliteral{28}{\isacharparenleft}}list{\isaliteral{5F}{\isacharunderscore}}of{\isaliteral{5F}{\isacharunderscore}}dlist\ dxs{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}fact\ list{\isaliteral{5F}{\isacharunderscore}}of{\isaliteral{5F}{\isacharunderscore}}dlist{\isaliteral{5F}{\isacharunderscore}}remove{\isaliteral{29}{\isacharparenright}}%
-\endisatagquote
-{\isafoldquote}%
-%
-\isadelimquote
-%
-\endisadelimquote
-%
-\begin{isamarkuptext}%
-\noindent Then the corresponding code is as follows:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\isatagquotetypewriter
-%
-\begin{isamarkuptext}%
-module\ Example\ where\ {\isaliteral{7B}{\isacharbraceleft}}\isanewline
-\isanewline
-import\ Prelude\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2F}{\isacharslash}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3C}{\isacharless}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3C}{\isacharless}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2B}{\isacharplus}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2D}{\isacharminus}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2F}{\isacharslash}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\isanewline
-\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{3E}{\isachargreater}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3C}{\isacharless}}{\isaliteral{3C}{\isacharless}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{26}{\isacharampersand}}{\isaliteral{26}{\isacharampersand}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5E}{\isacharcircum}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5E}{\isacharcircum}}{\isaliteral{5E}{\isacharcircum}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2E}{\isachardot}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{24}{\isachardollar}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{24}{\isachardollar}}{\isaliteral{21}{\isacharbang}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2B}{\isacharplus}}{\isaliteral{2B}{\isacharplus}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\isanewline
-\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{21}{\isacharbang}}{\isaliteral{21}{\isacharbang}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ Eq{\isaliteral{2C}{\isacharcomma}}\ error{\isaliteral{2C}{\isacharcomma}}\ id{\isaliteral{2C}{\isacharcomma}}\ return{\isaliteral{2C}{\isacharcomma}}\ not{\isaliteral{2C}{\isacharcomma}}\ fst{\isaliteral{2C}{\isacharcomma}}\ snd{\isaliteral{2C}{\isacharcomma}}\ map{\isaliteral{2C}{\isacharcomma}}\ filter{\isaliteral{2C}{\isacharcomma}}\ concat{\isaliteral{2C}{\isacharcomma}}\isanewline
-\ \ concatMap{\isaliteral{2C}{\isacharcomma}}\ reverse{\isaliteral{2C}{\isacharcomma}}\ zip{\isaliteral{2C}{\isacharcomma}}\ null{\isaliteral{2C}{\isacharcomma}}\ takeWhile{\isaliteral{2C}{\isacharcomma}}\ dropWhile{\isaliteral{2C}{\isacharcomma}}\ all{\isaliteral{2C}{\isacharcomma}}\ any{\isaliteral{2C}{\isacharcomma}}\ Integer{\isaliteral{2C}{\isacharcomma}}\isanewline
-\ \ negate{\isaliteral{2C}{\isacharcomma}}\ abs{\isaliteral{2C}{\isacharcomma}}\ divMod{\isaliteral{2C}{\isacharcomma}}\ String{\isaliteral{2C}{\isacharcomma}}\ Bool{\isaliteral{28}{\isacharparenleft}}True{\isaliteral{2C}{\isacharcomma}}\ False{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ Maybe{\isaliteral{28}{\isacharparenleft}}Nothing{\isaliteral{2C}{\isacharcomma}}\ Just{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-import\ qualified\ Prelude{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-newtype\ Dlist\ a\ {\isaliteral{3D}{\isacharequal}}\ Dlist\ {\isaliteral{5B}{\isacharbrackleft}}a{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-empty\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ forall\ a{\isaliteral{2E}{\isachardot}}\ Dlist\ a{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-empty\ {\isaliteral{3D}{\isacharequal}}\ Dlist\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-member\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ forall\ a{\isaliteral{2E}{\isachardot}}\ {\isaliteral{28}{\isacharparenleft}}Eq\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{5B}{\isacharbrackleft}}a{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ Bool{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-member\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ y\ {\isaliteral{3D}{\isacharequal}}\ False{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-member\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3A}{\isacharcolon}}\ xs{\isaliteral{29}{\isacharparenright}}\ y\ {\isaliteral{3D}{\isacharequal}}\ x\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}\ y\ {\isaliteral{7C}{\isacharbar}}{\isaliteral{7C}{\isacharbar}}\ member\ xs\ y{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-insert\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ forall\ a{\isaliteral{2E}{\isachardot}}\ {\isaliteral{28}{\isacharparenleft}}Eq\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{5B}{\isacharbrackleft}}a{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{5B}{\isacharbrackleft}}a{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-insert\ x\ xs\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}if\ member\ xs\ x\ then\ xs\ else\ x\ {\isaliteral{3A}{\isacharcolon}}\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-list{\isaliteral{5F}{\isacharunderscore}}of{\isaliteral{5F}{\isacharunderscore}}dlist\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ forall\ a{\isaliteral{2E}{\isachardot}}\ Dlist\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{5B}{\isacharbrackleft}}a{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-list{\isaliteral{5F}{\isacharunderscore}}of{\isaliteral{5F}{\isacharunderscore}}dlist\ {\isaliteral{28}{\isacharparenleft}}Dlist\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-inserta\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ forall\ a{\isaliteral{2E}{\isachardot}}\ {\isaliteral{28}{\isacharparenleft}}Eq\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ Dlist\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ Dlist\ a{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-inserta\ x\ dxs\ {\isaliteral{3D}{\isacharequal}}\ Dlist\ {\isaliteral{28}{\isacharparenleft}}insert\ x\ {\isaliteral{28}{\isacharparenleft}}list{\isaliteral{5F}{\isacharunderscore}}of{\isaliteral{5F}{\isacharunderscore}}dlist\ dxs{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-remove{\isadigit{1}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ forall\ a{\isaliteral{2E}{\isachardot}}\ {\isaliteral{28}{\isacharparenleft}}Eq\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{5B}{\isacharbrackleft}}a{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{5B}{\isacharbrackleft}}a{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-remove{\isadigit{1}}\ x\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-remove{\isadigit{1}}\ x\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{3A}{\isacharcolon}}\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}if\ x\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}\ y\ then\ xs\ else\ y\ {\isaliteral{3A}{\isacharcolon}}\ remove{\isadigit{1}}\ x\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-remove\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ forall\ a{\isaliteral{2E}{\isachardot}}\ {\isaliteral{28}{\isacharparenleft}}Eq\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ Dlist\ a\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{3E}{\isachargreater}}\ Dlist\ a{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-remove\ x\ dxs\ {\isaliteral{3D}{\isacharequal}}\ Dlist\ {\isaliteral{28}{\isacharparenleft}}remove{\isadigit{1}}\ x\ {\isaliteral{28}{\isacharparenleft}}list{\isaliteral{5F}{\isacharunderscore}}of{\isaliteral{5F}{\isacharunderscore}}dlist\ dxs{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\isanewline
-{\isaliteral{7D}{\isacharbraceright}}\isanewline%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagquotetypewriter
-{\isafoldquotetypewriter}%
-%
-\isadelimquotetypewriter
-%
-\endisadelimquotetypewriter
-%
-\begin{isamarkuptext}%
-Typical data structures implemented by representations involving
- invariants are available in the library, theory \isa{Mapping}
- specifies key-value-mappings (type \isa{{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ mapping});
- these can be implemented by red-black-trees (theory \isa{RBT}).%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-\isacommand{end}\isamarkupfalse%
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-\isanewline
-\isanewline
-\end{isabellebody}%
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End:
--- a/doc-src/Codegen/Thy/examples/Example.hs Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,28 +0,0 @@
-{-# LANGUAGE EmptyDataDecls, RankNTypes, ScopedTypeVariables #-}
-
-module Example where {
-
-import Prelude ((==), (/=), (<), (<=), (>=), (>), (+), (-), (*), (/),
- (**), (>>=), (>>), (=<<), (&&), (||), (^), (^^), (.), ($), ($!), (++),
- (!!), Eq, error, id, return, not, fst, snd, map, filter, concat,
- concatMap, reverse, zip, null, takeWhile, dropWhile, all, any, Integer,
- negate, abs, divMod, String, Bool(True, False), Maybe(Nothing, Just));
-import qualified Prelude;
-
-data Queue a = AQueue [a] [a];
-
-empty :: forall a. Queue a;
-empty = AQueue [] [];
-
-dequeue :: forall a. Queue a -> (Maybe a, Queue a);
-dequeue (AQueue [] []) = (Nothing, AQueue [] []);
-dequeue (AQueue xs (y : ys)) = (Just y, AQueue xs ys);
-dequeue (AQueue (v : va) []) =
- let {
- (y : ys) = reverse (v : va);
- } in (Just y, AQueue [] ys);
-
-enqueue :: forall a. a -> Queue a -> Queue a;
-enqueue x (AQueue xs ys) = AQueue (x : xs) ys;
-
-}
--- a/doc-src/Codegen/Thy/examples/example.ML Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,30 +0,0 @@
-structure Example : sig
- val fold : ('a -> 'b -> 'b) -> 'a list -> 'b -> 'b
- val rev : 'a list -> 'a list
- datatype 'a queue = AQueue of 'a list * 'a list
- val empty : 'a queue
- val dequeue : 'a queue -> 'a option * 'a queue
- val enqueue : 'a -> 'a queue -> 'a queue
-end = struct
-
-fun fold f (x :: xs) s = fold f xs (f x s)
- | fold f [] s = s;
-
-fun rev xs = fold (fn a => fn b => a :: b) xs [];
-
-datatype 'a queue = AQueue of 'a list * 'a list;
-
-val empty : 'a queue = AQueue ([], []);
-
-fun dequeue (AQueue ([], [])) = (NONE, AQueue ([], []))
- | dequeue (AQueue (xs, y :: ys)) = (SOME y, AQueue (xs, ys))
- | dequeue (AQueue (v :: va, [])) =
- let
- val y :: ys = rev (v :: va);
- in
- (SOME y, AQueue ([], ys))
- end;
-
-fun enqueue x (AQueue (xs, ys)) = AQueue (x :: xs, ys);
-
-end; (*struct Example*)
--- a/doc-src/Codegen/Thy/pictures/adaptation.tex Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,52 +0,0 @@
-
-\documentclass[12pt]{article}
-\usepackage{tikz}
-
-\begin{document}
-
-\thispagestyle{empty}
-\setlength{\fboxrule}{0.01pt}
-\setlength{\fboxsep}{4pt}
-
-\fcolorbox{white}{white}{
-
-\begin{tikzpicture}[scale = 0.5]
- \tikzstyle water=[color = blue, thick]
- \tikzstyle ice=[color = black, very thick, cap = round, join = round, fill = white]
- \tikzstyle process=[color = green, semithick, ->]
- \tikzstyle adaptation=[color = red, semithick, ->]
- \tikzstyle target=[color = black]
- \foreach \x in {0, ..., 24}
- \draw[style=water] (\x, 0.25) sin + (0.25, 0.25) cos + (0.25, -0.25) sin
- + (0.25, -0.25) cos + (0.25, 0.25);
- \draw[style=ice] (1, 0) --
- (3, 6) node[above, fill=white] {logic} -- (5, 0) -- cycle;
- \draw[style=ice] (9, 0) --
- (11, 6) node[above, fill=white] {intermediate language} -- (13, 0) -- cycle;
- \draw[style=ice] (15, -6) --
- (19, 6) node[above, fill=white] {target language} -- (23, -6) -- cycle;
- \draw[style=process]
- (3.5, 3) .. controls (7, 5) .. node[fill=white] {translation} (10.5, 3);
- \draw[style=process]
- (11.5, 3) .. controls (15, 5) .. node[fill=white] (serialisation) {serialisation} (18.5, 3);
- \node (adaptation) at (11, -2) [style=adaptation] {adaptation};
- \node at (19, 3) [rotate=90] {generated};
- \node at (19.5, -5) {language};
- \node at (19.5, -3) {library};
- \node (includes) at (19.5, -1) {includes};
- \node (reserved) at (16.5, -3) [rotate=72] {reserved}; % proper 71.57
- \draw[style=process]
- (includes) -- (serialisation);
- \draw[style=process]
- (reserved) -- (serialisation);
- \draw[style=adaptation]
- (adaptation) -- (serialisation);
- \draw[style=adaptation]
- (adaptation) -- (includes);
- \draw[style=adaptation]
- (adaptation) -- (reserved);
-\end{tikzpicture}
-
-}
-
-\end{document}
--- a/doc-src/Codegen/Thy/pictures/architecture.tex Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,50 +0,0 @@
-
-\documentclass[12pt]{article}
-\usepackage{tikz}
-\usetikzlibrary{shapes}
-\usetikzlibrary{arrows}
-
-\begin{document}
-
-\thispagestyle{empty}
-\setlength{\fboxrule}{0.01pt}
-\setlength{\fboxsep}{4pt}
-
-\fcolorbox{white}{white}{
-
-\newcommand{\sys}[1]{\emph{#1}}
-
-\begin{tikzpicture}[x = 4cm, y = 1cm]
- \tikzstyle positive=[color = black, fill = white];
- \tikzstyle negative=[color = white, fill = black];
- \tikzstyle entity=[rounded corners, draw, thick];
- \tikzstyle process=[ellipse, draw, thick];
- \tikzstyle arrow=[-stealth, semithick];
- \node (spec) at (0, 3) [entity, positive] {specification tools};
- \node (user) at (1, 3) [entity, positive] {user proofs};
- \node (spec_user_join) at (0.5, 3) [shape=coordinate] {};
- \node (raw) at (0.5, 4) [entity, positive] {raw code equations};
- \node (pre) at (1.5, 4) [process, positive] {preprocessing};
- \node (eqn) at (2.5, 4) [entity, positive] {code equations};
- \node (iml) at (0.5, 0) [entity, positive] {intermediate program};
- \node (seri) at (1.5, 0) [process, positive] {serialisation};
- \node (SML) at (2.5, 3) [entity, positive] {\sys{SML}};
- \node (OCaml) at (2.5, 2) [entity, positive] {\sys{OCaml}};
- \node (Haskell) at (2.5, 1) [entity, positive] {\sys{Haskell}};
- \node (Scala) at (2.5, 0) [entity, positive] {\sys{Scala}};
- \draw [semithick] (spec) -- (spec_user_join);
- \draw [semithick] (user) -- (spec_user_join);
- \draw [-diamond, semithick] (spec_user_join) -- (raw);
- \draw [arrow] (raw) -- (pre);
- \draw [arrow] (pre) -- (eqn);
- \draw [arrow] (eqn) -- node (transl) [process, positive] {translation} (iml);
- \draw [arrow] (iml) -- (seri);
- \draw [arrow] (seri) -- (SML);
- \draw [arrow] (seri) -- (OCaml);
- \draw [arrow] (seri) -- (Haskell);
- \draw [arrow] (seri) -- (Scala);
-\end{tikzpicture}
-
-}
-
-\end{document}
--- a/doc-src/Codegen/codegen.tex Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,53 +0,0 @@
-
-\documentclass[12pt,a4paper,fleqn]{article}
-\usepackage{latexsym,graphicx}
-\usepackage[refpage]{nomencl}
-\usepackage{multirow}
-\usepackage{../iman,../extra,../isar,../proof}
-\usepackage{../../lib/texinputs/isabelle,../../lib/texinputs/isabellesym}
-\usepackage{style}
-\usepackage{../pdfsetup}
-
-\hyphenation{Isabelle}
-\hyphenation{Isar}
-\isadroptag{theory}
-
-\title{\includegraphics[scale=0.5]{isabelle_isar}
- \\[4ex] Code generation from Isabelle/HOL theories}
-\author{\emph{Florian Haftmann with contributions from Lukas Bulwahn}}
-
-\begin{document}
-
-\maketitle
-
-\begin{abstract}
- \noindent This tutorial introduces the code generator facilities of Isabelle/HOL.
- They empower the user to turn HOL specifications into corresponding executable
- programs in the languages SML, OCaml, Haskell and Scala.
-\end{abstract}
-
-\thispagestyle{empty}\clearpage
-
-\pagenumbering{roman}
-\clearfirst
-
-\input{Thy/document/Introduction.tex}
-\input{Thy/document/Foundations.tex}
-\input{Thy/document/Refinement.tex}
-\input{Thy/document/Inductive_Predicate.tex}
-\input{Thy/document/Adaptation.tex}
-\input{Thy/document/Evaluation.tex}
-\input{Thy/document/Further.tex}
-
-\begingroup
-\bibliographystyle{plain} \small\raggedright\frenchspacing
-\bibliography{../manual}
-\endgroup
-
-\end{document}
-
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% End:
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Codegen/document/adaptation.tex Mon Aug 27 23:00:38 2012 +0200
@@ -0,0 +1,52 @@
+
+\documentclass[12pt]{article}
+\usepackage{tikz}
+
+\begin{document}
+
+\thispagestyle{empty}
+\setlength{\fboxrule}{0.01pt}
+\setlength{\fboxsep}{4pt}
+
+\fcolorbox{white}{white}{
+
+\begin{tikzpicture}[scale = 0.5]
+ \tikzstyle water=[color = blue, thick]
+ \tikzstyle ice=[color = black, very thick, cap = round, join = round, fill = white]
+ \tikzstyle process=[color = green, semithick, ->]
+ \tikzstyle adaptation=[color = red, semithick, ->]
+ \tikzstyle target=[color = black]
+ \foreach \x in {0, ..., 24}
+ \draw[style=water] (\x, 0.25) sin + (0.25, 0.25) cos + (0.25, -0.25) sin
+ + (0.25, -0.25) cos + (0.25, 0.25);
+ \draw[style=ice] (1, 0) --
+ (3, 6) node[above, fill=white] {logic} -- (5, 0) -- cycle;
+ \draw[style=ice] (9, 0) --
+ (11, 6) node[above, fill=white] {intermediate language} -- (13, 0) -- cycle;
+ \draw[style=ice] (15, -6) --
+ (19, 6) node[above, fill=white] {target language} -- (23, -6) -- cycle;
+ \draw[style=process]
+ (3.5, 3) .. controls (7, 5) .. node[fill=white] {translation} (10.5, 3);
+ \draw[style=process]
+ (11.5, 3) .. controls (15, 5) .. node[fill=white] (serialisation) {serialisation} (18.5, 3);
+ \node (adaptation) at (11, -2) [style=adaptation] {adaptation};
+ \node at (19, 3) [rotate=90] {generated};
+ \node at (19.5, -5) {language};
+ \node at (19.5, -3) {library};
+ \node (includes) at (19.5, -1) {includes};
+ \node (reserved) at (16.5, -3) [rotate=72] {reserved}; % proper 71.57
+ \draw[style=process]
+ (includes) -- (serialisation);
+ \draw[style=process]
+ (reserved) -- (serialisation);
+ \draw[style=adaptation]
+ (adaptation) -- (serialisation);
+ \draw[style=adaptation]
+ (adaptation) -- (includes);
+ \draw[style=adaptation]
+ (adaptation) -- (reserved);
+\end{tikzpicture}
+
+}
+
+\end{document}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Codegen/document/architecture.tex Mon Aug 27 23:00:38 2012 +0200
@@ -0,0 +1,50 @@
+
+\documentclass[12pt]{article}
+\usepackage{tikz}
+\usetikzlibrary{shapes}
+\usetikzlibrary{arrows}
+
+\begin{document}
+
+\thispagestyle{empty}
+\setlength{\fboxrule}{0.01pt}
+\setlength{\fboxsep}{4pt}
+
+\fcolorbox{white}{white}{
+
+\newcommand{\sys}[1]{\emph{#1}}
+
+\begin{tikzpicture}[x = 4cm, y = 1cm]
+ \tikzstyle positive=[color = black, fill = white];
+ \tikzstyle negative=[color = white, fill = black];
+ \tikzstyle entity=[rounded corners, draw, thick];
+ \tikzstyle process=[ellipse, draw, thick];
+ \tikzstyle arrow=[-stealth, semithick];
+ \node (spec) at (0, 3) [entity, positive] {specification tools};
+ \node (user) at (1, 3) [entity, positive] {user proofs};
+ \node (spec_user_join) at (0.5, 3) [shape=coordinate] {};
+ \node (raw) at (0.5, 4) [entity, positive] {raw code equations};
+ \node (pre) at (1.5, 4) [process, positive] {preprocessing};
+ \node (eqn) at (2.5, 4) [entity, positive] {code equations};
+ \node (iml) at (0.5, 0) [entity, positive] {intermediate program};
+ \node (seri) at (1.5, 0) [process, positive] {serialisation};
+ \node (SML) at (2.5, 3) [entity, positive] {\sys{SML}};
+ \node (OCaml) at (2.5, 2) [entity, positive] {\sys{OCaml}};
+ \node (Haskell) at (2.5, 1) [entity, positive] {\sys{Haskell}};
+ \node (Scala) at (2.5, 0) [entity, positive] {\sys{Scala}};
+ \draw [semithick] (spec) -- (spec_user_join);
+ \draw [semithick] (user) -- (spec_user_join);
+ \draw [-diamond, semithick] (spec_user_join) -- (raw);
+ \draw [arrow] (raw) -- (pre);
+ \draw [arrow] (pre) -- (eqn);
+ \draw [arrow] (eqn) -- node (transl) [process, positive] {translation} (iml);
+ \draw [arrow] (iml) -- (seri);
+ \draw [arrow] (seri) -- (SML);
+ \draw [arrow] (seri) -- (OCaml);
+ \draw [arrow] (seri) -- (Haskell);
+ \draw [arrow] (seri) -- (Scala);
+\end{tikzpicture}
+
+}
+
+\end{document}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Codegen/document/build Mon Aug 27 23:00:38 2012 +0200
@@ -0,0 +1,32 @@
+#!/bin/bash
+
+set -e
+
+FORMAT="$1"
+VARIANT="$2"
+
+"$ISABELLE_TOOL" logo -o isabelle_isar.pdf "Isar"
+"$ISABELLE_TOOL" logo -o isabelle_isar.eps "Isar"
+
+cp "$ISABELLE_HOME/doc-src/iman.sty" .
+cp "$ISABELLE_HOME/doc-src/extra.sty" .
+cp "$ISABELLE_HOME/doc-src/isar.sty" .
+cp "$ISABELLE_HOME/doc-src/proof.sty" .
+cp "$ISABELLE_HOME/doc-src/manual.bib" .
+
+for NAME in architecture adaptation
+do
+ latex "$NAME"
+ $ISABELLE_DVIPS -E -o "$NAME.eps" "$NAME.dvi"
+ $ISABELLE_EPSTOPDF "$NAME.eps"
+done
+
+"$ISABELLE_TOOL" latex -o sty
+"$ISABELLE_TOOL" latex -o "$FORMAT"
+"$ISABELLE_TOOL" latex -o bbl
+"$ISABELLE_TOOL" latex -o "$FORMAT"
+"$ISABELLE_TOOL" latex -o "$FORMAT"
+"$ISABELLE_HOME/doc-src/sedindex" root
+[ -f root.out ] && "$ISABELLE_HOME/doc-src/fixbookmarks" root.out
+"$ISABELLE_TOOL" latex -o "$FORMAT"
+"$ISABELLE_TOOL" latex -o "$FORMAT"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Codegen/document/root.tex Mon Aug 27 23:00:38 2012 +0200
@@ -0,0 +1,52 @@
+
+\documentclass[12pt,a4paper,fleqn]{article}
+\usepackage{latexsym,graphicx}
+\usepackage{multirow}
+\usepackage{iman,extra,isar,proof}
+\usepackage{isabelle,isabellesym}
+\usepackage{style}
+\usepackage{pdfsetup}
+
+\hyphenation{Isabelle}
+\hyphenation{Isar}
+\isadroptag{theory}
+
+\title{\includegraphics[scale=0.5]{isabelle_isar}
+ \\[4ex] Code generation from Isabelle/HOL theories}
+\author{\emph{Florian Haftmann with contributions from Lukas Bulwahn}}
+
+\begin{document}
+
+\maketitle
+
+\begin{abstract}
+ \noindent This tutorial introduces the code generator facilities of Isabelle/HOL.
+ They empower the user to turn HOL specifications into corresponding executable
+ programs in the languages SML, OCaml, Haskell and Scala.
+\end{abstract}
+
+\thispagestyle{empty}\clearpage
+
+\pagenumbering{roman}
+\clearfirst
+
+\input{Introduction.tex}
+\input{Foundations.tex}
+\input{Refinement.tex}
+\input{Inductive_Predicate.tex}
+\input{Adaptation.tex}
+\input{Evaluation.tex}
+\input{Further.tex}
+
+\begingroup
+\bibliographystyle{plain} \small\raggedright\frenchspacing
+\bibliography{manual}
+\endgroup
+
+\end{document}
+
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: t
+%%% End:
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Codegen/document/style.sty Mon Aug 27 23:00:38 2012 +0200
@@ -0,0 +1,75 @@
+
+%% toc
+\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1}
+\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}}
+
+%% paragraphs
+\setlength{\parindent}{1em}
+
+%% references
+\newcommand{\secref}[1]{\S\ref{#1}}
+\newcommand{\figref}[1]{figure~\ref{#1}}
+
+%% logical markup
+\newcommand{\strong}[1]{{\bfseries {#1}}}
+\newcommand{\qn}[1]{\emph{#1}}
+
+%% typographic conventions
+\newcommand{\qt}[1]{``{#1}''}
+\newcommand{\ditem}[1]{\item[\isastyletext #1]}
+
+%% quote environment
+\isakeeptag{quote}
+\renewenvironment{quote}
+ {\list{}{\leftmargin2em\rightmargin0pt}\parindent0pt\parskip0pt\item\relax}
+ {\endlist}
+\renewcommand{\isatagquote}{\begin{quote}}
+\renewcommand{\endisatagquote}{\end{quote}}
+\newcommand{\quotebreak}{\\[1.2ex]}
+
+%% typewriter text
+\newenvironment{typewriter}{\renewcommand{\isastyletext}{}%
+\renewcommand{\isadigit}[1]{{##1}}%
+\parindent0pt%
+\makeatletter\isa@parindent0pt\makeatother%
+\isabellestyle{tt}\isastyle%
+\fontsize{9pt}{9pt}\selectfont}{}
+
+\isakeeptag{quotetypewriter}
+\renewcommand{\isatagquotetypewriter}{\begin{quote}\begin{typewriter}}
+\renewcommand{\endisatagquotetypewriter}{\end{typewriter}\end{quote}}
+
+\isakeeptag{quotett}
+\renewcommand{\isatagquotett}{\begin{quote}\isabellestyle{tt}\isastyle}
+\renewcommand{\endisatagquotett}{\end{quote}}
+
+%% a trick
+\newcommand{\isasymSML}{SML}
+\newcommand{\isasymSMLdummy}{SML}
+
+%% presentation
+\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2}
+
+%% character detail
+\renewcommand{\isadigit}[1]{\isamath{#1}}
+\binperiod
+\underscoreoff
+
+%% format
+\pagestyle{headings}
+\isabellestyle{it}
+
+%% ml reference
+\newenvironment{mldecls}{\par\noindent\begingroup\footnotesize\def\isanewline{\\}\begin{tabular}{l}}{\end{tabular}\smallskip\endgroup}
+
+\isakeeptag{mlref}
+\renewcommand{\isatagmlref}{\subsection*{\makebox[0pt][r]{\fbox{\ML}~~}Reference}\begingroup\def\isastyletext{\rm}\small}
+\renewcommand{\endisatagmlref}{\endgroup}
+
+\isabellestyle{it}
+
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: "implementation"
+%%% End:
--- a/doc-src/Codegen/style.sty Mon Aug 27 22:31:16 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,75 +0,0 @@
-
-%% toc
-\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1}
-\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}}
-
-%% paragraphs
-\setlength{\parindent}{1em}
-
-%% references
-\newcommand{\secref}[1]{\S\ref{#1}}
-\newcommand{\figref}[1]{figure~\ref{#1}}
-
-%% logical markup
-\newcommand{\strong}[1]{{\bfseries {#1}}}
-\newcommand{\qn}[1]{\emph{#1}}
-
-%% typographic conventions
-\newcommand{\qt}[1]{``{#1}''}
-\newcommand{\ditem}[1]{\item[\isastyletext #1]}
-
-%% quote environment
-\isakeeptag{quote}
-\renewenvironment{quote}
- {\list{}{\leftmargin2em\rightmargin0pt}\parindent0pt\parskip0pt\item\relax}
- {\endlist}
-\renewcommand{\isatagquote}{\begin{quote}}
-\renewcommand{\endisatagquote}{\end{quote}}
-\newcommand{\quotebreak}{\\[1.2ex]}
-
-%% typewriter text
-\newenvironment{typewriter}{\renewcommand{\isastyletext}{}%
-\renewcommand{\isadigit}[1]{{##1}}%
-\parindent0pt%
-\makeatletter\isa@parindent0pt\makeatother%
-\isabellestyle{tt}\isastyle%
-\fontsize{9pt}{9pt}\selectfont}{}
-
-\isakeeptag{quotetypewriter}
-\renewcommand{\isatagquotetypewriter}{\begin{quote}\begin{typewriter}}
-\renewcommand{\endisatagquotetypewriter}{\end{typewriter}\end{quote}}
-
-\isakeeptag{quotett}
-\renewcommand{\isatagquotett}{\begin{quote}\isabellestyle{tt}\isastyle}
-\renewcommand{\endisatagquotett}{\end{quote}}
-
-%% a trick
-\newcommand{\isasymSML}{SML}
-\newcommand{\isasymSMLdummy}{SML}
-
-%% presentation
-\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2}
-
-%% character detail
-\renewcommand{\isadigit}[1]{\isamath{#1}}
-\binperiod
-\underscoreoff
-
-%% format
-\pagestyle{headings}
-\isabellestyle{it}
-
-%% ml reference
-\newenvironment{mldecls}{\par\noindent\begingroup\footnotesize\def\isanewline{\\}\begin{tabular}{l}}{\end{tabular}\smallskip\endgroup}
-
-\isakeeptag{mlref}
-\renewcommand{\isatagmlref}{\subsection*{\makebox[0pt][r]{\fbox{\ML}~~}Reference}\begingroup\def\isastyletext{\rm}\small}
-\renewcommand{\endisatagmlref}{\endgroup}
-
-\isabellestyle{it}
-
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "implementation"
-%%% End:
--- a/doc-src/ROOT Mon Aug 27 22:31:16 2012 +0200
+++ b/doc-src/ROOT Mon Aug 27 23:00:38 2012 +0200
@@ -7,10 +7,8 @@
"document/root.tex"
"document/style.sty"
-session Codegen (doc) in "Codegen/Thy" = "HOL-Library" +
- options [browser_info = false, document = false,
- document_dump = document, document_dump_mode = "tex",
- print_mode = "no_brackets,iff"]
+session Codegen (doc) in "Codegen" = "HOL-Library" +
+ options [document_variants = "codegen", print_mode = "no_brackets,iff"]
theories [document = false] Setup
theories
Introduction
@@ -20,6 +18,12 @@
Evaluation
Adaptation
Further
+ files
+ "document/adaptation.tex"
+ "document/architecture.tex"
+ "document/build"
+ "document/root.tex"
+ "document/style.sty"
session Functions (doc) in "Functions" = HOL +
options [document_variants = "functions"]