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