| author | wenzelm | 
| Sun, 28 Jan 2007 23:29:15 +0100 | |
| changeset 22201 | 6fe46a7259ec | 
| parent 22098 | 88be1b7775c8 | 
| child 25338 | 6eb185959aec | 
| permissions | -rw-r--r-- | 
| 11647 | 1 | (*<*) | 
| 16417 | 2 | theory Documents imports Main begin | 
| 11647 | 3 | (*>*) | 
| 4 | ||
| 12648 | 5 | section {* Concrete Syntax \label{sec:concrete-syntax} *}
 | 
| 12629 | 6 | |
| 7 | text {*
 | |
| 12766 | 8 | The core concept of Isabelle's framework for concrete syntax is that | 
| 9 |   of \bfindex{mixfix annotations}.  Associated with any kind of
 | |
| 10 | constant declaration, mixfixes affect both the grammar productions | |
| 11 | for the parser and output templates for the pretty printer. | |
| 12629 | 12 | |
| 12743 | 13 | In full generality, parser and pretty printer configuration is a | 
| 12766 | 14 |   subtle affair \cite{isabelle-ref}.  Your syntax specifications need
 | 
| 15 | to interact properly with the existing setup of Isabelle/Pure and | |
| 16 | Isabelle/HOL\@. To avoid creating ambiguities with existing | |
| 17 | elements, it is particularly important to give new syntactic | |
| 12764 | 18 | constructs the right precedence. | 
| 12629 | 19 | |
| 12670 | 20 | \medskip Subsequently we introduce a few simple syntax declaration | 
| 12743 | 21 | forms that already cover many common situations fairly well. | 
| 12629 | 22 | *} | 
| 23 | ||
| 24 | ||
| 12648 | 25 | subsection {* Infix Annotations *}
 | 
| 12629 | 26 | |
| 27 | text {*
 | |
| 12764 | 28 | Syntax annotations may be included wherever constants are declared, | 
| 12766 | 29 |   such as \isacommand{consts} and \isacommand{constdefs} --- and also
 | 
| 30 |   \isacommand{datatype}, which declares constructor operations.
 | |
| 31 | Type-constructors may be annotated as well, although this is less | |
| 32 |   frequently encountered in practice (the infix type @{text "\<times>"} comes
 | |
| 33 | to mind). | |
| 12629 | 34 | |
| 12645 | 35 |   Infix declarations\index{infix annotations} provide a useful special
 | 
| 12766 | 36 | case of mixfixes. The following example of the exclusive-or | 
| 37 | operation on boolean values illustrates typical infix declarations. | |
| 12629 | 38 | *} | 
| 39 | ||
| 40 | constdefs | |
| 41 | xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "[+]" 60) | |
| 42 | "A [+] B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)" | |
| 43 | ||
| 44 | text {*
 | |
| 12653 | 45 |   \noindent Now @{text "xor A B"} and @{text "A [+] B"} refer to the
 | 
| 46 | same expression internally. Any curried function with at least two | |
| 12766 | 47 | arguments may be given infix syntax. For partial applications with | 
| 48 |   fewer than two operands, there is a notation using the prefix~@{text
 | |
| 49 |   op}.  For instance, @{text xor} without arguments is represented as
 | |
| 50 |   @{text "op [+]"}; together with ordinary function application, this
 | |
| 12653 | 51 |   turns @{text "xor A"} into @{text "op [+] A"}.
 | 
| 12629 | 52 | |
| 12746 | 53 |   \medskip The keyword \isakeyword{infixl} seen above specifies an
 | 
| 54 |   infix operator that is nested to the \emph{left}: in iterated
 | |
| 55 | applications the more complex expression appears on the left-hand | |
| 12766 | 56 |   side, and @{term "A [+] B [+] C"} stands for @{text "(A [+] B) [+]
 | 
| 57 |   C"}.  Similarly, \isakeyword{infixr} means nesting to the
 | |
| 12746 | 58 |   \emph{right}, reading @{term "A [+] B [+] C"} as @{text "A [+] (B
 | 
| 12766 | 59 |   [+] C)"}.  A \emph{non-oriented} declaration via \isakeyword{infix}
 | 
| 60 |   would render @{term "A [+] B [+] C"} illegal, but demand explicit
 | |
| 61 | parentheses to indicate the intended grouping. | |
| 12743 | 62 | |
| 12746 | 63 |   The string @{text [source] "[+]"} in our annotation refers to the
 | 
| 64 | concrete syntax to represent the operator (a literal token), while | |
| 12764 | 65 |   the number @{text 60} determines the precedence of the construct:
 | 
| 12766 | 66 | the syntactic priorities of the arguments and result. Isabelle/HOL | 
| 67 | already uses up many popular combinations of ASCII symbols for its | |
| 68 |   own use, including both @{text "+"} and @{text "++"}.  Longer
 | |
| 69 | character combinations are more likely to be still available for | |
| 70 |   user extensions, such as our~@{text "[+]"}.
 | |
| 12629 | 71 | |
| 12766 | 72 | Operator precedences have a range of 0--1000. Very low or high | 
| 73 | priorities are reserved for the meta-logic. HOL syntax mainly uses | |
| 74 |   the range of 10--100: the equality infix @{text "="} is centered at
 | |
| 75 |   50; logical connectives (like @{text "\<or>"} and @{text "\<and>"}) are
 | |
| 76 |   below 50; algebraic ones (like @{text "+"} and @{text "*"}) are
 | |
| 77 | above 50. User syntax should strive to coexist with common HOL | |
| 78 | forms, or use the mostly unused range 100--900. | |
| 12629 | 79 | *} | 
| 80 | ||
| 12635 | 81 | |
| 12659 | 82 | subsection {* Mathematical Symbols \label{sec:syntax-symbols} *}
 | 
| 12629 | 83 | |
| 84 | text {*
 | |
| 12766 | 85 | Concrete syntax based on ASCII characters has inherent limitations. | 
| 86 | Mathematical notation demands a larger repertoire of glyphs. | |
| 87 | Several standards of extended character sets have been proposed over | |
| 88 | decades, but none has become universally available so far. Isabelle | |
| 89 |   has its own notion of \bfindex{symbols} as the smallest entities of
 | |
| 90 | source text, without referring to internal encodings. There are | |
| 91 | three kinds of such ``generalized characters'': | |
| 12635 | 92 | |
| 93 |   \begin{enumerate}
 | |
| 94 | ||
| 12653 | 95 | \item 7-bit ASCII characters | 
| 12635 | 96 | |
| 12653 | 97 | \item named symbols: \verb,\,\verb,<,$ident$\verb,>, | 
| 12629 | 98 | |
| 12653 | 99 | \item named control symbols: \verb,\,\verb,<^,$ident$\verb,>, | 
| 12635 | 100 | |
| 101 |   \end{enumerate}
 | |
| 102 | ||
| 14486 | 103 | Here $ident$ is any sequence of letters. | 
| 104 | This results in an infinite store of symbols, whose | |
| 12766 | 105 | interpretation is left to further front-end tools. For example, the | 
| 106 | user-interface of Proof~General + X-Symbol and the Isabelle document | |
| 107 |   processor (see \S\ref{sec:document-preparation}) display the
 | |
| 108 |   \verb,\,\verb,<forall>, symbol as~@{text \<forall>}.
 | |
| 12635 | 109 | |
| 110 | A list of standard Isabelle symbols is given in | |
| 12766 | 111 |   \cite[appendix~A]{isabelle-sys}.  You may introduce your own
 | 
| 12635 | 112 | interpretation of further symbols by configuring the appropriate | 
| 12653 | 113 |   front-end tool accordingly, e.g.\ by defining certain {\LaTeX}
 | 
| 114 |   macros (see also \S\ref{sec:doc-prep-symbols}).  There are also a
 | |
| 115 | few predefined control symbols, such as \verb,\,\verb,<^sub>, and | |
| 12635 | 116 | \verb,\,\verb,<^sup>, for sub- and superscript of the subsequent | 
| 12764 | 117 | printable symbol, respectively. For example, \verb,A\<^sup>\<star>, is | 
| 12670 | 118 |   output as @{text "A\<^sup>\<star>"}.
 | 
| 12635 | 119 | |
| 17183 | 120 | A number of symbols are considered letters by the Isabelle lexer and | 
| 121 | can be used as part of identifiers. These are the greek letters | |
| 122 |   @{text "\<alpha>"} (\verb+\+\verb+<alpha>+), @{text "\<beta>"}
 | |
| 123 |   (\verb+\+\verb+<beta>+), etc. (excluding @{text "\<lambda>"}),
 | |
| 124 |   special letters like @{text "\<A>"} (\verb+\+\verb+<A>+) and @{text
 | |
| 125 | "\<AA>"} (\verb+\+\verb+<AA>+), and the control symbols | |
| 126 | \verb+\+\verb+<^isub>+ and \verb+\+\verb+<^isup>+ for single letter | |
| 127 | sub and super scripts. This means that the input | |
| 14486 | 128 | |
| 129 | \medskip | |
| 17183 | 130 |   {\small\noindent \verb,\,\verb,<forall>\,\verb,<alpha>\<^isub>1.,~\verb,\,\verb,<alpha>\<^isub>1 = \,\verb,<Pi>\<^isup>\<A>,}
 | 
| 14486 | 131 | |
| 132 | \medskip | |
| 133 |   \noindent is recognized as the term @{term "\<forall>\<alpha>\<^isub>1. \<alpha>\<^isub>1 = \<Pi>\<^isup>\<A>"} 
 | |
| 17183 | 134 |   by Isabelle. Note that @{text "\<Pi>\<^isup>\<A>"} is a single
 | 
| 135 | syntactic entity, not an exponentiation. | |
| 14486 | 136 | |
| 17183 | 137 |   \medskip Replacing our previous definition of @{text xor} by the
 | 
| 138 | following specifies an Isabelle symbol for the new operator: | |
| 12629 | 139 | *} | 
| 140 | ||
| 141 | (*<*) | |
| 142 | hide const xor | |
| 22098 | 143 | setup {* Theory.add_path "version1" *}
 | 
| 12629 | 144 | (*>*) | 
| 145 | constdefs | |
| 146 | xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "\<oplus>" 60) | |
| 147 | "A \<oplus> B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)" | |
| 12635 | 148 | (*<*) | 
| 149 | local | |
| 150 | (*>*) | |
| 12629 | 151 | |
| 12635 | 152 | text {*
 | 
| 12653 | 153 | \noindent The X-Symbol package within Proof~General provides several | 
| 154 |   input methods to enter @{text \<oplus>} in the text.  If all fails one may
 | |
| 12766 | 155 | just type a named entity \verb,\,\verb,<oplus>, by hand; the | 
| 156 | corresponding symbol will be displayed after further input. | |
| 12635 | 157 | |
| 12766 | 158 | \medskip More flexible is to provide alternative syntax forms | 
| 159 |   through the \bfindex{print mode} concept~\cite{isabelle-ref}.  By
 | |
| 160 | convention, the mode of ``$xsymbols$'' is enabled whenever | |
| 161 |   Proof~General's X-Symbol mode or {\LaTeX} output is active.  Now
 | |
| 162 |   consider the following hybrid declaration of @{text xor}:
 | |
| 12635 | 163 | *} | 
| 164 | ||
| 165 | (*<*) | |
| 166 | hide const xor | |
| 22098 | 167 | setup {* Theory.add_path "version2" *}
 | 
| 12635 | 168 | (*>*) | 
| 169 | constdefs | |
| 170 | xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "[+]\<ignore>" 60) | |
| 171 | "A [+]\<ignore> B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)" | |
| 172 | ||
| 173 | syntax (xsymbols) | |
| 174 | xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "\<oplus>\<ignore>" 60) | |
| 12629 | 175 | (*<*) | 
| 176 | local | |
| 177 | (*>*) | |
| 178 | ||
| 12635 | 179 | text {*
 | 
| 12653 | 180 |   The \commdx{syntax} command introduced here acts like
 | 
| 12743 | 181 |   \isakeyword{consts}, but without declaring a logical constant.  The
 | 
| 12746 | 182 |   print mode specification of \isakeyword{syntax}, here @{text
 | 
| 183 | "(xsymbols)"}, is optional. Also note that its type merely serves | |
| 184 |   for syntactic purposes, and is \emph{not} checked for consistency
 | |
| 185 | with the real constant. | |
| 12635 | 186 | |
| 12672 | 187 |   \medskip We may now write @{text "A [+] B"} or @{text "A \<oplus> B"} in
 | 
| 12766 | 188 | input, while output uses the nicer syntax of $xsymbols$ whenever | 
| 12672 | 189 | that print mode is active. Such an arrangement is particularly | 
| 12766 | 190 | useful for interactive development, where users may type ASCII text | 
| 191 | and see mathematical symbols displayed during proofs. | |
| 12635 | 192 | *} | 
| 193 | ||
| 12629 | 194 | |
| 12648 | 195 | subsection {* Prefix Annotations *}
 | 
| 12629 | 196 | |
| 197 | text {*
 | |
| 12766 | 198 |   Prefix syntax annotations\index{prefix annotation} are another form
 | 
| 199 |   of mixfixes \cite{isabelle-ref}, without any template arguments or
 | |
| 200 | priorities --- just some literal syntax. The following example | |
| 201 | associates common symbols with the constructors of a datatype. | |
| 12629 | 202 | *} | 
| 203 | ||
| 204 | datatype currency = | |
| 205 |     Euro nat    ("\<euro>")
 | |
| 206 |   | Pounds nat  ("\<pounds>")
 | |
| 207 |   | Yen nat     ("\<yen>")
 | |
| 208 |   | Dollar nat  ("$")
 | |
| 209 | ||
| 210 | text {*
 | |
| 12653 | 211 | \noindent Here the mixfix annotations on the rightmost column happen | 
| 212 | to consist of a single Isabelle symbol each: \verb,\,\verb,<euro>,, | |
| 213 | \verb,\,\verb,<pounds>,, \verb,\,\verb,<yen>,, and \verb,$,. Recall | |
| 214 |   that a constructor like @{text Euro} actually is a function @{typ
 | |
| 12746 | 215 |   "nat \<Rightarrow> currency"}.  The expression @{text "Euro 10"} will be
 | 
| 12653 | 216 |   printed as @{term "\<euro> 10"}; only the head of the application is
 | 
| 12743 | 217 | subject to our concrete syntax. This rather simple form already | 
| 218 | achieves conformance with notational standards of the European | |
| 219 | Commission. | |
| 12629 | 220 | |
| 12766 | 221 |   Prefix syntax works the same way for \isakeyword{consts} or
 | 
| 12764 | 222 |   \isakeyword{constdefs}.
 | 
| 12651 | 223 | *} | 
| 224 | ||
| 225 | ||
| 226 | subsection {* Syntax Translations \label{sec:syntax-translations} *}
 | |
| 227 | ||
| 228 | text{*
 | |
| 12766 | 229 | Mixfix syntax annotations merely decorate particular constant | 
| 230 | application forms with concrete syntax, for instance replacing \ | |
| 231 |   @{text "xor A B"} by @{text "A \<oplus> B"}.  Occasionally, the
 | |
| 232 | relationship between some piece of notation and its internal form is | |
| 233 |   more complicated.  Here we need \bfindex{syntax translations}.
 | |
| 12651 | 234 | |
| 12764 | 235 |   Using the \isakeyword{syntax}\index{syntax (command)}, command we
 | 
| 236 | introduce uninterpreted notational elements. Then | |
| 237 |   \commdx{translations} relate input forms to complex logical
 | |
| 12766 | 238 | expressions. This provides a simple mechanism for syntactic macros; | 
| 239 | even heavier transformations may be written in ML | |
| 12670 | 240 |   \cite{isabelle-ref}.
 | 
| 12651 | 241 | |
| 12764 | 242 | \medskip A typical use of syntax translations is to introduce | 
| 12766 | 243 | relational notation for membership in a set of pair, replacing \ | 
| 244 |   @{text "(x, y) \<in> sim"} by @{text "x \<approx> y"}.
 | |
| 12635 | 245 | *} | 
| 246 | ||
| 247 | consts | |
| 12651 | 248 |   sim :: "('a \<times> 'a) set"
 | 
| 12629 | 249 | |
| 12651 | 250 | syntax | 
| 251 | "_sim" :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<approx>" 50) | |
| 252 | translations | |
| 253 | "x \<approx> y" \<rightleftharpoons> "(x, y) \<in> sim" | |
| 12629 | 254 | |
| 12651 | 255 | text {*
 | 
| 256 |   \noindent Here the name of the dummy constant @{text "_sim"} does
 | |
| 12766 | 257 | not matter, as long as it is not used elsewhere. Prefixing an | 
| 258 |   underscore is a common convention.  The \isakeyword{translations}
 | |
| 12651 | 259 | declaration already uses concrete syntax on the left-hand side; | 
| 260 |   internally we relate a raw application @{text "_sim x y"} with
 | |
| 261 |   @{text "(x, y) \<in> sim"}.
 | |
| 262 | ||
| 12653 | 263 | \medskip Another common application of syntax translations is to | 
| 12651 | 264 | provide variant versions of fundamental relational expressions, such | 
| 265 |   as @{text \<noteq>} for negated equalities.  The following declaration
 | |
| 266 | stems from Isabelle/HOL itself: | |
| 12629 | 267 | *} | 
| 268 | ||
| 12651 | 269 | syntax "_not_equal" :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<noteq>\<ignore>" 50) | 
| 270 | translations "x \<noteq>\<ignore> y" \<rightleftharpoons> "\<not> (x = y)" | |
| 12629 | 271 | |
| 12651 | 272 | text {*
 | 
| 273 | \noindent Normally one would introduce derived concepts like this | |
| 12653 | 274 |   within the logic, using \isakeyword{consts} + \isakeyword{defs}
 | 
| 275 |   instead of \isakeyword{syntax} + \isakeyword{translations}.  The
 | |
| 12651 | 276 | present formulation has the virtue that expressions are immediately | 
| 12665 | 277 | replaced by the ``definition'' upon parsing; the effect is reversed | 
| 278 | upon printing. | |
| 12651 | 279 | |
| 12766 | 280 | This sort of translation is appropriate when the defined concept is | 
| 281 | a trivial variation on an existing one. On the other hand, syntax | |
| 282 | translations do not scale up well to large hierarchies of concepts. | |
| 283 | Translations do not replace definitions! | |
| 12629 | 284 | *} | 
| 285 | ||
| 286 | ||
| 12653 | 287 | section {* Document Preparation \label{sec:document-preparation} *}
 | 
| 12629 | 288 | |
| 12645 | 289 | text {*
 | 
| 12653 | 290 |   Isabelle/Isar is centered around the concept of \bfindex{formal
 | 
| 12766 | 291 |   proof documents}\index{documents|bold}.  The outcome of a formal
 | 
| 292 | development effort is meant to be a human-readable record, presented | |
| 293 | as browsable PDF file or printed on paper. The overall document | |
| 294 | structure follows traditional mathematical articles, with sections, | |
| 295 | intermediate explanations, definitions, theorems and proofs. | |
| 12629 | 296 | |
| 12645 | 297 | \medskip The Isabelle document preparation system essentially acts | 
| 12670 | 298 |   as a front-end to {\LaTeX}.  After checking specifications and
 | 
| 299 | proofs formally, the theory sources are turned into typesetting | |
| 12766 | 300 | instructions in a schematic manner. This lets you write authentic | 
| 301 | reports on theory developments with little effort: many technical | |
| 302 | consistency checks are handled by the system. | |
| 12744 | 303 | |
| 304 | Here is an example to illustrate the idea of Isabelle document | |
| 305 | preparation. | |
| 12746 | 306 | *} | 
| 12744 | 307 | |
| 12746 | 308 | text_raw {* \begin{quotation} *}
 | 
| 309 | ||
| 310 | text {*
 | |
| 311 |   The following datatype definition of @{text "'a bintree"} models
 | |
| 312 |   binary trees with nodes being decorated by elements of type @{typ
 | |
| 313 | 'a}. | |
| 12744 | 314 | *} | 
| 315 | ||
| 316 | datatype 'a bintree = | |
| 12746 | 317 | Leaf | Branch 'a "'a bintree" "'a bintree" | 
| 12744 | 318 | |
| 319 | text {*
 | |
| 320 | \noindent The datatype induction rule generated here is of the form | |
| 12746 | 321 |   @{thm [indent = 1, display] bintree.induct [no_vars]}
 | 
| 322 | *} | |
| 12744 | 323 | |
| 12746 | 324 | text_raw {* \end{quotation} *}
 | 
| 325 | ||
| 326 | text {*
 | |
| 12766 | 327 | \noindent The above document output has been produced as follows: | 
| 12744 | 328 | |
| 329 |   \begin{ttbox}
 | |
| 330 |   text {\ttlbrace}*
 | |
| 331 |     The following datatype definition of {\at}{\ttlbrace}text "'a bintree"{\ttrbrace}
 | |
| 332 | models binary trees with nodes being decorated by elements | |
| 333 |     of type {\at}{\ttlbrace}typ 'a{\ttrbrace}.
 | |
| 334 |   *{\ttrbrace}
 | |
| 335 | ||
| 336 | datatype 'a bintree = | |
| 337 | Leaf | Branch 'a "'a bintree" "'a bintree" | |
| 12766 | 338 |   \end{ttbox}
 | 
| 339 |   \begin{ttbox}
 | |
| 12744 | 340 |   text {\ttlbrace}*
 | 
| 341 |     {\ttback}noindent The datatype induction rule generated here is
 | |
| 342 |     of the form {\at}{\ttlbrace}thm [display] bintree.induct [no_vars]{\ttrbrace}
 | |
| 343 |   *{\ttrbrace}
 | |
| 12766 | 344 |   \end{ttbox}\vspace{-\medskipamount}
 | 
| 12744 | 345 | |
| 12746 | 346 | \noindent Here we have augmented the theory by formal comments | 
| 12766 | 347 |   (using \isakeyword{text} blocks), the informal parts may again refer
 | 
| 348 | to formal entities by means of ``antiquotations'' (such as | |
| 12744 | 349 |   \texttt{\at}\verb,{text "'a bintree"}, or
 | 
| 12746 | 350 |   \texttt{\at}\verb,{typ 'a},), see also \S\ref{sec:doc-prep-text}.
 | 
| 12645 | 351 | *} | 
| 352 | ||
| 353 | ||
| 12648 | 354 | subsection {* Isabelle Sessions *}
 | 
| 12629 | 355 | |
| 356 | text {*
 | |
| 12653 | 357 | In contrast to the highly interactive mode of Isabelle/Isar theory | 
| 358 | development, the document preparation stage essentially works in | |
| 12670 | 359 |   batch-mode.  An Isabelle \bfindex{session} consists of a collection
 | 
| 12766 | 360 | of source files that may contribute to an output document. Each | 
| 361 | session is derived from a single parent, usually an object-logic | |
| 362 |   image like \texttt{HOL}.  This results in an overall tree structure,
 | |
| 363 | which is reflected by the output location in the file system | |
| 364 | (usually rooted at \verb,~/isabelle/browser_info,). | |
| 12645 | 365 | |
| 12683 | 366 | \medskip The easiest way to manage Isabelle sessions is via | 
| 12685 | 367 |   \texttt{isatool mkdir} (generates an initial session source setup)
 | 
| 368 |   and \texttt{isatool make} (run sessions controlled by
 | |
| 12683 | 369 |   \texttt{IsaMakefile}).  For example, a new session
 | 
| 370 |   \texttt{MySession} derived from \texttt{HOL} may be produced as
 | |
| 371 | follows: | |
| 372 | ||
| 373 | \begin{verbatim}
 | |
| 374 | isatool mkdir HOL MySession | |
| 375 | isatool make | |
| 376 | \end{verbatim}
 | |
| 377 | ||
| 12685 | 378 |   The \texttt{isatool make} job also informs about the file-system
 | 
| 379 | location of the ultimate results. The above dry run should be able | |
| 380 |   to produce some \texttt{document.pdf} (with dummy title, empty table
 | |
| 12743 | 381 | of contents etc.). Any failure at this stage usually indicates | 
| 17183 | 382 |   technical problems of the {\LaTeX} installation.
 | 
| 12683 | 383 | |
| 384 | \medskip The detailed arrangement of the session sources is as | |
| 12746 | 385 | follows. | 
| 12645 | 386 | |
| 387 |   \begin{itemize}
 | |
| 388 | ||
| 12670 | 389 |   \item Directory \texttt{MySession} holds the required theory files
 | 
| 390 |   $T@1$\texttt{.thy}, \dots, $T@n$\texttt{.thy}.
 | |
| 12645 | 391 | |
| 392 |   \item File \texttt{MySession/ROOT.ML} holds appropriate ML commands
 | |
| 393 | for loading all wanted theories, usually just | |
| 12665 | 394 |   ``\texttt{use_thy"$T@i$";}'' for any $T@i$ in leaf position of the
 | 
| 12670 | 395 | dependency graph. | 
| 12645 | 396 | |
| 397 |   \item Directory \texttt{MySession/document} contains everything
 | |
| 12653 | 398 |   required for the {\LaTeX} stage; only \texttt{root.tex} needs to be
 | 
| 399 | provided initially. | |
| 12645 | 400 | |
| 12653 | 401 |   The latter file holds appropriate {\LaTeX} code to commence a
 | 
| 402 | document (\verb,\documentclass, etc.), and to include the generated | |
| 12743 | 403 |   files $T@i$\texttt{.tex} for each theory.  Isabelle will generate a
 | 
| 404 |   file \texttt{session.tex} holding {\LaTeX} commands to include all
 | |
| 12746 | 405 | generated theory output files in topologically sorted order, so | 
| 406 |   \verb,\input{session}, in the body of \texttt{root.tex} does the job
 | |
| 407 | in most situations. | |
| 12653 | 408 | |
| 12681 | 409 |   \item \texttt{IsaMakefile} holds appropriate dependencies and
 | 
| 410 | invocations of Isabelle tools to control the batch job. In fact, | |
| 12746 | 411 |   several sessions may be managed by the same \texttt{IsaMakefile}.
 | 
| 12764 | 412 |   See the \emph{Isabelle System Manual} \cite{isabelle-sys} 
 | 
| 413 | for further details, especially on | |
| 12653 | 414 |   \texttt{isatool usedir} and \texttt{isatool make}.
 | 
| 12645 | 415 | |
| 416 |   \end{itemize}
 | |
| 417 | ||
| 12685 | 418 |   One may now start to populate the directory \texttt{MySession}, and
 | 
| 12766 | 419 |   the file \texttt{MySession/ROOT.ML} accordingly.  The file
 | 
| 420 |   \texttt{MySession/document/root.tex} should also be adapted at some
 | |
| 12685 | 421 | point; the default version is mostly self-explanatory. Note that | 
| 422 | \verb,\isabellestyle, enables fine-tuning of the general appearance | |
| 423 | of characters and mathematical symbols (see also | |
| 424 |   \S\ref{sec:doc-prep-symbols}).
 | |
| 12653 | 425 | |
| 12685 | 426 |   Especially observe the included {\LaTeX} packages \texttt{isabelle}
 | 
| 427 |   (mandatory), \texttt{isabellesym} (required for mathematical
 | |
| 12743 | 428 |   symbols), and the final \texttt{pdfsetup} (provides sane defaults
 | 
| 12764 | 429 |   for \texttt{hyperref}, including URL markup).  All three are
 | 
| 12743 | 430 | distributed with Isabelle. Further packages may be required in | 
| 12764 | 431 | particular applications, say for unusual mathematical symbols. | 
| 12645 | 432 | |
| 12746 | 433 |   \medskip Any additional files for the {\LaTeX} stage go into the
 | 
| 434 |   \texttt{MySession/document} directory as well.  In particular,
 | |
| 12766 | 435 |   adding a file named \texttt{root.bib} causes an automatic run of
 | 
| 436 |   \texttt{bibtex} to process a bibliographic database; see also
 | |
| 437 |   \texttt{isatool document} \cite{isabelle-sys}.
 | |
| 12645 | 438 | |
| 12653 | 439 | \medskip Any failure of the document preparation phase in an | 
| 12670 | 440 | Isabelle batch session leaves the generated sources in their target | 
| 12766 | 441 | location, identified by the accompanying error message. This lets | 
| 442 |   you trace {\LaTeX} problems with the generated files at hand.
 | |
| 12645 | 443 | *} | 
| 444 | ||
| 445 | ||
| 12648 | 446 | subsection {* Structure Markup *}
 | 
| 12645 | 447 | |
| 12653 | 448 | text {*
 | 
| 449 | The large-scale structure of Isabelle documents follows existing | |
| 450 |   {\LaTeX} conventions, with chapters, sections, subsubsections etc.
 | |
| 451 |   The Isar language includes separate \bfindex{markup commands}, which
 | |
| 12681 | 452 | do not affect the formal meaning of a theory (or proof), but result | 
| 12665 | 453 |   in corresponding {\LaTeX} elements.
 | 
| 12645 | 454 | |
| 12665 | 455 | There are separate markup commands depending on the textual context: | 
| 456 |   in header position (just before \isakeyword{theory}), within the
 | |
| 457 | theory body, or within a proof. The header needs to be treated | |
| 458 | specially here, since ordinary theory and proof commands may only | |
| 459 |   occur \emph{after} the initial \isakeyword{theory} specification.
 | |
| 12645 | 460 | |
| 12665 | 461 | \medskip | 
| 12645 | 462 | |
| 463 |   \begin{tabular}{llll}
 | |
| 464 | header & theory & proof & default meaning \\\hline | |
| 465 |     & \commdx{chapter} & & \verb,\chapter, \\
 | |
| 466 |   \commdx{header} & \commdx{section} & \commdx{sect} & \verb,\section, \\
 | |
| 467 |     & \commdx{subsection} & \commdx{subsect} & \verb,\subsection, \\
 | |
| 468 |     & \commdx{subsubsection} & \commdx{subsubsect} & \verb,\subsubsection, \\
 | |
| 469 |   \end{tabular}
 | |
| 470 | ||
| 471 | \medskip | |
| 472 | ||
| 473 | From the Isabelle perspective, each markup command takes a single | |
| 12746 | 474 |   $text$ argument (delimited by \verb,",~@{text \<dots>}~\verb,", or
 | 
| 475 |   \verb,{,\verb,*,~@{text \<dots>}~\verb,*,\verb,},).  After stripping any
 | |
| 12645 | 476 |   surrounding white space, the argument is passed to a {\LaTeX} macro
 | 
| 12766 | 477 |   \verb,\isamarkupXYZ, for command \isakeyword{XYZ}.  These macros are
 | 
| 478 | defined in \verb,isabelle.sty, according to the meaning given in the | |
| 479 | rightmost column above. | |
| 12645 | 480 | |
| 481 | \medskip The following source fragment illustrates structure markup | |
| 12653 | 482 |   of a theory.  Note that {\LaTeX} labels may be included inside of
 | 
| 483 | section headings as well. | |
| 12645 | 484 | |
| 485 |   \begin{ttbox}
 | |
| 486 |   header {\ttlbrace}* Some properties of Foo Bar elements *{\ttrbrace}
 | |
| 487 | ||
| 15136 | 488 | theory Foo_Bar | 
| 15141 | 489 | imports Main | 
| 15136 | 490 | begin | 
| 12645 | 491 | |
| 492 |   subsection {\ttlbrace}* Basic definitions *{\ttrbrace}
 | |
| 493 | ||
| 494 | consts | |
| 495 | foo :: \dots | |
| 496 | bar :: \dots | |
| 12648 | 497 | |
| 12645 | 498 | defs \dots | 
| 12648 | 499 | |
| 12645 | 500 |   subsection {\ttlbrace}* Derived rules *{\ttrbrace}
 | 
| 501 | ||
| 502 | lemma fooI: \dots | |
| 503 | lemma fooE: \dots | |
| 504 | ||
| 12648 | 505 |   subsection {\ttlbrace}* Main theorem {\ttback}label{\ttlbrace}sec:main-theorem{\ttrbrace} *{\ttrbrace}
 | 
| 12645 | 506 | |
| 507 | theorem main: \dots | |
| 508 | ||
| 509 | end | |
| 12766 | 510 |   \end{ttbox}\vspace{-\medskipamount}
 | 
| 12645 | 511 | |
| 12766 | 512 | You may occasionally want to change the meaning of markup commands, | 
| 513 |   say via \verb,\renewcommand, in \texttt{root.tex}.  For example,
 | |
| 514 | \verb,\isamarkupheader, is a good candidate for some tuning. We | |
| 515 | could move it up in the hierarchy to become \verb,\chapter,. | |
| 12645 | 516 | |
| 517 | \begin{verbatim}
 | |
| 518 |   \renewcommand{\isamarkupheader}[1]{\chapter{#1}}
 | |
| 519 | \end{verbatim}
 | |
| 520 | ||
| 12766 | 521 | \noindent Now we must change the document class given in | 
| 522 |   \texttt{root.tex} to something that supports chapters.  A suitable
 | |
| 523 |   command is \verb,\documentclass{report},.
 | |
| 12645 | 524 | |
| 12648 | 525 |   \medskip The {\LaTeX} macro \verb,\isabellecontext, is maintained to
 | 
| 526 | hold the name of the current theory context. This is particularly | |
| 12653 | 527 | useful for document headings: | 
| 12645 | 528 | |
| 529 | \begin{verbatim}
 | |
| 12653 | 530 |   \renewcommand{\isamarkupheader}[1]
 | 
| 12645 | 531 |   {\chapter{#1}\markright{THEORY~\isabellecontext}}
 | 
| 532 | \end{verbatim}
 | |
| 533 | ||
| 534 | \noindent Make sure to include something like | |
| 12648 | 535 |   \verb,\pagestyle{headings}, in \texttt{root.tex}; the document
 | 
| 12764 | 536 | should have more than two pages to show the effect. | 
| 12645 | 537 | *} | 
| 538 | ||
| 539 | ||
| 12744 | 540 | subsection {* Formal Comments and Antiquotations \label{sec:doc-prep-text} *}
 | 
| 12645 | 541 | |
| 542 | text {*
 | |
| 12744 | 543 |   Isabelle \bfindex{source comments}, which are of the form
 | 
| 12746 | 544 |   \verb,(,\verb,*,~@{text \<dots>}~\verb,*,\verb,),, essentially act like
 | 
| 545 | white space and do not really contribute to the content. They | |
| 546 | mainly serve technical purposes to mark certain oddities in the raw | |
| 547 |   input text.  In contrast, \bfindex{formal comments} are portions of
 | |
| 548 | text that are associated with formal Isabelle/Isar commands | |
| 12681 | 549 |   (\bfindex{marginal comments}), or as standalone paragraphs within a
 | 
| 12665 | 550 |   theory or proof context (\bfindex{text blocks}).
 | 
| 12659 | 551 | |
| 552 | \medskip Marginal comments are part of each command's concrete | |
| 12670 | 553 |   syntax \cite{isabelle-ref}; the common form is ``\verb,--,~$text$''
 | 
| 12746 | 554 |   where $text$ is delimited by \verb,",@{text \<dots>}\verb,", or
 | 
| 555 |   \verb,{,\verb,*,~@{text \<dots>}~\verb,*,\verb,}, as before.  Multiple
 | |
| 12670 | 556 | marginal comments may be given at the same time. Here is a simple | 
| 557 | example: | |
| 12665 | 558 | *} | 
| 559 | ||
| 560 | lemma "A --> A" | |
| 561 | -- "a triviality of propositional logic" | |
| 562 | -- "(should not really bother)" | |
| 563 | by (rule impI) -- "implicit assumption step involved here" | |
| 564 | ||
| 565 | text {*
 | |
| 566 | \noindent The above output has been produced as follows: | |
| 12659 | 567 | |
| 568 | \begin{verbatim}
 | |
| 569 | lemma "A --> A" | |
| 570 | -- "a triviality of propositional logic" | |
| 571 | -- "(should not really bother)" | |
| 572 | by (rule impI) -- "implicit assumption step involved here" | |
| 573 | \end{verbatim}
 | |
| 574 | ||
| 12670 | 575 |   From the {\LaTeX} viewpoint, ``\verb,--,'' acts like a markup
 | 
| 576 | command, associated with the macro \verb,\isamarkupcmt, (taking a | |
| 577 | single argument). | |
| 12659 | 578 | |
| 12665 | 579 |   \medskip Text blocks are introduced by the commands \bfindex{text}
 | 
| 580 |   and \bfindex{txt}, for theory and proof contexts, respectively.
 | |
| 581 | Each takes again a single $text$ argument, which is interpreted as a | |
| 582 |   free-form paragraph in {\LaTeX} (surrounded by some additional
 | |
| 12670 | 583 | vertical space). This behavior may be changed by redefining the | 
| 584 |   {\LaTeX} environments of \verb,isamarkuptext, or
 | |
| 585 | \verb,isamarkuptxt,, respectively (via \verb,\renewenvironment,) The | |
| 586 | text style of the body is determined by \verb,\isastyletext, and | |
| 587 | \verb,\isastyletxt,; the default setup uses a smaller font within | |
| 12746 | 588 | proofs. This may be changed as follows: | 
| 589 | ||
| 590 | \begin{verbatim}
 | |
| 591 |   \renewcommand{\isastyletxt}{\isastyletext}
 | |
| 592 | \end{verbatim}
 | |
| 12659 | 593 | |
| 12766 | 594 | \medskip The $text$ part of Isabelle markup commands essentially | 
| 595 |   inserts \emph{quoted material} into a formal text, mainly for
 | |
| 596 |   instruction of the reader.  An \bfindex{antiquotation} is again a
 | |
| 597 | formal object embedded into such an informal portion. The | |
| 598 | interpretation of antiquotations is limited to some well-formedness | |
| 599 | checks, with the result being pretty printed to the resulting | |
| 600 | document. Quoted text blocks together with antiquotations provide | |
| 601 | an attractive means of referring to formal entities, with good | |
| 602 | confidence in getting the technical details right (especially syntax | |
| 603 | and types). | |
| 12659 | 604 | |
| 12665 | 605 | The general syntax of antiquotations is as follows: | 
| 12659 | 606 |   \texttt{{\at}{\ttlbrace}$name$ $arguments${\ttrbrace}}, or
 | 
| 607 |   \texttt{{\at}{\ttlbrace}$name$ [$options$] $arguments${\ttrbrace}}
 | |
| 12665 | 608 | for a comma-separated list of options consisting of a $name$ or | 
| 12766 | 609 |   \texttt{$name$=$value$} each.  The syntax of $arguments$ depends on
 | 
| 610 | the kind of antiquotation, it generally follows the same conventions | |
| 611 | for types, terms, or theorems as in the formal part of a theory. | |
| 12645 | 612 | |
| 12766 | 613 | \medskip This sentence demonstrates quotations and antiquotations: | 
| 614 |   @{term "%x y. x"} is a well-typed term.
 | |
| 12659 | 615 | |
| 12764 | 616 | \medskip\noindent The output above was produced as follows: | 
| 12659 | 617 |   \begin{ttbox}
 | 
| 618 | text {\ttlbrace}*
 | |
| 12764 | 619 | This sentence demonstrates quotations and antiquotations: | 
| 12659 | 620 |   {\at}{\ttlbrace}term "%x y. x"{\ttrbrace} is a well-typed term.
 | 
| 621 | *{\ttrbrace}
 | |
| 12766 | 622 |   \end{ttbox}\vspace{-\medskipamount}
 | 
| 12659 | 623 | |
| 12764 | 624 | The notational change from the ASCII character~\verb,%, to the | 
| 12766 | 625 |   symbol~@{text \<lambda>} reveals that Isabelle printed this term, after
 | 
| 626 | parsing and type-checking. Document preparation enables symbolic | |
| 627 | output by default. | |
| 12659 | 628 | |
| 16523 | 629 | \medskip The next example includes an option to show the type of all | 
| 630 | variables. The antiquotation | |
| 12766 | 631 |   \texttt{{\at}}\verb,{term [show_types] "%x y. x"}, produces the
 | 
| 632 |   output @{term [show_types] "%x y. x"}.  Type inference has figured
 | |
| 633 | out the most general typings in the present theory context. Terms | |
| 634 | may acquire different typings due to constraints imposed by their | |
| 635 | environment; within a proof, for example, variables are given the | |
| 636 | same types as they have in the main goal statement. | |
| 12659 | 637 | |
| 12764 | 638 | \medskip Several further kinds of antiquotations and options are | 
| 12665 | 639 |   available \cite{isabelle-sys}.  Here are a few commonly used
 | 
| 12670 | 640 | combinations: | 
| 12659 | 641 | |
| 642 | \medskip | |
| 12651 | 643 | |
| 12659 | 644 |   \begin{tabular}{ll}
 | 
| 645 |   \texttt{\at}\verb,{typ,~$\tau$\verb,}, & print type $\tau$ \\
 | |
| 646 |   \texttt{\at}\verb,{term,~$t$\verb,}, & print term $t$ \\
 | |
| 647 |   \texttt{\at}\verb,{prop,~$\phi$\verb,}, & print proposition $\phi$ \\
 | |
| 12665 | 648 |   \texttt{\at}\verb,{prop [display],~$\phi$\verb,}, & print large proposition $\phi$ (with linebreaks) \\
 | 
| 12659 | 649 |   \texttt{\at}\verb,{prop [source],~$\phi$\verb,}, & check proposition $\phi$, print its input \\
 | 
| 650 |   \texttt{\at}\verb,{thm,~$a$\verb,}, & print fact $a$ \\
 | |
| 651 |   \texttt{\at}\verb,{thm,~$a$~\verb,[no_vars]}, & print fact $a$, fixing schematic variables \\
 | |
| 12746 | 652 |   \texttt{\at}\verb,{thm [source],~$a$\verb,}, & check availability of fact $a$, print its name \\
 | 
| 12659 | 653 |   \texttt{\at}\verb,{text,~$s$\verb,}, & print uninterpreted text $s$ \\
 | 
| 654 |   \end{tabular}
 | |
| 655 | ||
| 656 | \medskip | |
| 657 | ||
| 12665 | 658 |   Note that \attrdx{no_vars} given above is \emph{not} an
 | 
| 659 | antiquotation option, but an attribute of the theorem argument given | |
| 660 | here. This might be useful with a diagnostic command like | |
| 661 |   \isakeyword{thm}, too.
 | |
| 12659 | 662 | |
| 12665 | 663 |   \medskip The \texttt{\at}\verb,{text, $s$\verb,}, antiquotation is
 | 
| 12659 | 664 | particularly interesting. Embedding uninterpreted text within an | 
| 12665 | 665 | informal body might appear useless at first sight. Here the key | 
| 666 | virtue is that the string $s$ is processed as Isabelle output, | |
| 667 | interpreting Isabelle symbols appropriately. | |
| 12659 | 668 | |
| 12665 | 669 |   For example, \texttt{\at}\verb,{text "\<forall>\<exists>"}, produces @{text
 | 
| 670 | "\<forall>\<exists>"}, according to the standard interpretation of these symbol | |
| 671 |   (cf.\ \S\ref{sec:doc-prep-symbols}).  Thus we achieve consistent
 | |
| 12659 | 672 | mathematical notation in both the formal and informal parts of the | 
| 12766 | 673 | document very easily, independently of the term language of | 
| 674 |   Isabelle.  Manual {\LaTeX} code would leave more control over the
 | |
| 675 | typesetting, but is also slightly more tedious. | |
| 12645 | 676 | *} | 
| 677 | ||
| 678 | ||
| 12674 | 679 | subsection {* Interpretation of Symbols \label{sec:doc-prep-symbols} *}
 | 
| 12645 | 680 | |
| 681 | text {*
 | |
| 12665 | 682 |   As has been pointed out before (\S\ref{sec:syntax-symbols}),
 | 
| 12670 | 683 | Isabelle symbols are the smallest syntactic entities --- a | 
| 12681 | 684 | straightforward generalization of ASCII characters. While Isabelle | 
| 12665 | 685 | does not impose any interpretation of the infinite collection of | 
| 12764 | 686 |   named symbols, {\LaTeX} documents use canonical glyphs for certain
 | 
| 12670 | 687 |   standard symbols \cite[appendix~A]{isabelle-sys}.
 | 
| 12659 | 688 | |
| 12766 | 689 |   The {\LaTeX} code produced from Isabelle text follows a simple
 | 
| 690 | scheme. You can tune the final appearance by redefining certain | |
| 691 |   macros, say in \texttt{root.tex} of the document.
 | |
| 12670 | 692 | |
| 693 |   \begin{enumerate}
 | |
| 12659 | 694 | |
| 12670 | 695 |   \item 7-bit ASCII characters: letters \texttt{A\dots Z} and
 | 
| 12746 | 696 |   \texttt{a\dots z} are output directly, digits are passed as an
 | 
| 12670 | 697 | argument to the \verb,\isadigit, macro, other characters are | 
| 698 | replaced by specifically named macros of the form | |
| 12665 | 699 | \verb,\isacharXYZ,. | 
| 12659 | 700 | |
| 12766 | 701 | \item Named symbols: \verb,\,\verb,<XYZ>, is turned into | 
| 702 |   \verb,{\isasymXYZ},; note the additional braces.
 | |
| 12659 | 703 | |
| 12766 | 704 | \item Named control symbols: \verb,\,\verb,<^XYZ>, is turned into | 
| 705 | \verb,\isactrlXYZ,; subsequent symbols may act as arguments if the | |
| 706 | control macro is defined accordingly. | |
| 12670 | 707 | |
| 12659 | 708 |   \end{enumerate}
 | 
| 12665 | 709 | |
| 12764 | 710 |   You may occasionally wish to give new {\LaTeX} interpretations of
 | 
| 711 | named symbols. This merely requires an appropriate definition of | |
| 12766 | 712 | \verb,\isasymXYZ,, for \verb,\,\verb,<XYZ>, (see | 
| 12746 | 713 |   \texttt{isabelle.sty} for working examples).  Control symbols are
 | 
| 714 | slightly more difficult to get right, though. | |
| 12665 | 715 | |
| 716 | \medskip The \verb,\isabellestyle, macro provides a high-level | |
| 717 | interface to tune the general appearance of individual symbols. For | |
| 12670 | 718 |   example, \verb,\isabellestyle{it}, uses the italics text style to
 | 
| 719 |   mimic the general appearance of the {\LaTeX} math mode; double
 | |
| 12743 | 720 | quotes are not printed at all. The resulting quality of typesetting | 
| 721 | is quite good, so this should be the default style for work that | |
| 722 | gets distributed to a broader audience. | |
| 12645 | 723 | *} | 
| 724 | ||
| 725 | ||
| 12653 | 726 | subsection {* Suppressing Output \label{sec:doc-prep-suppress} *}
 | 
| 12645 | 727 | |
| 728 | text {*
 | |
| 12748 | 729 |   By default, Isabelle's document system generates a {\LaTeX} file for
 | 
| 730 | each theory that gets loaded while running the session. The | |
| 731 |   generated \texttt{session.tex} will include all of these in order of
 | |
| 732 | appearance, which in turn gets included by the standard | |
| 12743 | 733 |   \texttt{root.tex}.  Certainly one may change the order or suppress
 | 
| 12746 | 734 |   unwanted theories by ignoring \texttt{session.tex} and load
 | 
| 735 |   individual files directly in \texttt{root.tex}.  On the other hand,
 | |
| 736 | such an arrangement requires additional maintenance whenever the | |
| 737 | collection of theories changes. | |
| 12648 | 738 | |
| 739 | Alternatively, one may tune the theory loading process in | |
| 12653 | 740 |   \texttt{ROOT.ML} itself: traversal of the theory dependency graph
 | 
| 12670 | 741 | may be fine-tuned by adding \verb,use_thy, invocations, although | 
| 742 | topological sorting still has to be observed. Moreover, the ML | |
| 743 | operator \verb,no_document, temporarily disables document generation | |
| 12766 | 744 | while executing a theory loader command. Its usage is like this: | 
| 12648 | 745 | |
| 746 | \begin{verbatim}
 | |
| 12665 | 747 | no_document use_thy "T"; | 
| 12648 | 748 | \end{verbatim}
 | 
| 12645 | 749 | |
| 17183 | 750 | \medskip Theory output may be suppressed more selectively, either | 
| 751 |   via \bfindex{tagged command regions} or \bfindex{ignored material}.
 | |
| 12648 | 752 | |
| 17183 | 753 | Tagged command regions works by annotating commands with named tags, | 
| 754 |   which correspond to certain {\LaTeX} markup that tells how to treat
 | |
| 755 | particular parts of a document when doing the actual type-setting. | |
| 756 | By default, certain Isabelle/Isar commands are implicitly marked up | |
| 757 |   using the predefined tags ``\emph{theory}'' (for theory begin and
 | |
| 758 |   end), ``\emph{proof}'' (for proof commands), and ``\emph{ML}'' (for
 | |
| 759 | commands involving ML code). Users may add their own tags using the | |
| 760 |   \verb,%,\emph{tag} notation right after a command name.  In the
 | |
| 761 | subsequent example we hide a particularly irrelevant proof: | |
| 762 | *} | |
| 12648 | 763 | |
| 17183 | 764 | lemma "x = x" by %invisible (simp) | 
| 12648 | 765 | |
| 17183 | 766 | text {*
 | 
| 767 | The original source has been ``\verb,lemma "x = x" by %invisible (simp),''. | |
| 768 | Tags observe the structure of proofs; adjacent commands with the | |
| 769 | same tag are joined into a single region. The Isabelle document | |
| 770 | preparation system allows the user to specify how to interpret a | |
| 771 | tagged region, in order to keep, drop, or fold the corresponding | |
| 772 |   parts of the document.  See the \emph{Isabelle System Manual}
 | |
| 773 |   \cite{isabelle-sys} for further details, especially on
 | |
| 774 |   \texttt{isatool usedir} and \texttt{isatool document}.
 | |
| 12648 | 775 | |
| 17183 | 776 | Ignored material is specified by delimiting the original formal | 
| 777 | source with special source comments | |
| 778 | \verb,(,\verb,*,\verb,<,\verb,*,\verb,), and | |
| 779 | \verb,(,\verb,*,\verb,>,\verb,*,\verb,),. These parts are stripped | |
| 780 | before the type-setting phase, without affecting the formal checking | |
| 781 | of the theory, of course. For example, we may hide parts of a proof | |
| 782 | that seem unfit for general public inspection. The following | |
| 783 | ``fully automatic'' proof is actually a fake: | |
| 12651 | 784 | *} | 
| 785 | ||
| 786 | lemma "x \<noteq> (0::int) \<Longrightarrow> 0 < x * x" | |
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
13439diff
changeset | 787 | by (auto(*<*)simp add: zero_less_mult_iff(*>*)) | 
| 12651 | 788 | |
| 789 | text {*
 | |
| 17183 | 790 | \noindent The real source of the proof has been as follows: | 
| 12651 | 791 | |
| 792 | \begin{verbatim}
 | |
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
13439diff
changeset | 793 | by (auto(*<*)simp add: zero_less_mult_iff(*>*)) | 
| 12659 | 794 | \end{verbatim}
 | 
| 795 | %(* | |
| 12651 | 796 | |
| 12766 | 797 | \medskip Suppressing portions of printed text demands care. You | 
| 798 | should not misrepresent the underlying theory development. It is | |
| 799 | easy to invalidate the visible text by hiding references to | |
| 17183 | 800 | questionable axioms, for example. | 
| 12629 | 801 | *} | 
| 802 | ||
| 11647 | 803 | (*<*) | 
| 804 | end | |
| 805 | (*>*) |