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