doc-src/TutorialI/Documents/Documents.thy
author wenzelm
Mon, 07 Jan 2002 23:57:14 +0100
changeset 12659 2aa05eb15bd2
parent 12653 a55c066624eb
child 12665 61e53dbac238
permissions -rw-r--r--
getting close to completion;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11647
0538cb0f7999 initial setup for chapter on document preparation;
wenzelm
parents:
diff changeset
     1
(*<*)
0538cb0f7999 initial setup for chapter on document preparation;
wenzelm
parents:
diff changeset
     2
theory Documents = Main:
0538cb0f7999 initial setup for chapter on document preparation;
wenzelm
parents:
diff changeset
     3
(*>*)
0538cb0f7999 initial setup for chapter on document preparation;
wenzelm
parents:
diff changeset
     4
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
     5
section {* Concrete Syntax \label{sec:concrete-syntax} *}
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
     6
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
     7
text {*
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
     8
  Concerning Isabelle's ``inner'' language of simply-typed @{text
12653
wenzelm
parents: 12651
diff changeset
     9
  \<lambda>}-calculus, the core concept of Isabelle's elaborate
wenzelm
parents: 12651
diff changeset
    10
  infrastructure for concrete syntax is that of general
wenzelm
parents: 12651
diff changeset
    11
  \bfindex{mixfix annotations}.  Associated with any kind of constant
wenzelm
parents: 12651
diff changeset
    12
  declaration, mixfixes affect both the grammar productions for the
wenzelm
parents: 12651
diff changeset
    13
  parser and output templates for the pretty printer.
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    14
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    15
  In full generality, the whole affair of parser and pretty printer
12653
wenzelm
parents: 12651
diff changeset
    16
  configuration is rather subtle, see \cite{isabelle-ref} for further
wenzelm
parents: 12651
diff changeset
    17
  details.  Any syntax specifications given by end-users need to
wenzelm
parents: 12651
diff changeset
    18
  interact properly with the existing setup of Isabelle/Pure and
wenzelm
parents: 12651
diff changeset
    19
  Isabelle/HOL.  It is particularly important to get the precedence of
wenzelm
parents: 12651
diff changeset
    20
  new syntactic constructs right, avoiding ambiguities with existing
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    21
  elements.
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    22
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    23
  \medskip Subsequently we introduce a few simple declaration forms
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    24
  that already cover the most common situations fairly well.
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    25
*}
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    26
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    27
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
    28
subsection {* Infix Annotations *}
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    29
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    30
text {*
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    31
  Syntax annotations may be included wherever constants are declared
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    32
  directly or indirectly, including \isacommand{consts},
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    33
  \isacommand{constdefs}, or \isacommand{datatype} (for the
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    34
  constructor operations).  Type-constructors may be annotated as
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    35
  well, although this is less frequently encountered in practice
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    36
  (@{text "*"} and @{text "+"} types may come to mind).
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    37
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
    38
  Infix declarations\index{infix annotations} provide a useful special
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
    39
  case of mixfixes, where users need not care about the full details
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
    40
  of priorities, nesting, spacing, etc.  The subsequent example of the
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
    41
  exclusive-or operation on boolean values illustrates typical infix
12653
wenzelm
parents: 12651
diff changeset
    42
  declarations arising in practice.
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    43
*}
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    44
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    45
constdefs
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    46
  xor :: "bool \<Rightarrow> bool \<Rightarrow> bool"    (infixl "[+]" 60)
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    47
  "A [+] B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)"
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    48
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    49
text {*
12653
wenzelm
parents: 12651
diff changeset
    50
  \noindent Now @{text "xor A B"} and @{text "A [+] B"} refer to the
wenzelm
parents: 12651
diff changeset
    51
  same expression internally.  Any curried function with at least two
wenzelm
parents: 12651
diff changeset
    52
  arguments may be associated with infix syntax.  For partial
wenzelm
parents: 12651
diff changeset
    53
  applications with less than two operands there is a special notation
wenzelm
parents: 12651
diff changeset
    54
  with \isa{op} prefix: @{text xor} without arguments is represented
wenzelm
parents: 12651
diff changeset
    55
  as @{text "op [+]"}; together with plain prefix application this
wenzelm
parents: 12651
diff changeset
    56
  turns @{text "xor A"} into @{text "op [+] A"}.
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    57
12653
wenzelm
parents: 12651
diff changeset
    58
  \medskip The string @{text [source] "[+]"} in the above annotation
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    59
  refers to the bit of concrete syntax to represent the operator,
12653
wenzelm
parents: 12651
diff changeset
    60
  while the number @{text 60} determines the precedence of the
wenzelm
parents: 12651
diff changeset
    61
  construct (i.e.\ the syntactic priorities of the arguments and
wenzelm
parents: 12651
diff changeset
    62
  result).
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    63
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    64
  As it happens, Isabelle/HOL already spends many popular combinations
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    65
  of ASCII symbols for its own use, including both @{text "+"} and
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    66
  @{text "++"}.  Slightly more awkward combinations like the present
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    67
  @{text "[+]"} tend to be available for user extensions.  The current
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    68
  arrangement of inner syntax may be inspected via
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    69
  \commdx{print\protect\_syntax}, albeit its output is enormous.
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    70
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    71
  Operator precedence also needs some special considerations.  The
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    72
  admissible range is 0--1000.  Very low or high priorities are
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    73
  basically reserved for the meta-logic.  Syntax of Isabelle/HOL
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    74
  mainly uses the range of 10--100: the equality infix @{text "="} is
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    75
  centered at 50, logical connectives (like @{text "\<or>"} and @{text
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    76
  "\<and>"}) are below 50, and algebraic ones (like @{text "+"} and @{text
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    77
  "*"}) above 50.  User syntax should strive to coexist with common
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    78
  HOL forms, or use the mostly unused range 100--900.
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    79
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    80
  \medskip The keyword \isakeyword{infixl} specifies an operator that
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    81
  is nested to the \emph{left}: in iterated applications the more
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    82
  complex expression appears on the left-hand side: @{term "A [+] B
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    83
  [+] C"} stands for @{text "(A [+] B) [+] C"}.  Similarly,
12653
wenzelm
parents: 12651
diff changeset
    84
  \isakeyword{infixr} specifies to nesting to the \emph{right},
wenzelm
parents: 12651
diff changeset
    85
  reading @{term "A [+] B [+] C"} as @{text "A [+] (B [+] C)"}.  In
wenzelm
parents: 12651
diff changeset
    86
  contrast, a \emph{non-oriented} declaration via \isakeyword{infix}
wenzelm
parents: 12651
diff changeset
    87
  would have rendered @{term "A [+] B [+] C"} illegal, but demand
wenzelm
parents: 12651
diff changeset
    88
  explicit parentheses about the intended grouping.
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    89
*}
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    90
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
    91
12659
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
    92
subsection {* Mathematical Symbols \label{sec:syntax-symbols} *}
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    93
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
    94
text {*
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
    95
  Concrete syntax based on plain ASCII characters has its inherent
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
    96
  limitations.  Rich mathematical notation demands a larger repertoire
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
    97
  of symbols.  Several standards of extended character sets have been
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
    98
  proposed over decades, but none has become universally available so
12653
wenzelm
parents: 12651
diff changeset
    99
  far, not even Unicode\index{Unicode}.  Isabelle supports a generic
wenzelm
parents: 12651
diff changeset
   100
  notion of \bfindex{symbols} as the smallest entities of source text,
wenzelm
parents: 12651
diff changeset
   101
  without referring to internal encodings.
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   102
12653
wenzelm
parents: 12651
diff changeset
   103
  There are three kinds of such ``generalized characters'' available:
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   104
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   105
  \begin{enumerate}
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   106
12653
wenzelm
parents: 12651
diff changeset
   107
  \item 7-bit ASCII characters
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   108
12653
wenzelm
parents: 12651
diff changeset
   109
  \item named symbols: \verb,\,\verb,<,$ident$\verb,>,
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   110
12653
wenzelm
parents: 12651
diff changeset
   111
  \item named control symbols: \verb,\,\verb,<^,$ident$\verb,>,
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   112
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   113
  \end{enumerate}
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   114
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   115
  Here $ident$ may be any identifier according to the usual Isabelle
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   116
  conventions.  This results in an infinite store of symbols, whose
12653
wenzelm
parents: 12651
diff changeset
   117
  interpretation is left to further front-end tools.  For example,
wenzelm
parents: 12651
diff changeset
   118
  both by the user-interface of Proof~General + X-Symbol and the
wenzelm
parents: 12651
diff changeset
   119
  Isabelle document processor (see \S\ref{sec:document-preparation})
wenzelm
parents: 12651
diff changeset
   120
  display the \verb,\,\verb,<forall>, symbol really as ``$\forall$''.
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   121
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   122
  A list of standard Isabelle symbols is given in
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   123
  \cite[appendix~A]{isabelle-sys}.  Users may introduce their own
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   124
  interpretation of further symbols by configuring the appropriate
12653
wenzelm
parents: 12651
diff changeset
   125
  front-end tool accordingly, e.g.\ by defining certain {\LaTeX}
wenzelm
parents: 12651
diff changeset
   126
  macros (see also \S\ref{sec:doc-prep-symbols}).  There are also a
wenzelm
parents: 12651
diff changeset
   127
  few predefined control symbols, such as \verb,\,\verb,<^sub>, and
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   128
  \verb,\,\verb,<^sup>, for sub- and superscript of the subsequent
12653
wenzelm
parents: 12651
diff changeset
   129
  (printable) symbol, respectively.  For example, \verb,A\<^sup>\<star>, is
wenzelm
parents: 12651
diff changeset
   130
  shown as ``@{text "A\<^sup>\<star>"}''.
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   131
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   132
  \medskip The following version of our @{text xor} definition uses a
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   133
  standard Isabelle symbol to achieve typographically pleasing output.
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   134
*}
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   135
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   136
(*<*)
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   137
hide const xor
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   138
ML_setup {* Context.>> (Theory.add_path "1") *}
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   139
(*>*)
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   140
constdefs
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   141
  xor :: "bool \<Rightarrow> bool \<Rightarrow> bool"    (infixl "\<oplus>" 60)
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   142
  "A \<oplus> B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)"
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   143
(*<*)
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   144
local
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   145
(*>*)
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   146
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   147
text {*
12653
wenzelm
parents: 12651
diff changeset
   148
  \noindent The X-Symbol package within Proof~General provides several
wenzelm
parents: 12651
diff changeset
   149
  input methods to enter @{text \<oplus>} in the text.  If all fails one may
wenzelm
parents: 12651
diff changeset
   150
  just type \verb,\,\verb,<oplus>, by hand; the display will be
wenzelm
parents: 12651
diff changeset
   151
  adapted immediately after continuing input.
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   152
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   153
  \medskip A slightly more refined scheme is to provide alternative
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   154
  syntax via the \bfindex{print mode} concept of Isabelle (see also
12653
wenzelm
parents: 12651
diff changeset
   155
  \cite{isabelle-ref}).  By convention, the mode of ``$xsymbols$'' is
wenzelm
parents: 12651
diff changeset
   156
  enabled whenever Proof~General's X-Symbol mode (or {\LaTeX} output)
wenzelm
parents: 12651
diff changeset
   157
  is active.  Consider the following hybrid declaration of @{text
wenzelm
parents: 12651
diff changeset
   158
  xor}.
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   159
*}
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   160
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   161
(*<*)
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   162
hide const xor
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   163
ML_setup {* Context.>> (Theory.add_path "2") *}
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   164
(*>*)
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   165
constdefs
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   166
  xor :: "bool \<Rightarrow> bool \<Rightarrow> bool"    (infixl "[+]\<ignore>" 60)
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   167
  "A [+]\<ignore> B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)"
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   168
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   169
syntax (xsymbols)
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   170
  xor :: "bool \<Rightarrow> bool \<Rightarrow> bool"    (infixl "\<oplus>\<ignore>" 60)
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   171
(*<*)
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   172
local
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   173
(*>*)
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   174
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   175
text {*
12653
wenzelm
parents: 12651
diff changeset
   176
  The \commdx{syntax} command introduced here acts like
wenzelm
parents: 12651
diff changeset
   177
  \isakeyword{consts}, but without declaring a logical constant; an
wenzelm
parents: 12651
diff changeset
   178
  optional print mode specification may be given, too.  Note that the
wenzelm
parents: 12651
diff changeset
   179
  type declaration given here merely serves for syntactic purposes,
wenzelm
parents: 12651
diff changeset
   180
  and is not checked for consistency with the real constant.
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   181
12653
wenzelm
parents: 12651
diff changeset
   182
  \medskip We may now write either @{text "[+]"} or @{text "\<oplus>"} in
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   183
  input, while output uses the nicer syntax of $xsymbols$, provided
12653
wenzelm
parents: 12651
diff changeset
   184
  that print mode is presently active.  Such an arrangement is
wenzelm
parents: 12651
diff changeset
   185
  particularly useful for interactive development, where users may
wenzelm
parents: 12651
diff changeset
   186
  type plain ASCII text, but gain improved visual feedback from the
wenzelm
parents: 12651
diff changeset
   187
  system (say in current goal output).
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   188
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   189
  \begin{warn}
12653
wenzelm
parents: 12651
diff changeset
   190
  Alternative syntax declarations are apt to result in varying
wenzelm
parents: 12651
diff changeset
   191
  occurrences of concrete syntax in the input sources.  Isabelle
wenzelm
parents: 12651
diff changeset
   192
  provides no systematic way to convert alternative syntax expressions
wenzelm
parents: 12651
diff changeset
   193
  back and forth; print modes only affect situations where formal
wenzelm
parents: 12651
diff changeset
   194
  entities are pretty printed by the Isabelle process (e.g.\ output of
wenzelm
parents: 12651
diff changeset
   195
  terms and types), but not the original theory text.
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   196
  \end{warn}
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   197
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   198
  \medskip The following variant makes the alternative @{text \<oplus>}
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   199
  notation only available for output.  Thus we may enforce input
12653
wenzelm
parents: 12651
diff changeset
   200
  sources to refer to plain ASCII only, but effectively disable
wenzelm
parents: 12651
diff changeset
   201
  cut-and-paste from output as well.
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   202
*}
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   203
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   204
syntax (xsymbols output)
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   205
  xor :: "bool \<Rightarrow> bool \<Rightarrow> bool"    (infixl "\<oplus>\<ignore>" 60)
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   206
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   207
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   208
subsection {* Prefix Annotations *}
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   209
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   210
text {*
12653
wenzelm
parents: 12651
diff changeset
   211
  Prefix syntax annotations\index{prefix annotation} are just another
wenzelm
parents: 12651
diff changeset
   212
  degenerate form of general mixfixes \cite{isabelle-ref}, without any
wenzelm
parents: 12651
diff changeset
   213
  template arguments or priorities --- just some bits of literal
wenzelm
parents: 12651
diff changeset
   214
  syntax.  The following example illustrates this idea idea by
wenzelm
parents: 12651
diff changeset
   215
  associating common symbols with the constructors of a datatype.
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   216
*}
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   217
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   218
datatype currency =
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   219
    Euro nat    ("\<euro>")
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   220
  | Pounds nat  ("\<pounds>")
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   221
  | Yen nat     ("\<yen>")
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   222
  | Dollar nat  ("$")
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   223
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   224
text {*
12653
wenzelm
parents: 12651
diff changeset
   225
  \noindent Here the mixfix annotations on the rightmost column happen
wenzelm
parents: 12651
diff changeset
   226
  to consist of a single Isabelle symbol each: \verb,\,\verb,<euro>,,
wenzelm
parents: 12651
diff changeset
   227
  \verb,\,\verb,<pounds>,, \verb,\,\verb,<yen>,, and \verb,$,.  Recall
wenzelm
parents: 12651
diff changeset
   228
  that a constructor like @{text Euro} actually is a function @{typ
wenzelm
parents: 12651
diff changeset
   229
  "nat \<Rightarrow> currency"}.  An expression like @{text "Euro 10"} will be
wenzelm
parents: 12651
diff changeset
   230
  printed as @{term "\<euro> 10"}; only the head of the application is
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   231
  subject to our concrete syntax.  This simple form already achieves
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   232
  conformance with notational standards of the European Commission.
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   233
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   234
  \medskip Certainly, the same idea of prefix syntax also works for
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   235
  \isakeyword{consts}, \isakeyword{constdefs} etc.  The slightly
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   236
  unusual syntax declaration below decorates the existing @{typ
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   237
  currency} type with the international currency symbol @{text \<currency>}
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   238
  (\verb,\,\verb,<currency>,).
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   239
*}
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   240
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   241
syntax
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   242
  currency :: type    ("\<currency>")
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   243
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   244
text {*
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   245
  \noindent Here @{typ type} refers to the builtin syntactic category
12653
wenzelm
parents: 12651
diff changeset
   246
  of Isabelle types.  We may now write down funny things like @{text
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   247
  "\<euro> :: nat \<Rightarrow> \<currency>"}, for example.
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   248
*}
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   249
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   250
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   251
subsection {* Syntax Translations \label{sec:syntax-translations} *}
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   252
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   253
text{*
12653
wenzelm
parents: 12651
diff changeset
   254
  Mixfix syntax annotations work well for those situations where a
wenzelm
parents: 12651
diff changeset
   255
  particular constant application forms need to be decorated by
wenzelm
parents: 12651
diff changeset
   256
  concrete syntax; just reconsider @{text "xor A B"} versus @{text "A
wenzelm
parents: 12651
diff changeset
   257
  \<oplus> B"} covered before.  Occasionally, the relationship between some
wenzelm
parents: 12651
diff changeset
   258
  piece of notation and its internal form is slightly more involved.
wenzelm
parents: 12651
diff changeset
   259
  Here the concept of \bfindex{syntax translations} enters the scene.
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   260
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   261
  Using the raw \isakeyword{syntax}\index{syntax (command)} command we
12653
wenzelm
parents: 12651
diff changeset
   262
  may introduce uninterpreted notational elements, while
wenzelm
parents: 12651
diff changeset
   263
  \commdx{translations} relates the input forms with more complex
wenzelm
parents: 12651
diff changeset
   264
  logical expressions.  This essentially provides a simple mechanism
wenzelm
parents: 12651
diff changeset
   265
  for for syntactic macros; even heavier transformations may be
wenzelm
parents: 12651
diff changeset
   266
  programmed in ML \cite{isabelle-ref}.
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   267
12653
wenzelm
parents: 12651
diff changeset
   268
  \medskip A typical example of syntax translations is to decorate
wenzelm
parents: 12651
diff changeset
   269
  relational expressions (set membership of tuples) with nice symbolic
wenzelm
parents: 12651
diff changeset
   270
  notation, such as @{text "(x, y) \<in> sim"} versus @{text "x \<approx> y"}.
12635
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   271
*}
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   272
e2d44df29c94 more on concrete syntax;
wenzelm
parents: 12629
diff changeset
   273
consts
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   274
  sim :: "('a \<times> 'a) set"
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   275
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   276
syntax
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   277
  "_sim" :: "'a \<Rightarrow> 'a \<Rightarrow> bool"    (infix "\<approx>" 50)
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   278
translations
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   279
  "x \<approx> y" \<rightleftharpoons> "(x, y) \<in> sim"
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   280
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   281
text {*
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   282
  \noindent Here the name of the dummy constant @{text "_sim"} does
12653
wenzelm
parents: 12651
diff changeset
   283
  not really matter, as long as it is not used elsewhere.  Prefixing
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   284
  an underscore is a common convention.  The \isakeyword{translations}
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   285
  declaration already uses concrete syntax on the left-hand side;
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   286
  internally we relate a raw application @{text "_sim x y"} with
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   287
  @{text "(x, y) \<in> sim"}.
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   288
12653
wenzelm
parents: 12651
diff changeset
   289
  \medskip Another common application of syntax translations is to
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   290
  provide variant versions of fundamental relational expressions, such
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   291
  as @{text \<noteq>} for negated equalities.  The following declaration
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   292
  stems from Isabelle/HOL itself:
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   293
*}
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   294
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   295
syntax "_not_equal" :: "'a \<Rightarrow> 'a \<Rightarrow> bool"    (infixl "\<noteq>\<ignore>" 50)
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   296
translations "x \<noteq>\<ignore> y" \<rightleftharpoons> "\<not> (x = y)"
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   297
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   298
text {*
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   299
  \noindent Normally one would introduce derived concepts like this
12653
wenzelm
parents: 12651
diff changeset
   300
  within the logic, using \isakeyword{consts} + \isakeyword{defs}
wenzelm
parents: 12651
diff changeset
   301
  instead of \isakeyword{syntax} + \isakeyword{translations}.  The
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   302
  present formulation has the virtue that expressions are immediately
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   303
  replaced by its ``definition'' upon parsing; the effect is reversed
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   304
  upon printing.  Internally, @{text"\<noteq>"} never appears.
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   305
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   306
  Simulating definitions via translations is adequate for very basic
12653
wenzelm
parents: 12651
diff changeset
   307
  logical concepts, where a new representation is a trivial variation
wenzelm
parents: 12651
diff changeset
   308
  on an existing one.  On the other hand, syntax translations do not
wenzelm
parents: 12651
diff changeset
   309
  scale up well to large hierarchies of concepts built on each other.
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   310
*}
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   311
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   312
12653
wenzelm
parents: 12651
diff changeset
   313
section {* Document Preparation \label{sec:document-preparation} *}
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   314
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   315
text {*
12653
wenzelm
parents: 12651
diff changeset
   316
  Isabelle/Isar is centered around the concept of \bfindex{formal
wenzelm
parents: 12651
diff changeset
   317
  proof documents}\index{documents|bold}.  Ultimately the result of
wenzelm
parents: 12651
diff changeset
   318
  the user's theory development efforts is meant to be a
wenzelm
parents: 12651
diff changeset
   319
  human-readable record, presented as a browsable PDF file or printed
wenzelm
parents: 12651
diff changeset
   320
  on paper.  The overall document structure follows traditional
wenzelm
parents: 12651
diff changeset
   321
  mathematical articles, with sectioning, intermediate explanations,
wenzelm
parents: 12651
diff changeset
   322
  definitions, theorem statements and proofs.
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   323
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   324
  The Isar proof language \cite{Wenzel-PhD}, which is not covered in
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   325
  this book, admits to write formal proof texts that are acceptable
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   326
  both to the machine and human readers at the same time.  Thus
12653
wenzelm
parents: 12651
diff changeset
   327
  marginal comments and explanations may be kept at a minimum.  Even
wenzelm
parents: 12651
diff changeset
   328
  without proper coverage of human-readable proofs, Isabelle document
wenzelm
parents: 12651
diff changeset
   329
  is very useful to produce formally derived texts (elaborating on the
wenzelm
parents: 12651
diff changeset
   330
  specifications etc.).  Unstructured proof scripts given here may be
wenzelm
parents: 12651
diff changeset
   331
  just ignored by readers, or intentionally suppressed from the text
wenzelm
parents: 12651
diff changeset
   332
  by the writer (see also \S\ref{sec:doc-prep-suppress}).
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   333
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   334
  \medskip The Isabelle document preparation system essentially acts
12653
wenzelm
parents: 12651
diff changeset
   335
  like a formal front-end to {\LaTeX}.  After checking specifications
wenzelm
parents: 12651
diff changeset
   336
  and proofs, the theory sources are turned into typesetting
wenzelm
parents: 12651
diff changeset
   337
  instructions in a well-defined manner.  This enables users to write
wenzelm
parents: 12651
diff changeset
   338
  authentic reports on formal theory developments with little
wenzelm
parents: 12651
diff changeset
   339
  additional effort, the most tedious consistency checks are handled
wenzelm
parents: 12651
diff changeset
   340
  by the system.
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   341
*}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   342
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   343
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   344
subsection {* Isabelle Sessions *}
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   345
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   346
text {*
12653
wenzelm
parents: 12651
diff changeset
   347
  In contrast to the highly interactive mode of Isabelle/Isar theory
wenzelm
parents: 12651
diff changeset
   348
  development, the document preparation stage essentially works in
wenzelm
parents: 12651
diff changeset
   349
  batch-mode.  An Isabelle \bfindex{session} essentially consists of a
wenzelm
parents: 12651
diff changeset
   350
  collection of theory source files that contribute to a single output
wenzelm
parents: 12651
diff changeset
   351
  document eventually.  Session is derived from a single parent each
wenzelm
parents: 12651
diff changeset
   352
  (usually an object-logic image like \texttt{HOL}), resulting in an
wenzelm
parents: 12651
diff changeset
   353
  overall tree structure that is reflected in the output location
wenzelm
parents: 12651
diff changeset
   354
  within the file system (usually rooted at
wenzelm
parents: 12651
diff changeset
   355
  \verb,~/isabelle/browser_info,).
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   356
12653
wenzelm
parents: 12651
diff changeset
   357
  Here is the canonical arrangement of sources of a session called
wenzelm
parents: 12651
diff changeset
   358
  \texttt{MySession}:
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   359
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   360
  \begin{itemize}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   361
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   362
  \item Directory \texttt{MySession} contains the required theory
12653
wenzelm
parents: 12651
diff changeset
   363
  files ($A@1$\texttt{.thy}, \dots, $A@n$\texttt{.thy}).
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   364
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   365
  \item File \texttt{MySession/ROOT.ML} holds appropriate ML commands
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   366
  for loading all wanted theories, usually just
12653
wenzelm
parents: 12651
diff changeset
   367
  ``\texttt{use_thy~"$A@i$";}'' for any $A@i$ in leaf position of the
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   368
  theory dependency graph.
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   369
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   370
  \item Directory \texttt{MySession/document} contains everything
12653
wenzelm
parents: 12651
diff changeset
   371
  required for the {\LaTeX} stage; only \texttt{root.tex} needs to be
wenzelm
parents: 12651
diff changeset
   372
  provided initially.
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   373
12653
wenzelm
parents: 12651
diff changeset
   374
  The latter file holds appropriate {\LaTeX} code to commence a
wenzelm
parents: 12651
diff changeset
   375
  document (\verb,\documentclass, etc.), and to include the generated
wenzelm
parents: 12651
diff changeset
   376
  files $A@i$\texttt{.tex} for each theory.  The generated
wenzelm
parents: 12651
diff changeset
   377
  \texttt{session.tex} will hold {\LaTeX} commands to include all
wenzelm
parents: 12651
diff changeset
   378
  theory output files in topologically sorted order, so
wenzelm
parents: 12651
diff changeset
   379
  \verb,\input{session}, in \texttt{root.tex} will do it in most
wenzelm
parents: 12651
diff changeset
   380
  situations.
wenzelm
parents: 12651
diff changeset
   381
wenzelm
parents: 12651
diff changeset
   382
  \item In addition, \texttt{IsaMakefile} outside of the directory
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   383
  \texttt{MySession} holds appropriate dependencies and invocations of
12653
wenzelm
parents: 12651
diff changeset
   384
  Isabelle tools to control the batch job.  In fact, several sessions
wenzelm
parents: 12651
diff changeset
   385
  may be controlled by the same \texttt{IsaMakefile}.  See also
wenzelm
parents: 12651
diff changeset
   386
  \cite{isabelle-sys} for further details, especially on
wenzelm
parents: 12651
diff changeset
   387
  \texttt{isatool usedir} and \texttt{isatool make}.
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   388
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   389
  \end{itemize}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   390
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   391
  With everything put in its proper place, \texttt{isatool make}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   392
  should be sufficient to process the Isabelle session completely,
12653
wenzelm
parents: 12651
diff changeset
   393
  with the generated document appearing in its proper place.
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   394
12653
wenzelm
parents: 12651
diff changeset
   395
  \medskip In practice, users may want to have \texttt{isatool mkdir}
wenzelm
parents: 12651
diff changeset
   396
  generate an initial working setup without further ado.  For example,
wenzelm
parents: 12651
diff changeset
   397
  an empty session \texttt{MySession} derived from \texttt{HOL} may be
wenzelm
parents: 12651
diff changeset
   398
  produced as follows:
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   399
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   400
\begin{verbatim}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   401
  isatool mkdir HOL MySession
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   402
  isatool make
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   403
\end{verbatim}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   404
12653
wenzelm
parents: 12651
diff changeset
   405
  This processes the session with sensible default options, including
wenzelm
parents: 12651
diff changeset
   406
  verbose mode to tell the user where the ultimate results will
wenzelm
parents: 12651
diff changeset
   407
  appear.  The above dry run should produce should already be able to
wenzelm
parents: 12651
diff changeset
   408
  produce a single page of output (with a dummy title, empty table of
wenzelm
parents: 12651
diff changeset
   409
  contents etc.).  Any failure at that stage is likely to indicate
wenzelm
parents: 12651
diff changeset
   410
  technical problems with the user's {\LaTeX}
wenzelm
parents: 12651
diff changeset
   411
  installation.\footnote{Especially make sure that \texttt{pdflatex}
wenzelm
parents: 12651
diff changeset
   412
  is present; if all fails one may fall back on DVI output by changing
wenzelm
parents: 12651
diff changeset
   413
  \texttt{usedir} options \cite{isabelle-sys}.}
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   414
12653
wenzelm
parents: 12651
diff changeset
   415
  \medskip One may now start to populate the directory
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   416
  \texttt{MySession}, and the file \texttt{MySession/ROOT.ML}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   417
  accordingly.  \texttt{MySession/document/root.tex} should be also
12653
wenzelm
parents: 12651
diff changeset
   418
  adapted at some point; the default version is mostly
12659
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   419
  self-explanatory.  Note that the \verb,\isabellestyle, enables
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   420
  fine-tuning of the general appearance of characters and mathematical
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   421
  symbols (see also \S\ref{sec:doc-prep-symbols}).
12653
wenzelm
parents: 12651
diff changeset
   422
wenzelm
parents: 12651
diff changeset
   423
  Especially note the standard inclusion of {\LaTeX} packages
wenzelm
parents: 12651
diff changeset
   424
  \texttt{isabelle} (mandatory), and \texttt{isabellesym} (required
wenzelm
parents: 12651
diff changeset
   425
  for mathematical symbols), and the final \texttt{pdfsetup} (provides
wenzelm
parents: 12651
diff changeset
   426
  handsome defaults for \texttt{hyperref}, including URL markup).
wenzelm
parents: 12651
diff changeset
   427
  Further {\LaTeX} packages further packages may required in
wenzelm
parents: 12651
diff changeset
   428
  particular applications, e.g.\ for unusual Isabelle symbols.
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   429
12653
wenzelm
parents: 12651
diff changeset
   430
  \medskip Further auxiliary files for the {\LaTeX} stage should be
wenzelm
parents: 12651
diff changeset
   431
  included in the \texttt{MySession/document} directory, e.g.\
wenzelm
parents: 12651
diff changeset
   432
  additional {\TeX} sources or graphics.  In particular, adding
wenzelm
parents: 12651
diff changeset
   433
  \texttt{root.bib} here (with that specific name) causes an automatic
wenzelm
parents: 12651
diff changeset
   434
  run of \texttt{bibtex} to process a bibliographic database; see for
wenzelm
parents: 12651
diff changeset
   435
  further commodities \texttt{isatool document} covered in
wenzelm
parents: 12651
diff changeset
   436
  \cite{isabelle-sys}.
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   437
12653
wenzelm
parents: 12651
diff changeset
   438
  \medskip Any failure of the document preparation phase in an
wenzelm
parents: 12651
diff changeset
   439
  Isabelle batch session leaves the generated sources in there target
wenzelm
parents: 12651
diff changeset
   440
  location (as pointed out by the accompanied error message).  In case
wenzelm
parents: 12651
diff changeset
   441
  of {\LaTeX} errors, users may trace error messages at the file
wenzelm
parents: 12651
diff changeset
   442
  position of the generated text.
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   443
*}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   444
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   445
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   446
subsection {* Structure Markup *}
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   447
12653
wenzelm
parents: 12651
diff changeset
   448
text {*
wenzelm
parents: 12651
diff changeset
   449
  The large-scale structure of Isabelle documents follows existing
wenzelm
parents: 12651
diff changeset
   450
  {\LaTeX} conventions, with chapters, sections, subsubsections etc.
wenzelm
parents: 12651
diff changeset
   451
  The Isar language includes separate \bfindex{markup commands}, which
wenzelm
parents: 12651
diff changeset
   452
  do not effect the formal content of a theory (or proof), but result
wenzelm
parents: 12651
diff changeset
   453
  in corresponding {\LaTeX} elements issued to the output.
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   454
12653
wenzelm
parents: 12651
diff changeset
   455
  There are separate markup commands for different formal contexts: in
wenzelm
parents: 12651
diff changeset
   456
  header position (just before a \isakeyword{theory} command), within
wenzelm
parents: 12651
diff changeset
   457
  the theory body, or within a proof.  Note that the header needs to
wenzelm
parents: 12651
diff changeset
   458
  be treated specially, since ordinary theory and proof commands may
wenzelm
parents: 12651
diff changeset
   459
  only occur \emph{after} the initial \isakeyword{theory}
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   460
  specification.
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   461
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   462
  \smallskip
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   463
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   464
  \begin{tabular}{llll}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   465
  header & theory & proof & default meaning \\\hline
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   466
    & \commdx{chapter} & & \verb,\chapter, \\
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   467
  \commdx{header} & \commdx{section} & \commdx{sect} & \verb,\section, \\
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   468
    & \commdx{subsection} & \commdx{subsect} & \verb,\subsection, \\
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   469
    & \commdx{subsubsection} & \commdx{subsubsect} & \verb,\subsubsection, \\
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   470
  \end{tabular}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   471
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   472
  \medskip
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   473
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   474
  From the Isabelle perspective, each markup command takes a single
12659
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   475
  $text$ argument (delimited by \verb,",\dots\verb,", or
12653
wenzelm
parents: 12651
diff changeset
   476
  \verb,{,\verb,*,~\dots~\verb,*,\verb,},).  After stripping any
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   477
  surrounding white space, the argument is passed to a {\LaTeX} macro
12653
wenzelm
parents: 12651
diff changeset
   478
  \verb,\isamarkupXYZ, for any command \isakeyword{XYZ}.  These macros
wenzelm
parents: 12651
diff changeset
   479
  are defined in \verb,isabelle.sty, according to the meaning given in
wenzelm
parents: 12651
diff changeset
   480
  the rightmost column above.
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   481
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   482
  \medskip The following source fragment illustrates structure markup
12653
wenzelm
parents: 12651
diff changeset
   483
  of a theory.  Note that {\LaTeX} labels may be included inside of
wenzelm
parents: 12651
diff changeset
   484
  section headings as well.
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   485
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   486
  \begin{ttbox}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   487
  header {\ttlbrace}* Some properties of Foo Bar elements *{\ttrbrace}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   488
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   489
  theory Foo_Bar = Main:
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   490
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   491
  subsection {\ttlbrace}* Basic definitions *{\ttrbrace}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   492
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   493
  consts
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   494
    foo :: \dots
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   495
    bar :: \dots
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   496
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   497
  defs \dots
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   498
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   499
  subsection {\ttlbrace}* Derived rules *{\ttrbrace}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   500
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   501
  lemma fooI: \dots
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   502
  lemma fooE: \dots
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   503
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   504
  subsection {\ttlbrace}* Main theorem {\ttback}label{\ttlbrace}sec:main-theorem{\ttrbrace} *{\ttrbrace}
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   505
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   506
  theorem main: \dots
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   507
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   508
  end
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   509
  \end{ttbox}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   510
12653
wenzelm
parents: 12651
diff changeset
   511
  Users may occasionally want to change the meaning of markup
wenzelm
parents: 12651
diff changeset
   512
  commands, say via \verb,\renewcommand, in \texttt{root.tex}.  The
wenzelm
parents: 12651
diff changeset
   513
  \verb,\isamarkupheader, is a good candidate for some adaption, e.g.\
wenzelm
parents: 12651
diff changeset
   514
  moving it up in the hierarchy to become \verb,\chapter,.
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   515
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   516
\begin{verbatim}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   517
  \renewcommand{\isamarkupheader}[1]{\chapter{#1}}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   518
\end{verbatim}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   519
12653
wenzelm
parents: 12651
diff changeset
   520
  \noindent Certainly, this requires to change the default
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   521
  \verb,\documentclass{article}, in \texttt{root.tex} to something
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   522
  that supports the notion of chapters in the first place, e.g.\
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   523
  \verb,\documentclass{report},.
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   524
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   525
  \medskip The {\LaTeX} macro \verb,\isabellecontext, is maintained to
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   526
  hold the name of the current theory context.  This is particularly
12653
wenzelm
parents: 12651
diff changeset
   527
  useful for document headings:
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   528
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   529
\begin{verbatim}
12653
wenzelm
parents: 12651
diff changeset
   530
  \renewcommand{\isamarkupheader}[1]
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   531
  {\chapter{#1}\markright{THEORY~\isabellecontext}}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   532
\end{verbatim}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   533
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   534
  \noindent Make sure to include something like
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   535
  \verb,\pagestyle{headings}, in \texttt{root.tex}; the document
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   536
  should have more than 2 pages to show the effect.
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   537
*}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   538
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   539
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   540
subsection {* Formal Comments and Antiquotations *}
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   541
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   542
text {*
12659
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   543
  FIXME check
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   544
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   545
  Source comments of the form \verb,(,\verb,*,~\dots~\verb,*,\verb,),
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   546
  essentially act like white space and do not contribute to the formal
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   547
  text at all.  They mainly serve technical purposes to mark certain
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   548
  oddities or problems with the raw sources.
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   549
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   550
  In contrast, \bfindex{formal comments} are portions of text that are
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   551
  associated with formal Isabelle/Isar commands (\bfindex{marginal
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   552
  comments}), or even as stanalone paragraphs positioned explicitly
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   553
  within a theory or proof context (\bfindex{text blocks}).
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   554
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   555
  \medskip Marginal comments are part of each command's concrete
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   556
  syntax \cite{isabelle-ref}; the common form \verb,--,~text, where
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   557
  $text$, delimited by \verb,",\dots\verb,", or
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   558
  \verb,{,\verb,*,~\dots~\verb,*,\verb,}, as usual.  Multiple marginal
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   559
  comments may be given at the same time.  Here is a simple example:
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   560
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   561
\begin{verbatim}
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   562
  lemma "A --> A"
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   563
    -- "a triviality of propositional logic"
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   564
    -- "(should not really bother)"
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   565
    by (rule impI) -- "implicit assumption step involved here"
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   566
\end{verbatim}
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   567
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   568
  From the {\LaTeX} view, \verb,--, acts like a markup command, the
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   569
  corresponding macro is \verb,\isamarkupcmt, (of a single argument).
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   570
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   571
  \medskip The commands \bfindex{text} and \bfindex{txt} both
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   572
  introduce a text block (for theory and proof contexts,
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   573
  respectively), taking a single $text$ argument.  The {\LaTeX} output
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   574
  appears as a free-form text, surrounded with separate paragraphs and
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   575
  additional vertical spacing.  This behavior is determined by the
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   576
  {\LaTeX} environment definitions \verb,isamarkuptext, and
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   577
  \verb,isamarkuptxt,, respectively.  This may be changed via
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   578
  \verb,\renewenvironment,, but it is often sufficient to redefine
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   579
  \verb,\isastyletext, or \verb,\isastyletxt, only, which determine
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   580
  the body text style.
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   581
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   582
  \medskip The $text$ part of each of the various markup commands
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   583
  considered so far essentially inserts \emph{quoted} material within
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   584
  a formal text, essentially for instruction of the reader only
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   585
  (arbitrary {\LaTeX} macros may be included).
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   586
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   587
  An \bfindex{antiquotation} is again a formal object that has been
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   588
  embedded into such an informal portion of text.  Typically, the
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   589
  interpretation of an antiquotation is limited to well-formedness
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   590
  checks, with the result being pretty printed into the resulting
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   591
  document output.  So quoted text blocks together with antiquotations
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   592
  provide very handsome means to reference formal entities within
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   593
  informal expositions, with a high level of confidence in the
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   594
  technical details (especially syntax and types).
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   595
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   596
  The general format of antiquotations is as follows:
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   597
  \texttt{{\at}{\ttlbrace}$name$ $arguments${\ttrbrace}}, or
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   598
  \texttt{{\at}{\ttlbrace}$name$ [$options$] $arguments${\ttrbrace}}
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   599
  for a comma-separated list of $name$ or assignments
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   600
  \texttt{$name$=$value$} of $options$ (see \cite{isabelle-isar-ref}
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   601
  for details).  The syntax of $arguments$ depends on the
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   602
  antiquotation name, it generally follows the same conventions for
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   603
  types, terms, or theorems as in the formal part of a theory.
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   604
12659
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   605
  \medskip Here is an example use of the quotation-antiquotation
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   606
  technique: @{term "%x y. x"} is a well-typed term.
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   607
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   608
  \medskip This output has been produced as follows:
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   609
  \begin{ttbox}
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   610
text {\ttlbrace}*
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   611
  Here is an example use of the quotation-antiquotation technique:
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   612
  {\at}{\ttlbrace}term "%x y. x"{\ttrbrace} is a well-typed term.
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   613
*{\ttrbrace}
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   614
  \end{ttbox}
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   615
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   616
  From the notational change of the ASCII character \verb,%, to the
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   617
  symbol @{text \<lambda>} we see that the term really got printed by the
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   618
  system --- recall that the document preparation system enables
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   619
  symbolic output by default.
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   620
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   621
  \medskip In the following example we include an option to enable the
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   622
  \verb,show_types, flag of Isabelle: the antiquotation
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   623
  \texttt{{\at}}\verb,{term [show_types] "%x y. x"}, produces @{term
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   624
  [show_types] "%x y. x"}.  Here type-inference has figured out the
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   625
  most general typings in the present (theory) context.  Within a
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   626
  proof, one may get different results according to typings that have
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   627
  already been figured out in the text so far, say some fixed
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   628
  variables in the main statement given before hand.
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   629
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   630
  \medskip Several further kinds of antiquotations (and options) are
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   631
  available \cite{isabelle-sys}.  The most commonly used combinations
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   632
  are as follows:
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   633
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   634
  \medskip
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   635
12659
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   636
  \begin{tabular}{ll}
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   637
  \texttt{\at}\verb,{typ,~$\tau$\verb,}, & print type $\tau$ \\
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   638
  \texttt{\at}\verb,{term,~$t$\verb,}, & print term $t$ \\
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   639
  \texttt{\at}\verb,{prop,~$\phi$\verb,}, & print proposition $\phi$ \\
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   640
  \texttt{\at}\verb,{prop [source],~$\phi$\verb,}, & check proposition $\phi$, print its input \\
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   641
  \texttt{\at}\verb,{thm,~$a$\verb,}, & print fact $a$ \\
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   642
  \texttt{\at}\verb,{thm,~$a$~\verb,[no_vars]}, & print fact $a$, fixing schematic variables \\
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   643
  \texttt{\at}\verb,{thm [source],~$a$\verb,}, & check availability of fact $a$, print its name \\
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   644
  \texttt{\at}\verb,{text,~$s$\verb,}, & print uninterpreted text $s$ \\
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   645
  \end{tabular}
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   646
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   647
  \medskip
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   648
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   649
  Note that \verb,no_vars, given above is \emph{not} an option, but
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   650
  merely an attribute of the theorem argument given here.
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   651
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   652
  The \texttt{\at}\verb,{text, $s$\verb,}, antiquotation is
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   653
  particularly interesting.  Embedding uninterpreted text within an
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   654
  informal text body might appear useless at first sight.  Here the
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   655
  key virtue is that the string $s$ is processed as proper Isabelle
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   656
  output, interpreting Isabelle symbols (\S\ref{sec:syntax-symbols})
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   657
  appropriately.
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   658
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   659
  For example, \texttt{\at}\verb,{text "\<forall>\<exists>"}, produces @{text "\<forall>\<exists>"},
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   660
  according to the standard interpretation of these symbol (cf.\
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   661
  \S\ref{sec:doc-prep-symbols}).  Thus we achieve consistent
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   662
  mathematical notation in both the formal and informal parts of the
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   663
  document very easily.  Manual {\LaTeX} code leaves more control over
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   664
  the type-setting, but is slightly more tedious.  Also note that
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   665
  Isabelle symbols may be also displayed within the editor buffer of
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   666
  Proof~General.
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   667
*}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   668
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   669
12659
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   670
subsection {* Interpretation of symbols \label{sec:doc-prep-symbols} *}
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   671
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   672
text {*
12659
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   673
  FIXME tune
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   674
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   675
  According to \S\ref{sec:syntax-symbols}, the smalles syntactic
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   676
  entities of Isabelle text are symbols, a straight-forward
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   677
  generalization of ASCII characters.  Concerning document
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   678
  preperation, symbols are translated uniformly to {\LaTeX} code as
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   679
  follows.
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   680
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   681
  \begin{enumerate} \item 7-bit ASCII characters: letters
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   682
  \texttt{A\dots Z} and \texttt{a\dots z} are output verbatim, digits
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   683
  are passed as an argument to the \verb,\isadigit, macro, other
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   684
  characters are replaced by specific macros \verb,\isacharXYZ, (see
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   685
  also \texttt{isabelle.sty}).
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   686
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   687
  \item Named symbols: \verb,\,\verb,<,$XYZ$\verb,>, become
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   688
  \verb,\{\isasym,$XYZ$\verb,}, each (note the additional braces).
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   689
  See \cite[appendix~A]{isabelle-sys} and \texttt{isabellesym.sty} for
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   690
  the collection of predefined standard symbols.
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   691
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   692
  \item Named control symbols: \verb,\,\verb,<^,$XYZ$\verb,>, become
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   693
  \verb,\isactrl,$XYZ$; subsequent symbols may act as arguments, if
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   694
  the corresponding macro is defined accordingly.
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   695
  \end{enumerate}
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   696
*}
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   697
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   698
12653
wenzelm
parents: 12651
diff changeset
   699
subsection {* Suppressing Output \label{sec:doc-prep-suppress} *}
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   700
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   701
text {*
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   702
  By default Isabelle's document system generates a {\LaTeX} source
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   703
  file for each theory that happens to get loaded during the session.
12653
wenzelm
parents: 12651
diff changeset
   704
  The generated \texttt{session.tex} will include all of these in
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   705
  order of appearance, which in turn gets included by the standard
12653
wenzelm
parents: 12651
diff changeset
   706
  \texttt{root.tex}.  Certainly one may change the order of appearance
wenzelm
parents: 12651
diff changeset
   707
  or suppress unwanted theories by ignoring \texttt{session.tex} and
wenzelm
parents: 12651
diff changeset
   708
  include individual files in \texttt{root.tex} by hand.  On the other
wenzelm
parents: 12651
diff changeset
   709
  hand, such an arrangement requires additional maintenance chores
wenzelm
parents: 12651
diff changeset
   710
  whenever the collection of theories changes.
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   711
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   712
  Alternatively, one may tune the theory loading process in
12653
wenzelm
parents: 12651
diff changeset
   713
  \texttt{ROOT.ML} itself: traversal of the theory dependency graph
wenzelm
parents: 12651
diff changeset
   714
  may be fine-tuned by adding further \verb,use_thy, invocations,
wenzelm
parents: 12651
diff changeset
   715
  although topological sorting still has to be observed.  Moreover,
wenzelm
parents: 12651
diff changeset
   716
  the ML operator \verb,no_document, temporarily disables document
wenzelm
parents: 12651
diff changeset
   717
  generation while executing a theory loader command; its usage is
wenzelm
parents: 12651
diff changeset
   718
  like this:
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   719
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   720
\begin{verbatim}
12653
wenzelm
parents: 12651
diff changeset
   721
  no_document use_thy "A";
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   722
\end{verbatim}
12645
3af5de958a1a some text on document preparation;
wenzelm
parents: 12642
diff changeset
   723
12653
wenzelm
parents: 12651
diff changeset
   724
  \medskip Theory output may be also suppressed in smaller portions as
wenzelm
parents: 12651
diff changeset
   725
  well.  For example, research papers or slides usually do not include
wenzelm
parents: 12651
diff changeset
   726
  the formal content in full.  In order to delimit \bfindex{ignored
wenzelm
parents: 12651
diff changeset
   727
  material} special source comments
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   728
  \verb,(,\verb,*,\verb,<,\verb,*,\verb,), and
12653
wenzelm
parents: 12651
diff changeset
   729
  \verb,(,\verb,*,\verb,>,\verb,*,\verb,), may be included in the
wenzelm
parents: 12651
diff changeset
   730
  text.  Only the document preparation system is affected, the formal
wenzelm
parents: 12651
diff changeset
   731
  checking the theory is performed as before.
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   732
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   733
  In the following example we suppress the slightly formalistic
12653
wenzelm
parents: 12651
diff changeset
   734
  \isakeyword{theory} + \isakeyword{end} surroundings a theory.
12648
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   735
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   736
  \medskip
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   737
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   738
  \begin{tabular}{l}
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   739
  \verb,(,\verb,*,\verb,<,\verb,*,\verb,), \\
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   740
  \texttt{theory A = Main:} \\
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   741
  \verb,(,\verb,*,\verb,>,\verb,*,\verb,), \\
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   742
  ~~$\vdots$ \\
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   743
  \verb,(,\verb,*,\verb,<,\verb,*,\verb,), \\
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   744
  \texttt{end} \\
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   745
  \verb,(,\verb,*,\verb,>,\verb,*,\verb,), \\
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   746
  \end{tabular}
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   747
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   748
  \medskip
16d4b8c09086 some more text;
wenzelm
parents: 12645
diff changeset
   749
12653
wenzelm
parents: 12651
diff changeset
   750
  Text may be suppressed in a fine grained manner.  For example, we
wenzelm
parents: 12651
diff changeset
   751
  may even drop vital parts of a formal proof, pretending that things
wenzelm
parents: 12651
diff changeset
   752
  have been simpler than in reality.  For example, the following
wenzelm
parents: 12651
diff changeset
   753
  ``fully automatic'' proof is actually a fake:
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   754
*}
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   755
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   756
lemma "x \<noteq> (0::int) \<Longrightarrow> 0 < x * x"
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   757
  by (auto(*<*)simp add: int_less_le(*>*))
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   758
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   759
text {*
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   760
  \noindent Here the real source of the proof has been as follows:
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   761
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   762
\begin{verbatim}
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   763
  by (auto(*<*)simp add: int_less_le(*>*))
12659
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   764
\end{verbatim}
2aa05eb15bd2 getting close to completion;
wenzelm
parents: 12653
diff changeset
   765
%(*
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   766
12653
wenzelm
parents: 12651
diff changeset
   767
  \medskip Ignoring portions of printed does demand some care by the
wenzelm
parents: 12651
diff changeset
   768
  user.  First of all, the writer is responsible not to obfuscate the
wenzelm
parents: 12651
diff changeset
   769
  underlying formal development in an unduly manner.  It is fairly
wenzelm
parents: 12651
diff changeset
   770
  easy to invalidate the remaining visible text, e.g.\ by referencing
wenzelm
parents: 12651
diff changeset
   771
  questionable formal items (strange definitions, arbitrary axioms
wenzelm
parents: 12651
diff changeset
   772
  etc.) that have been hidden from sight beforehand.
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   773
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   774
  Some minor technical subtleties of the
12653
wenzelm
parents: 12651
diff changeset
   775
  \verb,(,\verb,*,\verb,<,\verb,*,\verb,),-\verb,(,\verb,*,\verb,>,\verb,*,\verb,),
wenzelm
parents: 12651
diff changeset
   776
  elements need to be kept in mind as well, since the system performs
wenzelm
parents: 12651
diff changeset
   777
  little sanity checks here.  Arguments of markup commands and formal
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   778
  comments must not be hidden, otherwise presentation fails.  Open and
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   779
  close parentheses need to be inserted carefully; it is fairly easy
12653
wenzelm
parents: 12651
diff changeset
   780
  to hide the wrong parts, especially after rearranging the sources.
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   781
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   782
  \medskip Authentic reports of formal theories, say as part of a
12653
wenzelm
parents: 12651
diff changeset
   783
  library, usually should refrain from suppressing parts of the text
12651
930df4604b36 some more ...;
wenzelm
parents: 12648
diff changeset
   784
  at all.  Other users may need the full information for their own
12653
wenzelm
parents: 12651
diff changeset
   785
  derivative work.  If a particular formalization appears inadequate
wenzelm
parents: 12651
diff changeset
   786
  for general public coverage, it is often more appropriate to think
wenzelm
parents: 12651
diff changeset
   787
  of a better way in the first place.
12629
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   788
*}
281aa36829d8 beginnings of concrete syntax;
wenzelm
parents: 11647
diff changeset
   789
11647
0538cb0f7999 initial setup for chapter on document preparation;
wenzelm
parents:
diff changeset
   790
(*<*)
0538cb0f7999 initial setup for chapter on document preparation;
wenzelm
parents:
diff changeset
   791
end
0538cb0f7999 initial setup for chapter on document preparation;
wenzelm
parents:
diff changeset
   792
(*>*)