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