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