doc-src/IsarRef/Thy/Document_Preparation.thy
author wenzelm
Tue, 10 Jun 2008 19:15:20 +0200
changeset 27127 cd6617d57a16
parent 27053 d58b0fd31b59
child 27453 eecd9d84e41b
permissions -rw-r--r--
added HOL/Tools/induct_tacs.ML;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     1
(* $Id$ *)
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     2
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     3
theory Document_Preparation
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     4
imports Main
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     5
begin
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     6
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     7
chapter {* Document preparation \label{ch:document-prep} *}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     8
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
     9
text {*
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    10
  Isabelle/Isar provides a simple document preparation system based on
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    11
  existing {PDF-\LaTeX} technology, with full support of hyper-links
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    12
  (both local references and URLs) and bookmarks.  Thus the results
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    13
  are equally well suited for WWW browsing and as printed copies.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    14
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    15
  \medskip Isabelle generates {\LaTeX} output as part of the run of a
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    16
  \emph{logic session} (see also \cite{isabelle-sys}).  Getting
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    17
  started with a working configuration for common situations is quite
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    18
  easy by using the Isabelle @{verbatim mkdir} and @{verbatim make}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    19
  tools.  First invoke
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    20
\begin{ttbox}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    21
  isatool mkdir Foo
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    22
\end{ttbox}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    23
  to initialize a separate directory for session @{verbatim Foo} ---
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    24
  it is safe to experiment, since @{verbatim "isatool mkdir"} never
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    25
  overwrites existing files.  Ensure that @{verbatim "Foo/ROOT.ML"}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    26
  holds ML commands to load all theories required for this session;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    27
  furthermore @{verbatim "Foo/document/root.tex"} should include any
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    28
  special {\LaTeX} macro packages required for your document (the
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    29
  default is usually sufficient as a start).
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    30
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    31
  The session is controlled by a separate @{verbatim IsaMakefile}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    32
  (with crude source dependencies by default).  This file is located
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    33
  one level up from the @{verbatim Foo} directory location.  Now
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    34
  invoke
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    35
\begin{ttbox}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    36
  isatool make Foo
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    37
\end{ttbox}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    38
  to run the @{verbatim Foo} session, with browser information and
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    39
  document preparation enabled.  Unless any errors are reported by
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    40
  Isabelle or {\LaTeX}, the output will appear inside the directory
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    41
  @{verbatim ISABELLE_BROWSER_INFO}, as reported by the batch job in
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    42
  verbose mode.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    43
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    44
  \medskip You may also consider to tune the @{verbatim usedir}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    45
  options in @{verbatim IsaMakefile}, for example to change the output
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    46
  format from @{verbatim pdf} to @{verbatim dvi}, or activate the
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    47
  @{verbatim "-D"} option to retain a second copy of the generated
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    48
  {\LaTeX} sources.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    49
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    50
  \medskip See \emph{The Isabelle System Manual} \cite{isabelle-sys}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    51
  for further details on Isabelle logic sessions and theory
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    52
  presentation.  The Isabelle/HOL tutorial \cite{isabelle-hol-book}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    53
  also covers theory presentation issues.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    54
*}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    55
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    56
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    57
section {* Markup commands \label{sec:markup} *}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    58
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    59
text {*
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    60
  \begin{matharray}{rcl}
27049
5072d6c77baa moved header command here;
wenzelm
parents: 27043
diff changeset
    61
    @{command_def "header"} & : & \isarkeep{toplevel} \\[0.5ex]
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    62
    @{command_def "chapter"} & : & \isarkeep{local{\dsh}theory} \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    63
    @{command_def "section"} & : & \isarkeep{local{\dsh}theory} \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    64
    @{command_def "subsection"} & : & \isarkeep{local{\dsh}theory} \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    65
    @{command_def "subsubsection"} & : & \isarkeep{local{\dsh}theory} \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    66
    @{command_def "text"} & : & \isarkeep{local{\dsh}theory} \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    67
    @{command_def "text_raw"} & : & \isarkeep{local{\dsh}theory} \\[0.5ex]
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    68
    @{command_def "sect"} & : & \isartrans{proof}{proof} \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    69
    @{command_def "subsect"} & : & \isartrans{proof}{proof} \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    70
    @{command_def "subsubsect"} & : & \isartrans{proof}{proof} \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    71
    @{command_def "txt"} & : & \isartrans{proof}{proof} \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    72
    @{command_def "txt_raw"} & : & \isartrans{proof}{proof} \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    73
  \end{matharray}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    74
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    75
  Apart from formal comments (see \secref{sec:comments}), markup
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    76
  commands provide a structured way to insert text into the document
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    77
  generated from a theory (see \cite{isabelle-sys} for more
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    78
  information on Isabelle's document preparation tools).
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    79
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    80
  \begin{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    81
    ('chapter' | 'section' | 'subsection' | 'subsubsection' | 'text') target? text
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    82
    ;
27049
5072d6c77baa moved header command here;
wenzelm
parents: 27043
diff changeset
    83
    ('header' | 'text\_raw' | 'sect' | 'subsect' | 'subsubsect' | 'txt' | 'txt\_raw') text
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    84
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    85
  \end{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    86
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    87
  \begin{descr}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    88
27049
5072d6c77baa moved header command here;
wenzelm
parents: 27043
diff changeset
    89
  \item [@{command "header"}~@{text "text"}] provides plain text
5072d6c77baa moved header command here;
wenzelm
parents: 27043
diff changeset
    90
  markup just preceding the formal beginning of a theory.  In actual
5072d6c77baa moved header command here;
wenzelm
parents: 27043
diff changeset
    91
  document preparation the corresponding {\LaTeX} macro @{verbatim
5072d6c77baa moved header command here;
wenzelm
parents: 27043
diff changeset
    92
  "\\isamarkupheader"} may be redefined to produce chapter or section
27053
wenzelm
parents: 27049
diff changeset
    93
  headings.
27049
5072d6c77baa moved header command here;
wenzelm
parents: 27043
diff changeset
    94
  
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    95
  \item [@{command "chapter"}, @{command "section"}, @{command
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    96
  "subsection"}, and @{command "subsubsection"}] mark chapter and
27053
wenzelm
parents: 27049
diff changeset
    97
  section headings.  The corresponding {\LaTeX} macros are @{verbatim
wenzelm
parents: 27049
diff changeset
    98
  "\\isamarkupchapter"}, @{verbatim "\\isamarkupsection"} etc.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
    99
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   100
  \item [@{command "text"} and @{command "txt"}] specify paragraphs of
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   101
  plain text.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   102
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   103
  \item [@{command "text_raw"} and @{command "txt_raw"}] insert
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   104
  {\LaTeX} source into the output, without additional markup.  Thus
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   105
  the full range of document manipulations becomes available.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   106
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   107
  \end{descr}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   108
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   109
  The @{text "text"} argument of these markup commands (except for
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   110
  @{command "text_raw"}) may contain references to formal entities
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   111
  (``antiquotations'', see also \secref{sec:antiq}).  These are
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   112
  interpreted in the present theory context, or the named @{text
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   113
  "target"}.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   114
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   115
  Any of these markup elements corresponds to a {\LaTeX} command with
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   116
  the name prefixed by @{verbatim "\\isamarkup"}.  For the sectioning
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   117
  commands this is a plain macro with a single argument, e.g.\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   118
  @{verbatim "\\isamarkupchapter{"}@{text "\<dots>"}@{verbatim "}"} for
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   119
  @{command "chapter"}.  The @{command "text"} markup results in a
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   120
  {\LaTeX} environment @{verbatim "\\begin{isamarkuptext}"} @{text
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   121
  "\<dots>"} @{verbatim "\\end{isamarkuptext}"}, while @{command "text_raw"}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   122
  causes the text to be inserted directly into the {\LaTeX} source.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   123
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   124
  \medskip The proof markup commands closely resemble those for theory
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   125
  specifications, but have a different formal status and produce
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   126
  different {\LaTeX} macros.  Also note that the @{command_ref
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   127
  "header"} declaration (see \secref{sec:begin-thy}) admits to insert
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   128
  section markup just preceding the actual theory definition.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   129
*}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   130
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   131
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   132
section {* Antiquotations \label{sec:antiq} *}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   133
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   134
text {*
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   135
  \begin{matharray}{rcl}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   136
    @{antiquotation_def "theory"} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   137
    @{antiquotation_def "thm"} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   138
    @{antiquotation_def "prop"} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   139
    @{antiquotation_def "term"} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   140
    @{antiquotation_def const} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   141
    @{antiquotation_def abbrev} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   142
    @{antiquotation_def typeof} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   143
    @{antiquotation_def typ} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   144
    @{antiquotation_def thm_style} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   145
    @{antiquotation_def term_style} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   146
    @{antiquotation_def "text"} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   147
    @{antiquotation_def goals} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   148
    @{antiquotation_def subgoals} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   149
    @{antiquotation_def prf} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   150
    @{antiquotation_def full_prf} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   151
    @{antiquotation_def ML} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   152
    @{antiquotation_def ML_type} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   153
    @{antiquotation_def ML_struct} & : & \isarantiq \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   154
  \end{matharray}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   155
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   156
  The text body of formal comments (see also \secref{sec:comments})
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   157
  may contain antiquotations of logical entities, such as theorems,
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   158
  terms and types, which are to be presented in the final output
27053
wenzelm
parents: 27049
diff changeset
   159
  produced by the Isabelle document preparation system.
27043
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   160
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   161
  Thus embedding of ``@{text "@{term [show_types] \"f x = a + x\"}"}''
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   162
  within a text block would cause
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   163
  \isa{{\isacharparenleft}f{\isasymColon}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}a{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharplus}\ x} to appear in the final {\LaTeX} document.  Also note that theorem
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   164
  antiquotations may involve attributes as well.  For example,
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   165
  @{text "@{thm sym [no_vars]}"} would print the theorem's
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   166
  statement where all schematic variables have been replaced by fixed
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   167
  ones, which are easier to read.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   168
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   169
  \begin{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   170
    atsign lbrace antiquotation rbrace
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   171
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   172
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   173
    antiquotation:
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   174
      'theory' options name |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   175
      'thm' options thmrefs |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   176
      'prop' options prop |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   177
      'term' options term |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   178
      'const' options term |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   179
      'abbrev' options term |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   180
      'typeof' options term |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   181
      'typ' options type |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   182
      'thm\_style' options name thmref |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   183
      'term\_style' options name term |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   184
      'text' options name |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   185
      'goals' options |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   186
      'subgoals' options |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   187
      'prf' options thmrefs |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   188
      'full\_prf' options thmrefs |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   189
      'ML' options name |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   190
      'ML\_type' options name |
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   191
      'ML\_struct' options name
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   192
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   193
    options: '[' (option * ',') ']'
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   194
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   195
    option: name | name '=' name
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   196
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   197
  \end{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   198
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   199
  Note that the syntax of antiquotations may \emph{not} include source
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   200
  comments @{verbatim "(*"}~@{text "\<dots>"}~@{verbatim "*)"} or verbatim
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   201
  text @{verbatim "{"}@{verbatim "*"}~@{text "\<dots>"}~@{verbatim
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   202
  "*"}@{verbatim "}"}.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   203
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   204
  \begin{descr}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   205
  
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   206
  \item [@{text "@{theory A}"}] prints the name @{text "A"}, which is
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   207
  guaranteed to refer to a valid ancestor theory in the current
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   208
  context.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   209
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   210
  \item [@{text "@{thm a\<^sub>1 \<dots> a\<^sub>n}"}] prints theorems
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   211
  @{text "a\<^sub>1 \<dots> a\<^sub>n"}.  Note that attribute specifications
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   212
  may be included as well (see also \secref{sec:syn-att}); the
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   213
  @{attribute_ref no_vars} rule (see \secref{sec:misc-meth-att}) would
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   214
  be particularly useful to suppress printing of schematic variables.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   215
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   216
  \item [@{text "@{prop \<phi>}"}] prints a well-typed proposition @{text
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   217
  "\<phi>"}.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   218
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   219
  \item [@{text "@{term t}"}] prints a well-typed term @{text "t"}.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   220
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   221
  \item [@{text "@{const c}"}] prints a logical or syntactic constant
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   222
  @{text "c"}.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   223
  
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   224
  \item [@{text "@{abbrev c x\<^sub>1 \<dots> x\<^sub>n}"}] prints a constant
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   225
  abbreviation @{text "c x\<^sub>1 \<dots> x\<^sub>n \<equiv> rhs"} as defined in
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   226
  the current context.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   227
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   228
  \item [@{text "@{typeof t}"}] prints the type of a well-typed term
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   229
  @{text "t"}.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   230
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   231
  \item [@{text "@{typ \<tau>}"}] prints a well-formed type @{text "\<tau>"}.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   232
  
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   233
  \item [@{text "@{thm_style s a}"}] prints theorem @{text a},
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   234
  previously applying a style @{text s} to it (see below).
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   235
  
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   236
  \item [@{text "@{term_style s t}"}] prints a well-typed term @{text
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   237
  t} after applying a style @{text s} to it (see below).
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   238
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   239
  \item [@{text "@{text s}"}] prints uninterpreted source text @{text
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   240
  s}.  This is particularly useful to print portions of text according
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   241
  to the Isabelle {\LaTeX} output style, without demanding
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   242
  well-formedness (e.g.\ small pieces of terms that should not be
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   243
  parsed or type-checked yet).
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   244
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   245
  \item [@{text "@{goals}"}] prints the current \emph{dynamic} goal
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   246
  state.  This is mainly for support of tactic-emulation scripts
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   247
  within Isar --- presentation of goal states does not conform to
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   248
  actual human-readable proof documents.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   249
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   250
  Please do not include goal states into document output unless you
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   251
  really know what you are doing!
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   252
  
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   253
  \item [@{text "@{subgoals}"}] is similar to @{text "@{goals}"}, but
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   254
  does not print the main goal.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   255
  
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   256
  \item [@{text "@{prf a\<^sub>1 \<dots> a\<^sub>n}"}] prints the (compact)
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   257
  proof terms corresponding to the theorems @{text "a\<^sub>1 \<dots>
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   258
  a\<^sub>n"}. Note that this requires proof terms to be switched on
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   259
  for the current object logic (see the ``Proof terms'' section of the
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   260
  Isabelle reference manual for information on how to do this).
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   261
  
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   262
  \item [@{text "@{full_prf a\<^sub>1 \<dots> a\<^sub>n}"}] is like @{text
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   263
  "@{prf a\<^sub>1 \<dots> a\<^sub>n}"}, but displays the full proof terms,
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   264
  i.e.\ also displays information omitted in the compact proof term,
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   265
  which is denoted by ``@{text _}'' placeholders there.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   266
  
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   267
  \item [@{text "@{ML s}"}, @{text "@{ML_type s}"}, and @{text
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   268
  "@{ML_struct s}"}] check text @{text s} as ML value, type, and
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   269
  structure, respectively.  The source is displayed verbatim.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   270
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   271
  \end{descr}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   272
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   273
  \medskip The following standard styles for use with @{text
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   274
  thm_style} and @{text term_style} are available:
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   275
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   276
  \begin{descr}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   277
  
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   278
  \item [@{text lhs}] extracts the first argument of any application
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   279
  form with at least two arguments -- typically meta-level or
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   280
  object-level equality, or any other binary relation.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   281
  
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   282
  \item [@{text rhs}] is like @{text lhs}, but extracts the second
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   283
  argument.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   284
  
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   285
  \item [@{text "concl"}] extracts the conclusion @{text C} from a rule
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   286
  in Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   287
  
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   288
  \item [@{text "prem1"}, \dots, @{text "prem9"}] extract premise
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   289
  number @{text "1, \<dots>, 9"}, respectively, from from a rule in
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   290
  Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   291
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   292
  \end{descr}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   293
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   294
  \medskip
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   295
  The following options are available to tune the output.  Note that most of
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   296
  these coincide with ML flags of the same names (see also \cite{isabelle-ref}).
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   297
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   298
  \begin{descr}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   299
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   300
  \item[@{text "show_types = bool"} and @{text "show_sorts = bool"}]
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   301
  control printing of explicit type and sort constraints.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   302
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   303
  \item[@{text "show_structs = bool"}] controls printing of implicit
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   304
  structures.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   305
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   306
  \item[@{text "long_names = bool"}] forces names of types and
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   307
  constants etc.\ to be printed in their fully qualified internal
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   308
  form.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   309
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   310
  \item[@{text "short_names = bool"}] forces names of types and
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   311
  constants etc.\ to be printed unqualified.  Note that internalizing
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   312
  the output again in the current context may well yield a different
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   313
  result.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   314
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   315
  \item[@{text "unique_names = bool"}] determines whether the printed
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   316
  version of qualified names should be made sufficiently long to avoid
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   317
  overlap with names declared further back.  Set to @{text false} for
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   318
  more concise output.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   319
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   320
  \item[@{text "eta_contract = bool"}] prints terms in @{text
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   321
  \<eta>}-contracted form.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   322
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   323
  \item[@{text "display = bool"}] indicates if the text is to be
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   324
  output as multi-line ``display material'', rather than a small piece
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   325
  of text without line breaks (which is the default).
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   326
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   327
  \item[@{text "break = bool"}] controls line breaks in non-display
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   328
  material.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   329
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   330
  \item[@{text "quotes = bool"}] indicates if the output should be
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   331
  enclosed in double quotes.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   332
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   333
  \item[@{text "mode = name"}] adds @{text name} to the print mode to
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   334
  be used for presentation (see also \cite{isabelle-ref}).  Note that
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   335
  the standard setup for {\LaTeX} output is already present by
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   336
  default, including the modes @{text latex} and @{text xsymbols}.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   337
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   338
  \item[@{text "margin = nat"} and @{text "indent = nat"}] change the
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   339
  margin or indentation for pretty printing of display material.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   340
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   341
  \item[@{text "source = bool"}] prints the source text of the
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   342
  antiquotation arguments, rather than the actual value.  Note that
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   343
  this does not affect well-formedness checks of @{antiquotation
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   344
  "thm"}, @{antiquotation "term"}, etc. (only the @{antiquotation
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   345
  "text"} antiquotation admits arbitrary output).
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   346
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   347
  \item[@{text "goals_limit = nat"}] determines the maximum number of
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   348
  goals to be printed.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   349
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   350
  \item[@{text "locale = name"}] specifies an alternative locale
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   351
  context used for evaluating and printing the subsequent argument.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   352
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   353
  \end{descr}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   354
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   355
  For boolean flags, ``@{text "name = true"}'' may be abbreviated as
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   356
  ``@{text name}''.  All of the above flags are disabled by default,
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   357
  unless changed from ML.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   358
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   359
  \medskip Note that antiquotations do not only spare the author from
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   360
  tedious typing of logical entities, but also achieve some degree of
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   361
  consistency-checking of informal explanations with formal
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   362
  developments: well-formedness of terms and types with respect to the
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   363
  current theory or proof context is ensured here.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   364
*}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   365
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   366
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   367
section {* Tagged commands \label{sec:tags} *}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   368
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   369
text {*
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   370
  Each Isabelle/Isar command may be decorated by presentation tags:
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   371
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   372
  \indexouternonterm{tags}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   373
  \begin{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   374
    tags: ( tag * )
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   375
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   376
    tag: '\%' (ident | string)
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   377
  \end{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   378
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   379
  The tags @{text "theory"}, @{text "proof"}, @{text "ML"} are already
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   380
  pre-declared for certain classes of commands:
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   381
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   382
 \medskip
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   383
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   384
  \begin{tabular}{ll}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   385
    @{text "theory"} & theory begin/end \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   386
    @{text "proof"} & all proof commands \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   387
    @{text "ML"} & all commands involving ML code \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   388
  \end{tabular}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   389
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   390
  \medskip The Isabelle document preparation system (see also
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   391
  \cite{isabelle-sys}) allows tagged command regions to be presented
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   392
  specifically, e.g.\ to fold proof texts, or drop parts of the text
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   393
  completely.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   394
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   395
  For example ``@{command "by"}~@{text "%invisible auto"}'' would
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   396
  cause that piece of proof to be treated as @{text invisible} instead
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   397
  of @{text "proof"} (the default), which may be either show or hidden
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   398
  depending on the document setup.  In contrast, ``@{command
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   399
  "by"}~@{text "%visible auto"}'' would force this text to be shown
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   400
  invariably.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   401
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   402
  Explicit tag specifications within a proof apply to all subsequent
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   403
  commands of the same level of nesting.  For example, ``@{command
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   404
  "proof"}~@{text "%visible \<dots>"}~@{command "qed"}'' would force the
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   405
  whole sub-proof to be typeset as @{text visible} (unless some of its
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   406
  parts are tagged differently).
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   407
*}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   408
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   409
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   410
section {* Draft presentation *}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   411
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   412
text {*
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   413
  \begin{matharray}{rcl}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   414
    @{command_def "display_drafts"}@{text "\<^sup>*"} & : & \isarkeep{\cdot} \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   415
    @{command_def "print_drafts"}@{text "\<^sup>*"} & : & \isarkeep{\cdot} \\
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   416
  \end{matharray}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   417
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   418
  \begin{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   419
    ('display\_drafts' | 'print\_drafts') (name +)
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   420
    ;
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   421
  \end{rail}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   422
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   423
  \begin{descr}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   424
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   425
  \item [@{command "display_drafts"}~@{text paths} and @{command
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   426
  "print_drafts"}~@{text paths}] perform simple output of a given list
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   427
  of raw source files.  Only those symbols that do not require
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   428
  additional {\LaTeX} packages are displayed properly, everything else
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   429
  is left verbatim.
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   430
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   431
  \end{descr}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   432
*}
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   433
3ff111ed85a1 added Document_Preparation;
wenzelm
parents:
diff changeset
   434
end