doc-src/IsarRef/Thy/Document_Preparation.thy
author wenzelm
Mon, 02 May 2011 01:05:50 +0200
changeset 43467 6c621a9d612a
parent 41137 ca132ef44944
child 43488 77d239840285
permissions -rw-r--r--
modernized rail diagrams using @{rail} antiquotation;
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@41048
    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@43467
    84
  @{rail "
wenzelm@43467
    85
    (@@{command chapter} | @@{command section} | @@{command subsection} |
wenzelm@43467
    86
      @@{command subsubsection} | @@{command text}) @{syntax target}? @{syntax text}
wenzelm@27043
    87
    ;
wenzelm@43467
    88
    (@@{command header} | @@{command text_raw} | @@{command sect} | @@{command subsect} |
wenzelm@43467
    89
      @@{command subsubsect} | @@{command txt} | @@{command txt_raw}) @{syntax text}
wenzelm@43467
    90
  "}
wenzelm@27043
    91
wenzelm@28760
    92
  \begin{description}
wenzelm@27043
    93
wenzelm@28760
    94
  \item @{command header} provides plain text markup just preceding
wenzelm@28747
    95
  the formal beginning of a theory.  The corresponding {\LaTeX} macro
wenzelm@28747
    96
  is @{verbatim "\\isamarkupheader"}, which acts like @{command
wenzelm@28747
    97
  section} by default.
wenzelm@27049
    98
  
wenzelm@28760
    99
  \item @{command chapter}, @{command section}, @{command subsection},
wenzelm@28760
   100
  and @{command subsubsection} mark chapter and section headings
wenzelm@28760
   101
  within the main theory body or local theory targets.  The
wenzelm@28747
   102
  corresponding {\LaTeX} macros are @{verbatim "\\isamarkupchapter"},
wenzelm@28749
   103
  @{verbatim "\\isamarkupsection"}, @{verbatim
wenzelm@28749
   104
  "\\isamarkupsubsection"} etc.
wenzelm@27043
   105
wenzelm@28760
   106
  \item @{command sect}, @{command subsect}, and @{command subsubsect}
wenzelm@28760
   107
  mark section headings within proofs.  The corresponding {\LaTeX}
wenzelm@28760
   108
  macros are @{verbatim "\\isamarkupsect"}, @{verbatim
wenzelm@28749
   109
  "\\isamarkupsubsect"} etc.
wenzelm@27043
   110
wenzelm@28760
   111
  \item @{command text} and @{command txt} specify paragraphs of plain
wenzelm@28760
   112
  text.  This corresponds to a {\LaTeX} environment @{verbatim
wenzelm@28747
   113
  "\\begin{isamarkuptext}"} @{text "\<dots>"} @{verbatim
wenzelm@28747
   114
  "\\end{isamarkuptext}"} etc.
wenzelm@28747
   115
wenzelm@28760
   116
  \item @{command text_raw} and @{command txt_raw} insert {\LaTeX}
wenzelm@28747
   117
  source into the output, without additional markup.  Thus the full
wenzelm@28747
   118
  range of document manipulations becomes available, at the risk of
wenzelm@28747
   119
  messing up document output.
wenzelm@27043
   120
wenzelm@28760
   121
  \end{description}
wenzelm@27043
   122
wenzelm@28747
   123
  Except for @{command "text_raw"} and @{command "txt_raw"}, the text
wenzelm@28747
   124
  passed to any of the above markup commands may refer to formal
wenzelm@28749
   125
  entities via \emph{document antiquotations}, see also
wenzelm@28749
   126
  \secref{sec:antiq}.  These are interpreted in the present theory or
wenzelm@28749
   127
  proof context, or the named @{text "target"}.
wenzelm@27043
   128
wenzelm@27043
   129
  \medskip The proof markup commands closely resemble those for theory
wenzelm@27043
   130
  specifications, but have a different formal status and produce
wenzelm@28749
   131
  different {\LaTeX} macros.  The default definitions coincide for
wenzelm@28749
   132
  analogous commands such as @{command section} and @{command sect}.
wenzelm@27043
   133
*}
wenzelm@27043
   134
wenzelm@27043
   135
wenzelm@28749
   136
section {* Document Antiquotations \label{sec:antiq} *}
wenzelm@27043
   137
wenzelm@27043
   138
text {*
wenzelm@27043
   139
  \begin{matharray}{rcl}
wenzelm@28761
   140
    @{antiquotation_def "theory"} & : & @{text antiquotation} \\
wenzelm@28761
   141
    @{antiquotation_def "thm"} & : & @{text antiquotation} \\
wenzelm@28761
   142
    @{antiquotation_def "lemma"} & : & @{text antiquotation} \\
wenzelm@28761
   143
    @{antiquotation_def "prop"} & : & @{text antiquotation} \\
wenzelm@28761
   144
    @{antiquotation_def "term"} & : & @{text antiquotation} \\
haftmann@32898
   145
    @{antiquotation_def term_type} & : & @{text antiquotation} \\
haftmann@32898
   146
    @{antiquotation_def typeof} & : & @{text antiquotation} \\
wenzelm@28761
   147
    @{antiquotation_def const} & : & @{text antiquotation} \\
wenzelm@28761
   148
    @{antiquotation_def abbrev} & : & @{text antiquotation} \\
wenzelm@28761
   149
    @{antiquotation_def typ} & : & @{text antiquotation} \\
haftmann@39551
   150
    @{antiquotation_def type} & : & @{text antiquotation} \\
haftmann@39551
   151
    @{antiquotation_def class} & : & @{text antiquotation} \\
wenzelm@28761
   152
    @{antiquotation_def "text"} & : & @{text antiquotation} \\
wenzelm@28761
   153
    @{antiquotation_def goals} & : & @{text antiquotation} \\
wenzelm@28761
   154
    @{antiquotation_def subgoals} & : & @{text antiquotation} \\
wenzelm@28761
   155
    @{antiquotation_def prf} & : & @{text antiquotation} \\
wenzelm@28761
   156
    @{antiquotation_def full_prf} & : & @{text antiquotation} \\
wenzelm@28761
   157
    @{antiquotation_def ML} & : & @{text antiquotation} \\
wenzelm@28761
   158
    @{antiquotation_def ML_type} & : & @{text antiquotation} \\
wenzelm@28761
   159
    @{antiquotation_def ML_struct} & : & @{text antiquotation} \\
wenzelm@41051
   160
    @{antiquotation_def "file"} & : & @{text antiquotation} \\
wenzelm@27043
   161
  \end{matharray}
wenzelm@27043
   162
wenzelm@28749
   163
  The overall content of an Isabelle/Isar theory may alternate between
wenzelm@28749
   164
  formal and informal text.  The main body consists of formal
wenzelm@28749
   165
  specification and proof commands, interspersed with markup commands
wenzelm@28749
   166
  (\secref{sec:markup}) or document comments (\secref{sec:comments}).
wenzelm@28749
   167
  The argument of markup commands quotes informal text to be printed
wenzelm@28749
   168
  in the resulting document, but may again refer to formal entities
wenzelm@28749
   169
  via \emph{document antiquotations}.
wenzelm@27043
   170
wenzelm@28749
   171
  For example, embedding of ``@{text [source=false] "@{term [show_types] \"f x = a + x\"}"}''
wenzelm@28749
   172
  within a text block makes
wenzelm@28749
   173
  \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
   174
wenzelm@28749
   175
  Antiquotations usually spare the author tedious typing of logical
wenzelm@28749
   176
  entities in full detail.  Even more importantly, some degree of
wenzelm@28749
   177
  consistency-checking between the main body of formal text and its
wenzelm@28749
   178
  informal explanation is achieved, since terms and types appearing in
wenzelm@28749
   179
  antiquotations are checked within the current theory or proof
wenzelm@28749
   180
  context.
wenzelm@27043
   181
wenzelm@43467
   182
  @{rail "
wenzelm@43467
   183
    '@{' antiquotation '}'
wenzelm@27043
   184
    ;
wenzelm@43467
   185
    @{syntax_def antiquotation}:
wenzelm@43467
   186
      @@{antiquotation theory} options @{syntax name} |
wenzelm@43467
   187
      @@{antiquotation thm} options styles @{syntax thmrefs} |
wenzelm@43467
   188
      @@{antiquotation lemma} options @{syntax prop} @'by' @{syntax method} |
wenzelm@43467
   189
      @@{antiquotation prop} options styles @{syntax prop} |
wenzelm@43467
   190
      @@{antiquotation term} options styles @{syntax term} |
wenzelm@43467
   191
      @@{antiquotation term_type} options styles @{syntax term} |
wenzelm@43467
   192
      @@{antiquotation typeof} options styles @{syntax term} |
wenzelm@43467
   193
      @@{antiquotation const} options @{syntax term} |
wenzelm@43467
   194
      @@{antiquotation abbrev} options @{syntax term} |
wenzelm@43467
   195
      @@{antiquotation typ} options @{syntax type} |
wenzelm@43467
   196
      @@{antiquotation type} options @{syntax name} |
wenzelm@43467
   197
      @@{antiquotation class} options @{syntax name} |
wenzelm@43467
   198
      @@{antiquotation text} options @{syntax name} |
wenzelm@43467
   199
      @@{antiquotation goals} options |
wenzelm@43467
   200
      @@{antiquotation subgoals} options |
wenzelm@43467
   201
      @@{antiquotation prf} options @{syntax thmrefs} |
wenzelm@43467
   202
      @@{antiquotation full_prf} options @{syntax thmrefs} |
wenzelm@43467
   203
      @@{antiquotation ML} options @{syntax name} |
wenzelm@43467
   204
      @@{antiquotation ML_type} options @{syntax name} |
wenzelm@43467
   205
      @@{antiquotation ML_struct} options @{syntax name} |
wenzelm@43467
   206
      @@{antiquotation \"file\"} options @{syntax name}
wenzelm@27043
   207
    ;
wenzelm@27043
   208
    options: '[' (option * ',') ']'
wenzelm@27043
   209
    ;
wenzelm@43467
   210
    option: @{syntax name} | @{syntax name} '=' @{syntax name}
wenzelm@27043
   211
    ;
haftmann@32892
   212
    styles: '(' (style + ',') ')'
haftmann@32891
   213
    ;
wenzelm@43467
   214
    style: (@{syntax name} +)
wenzelm@43467
   215
  "} %% FIXME check lemma
wenzelm@27043
   216
wenzelm@27043
   217
  Note that the syntax of antiquotations may \emph{not} include source
wenzelm@28749
   218
  comments @{verbatim "(*"}~@{text "\<dots>"}~@{verbatim "*)"} nor verbatim
wenzelm@27043
   219
  text @{verbatim "{"}@{verbatim "*"}~@{text "\<dots>"}~@{verbatim
wenzelm@27043
   220
  "*"}@{verbatim "}"}.
wenzelm@27043
   221
wenzelm@28760
   222
  \begin{description}
wenzelm@27043
   223
  
wenzelm@28760
   224
  \item @{text "@{theory A}"} prints the name @{text "A"}, which is
wenzelm@27043
   225
  guaranteed to refer to a valid ancestor theory in the current
wenzelm@27043
   226
  context.
wenzelm@27043
   227
wenzelm@28760
   228
  \item @{text "@{thm a\<^sub>1 \<dots> a\<^sub>n}"} prints theorems @{text "a\<^sub>1 \<dots> a\<^sub>n"}.
wenzelm@28749
   229
  Full fact expressions are allowed here, including attributes
wenzelm@28749
   230
  (\secref{sec:syn-att}).
wenzelm@27043
   231
wenzelm@28760
   232
  \item @{text "@{prop \<phi>}"} prints a well-typed proposition @{text
wenzelm@27043
   233
  "\<phi>"}.
wenzelm@27043
   234
wenzelm@28760
   235
  \item @{text "@{lemma \<phi> by m}"} proves a well-typed proposition
wenzelm@28749
   236
  @{text "\<phi>"} by method @{text m} and prints the original @{text "\<phi>"}.
haftmann@27453
   237
wenzelm@28760
   238
  \item @{text "@{term t}"} prints a well-typed term @{text "t"}.
wenzelm@27043
   239
haftmann@32898
   240
  \item @{text "@{term_type t}"} prints a well-typed term @{text "t"}
haftmann@32898
   241
  annotated with its type.
haftmann@32898
   242
haftmann@32898
   243
  \item @{text "@{typeof t}"} prints the type of a well-typed term
haftmann@32898
   244
  @{text "t"}.
haftmann@32898
   245
wenzelm@28760
   246
  \item @{text "@{const c}"} prints a logical or syntactic constant
wenzelm@27043
   247
  @{text "c"}.
wenzelm@27043
   248
  
wenzelm@28760
   249
  \item @{text "@{abbrev c x\<^sub>1 \<dots> x\<^sub>n}"} prints a constant abbreviation
wenzelm@28760
   250
  @{text "c x\<^sub>1 \<dots> x\<^sub>n \<equiv> rhs"} as defined in the current context.
haftmann@39551
   251
wenzelm@28760
   252
  \item @{text "@{typ \<tau>}"} prints a well-formed type @{text "\<tau>"}.
haftmann@39551
   253
wenzelm@39937
   254
  \item @{text "@{type \<kappa>}"} prints a (logical or syntactic) type
wenzelm@39937
   255
    constructor @{text "\<kappa>"}.
haftmann@39551
   256
haftmann@39551
   257
  \item @{text "@{class c}"} prints a class @{text c}.
haftmann@39551
   258
wenzelm@28760
   259
  \item @{text "@{text s}"} prints uninterpreted source text @{text
wenzelm@27043
   260
  s}.  This is particularly useful to print portions of text according
wenzelm@28749
   261
  to the Isabelle document style, without demanding well-formedness,
wenzelm@28749
   262
  e.g.\ small pieces of terms that should not be parsed or
wenzelm@28749
   263
  type-checked yet.
wenzelm@27043
   264
wenzelm@28760
   265
  \item @{text "@{goals}"} prints the current \emph{dynamic} goal
wenzelm@27043
   266
  state.  This is mainly for support of tactic-emulation scripts
wenzelm@28749
   267
  within Isar.  Presentation of goal states does not conform to the
wenzelm@28749
   268
  idea of human-readable proof documents!
wenzelm@27043
   269
wenzelm@28749
   270
  When explaining proofs in detail it is usually better to spell out
wenzelm@28749
   271
  the reasoning via proper Isar proof commands, instead of peeking at
wenzelm@28749
   272
  the internal machine configuration.
wenzelm@27043
   273
  
wenzelm@28760
   274
  \item @{text "@{subgoals}"} is similar to @{text "@{goals}"}, but
wenzelm@27043
   275
  does not print the main goal.
wenzelm@27043
   276
  
wenzelm@28760
   277
  \item @{text "@{prf a\<^sub>1 \<dots> a\<^sub>n}"} prints the (compact) proof terms
wenzelm@28749
   278
  corresponding to the theorems @{text "a\<^sub>1 \<dots> a\<^sub>n"}. Note that this
wenzelm@28749
   279
  requires proof terms to be switched on for the current logic
wenzelm@28749
   280
  session.
wenzelm@27043
   281
  
wenzelm@28760
   282
  \item @{text "@{full_prf a\<^sub>1 \<dots> a\<^sub>n}"} is like @{text "@{prf a\<^sub>1 \<dots>
wenzelm@28760
   283
  a\<^sub>n}"}, but prints the full proof terms, i.e.\ also displays
wenzelm@28760
   284
  information omitted in the compact proof term, which is denoted by
wenzelm@28760
   285
  ``@{text _}'' placeholders there.
wenzelm@27043
   286
  
wenzelm@28760
   287
  \item @{text "@{ML s}"}, @{text "@{ML_type s}"}, and @{text
wenzelm@28760
   288
  "@{ML_struct s}"} check text @{text s} as ML value, type, and
wenzelm@28749
   289
  structure, respectively.  The source is printed verbatim.
wenzelm@27043
   290
wenzelm@41051
   291
  \item @{text "@{file path}"} checks that @{text "path"} refers to a
wenzelm@41051
   292
  file (or directory) and prints it verbatim.
wenzelm@41051
   293
wenzelm@28760
   294
  \end{description}
wenzelm@28749
   295
*}
wenzelm@27043
   296
wenzelm@28749
   297
wenzelm@28749
   298
subsubsection {* Styled antiquotations *}
wenzelm@28749
   299
haftmann@32891
   300
text {* The antiquotations @{text thm}, @{text prop} and @{text
haftmann@32891
   301
  term} admit an extra \emph{style} specification to modify the
haftmann@32891
   302
  printed result.  A style is specified by a name with a possibly
haftmann@32891
   303
  empty number of arguments;  multiple styles can be sequenced with
haftmann@32891
   304
  commas.  The following standard styles are available:
wenzelm@27043
   305
wenzelm@28760
   306
  \begin{description}
wenzelm@27043
   307
  
wenzelm@28760
   308
  \item @{text lhs} extracts the first argument of any application
wenzelm@28749
   309
  form with at least two arguments --- typically meta-level or
wenzelm@27043
   310
  object-level equality, or any other binary relation.
wenzelm@27043
   311
  
wenzelm@28760
   312
  \item @{text rhs} is like @{text lhs}, but extracts the second
wenzelm@27043
   313
  argument.
wenzelm@27043
   314
  
wenzelm@28760
   315
  \item @{text "concl"} extracts the conclusion @{text C} from a rule
wenzelm@27043
   316
  in Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}.
wenzelm@27043
   317
  
haftmann@32891
   318
  \item @{text "prem"} @{text n} extract premise number
haftmann@32891
   319
  @{text "n"} from from a rule in Horn-clause
wenzelm@28760
   320
  normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}
wenzelm@27043
   321
wenzelm@28760
   322
  \end{description}
wenzelm@28749
   323
*}
wenzelm@27043
   324
wenzelm@28749
   325
wenzelm@28749
   326
subsubsection {* General options *}
wenzelm@28749
   327
wenzelm@28749
   328
text {* The following options are available to tune the printed output
wenzelm@28749
   329
  of antiquotations.  Note that many of these coincide with global ML
wenzelm@28749
   330
  flags of the same names.
wenzelm@27043
   331
wenzelm@28760
   332
  \begin{description}
wenzelm@27043
   333
wenzelm@30397
   334
  \item @{antiquotation_option_def show_types}~@{text "= bool"} and
wenzelm@30397
   335
  @{antiquotation_option_def show_sorts}~@{text "= bool"} control
wenzelm@30397
   336
  printing of explicit type and sort constraints.
wenzelm@27043
   337
wenzelm@30397
   338
  \item @{antiquotation_option_def show_structs}~@{text "= bool"}
wenzelm@30397
   339
  controls printing of implicit structures.
wenzelm@27043
   340
wenzelm@41137
   341
  \item @{antiquotation_option_def show_abbrevs}~@{text "= bool"}
wenzelm@41137
   342
  controls folding of abbreviations.
wenzelm@41137
   343
wenzelm@30397
   344
  \item @{antiquotation_option_def long_names}~@{text "= bool"} forces
wenzelm@30397
   345
  names of types and constants etc.\ to be printed in their fully
wenzelm@30397
   346
  qualified internal form.
wenzelm@27043
   347
wenzelm@30397
   348
  \item @{antiquotation_option_def short_names}~@{text "= bool"}
wenzelm@30397
   349
  forces names of types and constants etc.\ to be printed unqualified.
wenzelm@30397
   350
  Note that internalizing the output again in the current context may
wenzelm@30397
   351
  well yield a different result.
wenzelm@27043
   352
wenzelm@30397
   353
  \item @{antiquotation_option_def unique_names}~@{text "= bool"}
wenzelm@30397
   354
  determines whether the printed version of qualified names should be
wenzelm@30397
   355
  made sufficiently long to avoid overlap with names declared further
wenzelm@30397
   356
  back.  Set to @{text false} for more concise output.
wenzelm@27043
   357
wenzelm@30397
   358
  \item @{antiquotation_option_def eta_contract}~@{text "= bool"}
wenzelm@30397
   359
  prints terms in @{text \<eta>}-contracted form.
wenzelm@27043
   360
wenzelm@30397
   361
  \item @{antiquotation_option_def display}~@{text "= bool"} indicates
wenzelm@30397
   362
  if the text is to be output as multi-line ``display material'',
wenzelm@30397
   363
  rather than a small piece of text without line breaks (which is the
wenzelm@30397
   364
  default).
wenzelm@27043
   365
wenzelm@28749
   366
  In this mode the embedded entities are printed in the same style as
wenzelm@28749
   367
  the main theory text.
wenzelm@28749
   368
wenzelm@30397
   369
  \item @{antiquotation_option_def break}~@{text "= bool"} controls
wenzelm@30397
   370
  line breaks in non-display material.
wenzelm@27043
   371
wenzelm@30397
   372
  \item @{antiquotation_option_def quotes}~@{text "= bool"} indicates
wenzelm@30397
   373
  if the output should be enclosed in double quotes.
wenzelm@27043
   374
wenzelm@30397
   375
  \item @{antiquotation_option_def mode}~@{text "= name"} adds @{text
wenzelm@30397
   376
  name} to the print mode to be used for presentation.  Note that the
wenzelm@30397
   377
  standard setup for {\LaTeX} output is already present by default,
wenzelm@30397
   378
  including the modes @{text latex} and @{text xsymbols}.
wenzelm@27043
   379
wenzelm@30397
   380
  \item @{antiquotation_option_def margin}~@{text "= nat"} and
wenzelm@30397
   381
  @{antiquotation_option_def indent}~@{text "= nat"} change the margin
wenzelm@30397
   382
  or indentation for pretty printing of display material.
wenzelm@27043
   383
wenzelm@30397
   384
  \item @{antiquotation_option_def goals_limit}~@{text "= nat"}
wenzelm@30397
   385
  determines the maximum number of goals to be printed (for goal-based
wenzelm@30397
   386
  antiquotation).
wenzelm@27043
   387
wenzelm@30397
   388
  \item @{antiquotation_option_def source}~@{text "= bool"} prints the
wenzelm@30397
   389
  original source text of the antiquotation arguments, rather than its
wenzelm@30397
   390
  internal representation.  Note that formal checking of
wenzelm@30397
   391
  @{antiquotation "thm"}, @{antiquotation "term"}, etc. is still
wenzelm@30397
   392
  enabled; use the @{antiquotation "text"} antiquotation for unchecked
wenzelm@30397
   393
  output.
wenzelm@28749
   394
wenzelm@28749
   395
  Regular @{text "term"} and @{text "typ"} antiquotations with @{text
wenzelm@28749
   396
  "source = false"} involve a full round-trip from the original source
wenzelm@28749
   397
  to an internalized logical entity back to a source form, according
wenzelm@28749
   398
  to the syntax of the current context.  Thus the printed output is
wenzelm@28749
   399
  not under direct control of the author, it may even fluctuate a bit
wenzelm@28749
   400
  as the underlying theory is changed later on.
wenzelm@28749
   401
wenzelm@30397
   402
  In contrast, @{antiquotation_option_def source}~@{text "= true"}
wenzelm@30397
   403
  admits direct printing of the given source text, with the desirable
wenzelm@30397
   404
  well-formedness check in the background, but without modification of
wenzelm@30397
   405
  the printed text.
wenzelm@28749
   406
wenzelm@28760
   407
  \end{description}
wenzelm@27043
   408
wenzelm@27043
   409
  For boolean flags, ``@{text "name = true"}'' may be abbreviated as
wenzelm@27043
   410
  ``@{text name}''.  All of the above flags are disabled by default,
wenzelm@28749
   411
  unless changed from ML, say in the @{verbatim "ROOT.ML"} of the
wenzelm@28749
   412
  logic session.
wenzelm@27043
   413
*}
wenzelm@27043
   414
wenzelm@27043
   415
wenzelm@28750
   416
section {* Markup via command tags \label{sec:tags} *}
wenzelm@27043
   417
wenzelm@28750
   418
text {* Each Isabelle/Isar command may be decorated by additional
wenzelm@28750
   419
  presentation tags, to indicate some modification in the way it is
wenzelm@28750
   420
  printed in the document.
wenzelm@27043
   421
wenzelm@43467
   422
  @{rail "
wenzelm@43467
   423
    @{syntax_def tags}: ( tag * )
wenzelm@27043
   424
    ;
wenzelm@43467
   425
    tag: '%' (@{syntax ident} | @{syntax string})
wenzelm@43467
   426
  "}
wenzelm@27043
   427
wenzelm@28750
   428
  Some tags are pre-declared for certain classes of commands, serving
wenzelm@28750
   429
  as default markup if no tags are given in the text:
wenzelm@27043
   430
wenzelm@28750
   431
  \medskip
wenzelm@27043
   432
  \begin{tabular}{ll}
wenzelm@27043
   433
    @{text "theory"} & theory begin/end \\
wenzelm@27043
   434
    @{text "proof"} & all proof commands \\
wenzelm@27043
   435
    @{text "ML"} & all commands involving ML code \\
wenzelm@27043
   436
  \end{tabular}
wenzelm@27043
   437
wenzelm@28750
   438
  \medskip The Isabelle document preparation system
wenzelm@28750
   439
  \cite{isabelle-sys} allows tagged command regions to be presented
wenzelm@27043
   440
  specifically, e.g.\ to fold proof texts, or drop parts of the text
wenzelm@27043
   441
  completely.
wenzelm@27043
   442
wenzelm@28750
   443
  For example ``@{command "by"}~@{text "%invisible auto"}'' causes
wenzelm@28750
   444
  that piece of proof to be treated as @{text invisible} instead of
wenzelm@28750
   445
  @{text "proof"} (the default), which may be shown or hidden
wenzelm@27043
   446
  depending on the document setup.  In contrast, ``@{command
wenzelm@28750
   447
  "by"}~@{text "%visible auto"}'' forces this text to be shown
wenzelm@27043
   448
  invariably.
wenzelm@27043
   449
wenzelm@27043
   450
  Explicit tag specifications within a proof apply to all subsequent
wenzelm@27043
   451
  commands of the same level of nesting.  For example, ``@{command
wenzelm@28750
   452
  "proof"}~@{text "%visible \<dots>"}~@{command "qed"}'' forces the whole
wenzelm@28750
   453
  sub-proof to be typeset as @{text visible} (unless some of its parts
wenzelm@28750
   454
  are tagged differently).
wenzelm@28750
   455
wenzelm@28750
   456
  \medskip Command tags merely produce certain markup environments for
wenzelm@28750
   457
  type-setting.  The meaning of these is determined by {\LaTeX}
wenzelm@41048
   458
  macros, as defined in @{file "~~/lib/texinputs/isabelle.sty"} or
wenzelm@28750
   459
  by the document author.  The Isabelle document preparation tools
wenzelm@28750
   460
  also provide some high-level options to specify the meaning of
wenzelm@28750
   461
  arbitrary tags to ``keep'', ``drop'', or ``fold'' the corresponding
wenzelm@28750
   462
  parts of the text.  Logic sessions may also specify ``document
wenzelm@28750
   463
  versions'', where given tags are interpreted in some particular way.
wenzelm@28750
   464
  Again see \cite{isabelle-sys} for further details.
wenzelm@27043
   465
*}
wenzelm@27043
   466
wenzelm@27043
   467
wenzelm@27043
   468
section {* Draft presentation *}
wenzelm@27043
   469
wenzelm@27043
   470
text {*
wenzelm@27043
   471
  \begin{matharray}{rcl}
wenzelm@28761
   472
    @{command_def "display_drafts"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
wenzelm@28761
   473
    @{command_def "print_drafts"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
wenzelm@27043
   474
  \end{matharray}
wenzelm@27043
   475
wenzelm@43467
   476
  @{rail "
wenzelm@43467
   477
    (@@{command display_drafts} | @@{command print_drafts}) (@{syntax name} +)
wenzelm@43467
   478
wenzelm@43467
   479
  "}
wenzelm@27043
   480
wenzelm@28760
   481
  \begin{description}
wenzelm@27043
   482
wenzelm@28760
   483
  \item @{command "display_drafts"}~@{text paths} and @{command
wenzelm@28760
   484
  "print_drafts"}~@{text paths} perform simple output of a given list
wenzelm@27043
   485
  of raw source files.  Only those symbols that do not require
wenzelm@27043
   486
  additional {\LaTeX} packages are displayed properly, everything else
wenzelm@27043
   487
  is left verbatim.
wenzelm@27043
   488
wenzelm@28760
   489
  \end{description}
wenzelm@27043
   490
*}
wenzelm@27043
   491
wenzelm@27043
   492
end