doc-src/IsarRef/Thy/Document_Preparation.thy
author wenzelm
Tue, 03 May 2011 18:04:05 +0200
changeset 43529 8f5d5d71add0
parent 43522 e3fdb7c96be5
child 43542 04dfffda5671
permissions -rw-r--r--
some documentation of @{rail} antiquotation;
     1 theory Document_Preparation
     2 imports Base Main
     3 begin
     4 
     5 chapter {* Document preparation \label{ch:document-prep} *}
     6 
     7 text {* Isabelle/Isar provides a simple document preparation system
     8   based on regular {PDF-\LaTeX} technology, with full support for
     9   hyper-links and bookmarks.  Thus the results are well suited for WWW
    10   browsing and as printed copies.
    11 
    12   \medskip Isabelle generates {\LaTeX} output while running a
    13   \emph{logic session} (see also \cite{isabelle-sys}).  Getting
    14   started with a working configuration for common situations is quite
    15   easy by using the Isabelle @{verbatim mkdir} and @{verbatim make}
    16   tools.  First invoke
    17 \begin{ttbox}
    18   isabelle mkdir Foo
    19 \end{ttbox}
    20   to initialize a separate directory for session @{verbatim Foo} (it
    21   is safe to experiment, since @{verbatim "isabelle mkdir"} never
    22   overwrites existing files).  Ensure that @{verbatim "Foo/ROOT.ML"}
    23   holds ML commands to load all theories required for this session;
    24   furthermore @{verbatim "Foo/document/root.tex"} should include any
    25   special {\LaTeX} macro packages required for your document (the
    26   default is usually sufficient as a start).
    27 
    28   The session is controlled by a separate @{verbatim IsaMakefile}
    29   (with crude source dependencies by default).  This file is located
    30   one level up from the @{verbatim Foo} directory location.  Now
    31   invoke
    32 \begin{ttbox}
    33   isabelle make Foo
    34 \end{ttbox}
    35   to run the @{verbatim Foo} session, with browser information and
    36   document preparation enabled.  Unless any errors are reported by
    37   Isabelle or {\LaTeX}, the output will appear inside the directory
    38   defined by the @{verbatim ISABELLE_BROWSER_INFO} setting (as
    39   reported by the batch job in verbose mode).
    40 
    41   \medskip You may also consider to tune the @{verbatim usedir}
    42   options in @{verbatim IsaMakefile}, for example to switch the output
    43   format between @{verbatim pdf} and @{verbatim dvi}, or activate the
    44   @{verbatim "-D"} option to retain a second copy of the generated
    45   {\LaTeX} sources (for manual inspection or separate runs of
    46   @{executable latex}).
    47 
    48   \medskip See \emph{The Isabelle System Manual} \cite{isabelle-sys}
    49   for further details on Isabelle logic sessions and theory
    50   presentation.  The Isabelle/HOL tutorial \cite{isabelle-hol-book}
    51   also covers theory presentation to some extent.
    52 *}
    53 
    54 
    55 section {* Markup commands \label{sec:markup} *}
    56 
    57 text {*
    58   \begin{matharray}{rcl}
    59     @{command_def "header"} & : & @{text "toplevel \<rightarrow> toplevel"} \\[0.5ex]
    60     @{command_def "chapter"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
    61     @{command_def "section"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
    62     @{command_def "subsection"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
    63     @{command_def "subsubsection"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
    64     @{command_def "text"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
    65     @{command_def "text_raw"} & : & @{text "local_theory \<rightarrow> local_theory"} \\[0.5ex]
    66     @{command_def "sect"} & : & @{text "proof \<rightarrow> proof"} \\
    67     @{command_def "subsect"} & : & @{text "proof \<rightarrow> proof"} \\
    68     @{command_def "subsubsect"} & : & @{text "proof \<rightarrow> proof"} \\
    69     @{command_def "txt"} & : & @{text "proof \<rightarrow> proof"} \\
    70     @{command_def "txt_raw"} & : & @{text "proof \<rightarrow> proof"} \\
    71   \end{matharray}
    72 
    73   Markup commands provide a structured way to insert text into the
    74   document generated from a theory.  Each markup command takes a
    75   single @{syntax text} argument, which is passed as argument to a
    76   corresponding {\LaTeX} macro.  The default macros provided by
    77   @{file "~~/lib/texinputs/isabelle.sty"} can be redefined according
    78   to the needs of the underlying document and {\LaTeX} styles.
    79 
    80   Note that formal comments (\secref{sec:comments}) are similar to
    81   markup commands, but have a different status within Isabelle/Isar
    82   syntax.
    83 
    84   @{rail "
    85     (@@{command chapter} | @@{command section} | @@{command subsection} |
    86       @@{command subsubsection} | @@{command text}) @{syntax target}? @{syntax text}
    87     ;
    88     (@@{command header} | @@{command text_raw} | @@{command sect} | @@{command subsect} |
    89       @@{command subsubsect} | @@{command txt} | @@{command txt_raw}) @{syntax text}
    90   "}
    91 
    92   \begin{description}
    93 
    94   \item @{command header} provides plain text markup just preceding
    95   the formal beginning of a theory.  The corresponding {\LaTeX} macro
    96   is @{verbatim "\\isamarkupheader"}, which acts like @{command
    97   section} by default.
    98   
    99   \item @{command chapter}, @{command section}, @{command subsection},
   100   and @{command subsubsection} mark chapter and section headings
   101   within the main theory body or local theory targets.  The
   102   corresponding {\LaTeX} macros are @{verbatim "\\isamarkupchapter"},
   103   @{verbatim "\\isamarkupsection"}, @{verbatim
   104   "\\isamarkupsubsection"} etc.
   105 
   106   \item @{command sect}, @{command subsect}, and @{command subsubsect}
   107   mark section headings within proofs.  The corresponding {\LaTeX}
   108   macros are @{verbatim "\\isamarkupsect"}, @{verbatim
   109   "\\isamarkupsubsect"} etc.
   110 
   111   \item @{command text} and @{command txt} specify paragraphs of plain
   112   text.  This corresponds to a {\LaTeX} environment @{verbatim
   113   "\\begin{isamarkuptext}"} @{text "\<dots>"} @{verbatim
   114   "\\end{isamarkuptext}"} etc.
   115 
   116   \item @{command text_raw} and @{command txt_raw} insert {\LaTeX}
   117   source into the output, without additional markup.  Thus the full
   118   range of document manipulations becomes available, at the risk of
   119   messing up document output.
   120 
   121   \end{description}
   122 
   123   Except for @{command "text_raw"} and @{command "txt_raw"}, the text
   124   passed to any of the above markup commands may refer to formal
   125   entities via \emph{document antiquotations}, see also
   126   \secref{sec:antiq}.  These are interpreted in the present theory or
   127   proof context, or the named @{text "target"}.
   128 
   129   \medskip The proof markup commands closely resemble those for theory
   130   specifications, but have a different formal status and produce
   131   different {\LaTeX} macros.  The default definitions coincide for
   132   analogous commands such as @{command section} and @{command sect}.
   133 *}
   134 
   135 
   136 section {* Document Antiquotations \label{sec:antiq} *}
   137 
   138 text {*
   139   \begin{matharray}{rcl}
   140     @{antiquotation_def "theory"} & : & @{text antiquotation} \\
   141     @{antiquotation_def "thm"} & : & @{text antiquotation} \\
   142     @{antiquotation_def "lemma"} & : & @{text antiquotation} \\
   143     @{antiquotation_def "prop"} & : & @{text antiquotation} \\
   144     @{antiquotation_def "term"} & : & @{text antiquotation} \\
   145     @{antiquotation_def term_type} & : & @{text antiquotation} \\
   146     @{antiquotation_def typeof} & : & @{text antiquotation} \\
   147     @{antiquotation_def const} & : & @{text antiquotation} \\
   148     @{antiquotation_def abbrev} & : & @{text antiquotation} \\
   149     @{antiquotation_def typ} & : & @{text antiquotation} \\
   150     @{antiquotation_def type} & : & @{text antiquotation} \\
   151     @{antiquotation_def class} & : & @{text antiquotation} \\
   152     @{antiquotation_def "text"} & : & @{text antiquotation} \\
   153     @{antiquotation_def goals} & : & @{text antiquotation} \\
   154     @{antiquotation_def subgoals} & : & @{text antiquotation} \\
   155     @{antiquotation_def prf} & : & @{text antiquotation} \\
   156     @{antiquotation_def full_prf} & : & @{text antiquotation} \\
   157     @{antiquotation_def ML} & : & @{text antiquotation} \\
   158     @{antiquotation_def ML_type} & : & @{text antiquotation} \\
   159     @{antiquotation_def ML_struct} & : & @{text antiquotation} \\
   160     @{antiquotation_def "file"} & : & @{text antiquotation} \\
   161   \end{matharray}
   162 
   163   The overall content of an Isabelle/Isar theory may alternate between
   164   formal and informal text.  The main body consists of formal
   165   specification and proof commands, interspersed with markup commands
   166   (\secref{sec:markup}) or document comments (\secref{sec:comments}).
   167   The argument of markup commands quotes informal text to be printed
   168   in the resulting document, but may again refer to formal entities
   169   via \emph{document antiquotations}.
   170 
   171   For example, embedding of ``@{text [source=false] "@{term [show_types] \"f x = a + x\"}"}''
   172   within a text block makes
   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.
   174 
   175   Antiquotations usually spare the author tedious typing of logical
   176   entities in full detail.  Even more importantly, some degree of
   177   consistency-checking between the main body of formal text and its
   178   informal explanation is achieved, since terms and types appearing in
   179   antiquotations are checked within the current theory or proof
   180   context.
   181 
   182   @{rail "
   183     '@{' antiquotation '}'
   184     ;
   185     @{syntax_def antiquotation}:
   186       @@{antiquotation theory} options @{syntax name} |
   187       @@{antiquotation thm} options styles @{syntax thmrefs} |
   188       @@{antiquotation lemma} options @{syntax prop} @'by' @{syntax method} @{syntax method}? |
   189       @@{antiquotation prop} options styles @{syntax prop} |
   190       @@{antiquotation term} options styles @{syntax term} |
   191       @@{antiquotation term_type} options styles @{syntax term} |
   192       @@{antiquotation typeof} options styles @{syntax term} |
   193       @@{antiquotation const} options @{syntax term} |
   194       @@{antiquotation abbrev} options @{syntax term} |
   195       @@{antiquotation typ} options @{syntax type} |
   196       @@{antiquotation type} options @{syntax name} |
   197       @@{antiquotation class} options @{syntax name} |
   198       @@{antiquotation text} options @{syntax name} |
   199       @@{antiquotation goals} options |
   200       @@{antiquotation subgoals} options |
   201       @@{antiquotation prf} options @{syntax thmrefs} |
   202       @@{antiquotation full_prf} options @{syntax thmrefs} |
   203       @@{antiquotation ML} options @{syntax name} |
   204       @@{antiquotation ML_type} options @{syntax name} |
   205       @@{antiquotation ML_struct} options @{syntax name} |
   206       @@{antiquotation \"file\"} options @{syntax name}
   207     ;
   208     options: '[' (option * ',') ']'
   209     ;
   210     option: @{syntax name} | @{syntax name} '=' @{syntax name}
   211     ;
   212     styles: '(' (style + ',') ')'
   213     ;
   214     style: (@{syntax name} +)
   215   "}
   216 
   217   Note that the syntax of antiquotations may \emph{not} include source
   218   comments @{verbatim "(*"}~@{text "\<dots>"}~@{verbatim "*)"} nor verbatim
   219   text @{verbatim "{"}@{verbatim "*"}~@{text "\<dots>"}~@{verbatim
   220   "*"}@{verbatim "}"}.
   221 
   222   \begin{description}
   223   
   224   \item @{text "@{theory A}"} prints the name @{text "A"}, which is
   225   guaranteed to refer to a valid ancestor theory in the current
   226   context.
   227 
   228   \item @{text "@{thm a\<^sub>1 \<dots> a\<^sub>n}"} prints theorems @{text "a\<^sub>1 \<dots> a\<^sub>n"}.
   229   Full fact expressions are allowed here, including attributes
   230   (\secref{sec:syn-att}).
   231 
   232   \item @{text "@{prop \<phi>}"} prints a well-typed proposition @{text
   233   "\<phi>"}.
   234 
   235   \item @{text "@{lemma \<phi> by m}"} proves a well-typed proposition
   236   @{text "\<phi>"} by method @{text m} and prints the original @{text "\<phi>"}.
   237 
   238   \item @{text "@{term t}"} prints a well-typed term @{text "t"}.
   239 
   240   \item @{text "@{term_type t}"} prints a well-typed term @{text "t"}
   241   annotated with its type.
   242 
   243   \item @{text "@{typeof t}"} prints the type of a well-typed term
   244   @{text "t"}.
   245 
   246   \item @{text "@{const c}"} prints a logical or syntactic constant
   247   @{text "c"}.
   248   
   249   \item @{text "@{abbrev c x\<^sub>1 \<dots> x\<^sub>n}"} prints a constant abbreviation
   250   @{text "c x\<^sub>1 \<dots> x\<^sub>n \<equiv> rhs"} as defined in the current context.
   251 
   252   \item @{text "@{typ \<tau>}"} prints a well-formed type @{text "\<tau>"}.
   253 
   254   \item @{text "@{type \<kappa>}"} prints a (logical or syntactic) type
   255     constructor @{text "\<kappa>"}.
   256 
   257   \item @{text "@{class c}"} prints a class @{text c}.
   258 
   259   \item @{text "@{text s}"} prints uninterpreted source text @{text
   260   s}.  This is particularly useful to print portions of text according
   261   to the Isabelle document style, without demanding well-formedness,
   262   e.g.\ small pieces of terms that should not be parsed or
   263   type-checked yet.
   264 
   265   \item @{text "@{goals}"} prints the current \emph{dynamic} goal
   266   state.  This is mainly for support of tactic-emulation scripts
   267   within Isar.  Presentation of goal states does not conform to the
   268   idea of human-readable proof documents!
   269 
   270   When explaining proofs in detail it is usually better to spell out
   271   the reasoning via proper Isar proof commands, instead of peeking at
   272   the internal machine configuration.
   273   
   274   \item @{text "@{subgoals}"} is similar to @{text "@{goals}"}, but
   275   does not print the main goal.
   276   
   277   \item @{text "@{prf a\<^sub>1 \<dots> a\<^sub>n}"} prints the (compact) proof terms
   278   corresponding to the theorems @{text "a\<^sub>1 \<dots> a\<^sub>n"}. Note that this
   279   requires proof terms to be switched on for the current logic
   280   session.
   281   
   282   \item @{text "@{full_prf a\<^sub>1 \<dots> a\<^sub>n}"} is like @{text "@{prf a\<^sub>1 \<dots>
   283   a\<^sub>n}"}, but prints the full proof terms, i.e.\ also displays
   284   information omitted in the compact proof term, which is denoted by
   285   ``@{text _}'' placeholders there.
   286   
   287   \item @{text "@{ML s}"}, @{text "@{ML_type s}"}, and @{text
   288   "@{ML_struct s}"} check text @{text s} as ML value, type, and
   289   structure, respectively.  The source is printed verbatim.
   290 
   291   \item @{text "@{file path}"} checks that @{text "path"} refers to a
   292   file (or directory) and prints it verbatim.
   293 
   294   \end{description}
   295 *}
   296 
   297 
   298 subsubsection {* Styled antiquotations *}
   299 
   300 text {* The antiquotations @{text thm}, @{text prop} and @{text
   301   term} admit an extra \emph{style} specification to modify the
   302   printed result.  A style is specified by a name with a possibly
   303   empty number of arguments;  multiple styles can be sequenced with
   304   commas.  The following standard styles are available:
   305 
   306   \begin{description}
   307   
   308   \item @{text lhs} extracts the first argument of any application
   309   form with at least two arguments --- typically meta-level or
   310   object-level equality, or any other binary relation.
   311   
   312   \item @{text rhs} is like @{text lhs}, but extracts the second
   313   argument.
   314   
   315   \item @{text "concl"} extracts the conclusion @{text C} from a rule
   316   in Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}.
   317   
   318   \item @{text "prem"} @{text n} extract premise number
   319   @{text "n"} from from a rule in Horn-clause
   320   normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}
   321 
   322   \end{description}
   323 *}
   324 
   325 
   326 subsubsection {* General options *}
   327 
   328 text {* The following options are available to tune the printed output
   329   of antiquotations.  Note that many of these coincide with global ML
   330   flags of the same names.
   331 
   332   \begin{description}
   333 
   334   \item @{antiquotation_option_def show_types}~@{text "= bool"} and
   335   @{antiquotation_option_def show_sorts}~@{text "= bool"} control
   336   printing of explicit type and sort constraints.
   337 
   338   \item @{antiquotation_option_def show_structs}~@{text "= bool"}
   339   controls printing of implicit structures.
   340 
   341   \item @{antiquotation_option_def show_abbrevs}~@{text "= bool"}
   342   controls folding of abbreviations.
   343 
   344   \item @{antiquotation_option_def long_names}~@{text "= bool"} forces
   345   names of types and constants etc.\ to be printed in their fully
   346   qualified internal form.
   347 
   348   \item @{antiquotation_option_def short_names}~@{text "= bool"}
   349   forces names of types and constants etc.\ to be printed unqualified.
   350   Note that internalizing the output again in the current context may
   351   well yield a different result.
   352 
   353   \item @{antiquotation_option_def unique_names}~@{text "= bool"}
   354   determines whether the printed version of qualified names should be
   355   made sufficiently long to avoid overlap with names declared further
   356   back.  Set to @{text false} for more concise output.
   357 
   358   \item @{antiquotation_option_def eta_contract}~@{text "= bool"}
   359   prints terms in @{text \<eta>}-contracted form.
   360 
   361   \item @{antiquotation_option_def display}~@{text "= bool"} indicates
   362   if the text is to be output as multi-line ``display material'',
   363   rather than a small piece of text without line breaks (which is the
   364   default).
   365 
   366   In this mode the embedded entities are printed in the same style as
   367   the main theory text.
   368 
   369   \item @{antiquotation_option_def break}~@{text "= bool"} controls
   370   line breaks in non-display material.
   371 
   372   \item @{antiquotation_option_def quotes}~@{text "= bool"} indicates
   373   if the output should be enclosed in double quotes.
   374 
   375   \item @{antiquotation_option_def mode}~@{text "= name"} adds @{text
   376   name} to the print mode to be used for presentation.  Note that the
   377   standard setup for {\LaTeX} output is already present by default,
   378   including the modes @{text latex} and @{text xsymbols}.
   379 
   380   \item @{antiquotation_option_def margin}~@{text "= nat"} and
   381   @{antiquotation_option_def indent}~@{text "= nat"} change the margin
   382   or indentation for pretty printing of display material.
   383 
   384   \item @{antiquotation_option_def goals_limit}~@{text "= nat"}
   385   determines the maximum number of goals to be printed (for goal-based
   386   antiquotation).
   387 
   388   \item @{antiquotation_option_def source}~@{text "= bool"} prints the
   389   original source text of the antiquotation arguments, rather than its
   390   internal representation.  Note that formal checking of
   391   @{antiquotation "thm"}, @{antiquotation "term"}, etc. is still
   392   enabled; use the @{antiquotation "text"} antiquotation for unchecked
   393   output.
   394 
   395   Regular @{text "term"} and @{text "typ"} antiquotations with @{text
   396   "source = false"} involve a full round-trip from the original source
   397   to an internalized logical entity back to a source form, according
   398   to the syntax of the current context.  Thus the printed output is
   399   not under direct control of the author, it may even fluctuate a bit
   400   as the underlying theory is changed later on.
   401 
   402   In contrast, @{antiquotation_option source}~@{text "= true"}
   403   admits direct printing of the given source text, with the desirable
   404   well-formedness check in the background, but without modification of
   405   the printed text.
   406 
   407   \end{description}
   408 
   409   For boolean flags, ``@{text "name = true"}'' may be abbreviated as
   410   ``@{text name}''.  All of the above flags are disabled by default,
   411   unless changed from ML, say in the @{verbatim "ROOT.ML"} of the
   412   logic session.
   413 *}
   414 
   415 
   416 section {* Markup via command tags \label{sec:tags} *}
   417 
   418 text {* Each Isabelle/Isar command may be decorated by additional
   419   presentation tags, to indicate some modification in the way it is
   420   printed in the document.
   421 
   422   @{rail "
   423     @{syntax_def tags}: ( tag * )
   424     ;
   425     tag: '%' (@{syntax ident} | @{syntax string})
   426   "}
   427 
   428   Some tags are pre-declared for certain classes of commands, serving
   429   as default markup if no tags are given in the text:
   430 
   431   \medskip
   432   \begin{tabular}{ll}
   433     @{text "theory"} & theory begin/end \\
   434     @{text "proof"} & all proof commands \\
   435     @{text "ML"} & all commands involving ML code \\
   436   \end{tabular}
   437 
   438   \medskip The Isabelle document preparation system
   439   \cite{isabelle-sys} allows tagged command regions to be presented
   440   specifically, e.g.\ to fold proof texts, or drop parts of the text
   441   completely.
   442 
   443   For example ``@{command "by"}~@{text "%invisible auto"}'' causes
   444   that piece of proof to be treated as @{text invisible} instead of
   445   @{text "proof"} (the default), which may be shown or hidden
   446   depending on the document setup.  In contrast, ``@{command
   447   "by"}~@{text "%visible auto"}'' forces this text to be shown
   448   invariably.
   449 
   450   Explicit tag specifications within a proof apply to all subsequent
   451   commands of the same level of nesting.  For example, ``@{command
   452   "proof"}~@{text "%visible \<dots>"}~@{command "qed"}'' forces the whole
   453   sub-proof to be typeset as @{text visible} (unless some of its parts
   454   are tagged differently).
   455 
   456   \medskip Command tags merely produce certain markup environments for
   457   type-setting.  The meaning of these is determined by {\LaTeX}
   458   macros, as defined in @{file "~~/lib/texinputs/isabelle.sty"} or
   459   by the document author.  The Isabelle document preparation tools
   460   also provide some high-level options to specify the meaning of
   461   arbitrary tags to ``keep'', ``drop'', or ``fold'' the corresponding
   462   parts of the text.  Logic sessions may also specify ``document
   463   versions'', where given tags are interpreted in some particular way.
   464   Again see \cite{isabelle-sys} for further details.
   465 *}
   466 
   467 
   468 section {* Railroad diagrams *}
   469 
   470 text {*
   471   \begin{matharray}{rcl}
   472     @{antiquotation_def "rail"} & : & @{text antiquotation} \\
   473   \end{matharray}
   474 
   475   @{rail "'rail' string"}
   476 
   477   The @{antiquotation rail} antiquotation allows to include syntax
   478   diagrams into Isabelle documents.  {\LaTeX} requires the style file
   479   @{"file" "~~/lib/texinputs/pdfsetup.sty"}, which can be used via
   480   @{verbatim "\\usepackage{pdfsetup}"} in @{verbatim "root.tex"}, for
   481   example.
   482 
   483   The rail specification language is quoted here as Isabelle @{syntax
   484   string}; it has its own grammar given below.
   485 
   486   @{rail "
   487   rule? + ';'
   488   ;
   489   rule: ((identifier | @{syntax antiquotation}) ':')? body
   490   ;
   491   body: concatenation + '|'
   492   ;
   493   concatenation: ((atom '?'?) +) (('*' | '+') atom?)?
   494   ;
   495   atom: '(' body? ')' | identifier |
   496     '@'? (string | @{syntax antiquotation}) |
   497     '\\\\\\\\'
   498   "}
   499 
   500   The lexical syntax of @{text "identifier"} coincides with that of
   501   @{syntax ident} in regular Isabelle syntax, but @{text string} uses
   502   single quotes instead of double quotes of the standard @{syntax
   503   string} category, to avoid extra escapes.
   504 
   505   Each @{text rule} defines a formal language (with optional name),
   506   using a notation that is similar to EBNF or regular expressions with
   507   recursion.  The meaning and visual appearance of these rail language
   508   elements is illustrated by the following representative examples.
   509 
   510   \begin{itemize}
   511 
   512   \item Empty @{verbatim "()"}
   513 
   514   @{rail "()"}
   515 
   516   \item Nonterminal @{verbatim "A"}
   517 
   518   @{rail "A"}
   519 
   520   \item Nonterminal via Isabelle antiquotation
   521   @{verbatim "@{syntax method}"}
   522 
   523   @{rail "@{syntax method}"}
   524 
   525   \item Terminal @{verbatim "'xyz'"}
   526 
   527   @{rail "'xyz'"}
   528 
   529   \item Terminal in keyword style @{verbatim "@'xyz'"}
   530 
   531   @{rail "@'xyz'"}
   532 
   533   \item Terminal via Isabelle antiquotation
   534   @{verbatim "@@{method rule}"}
   535 
   536   @{rail "@@{method rule}"}
   537 
   538   \item Concatenation @{verbatim "A B C"}
   539 
   540   @{rail "A B C"}
   541 
   542   \item Linebreak @{verbatim "\\\\"} inside
   543   concatenation\footnote{Strictly speaking, this is only a single
   544   backslash, but the enclosing @{syntax string} syntax requires a
   545   second one for escaping.} @{verbatim "A B C \\\\ D E F"}
   546 
   547   @{rail "A B C \\ D E F"}
   548 
   549   \item Variants @{verbatim "A | B | C"}
   550 
   551   @{rail "A | B | C"}
   552 
   553   \item Option @{verbatim "A ?"}
   554 
   555   @{rail "A ?"}
   556 
   557   \item Repetition @{verbatim "A *"}
   558 
   559   @{rail "A *"}
   560 
   561   \item Repetition with separator @{verbatim "A * sep"}
   562 
   563   @{rail "A * sep"}
   564 
   565   \item Strict repetition @{verbatim "A +"}
   566 
   567   @{rail "A +"}
   568 
   569   \item Strict repetition with separator @{verbatim "A + sep"}
   570 
   571   @{rail "A + sep"}
   572 
   573   \end{itemize}
   574 *}
   575 
   576 
   577 section {* Draft presentation *}
   578 
   579 text {*
   580   \begin{matharray}{rcl}
   581     @{command_def "display_drafts"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
   582     @{command_def "print_drafts"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
   583   \end{matharray}
   584 
   585   @{rail "
   586     (@@{command display_drafts} | @@{command print_drafts}) (@{syntax name} +)
   587 
   588   "}
   589 
   590   \begin{description}
   591 
   592   \item @{command "display_drafts"}~@{text paths} and @{command
   593   "print_drafts"}~@{text paths} perform simple output of a given list
   594   of raw source files.  Only those symbols that do not require
   595   additional {\LaTeX} packages are displayed properly, everything else
   596   is left verbatim.
   597 
   598   \end{description}
   599 *}
   600 
   601 end