doc-src/IsarRef/Thy/Document_Preparation.thy
author bulwahn
Wed, 27 Jul 2011 20:28:00 +0200
changeset 44864 b141d7a3d4e3
parent 44499 1c43134ff988
child 47292 b03897da3c90
permissions -rw-r--r--
rudimentary documentation of the quotient package in the isar reference manual
     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   %% FIXME less monolithic presentation, move to individual sections!?
   183   @{rail "
   184     '@{' antiquotation '}'
   185     ;
   186     @{syntax_def antiquotation}:
   187       @@{antiquotation theory} options @{syntax name} |
   188       @@{antiquotation thm} options styles @{syntax thmrefs} |
   189       @@{antiquotation lemma} options @{syntax prop} @'by' @{syntax method} @{syntax method}? |
   190       @@{antiquotation prop} options styles @{syntax prop} |
   191       @@{antiquotation term} options styles @{syntax term} |
   192       @@{antiquotation (HOL) value} options styles @{syntax term} |
   193       @@{antiquotation term_type} options styles @{syntax term} |
   194       @@{antiquotation typeof} options styles @{syntax term} |
   195       @@{antiquotation const} options @{syntax term} |
   196       @@{antiquotation abbrev} options @{syntax term} |
   197       @@{antiquotation typ} options @{syntax type} |
   198       @@{antiquotation type} options @{syntax name} |
   199       @@{antiquotation class} options @{syntax name} |
   200       @@{antiquotation text} options @{syntax name} |
   201       @@{antiquotation goals} options |
   202       @@{antiquotation subgoals} options |
   203       @@{antiquotation prf} options @{syntax thmrefs} |
   204       @@{antiquotation full_prf} options @{syntax thmrefs} |
   205       @@{antiquotation ML} options @{syntax name} |
   206       @@{antiquotation ML_type} options @{syntax name} |
   207       @@{antiquotation ML_struct} options @{syntax name} |
   208       @@{antiquotation \"file\"} options @{syntax name}
   209     ;
   210     options: '[' (option * ',') ']'
   211     ;
   212     option: @{syntax name} | @{syntax name} '=' @{syntax name}
   213     ;
   214     styles: '(' (style + ',') ')'
   215     ;
   216     style: (@{syntax name} +)
   217   "}
   218 
   219   Note that the syntax of antiquotations may \emph{not} include source
   220   comments @{verbatim "(*"}~@{text "\<dots>"}~@{verbatim "*)"} nor verbatim
   221   text @{verbatim "{"}@{verbatim "*"}~@{text "\<dots>"}~@{verbatim
   222   "*"}@{verbatim "}"}.
   223 
   224   \begin{description}
   225   
   226   \item @{text "@{theory A}"} prints the name @{text "A"}, which is
   227   guaranteed to refer to a valid ancestor theory in the current
   228   context.
   229 
   230   \item @{text "@{thm a\<^sub>1 \<dots> a\<^sub>n}"} prints theorems @{text "a\<^sub>1 \<dots> a\<^sub>n"}.
   231   Full fact expressions are allowed here, including attributes
   232   (\secref{sec:syn-att}).
   233 
   234   \item @{text "@{prop \<phi>}"} prints a well-typed proposition @{text
   235   "\<phi>"}.
   236 
   237   \item @{text "@{lemma \<phi> by m}"} proves a well-typed proposition
   238   @{text "\<phi>"} by method @{text m} and prints the original @{text "\<phi>"}.
   239 
   240   \item @{text "@{term t}"} prints a well-typed term @{text "t"}.
   241   
   242   \item @{text "@{value t}"} evaluates a term @{text "t"} and prints
   243   its result, see also @{command_ref (HOL) value}.
   244 
   245   \item @{text "@{term_type t}"} prints a well-typed term @{text "t"}
   246   annotated with its type.
   247 
   248   \item @{text "@{typeof t}"} prints the type of a well-typed term
   249   @{text "t"}.
   250 
   251   \item @{text "@{const c}"} prints a logical or syntactic constant
   252   @{text "c"}.
   253   
   254   \item @{text "@{abbrev c x\<^sub>1 \<dots> x\<^sub>n}"} prints a constant abbreviation
   255   @{text "c x\<^sub>1 \<dots> x\<^sub>n \<equiv> rhs"} as defined in the current context.
   256 
   257   \item @{text "@{typ \<tau>}"} prints a well-formed type @{text "\<tau>"}.
   258 
   259   \item @{text "@{type \<kappa>}"} prints a (logical or syntactic) type
   260     constructor @{text "\<kappa>"}.
   261 
   262   \item @{text "@{class c}"} prints a class @{text c}.
   263 
   264   \item @{text "@{text s}"} prints uninterpreted source text @{text
   265   s}.  This is particularly useful to print portions of text according
   266   to the Isabelle document style, without demanding well-formedness,
   267   e.g.\ small pieces of terms that should not be parsed or
   268   type-checked yet.
   269 
   270   \item @{text "@{goals}"} prints the current \emph{dynamic} goal
   271   state.  This is mainly for support of tactic-emulation scripts
   272   within Isar.  Presentation of goal states does not conform to the
   273   idea of human-readable proof documents!
   274 
   275   When explaining proofs in detail it is usually better to spell out
   276   the reasoning via proper Isar proof commands, instead of peeking at
   277   the internal machine configuration.
   278   
   279   \item @{text "@{subgoals}"} is similar to @{text "@{goals}"}, but
   280   does not print the main goal.
   281   
   282   \item @{text "@{prf a\<^sub>1 \<dots> a\<^sub>n}"} prints the (compact) proof terms
   283   corresponding to the theorems @{text "a\<^sub>1 \<dots> a\<^sub>n"}. Note that this
   284   requires proof terms to be switched on for the current logic
   285   session.
   286   
   287   \item @{text "@{full_prf a\<^sub>1 \<dots> a\<^sub>n}"} is like @{text "@{prf a\<^sub>1 \<dots>
   288   a\<^sub>n}"}, but prints the full proof terms, i.e.\ also displays
   289   information omitted in the compact proof term, which is denoted by
   290   ``@{text _}'' placeholders there.
   291   
   292   \item @{text "@{ML s}"}, @{text "@{ML_type s}"}, and @{text
   293   "@{ML_struct s}"} check text @{text s} as ML value, type, and
   294   structure, respectively.  The source is printed verbatim.
   295 
   296   \item @{text "@{file path}"} checks that @{text "path"} refers to a
   297   file (or directory) and prints it verbatim.
   298 
   299   \end{description}
   300 *}
   301 
   302 
   303 subsection {* Styled antiquotations *}
   304 
   305 text {* The antiquotations @{text thm}, @{text prop} and @{text
   306   term} admit an extra \emph{style} specification to modify the
   307   printed result.  A style is specified by a name with a possibly
   308   empty number of arguments;  multiple styles can be sequenced with
   309   commas.  The following standard styles are available:
   310 
   311   \begin{description}
   312   
   313   \item @{text lhs} extracts the first argument of any application
   314   form with at least two arguments --- typically meta-level or
   315   object-level equality, or any other binary relation.
   316   
   317   \item @{text rhs} is like @{text lhs}, but extracts the second
   318   argument.
   319   
   320   \item @{text "concl"} extracts the conclusion @{text C} from a rule
   321   in Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}.
   322   
   323   \item @{text "prem"} @{text n} extract premise number
   324   @{text "n"} from from a rule in Horn-clause
   325   normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}
   326 
   327   \end{description}
   328 *}
   329 
   330 
   331 subsection {* General options *}
   332 
   333 text {* The following options are available to tune the printed output
   334   of antiquotations.  Note that many of these coincide with global ML
   335   flags of the same names.
   336 
   337   \begin{description}
   338 
   339   \item @{antiquotation_option_def show_types}~@{text "= bool"} and
   340   @{antiquotation_option_def show_sorts}~@{text "= bool"} control
   341   printing of explicit type and sort constraints.
   342 
   343   \item @{antiquotation_option_def show_structs}~@{text "= bool"}
   344   controls printing of implicit structures.
   345 
   346   \item @{antiquotation_option_def show_abbrevs}~@{text "= bool"}
   347   controls folding of abbreviations.
   348 
   349   \item @{antiquotation_option_def names_long}~@{text "= bool"} forces
   350   names of types and constants etc.\ to be printed in their fully
   351   qualified internal form.
   352 
   353   \item @{antiquotation_option_def names_short}~@{text "= bool"}
   354   forces names of types and constants etc.\ to be printed unqualified.
   355   Note that internalizing the output again in the current context may
   356   well yield a different result.
   357 
   358   \item @{antiquotation_option_def names_unique}~@{text "= bool"}
   359   determines whether the printed version of qualified names should be
   360   made sufficiently long to avoid overlap with names declared further
   361   back.  Set to @{text false} for more concise output.
   362 
   363   \item @{antiquotation_option_def eta_contract}~@{text "= bool"}
   364   prints terms in @{text \<eta>}-contracted form.
   365 
   366   \item @{antiquotation_option_def display}~@{text "= bool"} indicates
   367   if the text is to be output as multi-line ``display material'',
   368   rather than a small piece of text without line breaks (which is the
   369   default).
   370 
   371   In this mode the embedded entities are printed in the same style as
   372   the main theory text.
   373 
   374   \item @{antiquotation_option_def break}~@{text "= bool"} controls
   375   line breaks in non-display material.
   376 
   377   \item @{antiquotation_option_def quotes}~@{text "= bool"} indicates
   378   if the output should be enclosed in double quotes.
   379 
   380   \item @{antiquotation_option_def mode}~@{text "= name"} adds @{text
   381   name} to the print mode to be used for presentation.  Note that the
   382   standard setup for {\LaTeX} output is already present by default,
   383   including the modes @{text latex} and @{text xsymbols}.
   384 
   385   \item @{antiquotation_option_def margin}~@{text "= nat"} and
   386   @{antiquotation_option_def indent}~@{text "= nat"} change the margin
   387   or indentation for pretty printing of display material.
   388 
   389   \item @{antiquotation_option_def goals_limit}~@{text "= nat"}
   390   determines the maximum number of goals to be printed (for goal-based
   391   antiquotation).
   392 
   393   \item @{antiquotation_option_def source}~@{text "= bool"} prints the
   394   original source text of the antiquotation arguments, rather than its
   395   internal representation.  Note that formal checking of
   396   @{antiquotation "thm"}, @{antiquotation "term"}, etc. is still
   397   enabled; use the @{antiquotation "text"} antiquotation for unchecked
   398   output.
   399 
   400   Regular @{text "term"} and @{text "typ"} antiquotations with @{text
   401   "source = false"} involve a full round-trip from the original source
   402   to an internalized logical entity back to a source form, according
   403   to the syntax of the current context.  Thus the printed output is
   404   not under direct control of the author, it may even fluctuate a bit
   405   as the underlying theory is changed later on.
   406 
   407   In contrast, @{antiquotation_option source}~@{text "= true"}
   408   admits direct printing of the given source text, with the desirable
   409   well-formedness check in the background, but without modification of
   410   the printed text.
   411 
   412   \end{description}
   413 
   414   For boolean flags, ``@{text "name = true"}'' may be abbreviated as
   415   ``@{text name}''.  All of the above flags are disabled by default,
   416   unless changed from ML, say in the @{verbatim "ROOT.ML"} of the
   417   logic session.
   418 *}
   419 
   420 
   421 section {* Markup via command tags \label{sec:tags} *}
   422 
   423 text {* Each Isabelle/Isar command may be decorated by additional
   424   presentation tags, to indicate some modification in the way it is
   425   printed in the document.
   426 
   427   @{rail "
   428     @{syntax_def tags}: ( tag * )
   429     ;
   430     tag: '%' (@{syntax ident} | @{syntax string})
   431   "}
   432 
   433   Some tags are pre-declared for certain classes of commands, serving
   434   as default markup if no tags are given in the text:
   435 
   436   \medskip
   437   \begin{tabular}{ll}
   438     @{text "theory"} & theory begin/end \\
   439     @{text "proof"} & all proof commands \\
   440     @{text "ML"} & all commands involving ML code \\
   441   \end{tabular}
   442 
   443   \medskip The Isabelle document preparation system
   444   \cite{isabelle-sys} allows tagged command regions to be presented
   445   specifically, e.g.\ to fold proof texts, or drop parts of the text
   446   completely.
   447 
   448   For example ``@{command "by"}~@{text "%invisible auto"}'' causes
   449   that piece of proof to be treated as @{text invisible} instead of
   450   @{text "proof"} (the default), which may be shown or hidden
   451   depending on the document setup.  In contrast, ``@{command
   452   "by"}~@{text "%visible auto"}'' forces this text to be shown
   453   invariably.
   454 
   455   Explicit tag specifications within a proof apply to all subsequent
   456   commands of the same level of nesting.  For example, ``@{command
   457   "proof"}~@{text "%visible \<dots>"}~@{command "qed"}'' forces the whole
   458   sub-proof to be typeset as @{text visible} (unless some of its parts
   459   are tagged differently).
   460 
   461   \medskip Command tags merely produce certain markup environments for
   462   type-setting.  The meaning of these is determined by {\LaTeX}
   463   macros, as defined in @{file "~~/lib/texinputs/isabelle.sty"} or
   464   by the document author.  The Isabelle document preparation tools
   465   also provide some high-level options to specify the meaning of
   466   arbitrary tags to ``keep'', ``drop'', or ``fold'' the corresponding
   467   parts of the text.  Logic sessions may also specify ``document
   468   versions'', where given tags are interpreted in some particular way.
   469   Again see \cite{isabelle-sys} for further details.
   470 *}
   471 
   472 
   473 section {* Railroad diagrams *}
   474 
   475 text {*
   476   \begin{matharray}{rcl}
   477     @{antiquotation_def "rail"} & : & @{text antiquotation} \\
   478   \end{matharray}
   479 
   480   @{rail "'rail' string"}
   481 
   482   The @{antiquotation rail} antiquotation allows to include syntax
   483   diagrams into Isabelle documents.  {\LaTeX} requires the style file
   484   @{"file" "~~/lib/texinputs/pdfsetup.sty"}, which can be used via
   485   @{verbatim "\\usepackage{pdfsetup}"} in @{verbatim "root.tex"}, for
   486   example.
   487 
   488   The rail specification language is quoted here as Isabelle @{syntax
   489   string}; it has its own grammar given below.
   490 
   491   @{rail "
   492   rule? + ';'
   493   ;
   494   rule: ((identifier | @{syntax antiquotation}) ':')? body
   495   ;
   496   body: concatenation + '|'
   497   ;
   498   concatenation: ((atom '?'?) +) (('*' | '+') atom?)?
   499   ;
   500   atom: '(' body? ')' | identifier |
   501     '@'? (string | @{syntax antiquotation}) |
   502     '\\\\\\\\'
   503   "}
   504 
   505   The lexical syntax of @{text "identifier"} coincides with that of
   506   @{syntax ident} in regular Isabelle syntax, but @{text string} uses
   507   single quotes instead of double quotes of the standard @{syntax
   508   string} category, to avoid extra escapes.
   509 
   510   Each @{text rule} defines a formal language (with optional name),
   511   using a notation that is similar to EBNF or regular expressions with
   512   recursion.  The meaning and visual appearance of these rail language
   513   elements is illustrated by the following representative examples.
   514 
   515   \begin{itemize}
   516 
   517   \item Empty @{verbatim "()"}
   518 
   519   @{rail "()"}
   520 
   521   \item Nonterminal @{verbatim "A"}
   522 
   523   @{rail "A"}
   524 
   525   \item Nonterminal via Isabelle antiquotation
   526   @{verbatim "@{syntax method}"}
   527 
   528   @{rail "@{syntax method}"}
   529 
   530   \item Terminal @{verbatim "'xyz'"}
   531 
   532   @{rail "'xyz'"}
   533 
   534   \item Terminal in keyword style @{verbatim "@'xyz'"}
   535 
   536   @{rail "@'xyz'"}
   537 
   538   \item Terminal via Isabelle antiquotation
   539   @{verbatim "@@{method rule}"}
   540 
   541   @{rail "@@{method rule}"}
   542 
   543   \item Concatenation @{verbatim "A B C"}
   544 
   545   @{rail "A B C"}
   546 
   547   \item Linebreak @{verbatim "\\\\"} inside
   548   concatenation\footnote{Strictly speaking, this is only a single
   549   backslash, but the enclosing @{syntax string} syntax requires a
   550   second one for escaping.} @{verbatim "A B C \\\\ D E F"}
   551 
   552   @{rail "A B C \\ D E F"}
   553 
   554   \item Variants @{verbatim "A | B | C"}
   555 
   556   @{rail "A | B | C"}
   557 
   558   \item Option @{verbatim "A ?"}
   559 
   560   @{rail "A ?"}
   561 
   562   \item Repetition @{verbatim "A *"}
   563 
   564   @{rail "A *"}
   565 
   566   \item Repetition with separator @{verbatim "A * sep"}
   567 
   568   @{rail "A * sep"}
   569 
   570   \item Strict repetition @{verbatim "A +"}
   571 
   572   @{rail "A +"}
   573 
   574   \item Strict repetition with separator @{verbatim "A + sep"}
   575 
   576   @{rail "A + sep"}
   577 
   578   \end{itemize}
   579 *}
   580 
   581 
   582 section {* Draft presentation *}
   583 
   584 text {*
   585   \begin{matharray}{rcl}
   586     @{command_def "display_drafts"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
   587     @{command_def "print_drafts"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
   588   \end{matharray}
   589 
   590   @{rail "
   591     (@@{command display_drafts} | @@{command print_drafts}) (@{syntax name} +)
   592 
   593   "}
   594 
   595   \begin{description}
   596 
   597   \item @{command "display_drafts"}~@{text paths} and @{command
   598   "print_drafts"}~@{text paths} perform simple output of a given list
   599   of raw source files.  Only those symbols that do not require
   600   additional {\LaTeX} packages are displayed properly, everything else
   601   is left verbatim.
   602 
   603   \end{description}
   604 *}
   605 
   606 end