| author | haftmann | 
| Mon, 31 May 2010 17:29:28 +0200 | |
| changeset 37211 | 32a6f471f090 | 
| parent 34179 | 5490151d1052 | 
| child 37212 | b8e02ce2559f | 
| permissions | -rw-r--r-- | 
| 28447 | 1 | % | 
| 2 | \begin{isabellebody}%
 | |
| 3 | \def\isabellecontext{Program}%
 | |
| 4 | % | |
| 5 | \isadelimtheory | |
| 6 | % | |
| 7 | \endisadelimtheory | |
| 8 | % | |
| 9 | \isatagtheory | |
| 10 | \isacommand{theory}\isamarkupfalse%
 | |
| 11 | \ Program\isanewline | |
| 12 | \isakeyword{imports}\ Introduction\isanewline
 | |
| 13 | \isakeyword{begin}%
 | |
| 14 | \endisatagtheory | |
| 15 | {\isafoldtheory}%
 | |
| 16 | % | |
| 17 | \isadelimtheory | |
| 18 | % | |
| 19 | \endisadelimtheory | |
| 20 | % | |
| 21 | \isamarkupsection{Turning Theories into Programs \label{sec:program}%
 | |
| 22 | } | |
| 23 | \isamarkuptrue% | |
| 24 | % | |
| 25 | \isamarkupsubsection{The \isa{Isabelle{\isacharslash}HOL} default setup%
 | |
| 26 | } | |
| 27 | \isamarkuptrue% | |
| 28 | % | |
| 29 | \begin{isamarkuptext}%
 | |
| 30 | We have already seen how by default equations stemming from | |
| 34155 | 31 |   \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}, \hyperlink{command.primrec}{\mbox{\isa{\isacommand{primrec}}}} and \hyperlink{command.fun}{\mbox{\isa{\isacommand{fun}}}}
 | 
| 28447 | 32 | statements are used for code generation. This default behaviour | 
| 34155 | 33 | can be changed, e.g.\ by providing different code equations. | 
| 34 |   The customisations shown in this section are \emph{safe}
 | |
| 35 | as regards correctness: all programs that can be generated are partially | |
| 28447 | 36 | correct.% | 
| 37 | \end{isamarkuptext}%
 | |
| 38 | \isamarkuptrue% | |
| 39 | % | |
| 40 | \isamarkupsubsection{Selecting code equations%
 | |
| 41 | } | |
| 42 | \isamarkuptrue% | |
| 43 | % | |
| 44 | \begin{isamarkuptext}%
 | |
| 45 | Coming back to our introductory example, we | |
| 29560 | 46 |   could provide an alternative code equations for \isa{dequeue}
 | 
| 28447 | 47 | explicitly:% | 
| 48 | \end{isamarkuptext}%
 | |
| 49 | \isamarkuptrue% | |
| 50 | % | |
| 28564 | 51 | \isadelimquote | 
| 28447 | 52 | % | 
| 28564 | 53 | \endisadelimquote | 
| 28447 | 54 | % | 
| 28564 | 55 | \isatagquote | 
| 28447 | 56 | \isacommand{lemma}\isamarkupfalse%
 | 
| 28562 | 57 | \ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline
 | 
| 29798 | 58 | \ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline
 | 
| 59 | \ \ \ \ \ {\isacharparenleft}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ {\isacharparenleft}None{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline
 | |
| 60 | \ \ \ \ \ \ \ else\ dequeue\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 61 | \ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline
 | |
| 62 | \ \ \ \ \ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 28447 | 63 | \ \ \isacommand{by}\isamarkupfalse%
 | 
| 64 | \ {\isacharparenleft}cases\ xs{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\ {\isacharparenleft}cases\ {\isachardoublequoteopen}rev\ xs{\isachardoublequoteclose}{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}%
 | |
| 28564 | 65 | \endisatagquote | 
| 66 | {\isafoldquote}%
 | |
| 28447 | 67 | % | 
| 28564 | 68 | \isadelimquote | 
| 28447 | 69 | % | 
| 28564 | 70 | \endisadelimquote | 
| 28447 | 71 | % | 
| 72 | \begin{isamarkuptext}%
 | |
| 28562 | 73 | \noindent The annotation \isa{{\isacharbrackleft}code{\isacharbrackright}} is an \isa{Isar}
 | 
| 28447 | 74 |   \isa{attribute} which states that the given theorems should be
 | 
| 29560 | 75 |   considered as code equations for a \isa{fun} statement --
 | 
| 28447 | 76 | the corresponding constant is determined syntactically. The resulting code:% | 
| 77 | \end{isamarkuptext}%
 | |
| 78 | \isamarkuptrue% | |
| 79 | % | |
| 28564 | 80 | \isadelimquote | 
| 28447 | 81 | % | 
| 28564 | 82 | \endisadelimquote | 
| 28447 | 83 | % | 
| 28564 | 84 | \isatagquote | 
| 28447 | 85 | % | 
| 86 | \begin{isamarkuptext}%
 | |
| 28727 | 87 | \isatypewriter% | 
| 28447 | 88 | \noindent% | 
| 29297 | 89 | \hspace*{0pt}dequeue ::~forall a.~Queue a -> (Maybe a,~Queue a);\\
 | 
| 34179 | 90 | \hspace*{0pt}dequeue (AQueue xs (y :~ys)) = (Just y,~AQueue xs ys);\\
 | 
| 91 | \hspace*{0pt}dequeue (AQueue xs []) =\\
 | |
| 29798 | 92 | \hspace*{0pt} ~(if nulla xs then (Nothing,~AQueue [] [])\\
 | 
| 34179 | 93 | \hspace*{0pt} ~~~else dequeue (AQueue [] (rev xs)));%
 | 
| 28447 | 94 | \end{isamarkuptext}%
 | 
| 95 | \isamarkuptrue% | |
| 96 | % | |
| 28564 | 97 | \endisatagquote | 
| 98 | {\isafoldquote}%
 | |
| 28447 | 99 | % | 
| 28564 | 100 | \isadelimquote | 
| 28447 | 101 | % | 
| 28564 | 102 | \endisadelimquote | 
| 28447 | 103 | % | 
| 104 | \begin{isamarkuptext}%
 | |
| 105 | \noindent You may note that the equality test \isa{xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}} has been
 | |
| 106 |   replaced by the predicate \isa{null\ xs}.  This is due to the default
 | |
| 107 |   setup in the \qn{preprocessor} to be discussed further below (\secref{sec:preproc}).
 | |
| 108 | ||
| 109 | Changing the default constructor set of datatypes is also | |
| 29798 | 110 |   possible.  See \secref{sec:datatypes} for an example.
 | 
| 28447 | 111 | |
| 112 |   As told in \secref{sec:concept}, code generation is based
 | |
| 113 | on a structured collection of code theorems. | |
| 34155 | 114 | This collection | 
| 28447 | 115 |   may be inspected using the \hyperlink{command.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}} command:%
 | 
| 116 | \end{isamarkuptext}%
 | |
| 117 | \isamarkuptrue% | |
| 118 | % | |
| 28564 | 119 | \isadelimquote | 
| 28447 | 120 | % | 
| 28564 | 121 | \endisadelimquote | 
| 28447 | 122 | % | 
| 28564 | 123 | \isatagquote | 
| 28447 | 124 | \isacommand{code{\isacharunderscore}thms}\isamarkupfalse%
 | 
| 125 | \ dequeue% | |
| 28564 | 126 | \endisatagquote | 
| 127 | {\isafoldquote}%
 | |
| 28447 | 128 | % | 
| 28564 | 129 | \isadelimquote | 
| 28447 | 130 | % | 
| 28564 | 131 | \endisadelimquote | 
| 28447 | 132 | % | 
| 133 | \begin{isamarkuptext}%
 | |
| 29560 | 134 | \noindent prints a table with \emph{all} code equations
 | 
| 28447 | 135 |   for \isa{dequeue}, including
 | 
| 29560 | 136 |   \emph{all} code equations those equations depend
 | 
| 28447 | 137 | on recursively. | 
| 138 | ||
| 139 |   Similarly, the \hyperlink{command.code-deps}{\mbox{\isa{\isacommand{code{\isacharunderscore}deps}}}} command shows a graph
 | |
| 29560 | 140 | visualising dependencies between code equations.% | 
| 28447 | 141 | \end{isamarkuptext}%
 | 
| 142 | \isamarkuptrue% | |
| 143 | % | |
| 144 | \isamarkupsubsection{\isa{class} and \isa{instantiation}%
 | |
| 145 | } | |
| 146 | \isamarkuptrue% | |
| 147 | % | |
| 148 | \begin{isamarkuptext}%
 | |
| 149 | Concerning type classes and code generation, let us examine an example | |
| 150 | from abstract algebra:% | |
| 151 | \end{isamarkuptext}%
 | |
| 152 | \isamarkuptrue% | |
| 153 | % | |
| 28564 | 154 | \isadelimquote | 
| 28447 | 155 | % | 
| 28564 | 156 | \endisadelimquote | 
| 28447 | 157 | % | 
| 28564 | 158 | \isatagquote | 
| 28447 | 159 | \isacommand{class}\isamarkupfalse%
 | 
| 29798 | 160 | \ semigroup\ {\isacharequal}\isanewline
 | 
| 28447 | 161 | \ \ \isakeyword{fixes}\ mult\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymotimes}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline
 | 
| 162 | \ \ \isakeyword{assumes}\ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 163 | \isanewline | |
| 164 | \isacommand{class}\isamarkupfalse%
 | |
| 165 | \ monoid\ {\isacharequal}\ semigroup\ {\isacharplus}\isanewline
 | |
| 166 | \ \ \isakeyword{fixes}\ neutral\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isacharparenleft}{\isachardoublequoteopen}{\isasymone}{\isachardoublequoteclose}{\isacharparenright}\isanewline
 | |
| 167 | \ \ \isakeyword{assumes}\ neutl{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline
 | |
| 168 | \ \ \ \ \isakeyword{and}\ neutr{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline
 | |
| 169 | \isanewline | |
| 170 | \isacommand{instantiation}\isamarkupfalse%
 | |
| 171 | \ nat\ {\isacharcolon}{\isacharcolon}\ monoid\isanewline
 | |
| 172 | \isakeyword{begin}\isanewline
 | |
| 173 | \isanewline | |
| 174 | \isacommand{primrec}\isamarkupfalse%
 | |
| 175 | \ mult{\isacharunderscore}nat\ \isakeyword{where}\isanewline
 | |
| 176 | \ \ \ \ {\isachardoublequoteopen}{\isadigit{0}}\ {\isasymotimes}\ n\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}nat{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 177 | \ \ {\isacharbar}\ {\isachardoublequoteopen}Suc\ m\ {\isasymotimes}\ n\ {\isacharequal}\ n\ {\isacharplus}\ m\ {\isasymotimes}\ n{\isachardoublequoteclose}\isanewline
 | |
| 178 | \isanewline | |
| 179 | \isacommand{definition}\isamarkupfalse%
 | |
| 180 | \ neutral{\isacharunderscore}nat\ \isakeyword{where}\isanewline
 | |
| 181 | \ \ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline
 | |
| 182 | \isanewline | |
| 183 | \isacommand{lemma}\isamarkupfalse%
 | |
| 184 | \ add{\isacharunderscore}mult{\isacharunderscore}distrib{\isacharcolon}\isanewline
 | |
| 185 | \ \ \isakeyword{fixes}\ n\ m\ q\ {\isacharcolon}{\isacharcolon}\ nat\isanewline
 | |
| 186 | \ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isacharparenleft}n\ {\isacharplus}\ m{\isacharparenright}\ {\isasymotimes}\ q\ {\isacharequal}\ n\ {\isasymotimes}\ q\ {\isacharplus}\ m\ {\isasymotimes}\ q{\isachardoublequoteclose}\isanewline
 | |
| 187 | \ \ \isacommand{by}\isamarkupfalse%
 | |
| 188 | \ {\isacharparenleft}induct\ n{\isacharparenright}\ simp{\isacharunderscore}all\isanewline
 | |
| 189 | \isanewline | |
| 190 | \isacommand{instance}\isamarkupfalse%
 | |
| 191 | \ \isacommand{proof}\isamarkupfalse%
 | |
| 192 | \isanewline | |
| 193 | \ \ \isacommand{fix}\isamarkupfalse%
 | |
| 194 | \ m\ n\ q\ {\isacharcolon}{\isacharcolon}\ nat\isanewline
 | |
| 195 | \ \ \isacommand{show}\isamarkupfalse%
 | |
| 196 | \ {\isachardoublequoteopen}m\ {\isasymotimes}\ n\ {\isasymotimes}\ q\ {\isacharequal}\ m\ {\isasymotimes}\ {\isacharparenleft}n\ {\isasymotimes}\ q{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 197 | \ \ \ \ \isacommand{by}\isamarkupfalse%
 | |
| 198 | \ {\isacharparenleft}induct\ m{\isacharparenright}\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ add{\isacharunderscore}mult{\isacharunderscore}distrib{\isacharparenright}\isanewline
 | |
| 199 | \ \ \isacommand{show}\isamarkupfalse%
 | |
| 200 | \ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline
 | |
| 201 | \ \ \ \ \isacommand{by}\isamarkupfalse%
 | |
| 202 | \ {\isacharparenleft}simp\ add{\isacharcolon}\ neutral{\isacharunderscore}nat{\isacharunderscore}def{\isacharparenright}\isanewline
 | |
| 203 | \ \ \isacommand{show}\isamarkupfalse%
 | |
| 204 | \ {\isachardoublequoteopen}m\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ m{\isachardoublequoteclose}\isanewline
 | |
| 205 | \ \ \ \ \isacommand{by}\isamarkupfalse%
 | |
| 206 | \ {\isacharparenleft}induct\ m{\isacharparenright}\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ neutral{\isacharunderscore}nat{\isacharunderscore}def{\isacharparenright}\isanewline
 | |
| 207 | \isacommand{qed}\isamarkupfalse%
 | |
| 208 | \isanewline | |
| 209 | \isanewline | |
| 210 | \isacommand{end}\isamarkupfalse%
 | |
| 211 | % | |
| 28564 | 212 | \endisatagquote | 
| 213 | {\isafoldquote}%
 | |
| 28447 | 214 | % | 
| 28564 | 215 | \isadelimquote | 
| 28447 | 216 | % | 
| 28564 | 217 | \endisadelimquote | 
| 28447 | 218 | % | 
| 219 | \begin{isamarkuptext}%
 | |
| 220 | \noindent We define the natural operation of the natural numbers | |
| 221 | on monoids:% | |
| 222 | \end{isamarkuptext}%
 | |
| 223 | \isamarkuptrue% | |
| 224 | % | |
| 28564 | 225 | \isadelimquote | 
| 28447 | 226 | % | 
| 28564 | 227 | \endisadelimquote | 
| 28447 | 228 | % | 
| 28564 | 229 | \isatagquote | 
| 28447 | 230 | \isacommand{primrec}\isamarkupfalse%
 | 
| 231 | \ {\isacharparenleft}\isakeyword{in}\ monoid{\isacharparenright}\ pow\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
 | |
| 232 | \ \ \ \ {\isachardoublequoteopen}pow\ {\isadigit{0}}\ a\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline
 | |
| 233 | \ \ {\isacharbar}\ {\isachardoublequoteopen}pow\ {\isacharparenleft}Suc\ n{\isacharparenright}\ a\ {\isacharequal}\ a\ {\isasymotimes}\ pow\ n\ a{\isachardoublequoteclose}%
 | |
| 28564 | 234 | \endisatagquote | 
| 235 | {\isafoldquote}%
 | |
| 28447 | 236 | % | 
| 28564 | 237 | \isadelimquote | 
| 28447 | 238 | % | 
| 28564 | 239 | \endisadelimquote | 
| 28447 | 240 | % | 
| 241 | \begin{isamarkuptext}%
 | |
| 242 | \noindent This we use to define the discrete exponentiation function:% | |
| 243 | \end{isamarkuptext}%
 | |
| 244 | \isamarkuptrue% | |
| 245 | % | |
| 28564 | 246 | \isadelimquote | 
| 28447 | 247 | % | 
| 28564 | 248 | \endisadelimquote | 
| 28447 | 249 | % | 
| 28564 | 250 | \isatagquote | 
| 28447 | 251 | \isacommand{definition}\isamarkupfalse%
 | 
| 252 | \ bexp\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
 | |
| 253 | \ \ {\isachardoublequoteopen}bexp\ n\ {\isacharequal}\ pow\ n\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}%
 | |
| 28564 | 254 | \endisatagquote | 
| 255 | {\isafoldquote}%
 | |
| 28447 | 256 | % | 
| 28564 | 257 | \isadelimquote | 
| 28447 | 258 | % | 
| 28564 | 259 | \endisadelimquote | 
| 28447 | 260 | % | 
| 261 | \begin{isamarkuptext}%
 | |
| 34155 | 262 | \noindent The corresponding code in Haskell uses that language's native classes:% | 
| 28447 | 263 | \end{isamarkuptext}%
 | 
| 264 | \isamarkuptrue% | |
| 265 | % | |
| 28564 | 266 | \isadelimquote | 
| 28447 | 267 | % | 
| 28564 | 268 | \endisadelimquote | 
| 28447 | 269 | % | 
| 28564 | 270 | \isatagquote | 
| 28447 | 271 | % | 
| 272 | \begin{isamarkuptext}%
 | |
| 28727 | 273 | \isatypewriter% | 
| 28447 | 274 | \noindent% | 
| 28714 | 275 | \hspace*{0pt}module Example where {\char123}\\
 | 
| 276 | \hspace*{0pt}\\
 | |
| 30121 | 277 | \hspace*{0pt}data Nat = Zero{\char95}nat | Suc Nat;\\
 | 
| 28714 | 278 | \hspace*{0pt}\\
 | 
| 279 | \hspace*{0pt}class Semigroup a where {\char123}\\
 | |
| 280 | \hspace*{0pt} ~mult ::~a -> a -> a;\\
 | |
| 281 | \hspace*{0pt}{\char125};\\
 | |
| 282 | \hspace*{0pt}\\
 | |
| 34179 | 283 | \hspace*{0pt}class (Semigroup a) => Monoid a where {\char123}\\
 | 
| 28714 | 284 | \hspace*{0pt} ~neutral ::~a;\\
 | 
| 285 | \hspace*{0pt}{\char125};\\
 | |
| 286 | \hspace*{0pt}\\
 | |
| 34179 | 287 | \hspace*{0pt}pow ::~forall a.~(Monoid a) => Nat -> a -> a;\\
 | 
| 28714 | 288 | \hspace*{0pt}pow Zero{\char95}nat a = neutral;\\
 | 
| 34179 | 289 | \hspace*{0pt}pow (Suc n) a = mult a (pow n a);\\
 | 
| 28714 | 290 | \hspace*{0pt}\\
 | 
| 291 | \hspace*{0pt}plus{\char95}nat ::~Nat -> Nat -> Nat;\\
 | |
| 34179 | 292 | \hspace*{0pt}plus{\char95}nat (Suc m) n = plus{\char95}nat m (Suc n);\\
 | 
| 28714 | 293 | \hspace*{0pt}plus{\char95}nat Zero{\char95}nat n = n;\\
 | 
| 294 | \hspace*{0pt}\\
 | |
| 295 | \hspace*{0pt}neutral{\char95}nat ::~Nat;\\
 | |
| 296 | \hspace*{0pt}neutral{\char95}nat = Suc Zero{\char95}nat;\\
 | |
| 297 | \hspace*{0pt}\\
 | |
| 298 | \hspace*{0pt}mult{\char95}nat ::~Nat -> Nat -> Nat;\\
 | |
| 299 | \hspace*{0pt}mult{\char95}nat Zero{\char95}nat n = Zero{\char95}nat;\\
 | |
| 34179 | 300 | \hspace*{0pt}mult{\char95}nat (Suc m) n = plus{\char95}nat n (mult{\char95}nat m n);\\
 | 
| 28714 | 301 | \hspace*{0pt}\\
 | 
| 302 | \hspace*{0pt}instance Semigroup Nat where {\char123}\\
 | |
| 303 | \hspace*{0pt} ~mult = mult{\char95}nat;\\
 | |
| 304 | \hspace*{0pt}{\char125};\\
 | |
| 305 | \hspace*{0pt}\\
 | |
| 306 | \hspace*{0pt}instance Monoid Nat where {\char123}\\
 | |
| 307 | \hspace*{0pt} ~neutral = neutral{\char95}nat;\\
 | |
| 308 | \hspace*{0pt}{\char125};\\
 | |
| 309 | \hspace*{0pt}\\
 | |
| 310 | \hspace*{0pt}bexp ::~Nat -> Nat;\\
 | |
| 34179 | 311 | \hspace*{0pt}bexp n = pow n (Suc (Suc Zero{\char95}nat));\\
 | 
| 28714 | 312 | \hspace*{0pt}\\
 | 
| 313 | \hspace*{0pt}{\char125}%
 | |
| 28447 | 314 | \end{isamarkuptext}%
 | 
| 315 | \isamarkuptrue% | |
| 316 | % | |
| 28564 | 317 | \endisatagquote | 
| 318 | {\isafoldquote}%
 | |
| 28447 | 319 | % | 
| 28564 | 320 | \isadelimquote | 
| 28447 | 321 | % | 
| 28564 | 322 | \endisadelimquote | 
| 28447 | 323 | % | 
| 324 | \begin{isamarkuptext}%
 | |
| 325 | \noindent This is a convenient place to show how explicit dictionary construction | |
| 326 |   manifests in generated code (here, the same example in \isa{SML}):%
 | |
| 327 | \end{isamarkuptext}%
 | |
| 328 | \isamarkuptrue% | |
| 329 | % | |
| 28564 | 330 | \isadelimquote | 
| 28447 | 331 | % | 
| 28564 | 332 | \endisadelimquote | 
| 28447 | 333 | % | 
| 28564 | 334 | \isatagquote | 
| 28447 | 335 | % | 
| 336 | \begin{isamarkuptext}%
 | |
| 28727 | 337 | \isatypewriter% | 
| 28447 | 338 | \noindent% | 
| 34155 | 339 | \hspace*{0pt}structure Example :~sig\\
 | 
| 34179 | 340 | \hspace*{0pt} ~datatype nat = Zero{\char95}nat | Suc of nat\\
 | 
| 34155 | 341 | \hspace*{0pt} ~type 'a semigroup\\
 | 
| 34179 | 342 | \hspace*{0pt} ~val mult :~'a semigroup -> 'a -> 'a -> 'a\\
 | 
| 34155 | 343 | \hspace*{0pt} ~type 'a monoid\\
 | 
| 34179 | 344 | \hspace*{0pt} ~val semigroup{\char95}monoid :~'a monoid -> 'a semigroup\\
 | 
| 345 | \hspace*{0pt} ~val neutral :~'a monoid -> 'a\\
 | |
| 346 | \hspace*{0pt} ~val pow :~'a monoid -> nat -> 'a -> 'a\\
 | |
| 34155 | 347 | \hspace*{0pt} ~val plus{\char95}nat :~nat -> nat -> nat\\
 | 
| 348 | \hspace*{0pt} ~val neutral{\char95}nat :~nat\\
 | |
| 349 | \hspace*{0pt} ~val mult{\char95}nat :~nat -> nat -> nat\\
 | |
| 350 | \hspace*{0pt} ~val semigroup{\char95}nat :~nat semigroup\\
 | |
| 351 | \hspace*{0pt} ~val monoid{\char95}nat :~nat monoid\\
 | |
| 352 | \hspace*{0pt} ~val bexp :~nat -> nat\\
 | |
| 353 | \hspace*{0pt}end = struct\\
 | |
| 28714 | 354 | \hspace*{0pt}\\
 | 
| 30121 | 355 | \hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\
 | 
| 28714 | 356 | \hspace*{0pt}\\
 | 
| 34179 | 357 | \hspace*{0pt}type 'a semigroup = {\char123}mult :~'a -> 'a -> 'a{\char125};\\
 | 
| 358 | \hspace*{0pt}val mult = {\char35}mult :~'a semigroup -> 'a -> 'a -> 'a;\\
 | |
| 28714 | 359 | \hspace*{0pt}\\
 | 
| 34179 | 360 | \hspace*{0pt}type 'a monoid = {\char123}semigroup{\char95}monoid :~'a semigroup,~neutral :~'a{\char125};\\
 | 
| 361 | \hspace*{0pt}val semigroup{\char95}monoid = {\char35}semigroup{\char95}monoid :~'a monoid -> 'a semigroup;\\
 | |
| 362 | \hspace*{0pt}val neutral = {\char35}neutral :~'a monoid -> 'a;\\
 | |
| 28714 | 363 | \hspace*{0pt}\\
 | 
| 364 | \hspace*{0pt}fun pow A{\char95}~Zero{\char95}nat a = neutral A{\char95}\\
 | |
| 34179 | 365 | \hspace*{0pt} ~| pow A{\char95}~(Suc n) a = mult (semigroup{\char95}monoid A{\char95}) a (pow A{\char95}~n a);\\
 | 
| 28714 | 366 | \hspace*{0pt}\\
 | 
| 34179 | 367 | \hspace*{0pt}fun plus{\char95}nat (Suc m) n = plus{\char95}nat m (Suc n)\\
 | 
| 368 | \hspace*{0pt} ~| plus{\char95}nat Zero{\char95}nat n = n;\\
 | |
| 28714 | 369 | \hspace*{0pt}\\
 | 
| 33707 
68841fb382e0
dropped obsolete documentation; updated generated sources
 haftmann parents: 
32000diff
changeset | 370 | \hspace*{0pt}val neutral{\char95}nat :~nat = Suc Zero{\char95}nat;\\
 | 
| 28714 | 371 | \hspace*{0pt}\\
 | 
| 34179 | 372 | \hspace*{0pt}fun mult{\char95}nat Zero{\char95}nat n = Zero{\char95}nat\\
 | 
| 373 | \hspace*{0pt} ~| mult{\char95}nat (Suc m) n = plus{\char95}nat n (mult{\char95}nat m n);\\
 | |
| 28714 | 374 | \hspace*{0pt}\\
 | 
| 34179 | 375 | \hspace*{0pt}val semigroup{\char95}nat = {\char123}mult = mult{\char95}nat{\char125}~:~nat semigroup;\\
 | 
| 28714 | 376 | \hspace*{0pt}\\
 | 
| 34179 | 377 | \hspace*{0pt}val monoid{\char95}nat = {\char123}semigroup{\char95}monoid = semigroup{\char95}nat,~neutral = neutral{\char95}nat{\char125}\\
 | 
| 31544 | 378 | \hspace*{0pt} ~:~nat monoid;\\
 | 
| 28714 | 379 | \hspace*{0pt}\\
 | 
| 34179 | 380 | \hspace*{0pt}fun bexp n = pow monoid{\char95}nat n (Suc (Suc Zero{\char95}nat));\\
 | 
| 28714 | 381 | \hspace*{0pt}\\
 | 
| 29297 | 382 | \hspace*{0pt}end;~(*struct Example*)%
 | 
| 28447 | 383 | \end{isamarkuptext}%
 | 
| 384 | \isamarkuptrue% | |
| 385 | % | |
| 28564 | 386 | \endisatagquote | 
| 387 | {\isafoldquote}%
 | |
| 28447 | 388 | % | 
| 28564 | 389 | \isadelimquote | 
| 28447 | 390 | % | 
| 28564 | 391 | \endisadelimquote | 
| 28447 | 392 | % | 
| 393 | \begin{isamarkuptext}%
 | |
| 34155 | 394 | \noindent Note the parameters with trailing underscore (\verb|A_|), | 
| 28447 | 395 | which are the dictionary parameters.% | 
| 396 | \end{isamarkuptext}%
 | |
| 397 | \isamarkuptrue% | |
| 398 | % | |
| 399 | \isamarkupsubsection{The preprocessor \label{sec:preproc}%
 | |
| 400 | } | |
| 401 | \isamarkuptrue% | |
| 402 | % | |
| 403 | \begin{isamarkuptext}%
 | |
| 404 | Before selected function theorems are turned into abstract | |
| 405 | code, a chain of definitional transformation steps is carried | |
| 406 |   out: \emph{preprocessing}.  In essence, the preprocessor
 | |
| 407 |   consists of two components: a \emph{simpset} and \emph{function transformers}.
 | |
| 408 | ||
| 34155 | 409 |   The \emph{simpset} can apply the full generality of the
 | 
| 32000 | 410 | Isabelle simplifier. Due to the interpretation of theorems as code | 
| 411 | equations, rewrites are applied to the right hand side and the | |
| 412 | arguments of the left hand side of an equation, but never to the | |
| 413 | constant heading the left hand side. An important special case are | |
| 34155 | 414 |   \emph{unfold theorems},  which may be declared and removed using
 | 
| 32000 | 415 |   the \hyperlink{attribute.code-unfold}{\mbox{\isa{code{\isacharunderscore}unfold}}} or \emph{\hyperlink{attribute.code-unfold}{\mbox{\isa{code{\isacharunderscore}unfold}}} del}
 | 
| 34155 | 416 | attribute, respectively. | 
| 28447 | 417 | |
| 418 | Some common applications:% | |
| 419 | \end{isamarkuptext}%
 | |
| 420 | \isamarkuptrue% | |
| 421 | % | |
| 422 | \begin{itemize}
 | |
| 423 | % | |
| 424 | \begin{isamarkuptext}%
 | |
| 425 | \item replacing non-executable constructs by executable ones:% | |
| 426 | \end{isamarkuptext}%
 | |
| 427 | \isamarkuptrue% | |
| 428 | % | |
| 28564 | 429 | \isadelimquote | 
| 28447 | 430 | % | 
| 28564 | 431 | \endisadelimquote | 
| 28447 | 432 | % | 
| 28564 | 433 | \isatagquote | 
| 28447 | 434 | \isacommand{lemma}\isamarkupfalse%
 | 
| 32000 | 435 | \ {\isacharbrackleft}code{\isacharunderscore}inline{\isacharbrackright}{\isacharcolon}\isanewline
 | 
| 28447 | 436 | \ \ {\isachardoublequoteopen}x\ {\isasymin}\ set\ xs\ {\isasymlongleftrightarrow}\ x\ mem\ xs{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
 | 
| 437 | \ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all%
 | |
| 28564 | 438 | \endisatagquote | 
| 439 | {\isafoldquote}%
 | |
| 28447 | 440 | % | 
| 28564 | 441 | \isadelimquote | 
| 28447 | 442 | % | 
| 28564 | 443 | \endisadelimquote | 
| 28447 | 444 | % | 
| 445 | \begin{isamarkuptext}%
 | |
| 446 | \item eliminating superfluous constants:% | |
| 447 | \end{isamarkuptext}%
 | |
| 448 | \isamarkuptrue% | |
| 449 | % | |
| 28564 | 450 | \isadelimquote | 
| 28447 | 451 | % | 
| 28564 | 452 | \endisadelimquote | 
| 28447 | 453 | % | 
| 28564 | 454 | \isatagquote | 
| 28447 | 455 | \isacommand{lemma}\isamarkupfalse%
 | 
| 32000 | 456 | \ {\isacharbrackleft}code{\isacharunderscore}inline{\isacharbrackright}{\isacharcolon}\isanewline
 | 
| 28447 | 457 | \ \ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharequal}\ Suc\ {\isadigit{0}}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
 | 
| 458 | \ simp% | |
| 28564 | 459 | \endisatagquote | 
| 460 | {\isafoldquote}%
 | |
| 28447 | 461 | % | 
| 28564 | 462 | \isadelimquote | 
| 28447 | 463 | % | 
| 28564 | 464 | \endisadelimquote | 
| 28447 | 465 | % | 
| 466 | \begin{isamarkuptext}%
 | |
| 467 | \item replacing executable but inconvenient constructs:% | |
| 468 | \end{isamarkuptext}%
 | |
| 469 | \isamarkuptrue% | |
| 470 | % | |
| 28564 | 471 | \isadelimquote | 
| 28447 | 472 | % | 
| 28564 | 473 | \endisadelimquote | 
| 28447 | 474 | % | 
| 28564 | 475 | \isatagquote | 
| 28447 | 476 | \isacommand{lemma}\isamarkupfalse%
 | 
| 32000 | 477 | \ {\isacharbrackleft}code{\isacharunderscore}inline{\isacharbrackright}{\isacharcolon}\isanewline
 | 
| 28447 | 478 | \ \ {\isachardoublequoteopen}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongleftrightarrow}\ List{\isachardot}null\ xs{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
 | 
| 479 | \ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all%
 | |
| 28564 | 480 | \endisatagquote | 
| 481 | {\isafoldquote}%
 | |
| 28447 | 482 | % | 
| 28564 | 483 | \isadelimquote | 
| 28447 | 484 | % | 
| 28564 | 485 | \endisadelimquote | 
| 28447 | 486 | % | 
| 487 | \end{itemize}
 | |
| 488 | % | |
| 489 | \begin{isamarkuptext}%
 | |
| 490 | \noindent \emph{Function transformers} provide a very general interface,
 | |
| 491 | transforming a list of function theorems to another | |
| 492 | list of function theorems, provided that neither the heading | |
| 493 |   constant nor its type change.  The \isa{{\isadigit{0}}} / \isa{Suc}
 | |
| 494 | pattern elimination implemented in | |
| 495 |   theory \isa{Efficient{\isacharunderscore}Nat} (see \secref{eff_nat}) uses this
 | |
| 496 | interface. | |
| 497 | ||
| 498 | \noindent The current setup of the preprocessor may be inspected using | |
| 31254 | 499 |   the \hyperlink{command.print-codeproc}{\mbox{\isa{\isacommand{print{\isacharunderscore}codeproc}}}} command.
 | 
| 28447 | 500 |   \hyperlink{command.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}} provides a convenient
 | 
| 501 | mechanism to inspect the impact of a preprocessor setup | |
| 29560 | 502 | on code equations. | 
| 28447 | 503 | |
| 504 |   \begin{warn}
 | |
| 32000 | 505 | |
| 506 |     Attribute \hyperlink{attribute.code-unfold}{\mbox{\isa{code{\isacharunderscore}unfold}}} also applies to the
 | |
| 507 |     preprocessor of the ancient \isa{SML\ code\ generator}; in case
 | |
| 508 |     this is not what you intend, use \hyperlink{attribute.code-inline}{\mbox{\isa{code{\isacharunderscore}inline}}} instead.
 | |
| 28447 | 509 |   \end{warn}%
 | 
| 510 | \end{isamarkuptext}%
 | |
| 511 | \isamarkuptrue% | |
| 512 | % | |
| 513 | \isamarkupsubsection{Datatypes \label{sec:datatypes}%
 | |
| 514 | } | |
| 515 | \isamarkuptrue% | |
| 516 | % | |
| 517 | \begin{isamarkuptext}%
 | |
| 518 | Conceptually, any datatype is spanned by a set of | |
| 29798 | 519 |   \emph{constructors} of type \isa{{\isasymtau}\ {\isacharequal}\ {\isasymdots}\ {\isasymRightarrow}\ {\isasymkappa}\ {\isasymalpha}\isactrlisub {\isadigit{1}}\ {\isasymdots}\ {\isasymalpha}\isactrlisub n} where \isa{{\isacharbraceleft}{\isasymalpha}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymalpha}\isactrlisub n{\isacharbraceright}} is exactly the set of \emph{all} type variables in
 | 
| 520 |   \isa{{\isasymtau}}.  The HOL datatype package by default registers any new
 | |
| 521 | datatype in the table of datatypes, which may be inspected using the | |
| 522 |   \hyperlink{command.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}} command.
 | |
| 28447 | 523 | |
| 29798 | 524 | In some cases, it is appropriate to alter or extend this table. As | 
| 525 | an example, we will develop an alternative representation of the | |
| 526 |   queue example given in \secref{sec:intro}.  The amortised
 | |
| 527 | representation is convenient for generating code but exposes its | |
| 528 |   \qt{implementation} details, which may be cumbersome when proving
 | |
| 34155 | 529 | theorems about it. Therefore, here is a simple, straightforward | 
| 29798 | 530 | representation of queues:% | 
| 531 | \end{isamarkuptext}%
 | |
| 532 | \isamarkuptrue% | |
| 533 | % | |
| 534 | \isadelimquote | |
| 535 | % | |
| 536 | \endisadelimquote | |
| 537 | % | |
| 538 | \isatagquote | |
| 539 | \isacommand{datatype}\isamarkupfalse%
 | |
| 540 | \ {\isacharprime}a\ queue\ {\isacharequal}\ Queue\ {\isachardoublequoteopen}{\isacharprime}a\ list{\isachardoublequoteclose}\isanewline
 | |
| 541 | \isanewline | |
| 542 | \isacommand{definition}\isamarkupfalse%
 | |
| 543 | \ empty\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
 | |
| 544 | \ \ {\isachardoublequoteopen}empty\ {\isacharequal}\ Queue\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
 | |
| 545 | \isanewline | |
| 546 | \isacommand{primrec}\isamarkupfalse%
 | |
| 547 | \ enqueue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
 | |
| 548 | \ \ {\isachardoublequoteopen}enqueue\ x\ {\isacharparenleft}Queue\ xs{\isacharparenright}\ {\isacharequal}\ Queue\ {\isacharparenleft}xs\ {\isacharat}\ {\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 549 | \isanewline | |
| 550 | \isacommand{fun}\isamarkupfalse%
 | |
| 551 | \ dequeue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ option\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
 | |
| 552 | \ \ \ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}Queue\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}None{\isacharcomma}\ Queue\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 553 | \ \ {\isacharbar}\ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}Queue\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Some\ x{\isacharcomma}\ Queue\ xs{\isacharparenright}{\isachardoublequoteclose}%
 | |
| 554 | \endisatagquote | |
| 555 | {\isafoldquote}%
 | |
| 556 | % | |
| 557 | \isadelimquote | |
| 558 | % | |
| 559 | \endisadelimquote | |
| 560 | % | |
| 561 | \begin{isamarkuptext}%
 | |
| 562 | \noindent This we can use directly for proving; for executing, | |
| 563 | we provide an alternative characterisation:% | |
| 28447 | 564 | \end{isamarkuptext}%
 | 
| 565 | \isamarkuptrue% | |
| 566 | % | |
| 28564 | 567 | \isadelimquote | 
| 28447 | 568 | % | 
| 28564 | 569 | \endisadelimquote | 
| 28447 | 570 | % | 
| 28564 | 571 | \isatagquote | 
| 28447 | 572 | \isacommand{definition}\isamarkupfalse%
 | 
| 29798 | 573 | \ AQueue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
 | 
| 574 | \ \ {\isachardoublequoteopen}AQueue\ xs\ ys\ {\isacharequal}\ Queue\ {\isacharparenleft}ys\ {\isacharat}\ rev\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 28447 | 575 | \isanewline | 
| 29798 | 576 | \isacommand{code{\isacharunderscore}datatype}\isamarkupfalse%
 | 
| 577 | \ AQueue% | |
| 28564 | 578 | \endisatagquote | 
| 579 | {\isafoldquote}%
 | |
| 28447 | 580 | % | 
| 28564 | 581 | \isadelimquote | 
| 28447 | 582 | % | 
| 28564 | 583 | \endisadelimquote | 
| 28447 | 584 | % | 
| 585 | \begin{isamarkuptext}%
 | |
| 30227 | 586 | \noindent Here we define a \qt{constructor} \isa{AQueue} which
 | 
| 29798 | 587 |   is defined in terms of \isa{Queue} and interprets its arguments
 | 
| 588 |   according to what the \emph{content} of an amortised queue is supposed
 | |
| 589 | to be. Equipped with this, we are able to prove the following equations | |
| 590 |   for our primitive queue operations which \qt{implement} the simple
 | |
| 591 | queues in an amortised fashion:% | |
| 28447 | 592 | \end{isamarkuptext}%
 | 
| 593 | \isamarkuptrue% | |
| 594 | % | |
| 28564 | 595 | \isadelimquote | 
| 28447 | 596 | % | 
| 28564 | 597 | \endisadelimquote | 
| 28447 | 598 | % | 
| 28564 | 599 | \isatagquote | 
| 28447 | 600 | \isacommand{lemma}\isamarkupfalse%
 | 
| 29798 | 601 | \ empty{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline
 | 
| 602 | \ \ {\isachardoublequoteopen}empty\ {\isacharequal}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
 | |
| 603 | \ \ \isacommand{unfolding}\isamarkupfalse%
 | |
| 604 | \ AQueue{\isacharunderscore}def\ empty{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
 | |
| 605 | \ simp\isanewline | |
| 606 | \isanewline | |
| 607 | \isacommand{lemma}\isamarkupfalse%
 | |
| 608 | \ enqueue{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline
 | |
| 609 | \ \ {\isachardoublequoteopen}enqueue\ x\ {\isacharparenleft}AQueue\ xs\ ys{\isacharparenright}\ {\isacharequal}\ AQueue\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ ys{\isachardoublequoteclose}\isanewline
 | |
| 610 | \ \ \isacommand{unfolding}\isamarkupfalse%
 | |
| 611 | \ AQueue{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
 | |
| 612 | \ simp\isanewline | |
| 613 | \isanewline | |
| 614 | \isacommand{lemma}\isamarkupfalse%
 | |
| 615 | \ dequeue{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline
 | |
| 616 | \ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline
 | |
| 617 | \ \ \ \ {\isacharparenleft}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ {\isacharparenleft}None{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline
 | |
| 618 | \ \ \ \ else\ dequeue\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 619 | \ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 620 | \ \ \isacommand{unfolding}\isamarkupfalse%
 | |
| 621 | \ AQueue{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
 | |
| 622 | \ simp{\isacharunderscore}all%
 | |
| 28564 | 623 | \endisatagquote | 
| 624 | {\isafoldquote}%
 | |
| 28447 | 625 | % | 
| 28564 | 626 | \isadelimquote | 
| 28447 | 627 | % | 
| 28564 | 628 | \endisadelimquote | 
| 28447 | 629 | % | 
| 630 | \begin{isamarkuptext}%
 | |
| 29798 | 631 | \noindent For completeness, we provide a substitute for the | 
| 632 |   \isa{case} combinator on queues:%
 | |
| 28447 | 633 | \end{isamarkuptext}%
 | 
| 634 | \isamarkuptrue% | |
| 635 | % | |
| 28564 | 636 | \isadelimquote | 
| 28447 | 637 | % | 
| 28564 | 638 | \endisadelimquote | 
| 28447 | 639 | % | 
| 28564 | 640 | \isatagquote | 
| 29798 | 641 | \isacommand{lemma}\isamarkupfalse%
 | 
| 30227 | 642 | \ queue{\isacharunderscore}case{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline
 | 
| 643 | \ \ {\isachardoublequoteopen}queue{\isacharunderscore}case\ f\ {\isacharparenleft}AQueue\ xs\ ys{\isacharparenright}\ {\isacharequal}\ f\ {\isacharparenleft}ys\ {\isacharat}\ rev\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 29798 | 644 | \ \ \isacommand{unfolding}\isamarkupfalse%
 | 
| 30227 | 645 | \ AQueue{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
 | 
| 29798 | 646 | \ simp% | 
| 28564 | 647 | \endisatagquote | 
| 648 | {\isafoldquote}%
 | |
| 28447 | 649 | % | 
| 28564 | 650 | \isadelimquote | 
| 28447 | 651 | % | 
| 28564 | 652 | \endisadelimquote | 
| 28447 | 653 | % | 
| 654 | \begin{isamarkuptext}%
 | |
| 29798 | 655 | \noindent The resulting code looks as expected:% | 
| 28447 | 656 | \end{isamarkuptext}%
 | 
| 657 | \isamarkuptrue% | |
| 658 | % | |
| 28564 | 659 | \isadelimquote | 
| 28447 | 660 | % | 
| 28564 | 661 | \endisadelimquote | 
| 28447 | 662 | % | 
| 28564 | 663 | \isatagquote | 
| 28447 | 664 | % | 
| 665 | \begin{isamarkuptext}%
 | |
| 28727 | 666 | \isatypewriter% | 
| 28447 | 667 | \noindent% | 
| 34155 | 668 | \hspace*{0pt}structure Example :~sig\\
 | 
| 34179 | 669 | \hspace*{0pt} ~val foldl :~('a -> 'b -> 'a) -> 'a -> 'b list -> 'a\\
 | 
| 34155 | 670 | \hspace*{0pt} ~val rev :~'a list -> 'a list\\
 | 
| 671 | \hspace*{0pt} ~val null :~'a list -> bool\\
 | |
| 34179 | 672 | \hspace*{0pt} ~datatype 'a queue = AQueue of 'a list * 'a list\\
 | 
| 34155 | 673 | \hspace*{0pt} ~val empty :~'a queue\\
 | 
| 674 | \hspace*{0pt} ~val dequeue :~'a queue -> 'a option * 'a queue\\
 | |
| 675 | \hspace*{0pt} ~val enqueue :~'a -> 'a queue -> 'a queue\\
 | |
| 676 | \hspace*{0pt}end = struct\\
 | |
| 28714 | 677 | \hspace*{0pt}\\
 | 
| 29798 | 678 | \hspace*{0pt}fun foldl f a [] = a\\
 | 
| 34179 | 679 | \hspace*{0pt} ~| foldl f a (x ::~xs) = foldl f (f a x) xs;\\
 | 
| 29798 | 680 | \hspace*{0pt}\\
 | 
| 34179 | 681 | \hspace*{0pt}fun rev xs = foldl (fn xsa => fn x => x ::~xsa) [] xs;\\
 | 
| 29798 | 682 | \hspace*{0pt}\\
 | 
| 683 | \hspace*{0pt}fun null [] = true\\
 | |
| 34179 | 684 | \hspace*{0pt} ~| null (x ::~xs) = false;\\
 | 
| 28714 | 685 | \hspace*{0pt}\\
 | 
| 34179 | 686 | \hspace*{0pt}datatype 'a queue = AQueue of 'a list * 'a list;\\
 | 
| 29798 | 687 | \hspace*{0pt}\\
 | 
| 34179 | 688 | \hspace*{0pt}val empty :~'a queue = AQueue ([],~[]);\\
 | 
| 29798 | 689 | \hspace*{0pt}\\
 | 
| 34179 | 690 | \hspace*{0pt}fun dequeue (AQueue (xs,~y ::~ys)) = (SOME y,~AQueue (xs,~ys))\\
 | 
| 691 | \hspace*{0pt} ~| dequeue (AQueue (xs,~[])) =\\
 | |
| 692 | \hspace*{0pt} ~~~(if null xs then (NONE,~AQueue ([],~[]))\\
 | |
| 693 | \hspace*{0pt} ~~~~~else dequeue (AQueue ([],~rev xs)));\\
 | |
| 29798 | 694 | \hspace*{0pt}\\
 | 
| 34179 | 695 | \hspace*{0pt}fun enqueue x (AQueue (xs,~ys)) = AQueue (x ::~xs,~ys);\\
 | 
| 28714 | 696 | \hspace*{0pt}\\
 | 
| 29297 | 697 | \hspace*{0pt}end;~(*struct Example*)%
 | 
| 28447 | 698 | \end{isamarkuptext}%
 | 
| 699 | \isamarkuptrue% | |
| 700 | % | |
| 28564 | 701 | \endisatagquote | 
| 702 | {\isafoldquote}%
 | |
| 28447 | 703 | % | 
| 28564 | 704 | \isadelimquote | 
| 28447 | 705 | % | 
| 28564 | 706 | \endisadelimquote | 
| 28447 | 707 | % | 
| 708 | \begin{isamarkuptext}%
 | |
| 29798 | 709 | \noindent From this example, it can be glimpsed that using own | 
| 710 | constructor sets is a little delicate since it changes the set of | |
| 711 | valid patterns for values of that type. Without going into much | |
| 712 | detail, here some practical hints: | |
| 28447 | 713 | |
| 714 |   \begin{itemize}
 | |
| 29798 | 715 | |
| 716 | \item When changing the constructor set for datatypes, take care | |
| 30227 | 717 |       to provide alternative equations for the \isa{case} combinator.
 | 
| 29798 | 718 | |
| 719 | \item Values in the target language need not to be normalised -- | |
| 720 | different values in the target language may represent the same | |
| 721 | value in the logic. | |
| 722 | ||
| 723 | \item Usually, a good methodology to deal with the subtleties of | |
| 724 | pattern matching is to see the type as an abstract type: provide | |
| 725 | a set of operations which operate on the concrete representation | |
| 726 | of the type, and derive further operations by combinations of | |
| 727 | these primitive ones, without relying on a particular | |
| 728 | representation. | |
| 729 | ||
| 28447 | 730 |   \end{itemize}%
 | 
| 731 | \end{isamarkuptext}%
 | |
| 732 | \isamarkuptrue% | |
| 733 | % | |
| 30938 
c6c9359e474c
wellsortedness is no issue for a user manual any more
 haftmann parents: 
30227diff
changeset | 734 | \isamarkupsubsection{Equality%
 | 
| 28447 | 735 | } | 
| 736 | \isamarkuptrue% | |
| 737 | % | |
| 738 | \begin{isamarkuptext}%
 | |
| 739 | Surely you have already noticed how equality is treated | |
| 740 | by the code generator:% | |
| 741 | \end{isamarkuptext}%
 | |
| 742 | \isamarkuptrue% | |
| 743 | % | |
| 28564 | 744 | \isadelimquote | 
| 28447 | 745 | % | 
| 28564 | 746 | \endisadelimquote | 
| 28447 | 747 | % | 
| 28564 | 748 | \isatagquote | 
| 28447 | 749 | \isacommand{primrec}\isamarkupfalse%
 | 
| 750 | \ collect{\isacharunderscore}duplicates\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
 | |
| 751 | \ \ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ xs{\isachardoublequoteclose}\isanewline
 | |
| 752 | \ \ {\isacharbar}\ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharparenleft}z{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ z\ {\isasymin}\ set\ xs\isanewline
 | |
| 753 | \ \ \ \ \ \ then\ if\ z\ {\isasymin}\ set\ ys\isanewline
 | |
| 754 | \ \ \ \ \ \ \ \ then\ collect{\isacharunderscore}duplicates\ xs\ ys\ zs\isanewline
 | |
| 755 | \ \ \ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ xs\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs\isanewline
 | |
| 756 | \ \ \ \ \ \ else\ collect{\isacharunderscore}duplicates\ {\isacharparenleft}z{\isacharhash}xs{\isacharparenright}\ {\isacharparenleft}z{\isacharhash}ys{\isacharparenright}\ zs{\isacharparenright}{\isachardoublequoteclose}%
 | |
| 28564 | 757 | \endisatagquote | 
| 758 | {\isafoldquote}%
 | |
| 28447 | 759 | % | 
| 28564 | 760 | \isadelimquote | 
| 28447 | 761 | % | 
| 28564 | 762 | \endisadelimquote | 
| 28447 | 763 | % | 
| 764 | \begin{isamarkuptext}%
 | |
| 765 | \noindent The membership test during preprocessing is rewritten, | |
| 766 |   resulting in \isa{op\ mem}, which itself
 | |
| 767 | performs an explicit equality check.% | |
| 768 | \end{isamarkuptext}%
 | |
| 769 | \isamarkuptrue% | |
| 770 | % | |
| 28564 | 771 | \isadelimquote | 
| 28447 | 772 | % | 
| 28564 | 773 | \endisadelimquote | 
| 28447 | 774 | % | 
| 28564 | 775 | \isatagquote | 
| 28447 | 776 | % | 
| 777 | \begin{isamarkuptext}%
 | |
| 28727 | 778 | \isatypewriter% | 
| 28447 | 779 | \noindent% | 
| 34155 | 780 | \hspace*{0pt}structure Example :~sig\\
 | 
| 34179 | 781 | \hspace*{0pt} ~type 'a eq\\
 | 
| 782 | \hspace*{0pt} ~val eq :~'a eq -> 'a -> 'a -> bool\\
 | |
| 783 | \hspace*{0pt} ~val eqa :~'a eq -> 'a -> 'a -> bool\\
 | |
| 784 | \hspace*{0pt} ~val member :~'a eq -> 'a -> 'a list -> bool\\
 | |
| 34155 | 785 | \hspace*{0pt} ~val collect{\char95}duplicates :\\
 | 
| 34179 | 786 | \hspace*{0pt} ~~~'a eq -> 'a list -> 'a list -> 'a list -> 'a list\\
 | 
| 34155 | 787 | \hspace*{0pt}end = struct\\
 | 
| 28714 | 788 | \hspace*{0pt}\\
 | 
| 34179 | 789 | \hspace*{0pt}type 'a eq = {\char123}eq :~'a -> 'a -> bool{\char125};\\
 | 
| 790 | \hspace*{0pt}val eq = {\char35}eq :~'a eq -> 'a -> 'a -> bool;\\
 | |
| 28714 | 791 | \hspace*{0pt}\\
 | 
| 31045 | 792 | \hspace*{0pt}fun eqa A{\char95}~a b = eq A{\char95}~a b;\\
 | 
| 28714 | 793 | \hspace*{0pt}\\
 | 
| 794 | \hspace*{0pt}fun member A{\char95}~x [] = false\\
 | |
| 34179 | 795 | \hspace*{0pt} ~| member A{\char95}~x (y ::~ys) = eqa A{\char95}~x y orelse member A{\char95}~x ys;\\
 | 
| 28714 | 796 | \hspace*{0pt}\\
 | 
| 34179 | 797 | \hspace*{0pt}fun collect{\char95}duplicates A{\char95}~xs ys [] = xs\\
 | 
| 798 | \hspace*{0pt} ~| collect{\char95}duplicates A{\char95}~xs ys (z ::~zs) =\\
 | |
| 28714 | 799 | \hspace*{0pt} ~~~(if member A{\char95}~z xs\\
 | 
| 800 | \hspace*{0pt} ~~~~~then (if member A{\char95}~z ys then collect{\char95}duplicates A{\char95}~xs ys zs\\
 | |
| 34179 | 801 | \hspace*{0pt} ~~~~~~~~~~~~else collect{\char95}duplicates A{\char95}~xs (z ::~ys) zs)\\
 | 
| 802 | \hspace*{0pt} ~~~~~else collect{\char95}duplicates A{\char95}~(z ::~xs) (z ::~ys) zs);\\
 | |
| 28714 | 803 | \hspace*{0pt}\\
 | 
| 29297 | 804 | \hspace*{0pt}end;~(*struct Example*)%
 | 
| 28447 | 805 | \end{isamarkuptext}%
 | 
| 806 | \isamarkuptrue% | |
| 807 | % | |
| 28564 | 808 | \endisatagquote | 
| 809 | {\isafoldquote}%
 | |
| 28447 | 810 | % | 
| 28564 | 811 | \isadelimquote | 
| 28447 | 812 | % | 
| 28564 | 813 | \endisadelimquote | 
| 28447 | 814 | % | 
| 815 | \begin{isamarkuptext}%
 | |
| 816 | \noindent Obviously, polymorphic equality is implemented the Haskell | |
| 817 | way using a type class. How is this achieved? HOL introduces | |
| 818 |   an explicit class \isa{eq} with a corresponding operation
 | |
| 819 |   \isa{eq{\isacharunderscore}class{\isachardot}eq} such that \isa{eq{\isacharunderscore}class{\isachardot}eq\ {\isacharequal}\ op\ {\isacharequal}}.
 | |
| 820 | The preprocessing framework does the rest by propagating the | |
| 29560 | 821 |   \isa{eq} constraints through all dependent code equations.
 | 
| 28447 | 822 |   For datatypes, instances of \isa{eq} are implicitly derived
 | 
| 823 |   when possible.  For other types, you may instantiate \isa{eq}
 | |
| 33707 
68841fb382e0
dropped obsolete documentation; updated generated sources
 haftmann parents: 
32000diff
changeset | 824 | manually like any other type class.% | 
| 28447 | 825 | \end{isamarkuptext}%
 | 
| 826 | \isamarkuptrue% | |
| 827 | % | |
| 28462 | 828 | \isamarkupsubsection{Explicit partiality%
 | 
| 28447 | 829 | } | 
| 830 | \isamarkuptrue% | |
| 831 | % | |
| 832 | \begin{isamarkuptext}%
 | |
| 28462 | 833 | Partiality usually enters the game by partial patterns, as | 
| 834 | in the following example, again for amortised queues:% | |
| 835 | \end{isamarkuptext}%
 | |
| 836 | \isamarkuptrue% | |
| 837 | % | |
| 28564 | 838 | \isadelimquote | 
| 28462 | 839 | % | 
| 28564 | 840 | \endisadelimquote | 
| 28462 | 841 | % | 
| 28564 | 842 | \isatagquote | 
| 29798 | 843 | \isacommand{definition}\isamarkupfalse%
 | 
| 28462 | 844 | \ strict{\isacharunderscore}dequeue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
 | 
| 29798 | 845 | \ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue\ q\ {\isacharequal}\ {\isacharparenleft}case\ dequeue\ q\isanewline
 | 
| 846 | \ \ \ \ of\ {\isacharparenleft}Some\ x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 847 | \isanewline | |
| 848 | \isacommand{lemma}\isamarkupfalse%
 | |
| 849 | \ strict{\isacharunderscore}dequeue{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline
 | |
| 850 | \ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 851 | \ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline
 | |
| 852 | \ \ \ \ {\isacharparenleft}case\ rev\ xs\ of\ y\ {\isacharhash}\ ys\ {\isasymRightarrow}\ {\isacharparenleft}y{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ ys{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 853 | \ \ \isacommand{by}\isamarkupfalse%
 | |
| 854 | \ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ strict{\isacharunderscore}dequeue{\isacharunderscore}def\ dequeue{\isacharunderscore}AQueue\ split{\isacharcolon}\ list{\isachardot}splits{\isacharparenright}%
 | |
| 28564 | 855 | \endisatagquote | 
| 856 | {\isafoldquote}%
 | |
| 28462 | 857 | % | 
| 28564 | 858 | \isadelimquote | 
| 28462 | 859 | % | 
| 28564 | 860 | \endisadelimquote | 
| 28462 | 861 | % | 
| 862 | \begin{isamarkuptext}%
 | |
| 863 | \noindent In the corresponding code, there is no equation | |
| 30227 | 864 |   for the pattern \isa{AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}}:%
 | 
| 28462 | 865 | \end{isamarkuptext}%
 | 
| 866 | \isamarkuptrue% | |
| 867 | % | |
| 28564 | 868 | \isadelimquote | 
| 28462 | 869 | % | 
| 28564 | 870 | \endisadelimquote | 
| 28462 | 871 | % | 
| 28564 | 872 | \isatagquote | 
| 28462 | 873 | % | 
| 874 | \begin{isamarkuptext}%
 | |
| 28727 | 875 | \isatypewriter% | 
| 28462 | 876 | \noindent% | 
| 29297 | 877 | \hspace*{0pt}strict{\char95}dequeue ::~forall a.~Queue a -> (a,~Queue a);\\
 | 
| 34179 | 878 | \hspace*{0pt}strict{\char95}dequeue (AQueue xs []) =\\
 | 
| 31848 | 879 | \hspace*{0pt} ~let {\char123}\\
 | 
| 34179 | 880 | \hspace*{0pt} ~~~(y :~ys) = rev xs;\\
 | 
| 31848 | 881 | \hspace*{0pt} ~{\char125}~in (y,~AQueue [] ys);\\
 | 
| 34179 | 882 | \hspace*{0pt}strict{\char95}dequeue (AQueue xs (y :~ys)) = (y,~AQueue xs ys);%
 | 
| 28462 | 883 | \end{isamarkuptext}%
 | 
| 884 | \isamarkuptrue% | |
| 885 | % | |
| 28564 | 886 | \endisatagquote | 
| 887 | {\isafoldquote}%
 | |
| 28462 | 888 | % | 
| 28564 | 889 | \isadelimquote | 
| 28462 | 890 | % | 
| 28564 | 891 | \endisadelimquote | 
| 28462 | 892 | % | 
| 893 | \begin{isamarkuptext}%
 | |
| 894 | \noindent In some cases it is desirable to have this | |
| 895 |   pseudo-\qt{partiality} more explicitly, e.g.~as follows:%
 | |
| 896 | \end{isamarkuptext}%
 | |
| 897 | \isamarkuptrue% | |
| 898 | % | |
| 28564 | 899 | \isadelimquote | 
| 28462 | 900 | % | 
| 28564 | 901 | \endisadelimquote | 
| 28462 | 902 | % | 
| 28564 | 903 | \isatagquote | 
| 28462 | 904 | \isacommand{axiomatization}\isamarkupfalse%
 | 
| 905 | \ empty{\isacharunderscore}queue\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\isanewline
 | |
| 906 | \isanewline | |
| 29798 | 907 | \isacommand{definition}\isamarkupfalse%
 | 
| 28462 | 908 | \ strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
 | 
| 29798 | 909 | \ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue{\isacharprime}\ q\ {\isacharequal}\ {\isacharparenleft}case\ dequeue\ q\ of\ {\isacharparenleft}Some\ x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}x{\isacharcomma}\ q{\isacharprime}{\isacharparenright}\ {\isacharbar}\ {\isacharunderscore}\ {\isasymRightarrow}\ empty{\isacharunderscore}queue{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | 
| 28462 | 910 | \isanewline | 
| 29798 | 911 | \isacommand{lemma}\isamarkupfalse%
 | 
| 912 | \ strict{\isacharunderscore}dequeue{\isacharprime}{\isacharunderscore}AQueue\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline
 | |
| 913 | \ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ empty{\isacharunderscore}queue\isanewline
 | |
| 914 | \ \ \ \ \ else\ strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 915 | \ \ {\isachardoublequoteopen}strict{\isacharunderscore}dequeue{\isacharprime}\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline
 | |
| 916 | \ \ \ \ \ {\isacharparenleft}y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 917 | \ \ \isacommand{by}\isamarkupfalse%
 | |
| 918 | \ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ strict{\isacharunderscore}dequeue{\isacharprime}{\isacharunderscore}def\ dequeue{\isacharunderscore}AQueue\ split{\isacharcolon}\ list{\isachardot}splits{\isacharparenright}%
 | |
| 28564 | 919 | \endisatagquote | 
| 920 | {\isafoldquote}%
 | |
| 28462 | 921 | % | 
| 28564 | 922 | \isadelimquote | 
| 28462 | 923 | % | 
| 28564 | 924 | \endisadelimquote | 
| 28462 | 925 | % | 
| 926 | \begin{isamarkuptext}%
 | |
| 34155 | 927 | Observe that on the right hand side of the definition of \isa{strict{\isacharunderscore}dequeue{\isacharprime}}, the unspecified constant \isa{empty{\isacharunderscore}queue} occurs.
 | 
| 28462 | 928 | |
| 29798 | 929 | Normally, if constants without any code equations occur in a | 
| 930 | program, the code generator complains (since in most cases this is | |
| 34155 | 931 | indeed an error). But such constants can also be thought | 
| 932 | of as function definitions which always fail, | |
| 29798 | 933 | since there is never a successful pattern match on the left hand | 
| 934 | side. In order to categorise a constant into that category | |
| 935 |   explicitly, use \hyperlink{command.code-abort}{\mbox{\isa{\isacommand{code{\isacharunderscore}abort}}}}:%
 | |
| 28462 | 936 | \end{isamarkuptext}%
 | 
| 937 | \isamarkuptrue% | |
| 938 | % | |
| 28564 | 939 | \isadelimquote | 
| 28462 | 940 | % | 
| 28564 | 941 | \endisadelimquote | 
| 28462 | 942 | % | 
| 28564 | 943 | \isatagquote | 
| 28462 | 944 | \isacommand{code{\isacharunderscore}abort}\isamarkupfalse%
 | 
| 945 | \ empty{\isacharunderscore}queue%
 | |
| 28564 | 946 | \endisatagquote | 
| 947 | {\isafoldquote}%
 | |
| 28462 | 948 | % | 
| 28564 | 949 | \isadelimquote | 
| 28462 | 950 | % | 
| 28564 | 951 | \endisadelimquote | 
| 28462 | 952 | % | 
| 953 | \begin{isamarkuptext}%
 | |
| 954 | \noindent Then the code generator will just insert an error or | |
| 955 | exception at the appropriate position:% | |
| 956 | \end{isamarkuptext}%
 | |
| 957 | \isamarkuptrue% | |
| 958 | % | |
| 28564 | 959 | \isadelimquote | 
| 28462 | 960 | % | 
| 28564 | 961 | \endisadelimquote | 
| 28462 | 962 | % | 
| 28564 | 963 | \isatagquote | 
| 28462 | 964 | % | 
| 965 | \begin{isamarkuptext}%
 | |
| 28727 | 966 | \isatypewriter% | 
| 28462 | 967 | \noindent% | 
| 29297 | 968 | \hspace*{0pt}empty{\char95}queue ::~forall a.~a;\\
 | 
| 28714 | 969 | \hspace*{0pt}empty{\char95}queue = error {\char34}empty{\char95}queue{\char34};\\
 | 
| 970 | \hspace*{0pt}\\
 | |
| 29297 | 971 | \hspace*{0pt}strict{\char95}dequeue' ::~forall a.~Queue a -> (a,~Queue a);\\
 | 
| 34179 | 972 | \hspace*{0pt}strict{\char95}dequeue' (AQueue xs (y :~ys)) = (y,~AQueue xs ys);\\
 | 
| 973 | \hspace*{0pt}strict{\char95}dequeue' (AQueue xs []) =\\
 | |
| 29798 | 974 | \hspace*{0pt} ~(if nulla xs then empty{\char95}queue\\
 | 
| 34179 | 975 | \hspace*{0pt} ~~~else strict{\char95}dequeue' (AQueue [] (rev xs)));%
 | 
| 28462 | 976 | \end{isamarkuptext}%
 | 
| 977 | \isamarkuptrue% | |
| 978 | % | |
| 28564 | 979 | \endisatagquote | 
| 980 | {\isafoldquote}%
 | |
| 28462 | 981 | % | 
| 28564 | 982 | \isadelimquote | 
| 28462 | 983 | % | 
| 28564 | 984 | \endisadelimquote | 
| 28462 | 985 | % | 
| 986 | \begin{isamarkuptext}%
 | |
| 987 | \noindent This feature however is rarely needed in practice. | |
| 988 |   Note also that the \isa{HOL} default setup already declares
 | |
| 989 |   \isa{undefined} as \hyperlink{command.code-abort}{\mbox{\isa{\isacommand{code{\isacharunderscore}abort}}}}, which is most
 | |
| 990 | likely to be used in such situations.% | |
| 28447 | 991 | \end{isamarkuptext}%
 | 
| 992 | \isamarkuptrue% | |
| 993 | % | |
| 33926 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 994 | \isamarkupsubsection{Inductive Predicates%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 995 | } | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 996 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 997 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 998 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 999 | To execute inductive predicates, a special preprocessor, the predicate | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1000 | compiler, generates code equations from the introduction rules of the predicates. | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1001 | The mechanisms of this compiler are described in \cite{Berghofer-Bulwahn-Haftmann:2009:TPHOL}.
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1002 | Consider the simple predicate \isa{append} given by these two
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1003 | introduction rules:% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1004 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1005 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1006 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1007 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1008 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1009 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1010 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1011 | \isatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1012 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1013 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1014 | \isa{append\ {\isacharbrackleft}{\isacharbrackright}\ ys\ ys}\\
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1015 | \noindent\isa{append\ xs\ ys\ zs\ {\isasymLongrightarrow}\ append\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ ys\ {\isacharparenleft}x\ {\isacharhash}\ zs{\isacharparenright}}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1016 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1017 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1018 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1019 | \endisatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1020 | {\isafoldquote}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1021 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1022 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1023 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1024 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1025 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1026 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1027 | \noindent To invoke the compiler, simply use \indexdef{}{command}{code\_pred}\hypertarget{command.code-pred}{\hyperlink{command.code-pred}{\mbox{\isa{\isacommand{code{\isacharunderscore}pred}}}}}:%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1028 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1029 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1030 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1031 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1032 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1033 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1034 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1035 | \isatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1036 | \isacommand{code{\isacharunderscore}pred}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1037 | \ append\ \isacommand{{\isachardot}}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1038 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1039 | \endisatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1040 | {\isafoldquote}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1041 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1042 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1043 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1044 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1045 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1046 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1047 | \noindent The \hyperlink{command.code-pred}{\mbox{\isa{\isacommand{code{\isacharunderscore}pred}}}} command takes the name
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1048 | of the inductive predicate and then you put a period to discharge | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1049 | a trivial correctness proof. | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1050 | The compiler infers possible modes | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1051 | for the predicate and produces the derived code equations. | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1052 | Modes annotate which (parts of the) arguments are to be taken as input, | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1053 | and which output. Modes are similar to types, but use the notation \isa{i}
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1054 | for input and \isa{o} for output.
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1055 | |
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1056 | For \isa{append}, the compiler can infer the following modes:
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1057 | \begin{itemize}
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1058 | \item \isa{i\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ bool}
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1059 | \item \isa{i\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ bool}
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1060 | \item \isa{o\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ bool}
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1061 | \end{itemize}
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1062 | You can compute sets of predicates using \indexdef{}{command}{values}\hypertarget{command.values}{\hyperlink{command.values}{\mbox{\isa{\isacommand{values}}}}}:%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1063 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1064 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1065 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1066 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1067 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1068 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1069 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1070 | \isatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1071 | \isacommand{values}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1072 | \ {\isachardoublequoteopen}{\isacharbraceleft}zs{\isachardot}\ append\ {\isacharbrackleft}{\isacharparenleft}{\isadigit{1}}{\isacharcolon}{\isacharcolon}nat{\isacharparenright}{\isacharcomma}{\isadigit{2}}{\isacharcomma}{\isadigit{3}}{\isacharbrackright}\ {\isacharbrackleft}{\isadigit{4}}{\isacharcomma}{\isadigit{5}}{\isacharbrackright}\ zs{\isacharbraceright}{\isachardoublequoteclose}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1073 | \endisatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1074 | {\isafoldquote}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1075 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1076 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1077 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1078 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1079 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1080 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1081 | \noindent outputs \isa{{\isacharbraceleft}{\isacharbrackleft}{\isadigit{1}}{\isacharcomma}\ {\isadigit{2}}{\isacharcomma}\ {\isadigit{3}}{\isacharcomma}\ {\isadigit{4}}{\isacharcomma}\ {\isadigit{5}}{\isacharbrackright}{\isacharbraceright}}, and%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1082 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1083 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1084 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1085 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1086 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1087 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1088 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1089 | \isatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1090 | \isacommand{values}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1091 | \ {\isachardoublequoteopen}{\isacharbraceleft}{\isacharparenleft}xs{\isacharcomma}\ ys{\isacharparenright}{\isachardot}\ append\ xs\ ys\ {\isacharbrackleft}{\isacharparenleft}{\isadigit{2}}{\isacharcolon}{\isacharcolon}nat{\isacharparenright}{\isacharcomma}{\isadigit{3}}{\isacharbrackright}{\isacharbraceright}{\isachardoublequoteclose}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1092 | \endisatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1093 | {\isafoldquote}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1094 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1095 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1096 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1097 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1098 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1099 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1100 | \noindent outputs \isa{{\isacharbraceleft}{\isacharparenleft}{\isacharbrackleft}{\isacharbrackright}{\isacharcomma}\ {\isacharbrackleft}{\isadigit{2}}{\isacharcomma}\ {\isadigit{3}}{\isacharbrackright}{\isacharparenright}{\isacharcomma}\ {\isacharparenleft}{\isacharbrackleft}{\isadigit{2}}{\isacharbrackright}{\isacharcomma}\ {\isacharbrackleft}{\isadigit{3}}{\isacharbrackright}{\isacharparenright}{\isacharcomma}\ {\isacharparenleft}{\isacharbrackleft}{\isadigit{2}}{\isacharcomma}\ {\isadigit{3}}{\isacharbrackright}{\isacharcomma}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isacharbraceright}}.%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1101 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1102 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1103 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1104 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1105 | \noindent If you are only interested in the first elements of the set | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1106 | comprehension (with respect to a depth-first search on the introduction rules), you can | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1107 | pass an argument to | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1108 | \hyperlink{command.values}{\mbox{\isa{\isacommand{values}}}} to specify the number of elements you want:%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1109 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1110 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1111 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1112 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1113 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1114 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1115 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1116 | \isatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1117 | \isacommand{values}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1118 | \ {\isadigit{1}}\ {\isachardoublequoteopen}{\isacharbraceleft}{\isacharparenleft}xs{\isacharcomma}\ ys{\isacharparenright}{\isachardot}\ append\ xs\ ys\ {\isacharbrackleft}{\isacharparenleft}{\isadigit{1}}{\isacharcolon}{\isacharcolon}nat{\isacharparenright}{\isacharcomma}{\isadigit{2}}{\isacharcomma}{\isadigit{3}}{\isacharcomma}{\isadigit{4}}{\isacharbrackright}{\isacharbraceright}{\isachardoublequoteclose}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1119 | \isacommand{values}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1120 | \ {\isadigit{3}}\ {\isachardoublequoteopen}{\isacharbraceleft}{\isacharparenleft}xs{\isacharcomma}\ ys{\isacharparenright}{\isachardot}\ append\ xs\ ys\ {\isacharbrackleft}{\isacharparenleft}{\isadigit{1}}{\isacharcolon}{\isacharcolon}nat{\isacharparenright}{\isacharcomma}{\isadigit{2}}{\isacharcomma}{\isadigit{3}}{\isacharcomma}{\isadigit{4}}{\isacharbrackright}{\isacharbraceright}{\isachardoublequoteclose}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1121 | \endisatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1122 | {\isafoldquote}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1123 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1124 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1125 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1126 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1127 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1128 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1129 | \noindent The \hyperlink{command.values}{\mbox{\isa{\isacommand{values}}}} command can only compute set
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1130 | comprehensions for which a mode has been inferred. | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1131 | |
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1132 | The code equations for a predicate are made available as theorems with | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1133 |  the suffix \isa{equation}, and can be inspected with:%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1134 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1135 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1136 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1137 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1138 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1139 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1140 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1141 | \isatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1142 | \isacommand{thm}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1143 | \ append{\isachardot}equation%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1144 | \endisatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1145 | {\isafoldquote}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1146 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1147 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1148 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1149 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1150 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1151 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1152 | \noindent More advanced options are described in the following subsections.% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1153 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1154 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1155 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1156 | \isamarkupsubsubsection{Alternative names for functions%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1157 | } | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1158 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1159 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1160 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1161 | By default, the functions generated from a predicate are named after the predicate with the | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1162 | mode mangled into the name (e.g., \isa{append{\isacharunderscore}i{\isacharunderscore}i{\isacharunderscore}o}).
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1163 | You can specify your own names as follows:% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1164 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1165 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1166 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1167 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1168 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1169 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1170 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1171 | \isatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1172 | \isacommand{code{\isacharunderscore}pred}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1173 | \ {\isacharparenleft}modes{\isacharcolon}\ i\ {\isacharequal}{\isachargreater}\ i\ {\isacharequal}{\isachargreater}\ o\ {\isacharequal}{\isachargreater}\ bool\ as\ concat{\isacharcomma}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1174 | \ \ o\ {\isacharequal}{\isachargreater}\ o\ {\isacharequal}{\isachargreater}\ i\ {\isacharequal}{\isachargreater}\ bool\ as\ split{\isacharcomma}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1175 | \ \ i\ {\isacharequal}{\isachargreater}\ o\ {\isacharequal}{\isachargreater}\ i\ {\isacharequal}{\isachargreater}\ bool\ as\ suffix{\isacharparenright}\ append\ \isacommand{{\isachardot}}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1176 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1177 | \endisatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1178 | {\isafoldquote}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1179 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1180 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1181 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1182 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1183 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1184 | \isamarkupsubsubsection{Alternative introduction rules%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1185 | } | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1186 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1187 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1188 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1189 | Sometimes the introduction rules of an predicate are not executable because they contain | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1190 | non-executable constants or specific modes could not be inferred. | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1191 | It is also possible that the introduction rules yield a function that loops forever | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1192 | due to the execution in a depth-first search manner. | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1193 | Therefore, you can declare alternative introduction rules for predicates with the attribute | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1194 | \hyperlink{attribute.code-pred-intro}{\mbox{\isa{code{\isacharunderscore}pred{\isacharunderscore}intro}}}.
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1195 | For example, the transitive closure is defined by:% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1196 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1197 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1198 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1199 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1200 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1201 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1202 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1203 | \isatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1204 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1205 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1206 | \isa{r\ a\ b\ {\isasymLongrightarrow}\ r\isactrlsup {\isacharplus}\isactrlsup {\isacharplus}\ a\ b}\\
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1207 | \noindent \isa{{\isasymlbrakk}r\isactrlsup {\isacharplus}\isactrlsup {\isacharplus}\ a\ b{\isacharsemicolon}\ r\ b\ c{\isasymrbrakk}\ {\isasymLongrightarrow}\ r\isactrlsup {\isacharplus}\isactrlsup {\isacharplus}\ a\ c}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1208 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1209 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1210 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1211 | \endisatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1212 | {\isafoldquote}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1213 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1214 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1215 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1216 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1217 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1218 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1219 | \noindent These rules do not suit well for executing the transitive closure | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1220 | with the mode \isa{{\isacharparenleft}i\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ bool}, as the second rule will
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1221 | cause an infinite loop in the recursive call. | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1222 | This can be avoided using the following alternative rules which are | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1223 | declared to the predicate compiler by the attribute \hyperlink{attribute.code-pred-intro}{\mbox{\isa{code{\isacharunderscore}pred{\isacharunderscore}intro}}}:%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1224 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1225 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1226 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1227 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1228 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1229 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1230 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1231 | \isatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1232 | \isacommand{lemma}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1233 | \ {\isacharbrackleft}code{\isacharunderscore}pred{\isacharunderscore}intro{\isacharbrackright}{\isacharcolon}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1234 | \ \ {\isachardoublequoteopen}r\ a\ b\ {\isasymLongrightarrow}\ r\isactrlsup {\isacharplus}\isactrlsup {\isacharplus}\ a\ b{\isachardoublequoteclose}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1235 | \ \ {\isachardoublequoteopen}r\ a\ b\ {\isasymLongrightarrow}\ r\isactrlsup {\isacharplus}\isactrlsup {\isacharplus}\ b\ c\ {\isasymLongrightarrow}\ r\isactrlsup {\isacharplus}\isactrlsup {\isacharplus}\ a\ c{\isachardoublequoteclose}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1236 | \isacommand{by}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1237 | \ auto% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1238 | \endisatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1239 | {\isafoldquote}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1240 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1241 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1242 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1243 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1244 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1245 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1246 | \noindent After declaring all alternative rules for the transitive closure, | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1247 | you invoke \hyperlink{command.code-pred}{\mbox{\isa{\isacommand{code{\isacharunderscore}pred}}}} as usual.
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1248 | As you have declared alternative rules for the predicate, you are urged to prove that these | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1249 | introduction rules are complete, i.e., that you can derive an elimination rule for the | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1250 | alternative rules:% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1251 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1252 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1253 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1254 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1255 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1256 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1257 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1258 | \isatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1259 | \isacommand{code{\isacharunderscore}pred}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1260 | \ tranclp\isanewline | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1261 | \isacommand{proof}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1262 | \ {\isacharminus}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1263 | \ \ \isacommand{case}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1264 | \ tranclp\isanewline | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1265 | \ \ \isacommand{from}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1266 | \ this\ converse{\isacharunderscore}tranclpE{\isacharbrackleft}OF\ this{\isacharparenleft}{\isadigit{1}}{\isacharparenright}{\isacharbrackright}\ \isacommand{show}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1267 | \ thesis\ \isacommand{by}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1268 | \ metis\isanewline | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1269 | \isacommand{qed}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1270 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1271 | \endisatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1272 | {\isafoldquote}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1273 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1274 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1275 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1276 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1277 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1278 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1279 | \noindent Alternative rules can also be used for constants that have not | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1280 | been defined inductively. For example, the lexicographic order which | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1281 | is defined as:% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1282 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1283 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1284 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1285 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1286 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1287 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1288 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1289 | \isatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1290 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1291 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1292 | \begin{isabelle}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1293 | lexord\ r\ {\isasymequiv}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1294 | {\isacharbraceleft}{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}{\isachardot}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1295 | \isaindent{\ }{\isasymexists}a\ v{\isachardot}\ y\ {\isacharequal}\ x\ {\isacharat}\ a\ {\isacharhash}\ v\ {\isasymor}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1296 | \isaindent{\ {\isasymexists}a\ v{\isachardot}\ }{\isacharparenleft}{\isasymexists}u\ a\ b\ v\ w{\isachardot}\ {\isacharparenleft}a{\isacharcomma}\ b{\isacharparenright}\ {\isasymin}\ r\ {\isasymand}\ x\ {\isacharequal}\ u\ {\isacharat}\ a\ {\isacharhash}\ v\ {\isasymand}\ y\ {\isacharequal}\ u\ {\isacharat}\ b\ {\isacharhash}\ w{\isacharparenright}{\isacharbraceright}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1297 | \end{isabelle}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1298 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1299 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1300 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1301 | \endisatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1302 | {\isafoldquote}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1303 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1304 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1305 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1306 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1307 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1308 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1309 | \noindent To make it executable, you can derive the following two rules and prove the | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1310 | elimination rule:% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1311 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1312 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1313 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1314 | \isadelimproof | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1315 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1316 | \endisadelimproof | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1317 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1318 | \isatagproof | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1319 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1320 | \endisatagproof | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1321 | {\isafoldproof}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1322 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1323 | \isadelimproof | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1324 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1325 | \endisadelimproof | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1326 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1327 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1328 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1329 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1330 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1331 | \isatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1332 | \isacommand{lemma}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1333 | \ {\isacharbrackleft}code{\isacharunderscore}pred{\isacharunderscore}intro{\isacharbrackright}{\isacharcolon}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1334 | \ \ {\isachardoublequoteopen}append\ xs\ {\isacharparenleft}a\ {\isacharhash}\ v{\isacharparenright}\ ys\ {\isasymLongrightarrow}\ lexord\ r\ {\isacharparenleft}xs{\isacharcomma}\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1335 | \isacommand{lemma}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1336 | \ {\isacharbrackleft}code{\isacharunderscore}pred{\isacharunderscore}intro{\isacharbrackright}{\isacharcolon}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1337 | \ \ {\isachardoublequoteopen}append\ u\ {\isacharparenleft}a\ {\isacharhash}\ v{\isacharparenright}\ xs\ {\isasymLongrightarrow}\ append\ u\ {\isacharparenleft}b\ {\isacharhash}\ w{\isacharparenright}\ ys\ {\isasymLongrightarrow}\ r\ {\isacharparenleft}a{\isacharcomma}\ b{\isacharparenright}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1338 | \ \ {\isasymLongrightarrow}\ lexord\ r\ {\isacharparenleft}xs{\isacharcomma}\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1339 | \isacommand{code{\isacharunderscore}pred}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1340 | \ lexord% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1341 | \endisatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1342 | {\isafoldquote}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1343 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1344 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1345 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1346 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1347 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1348 | \isamarkupsubsubsection{Options for values%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1349 | } | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1350 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1351 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1352 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1353 | In the presence of higher-order predicates, multiple modes for some predicate could be inferred | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1354 | that are not disambiguated by the pattern of the set comprehension. | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1355 | To disambiguate the modes for the arguments of a predicate, you can state | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1356 | the modes explicitly in the \hyperlink{command.values}{\mbox{\isa{\isacommand{values}}}} command. 
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1357 | Consider the simple predicate \isa{succ}:%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1358 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1359 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1360 | \isacommand{inductive}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1361 | \ succ\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1362 | \isakeyword{where}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1363 | \ \ {\isachardoublequoteopen}succ\ {\isadigit{0}}\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1364 | {\isacharbar}\ {\isachardoublequoteopen}succ\ x\ y\ {\isasymLongrightarrow}\ succ\ {\isacharparenleft}Suc\ x{\isacharparenright}\ {\isacharparenleft}Suc\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1365 | \isanewline | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1366 | \isacommand{code{\isacharunderscore}pred}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1367 | \ succ% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1368 | \isadelimproof | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1369 | \ % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1370 | \endisadelimproof | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1371 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1372 | \isatagproof | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1373 | \isacommand{{\isachardot}}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1374 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1375 | \endisatagproof | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1376 | {\isafoldproof}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1377 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1378 | \isadelimproof | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1379 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1380 | \endisadelimproof | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1381 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1382 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1383 | \noindent For this, the predicate compiler can infer modes \isa{o\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ bool}, \isa{i\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ bool},
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1384 |   \isa{o\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ bool} and \isa{i\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ bool}.
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1385 | The invocation of \hyperlink{command.values}{\mbox{\isa{\isacommand{values}}}} \isa{{\isacharbraceleft}n{\isachardot}\ tranclp\ succ\ {\isadigit{1}}{\isadigit{0}}\ n{\isacharbraceright}} loops, as multiple
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1386 | modes for the predicate \isa{succ} are possible and here the first mode \isa{o\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ bool}
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1387 | is chosen. To choose another mode for the argument, you can declare the mode for the argument between | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1388 | the \hyperlink{command.values}{\mbox{\isa{\isacommand{values}}}} and the number of elements.%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1389 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1390 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1391 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1392 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1393 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1394 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1395 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1396 | \isatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1397 | \isacommand{values}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1398 | \ {\isacharbrackleft}mode{\isacharcolon}\ i\ {\isacharequal}{\isachargreater}\ o\ {\isacharequal}{\isachargreater}\ bool{\isacharbrackright}\ {\isadigit{2}}{\isadigit{0}}\ {\isachardoublequoteopen}{\isacharbraceleft}n{\isachardot}\ tranclp\ succ\ {\isadigit{1}}{\isadigit{0}}\ n{\isacharbraceright}{\isachardoublequoteclose}\isanewline
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1399 | \isacommand{values}\isamarkupfalse%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1400 | \ {\isacharbrackleft}mode{\isacharcolon}\ o\ {\isacharequal}{\isachargreater}\ i\ {\isacharequal}{\isachargreater}\ bool{\isacharbrackright}\ {\isadigit{1}}{\isadigit{0}}\ {\isachardoublequoteopen}{\isacharbraceleft}n{\isachardot}\ tranclp\ succ\ n\ {\isadigit{1}}{\isadigit{0}}{\isacharbraceright}{\isachardoublequoteclose}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1401 | \endisatagquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1402 | {\isafoldquote}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1403 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1404 | \isadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1405 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1406 | \endisadelimquote | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1407 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1408 | \isamarkupsubsubsection{Embedding into functional code within Isabelle/HOL%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1409 | } | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1410 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1411 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1412 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1413 | To embed the computation of an inductive predicate into functions that are defined in Isabelle/HOL, | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1414 | you have a number of options: | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1415 | \begin{itemize}
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1416 | \item You want to use the first-order predicate with the mode | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1417 | where all arguments are input. Then you can use the predicate directly, e.g. | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1418 | \begin{quote}
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1419 |   \isa{valid{\isacharunderscore}suffix\ ys\ zs\ {\isacharequal}}\\
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1420 |   \isa{{\isacharparenleft}if\ append\ {\isacharbrackleft}Suc\ {\isadigit{0}}{\isacharcomma}\ {\isadigit{2}}{\isacharbrackright}\ ys\ zs\ then\ Some\ ys\ else\ None{\isacharparenright}}
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1421 | \end{quote}
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1422 | \item If you know that the execution returns only one value (it is deterministic), then you can | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1423 |   use the combinator \isa{Predicate{\isachardot}the}, e.g., a functional concatenation of lists
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1424 | is defined with | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1425 | \begin{quote}
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1426 |   \isa{functional{\isacharunderscore}concat\ xs\ ys\ {\isacharequal}\ Predicate{\isachardot}the\ {\isacharparenleft}append{\isacharunderscore}i{\isacharunderscore}i{\isacharunderscore}o\ xs\ ys{\isacharparenright}}
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1427 | \end{quote}
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1428 | Note that if the evaluation does not return a unique value, it raises a run-time error | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1429 |   \isa{not{\isacharunderscore}unique}.
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1430 | \end{itemize}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1431 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1432 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1433 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1434 | \isamarkupsubsubsection{Further Examples%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1435 | } | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1436 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1437 | % | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1438 | \begin{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1439 | Further examples for compiling inductive predicates can be found in | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1440 | the \isa{HOL{\isacharslash}ex{\isacharslash}Predicate{\isacharunderscore}Compile{\isacharunderscore}ex} theory file.
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1441 | There are also some examples in the Archive of Formal Proofs, notably | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1442 | in the \isa{POPLmark{\isacharminus}deBruijn} and the \isa{FeatherweightJava} sessions.%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1443 | \end{isamarkuptext}%
 | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1444 | \isamarkuptrue% | 
| 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1445 | % | 
| 28447 | 1446 | \isadelimtheory | 
| 1447 | % | |
| 1448 | \endisadelimtheory | |
| 1449 | % | |
| 1450 | \isatagtheory | |
| 1451 | \isacommand{end}\isamarkupfalse%
 | |
| 1452 | % | |
| 1453 | \endisatagtheory | |
| 1454 | {\isafoldtheory}%
 | |
| 1455 | % | |
| 1456 | \isadelimtheory | |
| 1457 | % | |
| 1458 | \endisadelimtheory | |
| 1459 | \isanewline | |
| 33926 
dd017d9db05f
adding subsection about the predicate compiler to the code generator tutorial
 bulwahn parents: 
33707diff
changeset | 1460 | \end{isabellebody}%
 | 
| 28447 | 1461 | %%% Local Variables: | 
| 1462 | %%% mode: latex | |
| 1463 | %%% TeX-master: "root" | |
| 1464 | %%% End: |