doc-src/TutorialI/Documents/document/Documents.tex
changeset 12747 cc59ceb0bcb4
parent 12745 d7b4d8c9bf86
child 12749 c0ce43e809fd
equal deleted inserted replaced
12746:892af6538f3d 12747:cc59ceb0bcb4
    56   applications with less than two operands there is a special notation
    56   applications with less than two operands there is a special notation
    57   with \isa{op} prefix: \isa{xor} without arguments is represented
    57   with \isa{op} prefix: \isa{xor} without arguments is represented
    58   as \isa{op\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}}; together with plain prefix application this
    58   as \isa{op\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}}; together with plain prefix application this
    59   turns \isa{xor\ A} into \isa{op\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ A}.
    59   turns \isa{xor\ A} into \isa{op\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ A}.
    60 
    60 
    61   \medskip The keyword \isakeyword{infixl} specifies an infix operator
    61   \medskip The keyword \isakeyword{infixl} seen above specifies an
    62   that is nested to the \emph{left}: in iterated applications the more
    62   infix operator that is nested to the \emph{left}: in iterated
    63   complex expression appears on the left-hand side: \isa{A\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ B\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ C} stands for \isa{{\isacharparenleft}A\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ B{\isacharparenright}\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ C}.  Similarly,
    63   applications the more complex expression appears on the left-hand
    64   \isakeyword{infixr} specifies to nesting to the \emph{right},
    64   side, \isa{A\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ B\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ C} stands for \isa{{\isacharparenleft}A\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ B{\isacharparenright}\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ C}.
    65   reading \isa{A\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ B\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ C} as \isa{A\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ {\isacharparenleft}B\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ C{\isacharparenright}}.  In
    65   Similarly, \isakeyword{infixr} specifies nesting to the
    66   contrast, a \emph{non-oriented} declaration via \isakeyword{infix}
    66   \emph{right}, reading \isa{A\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ B\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ C} as \isa{A\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ {\isacharparenleft}B\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ C{\isacharparenright}}.  In contrast, a \emph{non-oriented} declaration via
    67   would render \isa{A\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ B\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ C} illegal, but demand explicit
    67   \isakeyword{infix} would render \isa{A\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ B\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ C} illegal, but
    68   parentheses to indicate the intended grouping.
    68   demand explicit parentheses to indicate the intended grouping.
    69 
    69 
    70   The string \isa{{\isachardoublequote}{\isacharbrackleft}{\isacharplus}{\isacharbrackright}{\isachardoublequote}} in the above annotation refers to
    70   The string \isa{{\isachardoublequote}{\isacharbrackleft}{\isacharplus}{\isacharbrackright}{\isachardoublequote}} in our annotation refers to the
    71   the concrete syntax to represent the operator (a literal token),
    71   concrete syntax to represent the operator (a literal token), while
    72   while the number \isa{{\isadigit{6}}{\isadigit{0}}} determines the precedence of the
    72   the number \isa{{\isadigit{6}}{\isadigit{0}}} determines the precedence of the construct
    73   construct (i.e.\ the syntactic priorities of the arguments and
    73   (i.e.\ the syntactic priorities of the arguments and result).  As it
    74   result).  As it happens, Isabelle/HOL already uses up many popular
    74   happens, Isabelle/HOL already uses up many popular combinations of
    75   combinations of ASCII symbols for its own use, including both \isa{{\isacharplus}} and \isa{{\isacharplus}{\isacharplus}}.  Slightly more awkward combinations like the
    75   ASCII symbols for its own use, including both \isa{{\isacharplus}} and \isa{{\isacharplus}{\isacharplus}}.  As a rule of thumb, more awkward character combinations are
    76   present \isa{{\isacharbrackleft}{\isacharplus}{\isacharbrackright}} tend to be available for user extensions.
    76   more likely to be still available for user extensions, just as our
       
    77   present \isa{{\isacharbrackleft}{\isacharplus}{\isacharbrackright}}.
    77 
    78 
    78   Operator precedence also needs some special considerations.  The
    79   Operator precedence also needs some special considerations.  The
    79   admissible range is 0--1000.  Very low or high priorities are
    80   admissible range is 0--1000.  Very low or high priorities are
    80   basically reserved for the meta-logic.  Syntax of Isabelle/HOL
    81   basically reserved for the meta-logic.  Syntax of Isabelle/HOL
    81   mainly uses the range of 10--100: the equality infix \isa{{\isacharequal}} is
    82   mainly uses the range of 10--100: the equality infix \isa{{\isacharequal}} is
   141 \noindent The X-Symbol package within Proof~General provides several
   142 \noindent The X-Symbol package within Proof~General provides several
   142   input methods to enter \isa{{\isasymoplus}} in the text.  If all fails one may
   143   input methods to enter \isa{{\isasymoplus}} in the text.  If all fails one may
   143   just type a named entity \verb,\,\verb,<oplus>, by hand; the display
   144   just type a named entity \verb,\,\verb,<oplus>, by hand; the display
   144   will be adapted immediately after continuing input.
   145   will be adapted immediately after continuing input.
   145 
   146 
   146   \medskip A slightly more refined scheme is to provide alternative
   147   \medskip A slightly more refined scheme of providing alternative
   147   syntax via the \bfindex{print mode} concept of Isabelle (see also
   148   syntax forms uses the \bfindex{print mode} concept of Isabelle (see
   148   \cite{isabelle-ref}).  By convention, the mode of ``$xsymbols$'' is
   149   also \cite{isabelle-ref}).  By convention, the mode of
   149   enabled whenever Proof~General's X-Symbol mode (or {\LaTeX} output)
   150   ``$xsymbols$'' is enabled whenever Proof~General's X-Symbol mode or
   150   is active.  Now consider the following hybrid declaration of \isa{xor}.%
   151   {\LaTeX} output is active.  Now consider the following hybrid
       
   152   declaration of \isa{xor}:%
   151 \end{isamarkuptext}%
   153 \end{isamarkuptext}%
   152 \isamarkuptrue%
   154 \isamarkuptrue%
   153 \isamarkupfalse%
   155 \isamarkupfalse%
   154 \isamarkupfalse%
   156 \isamarkupfalse%
   155 \isacommand{constdefs}\isanewline
   157 \isacommand{constdefs}\isanewline
   162 \isamarkupfalse%
   164 \isamarkupfalse%
   163 %
   165 %
   164 \begin{isamarkuptext}%
   166 \begin{isamarkuptext}%
   165 The \commdx{syntax} command introduced here acts like
   167 The \commdx{syntax} command introduced here acts like
   166   \isakeyword{consts}, but without declaring a logical constant.  The
   168   \isakeyword{consts}, but without declaring a logical constant.  The
   167   print mode specification (here \isa{{\isacharparenleft}xsymbols{\isacharparenright}}) limits the
   169   print mode specification of \isakeyword{syntax}, here \isa{{\isacharparenleft}xsymbols{\isacharparenright}}, is optional.  Also note that its type merely serves
   168   effect of the syntax annotation concerning output; that alternative
   170   for syntactic purposes, and is \emph{not} checked for consistency
   169   production available for input invariably.  Also note that the type
   171   with the real constant.
   170   declaration in \isakeyword{syntax} merely serves for syntactic
       
   171   purposes, and is \emph{not} checked for consistency with the real
       
   172   constant.
       
   173 
   172 
   174   \medskip We may now write \isa{A\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ B} or \isa{A\ {\isasymoplus}\ B} in
   173   \medskip We may now write \isa{A\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ B} or \isa{A\ {\isasymoplus}\ B} in
   175   input, while output uses the nicer syntax of $xsymbols$, provided
   174   input, while output uses the nicer syntax of $xsymbols$, provided
   176   that print mode is active.  Such an arrangement is particularly
   175   that print mode is active.  Such an arrangement is particularly
   177   useful for interactive development, where users may type plain ASCII
   176   useful for interactive development, where users may type plain ASCII
   199 %
   198 %
   200 \begin{isamarkuptext}%
   199 \begin{isamarkuptext}%
   201 \noindent Here the mixfix annotations on the rightmost column happen
   200 \noindent Here the mixfix annotations on the rightmost column happen
   202   to consist of a single Isabelle symbol each: \verb,\,\verb,<euro>,,
   201   to consist of a single Isabelle symbol each: \verb,\,\verb,<euro>,,
   203   \verb,\,\verb,<pounds>,, \verb,\,\verb,<yen>,, and \verb,$,.  Recall
   202   \verb,\,\verb,<pounds>,, \verb,\,\verb,<yen>,, and \verb,$,.  Recall
   204   that a constructor like \isa{Euro} actually is a function \isa{nat\ {\isasymRightarrow}\ currency}.  An expression like \isa{Euro\ {\isadigit{1}}{\isadigit{0}}} will be
   203   that a constructor like \isa{Euro} actually is a function \isa{nat\ {\isasymRightarrow}\ currency}.  The expression \isa{Euro\ {\isadigit{1}}{\isadigit{0}}} will be
   205   printed as \isa{{\isasymeuro}\ {\isadigit{1}}{\isadigit{0}}}; only the head of the application is
   204   printed as \isa{{\isasymeuro}\ {\isadigit{1}}{\isadigit{0}}}; only the head of the application is
   206   subject to our concrete syntax.  This rather simple form already
   205   subject to our concrete syntax.  This rather simple form already
   207   achieves conformance with notational standards of the European
   206   achieves conformance with notational standards of the European
   208   Commission.
   207   Commission.
   209 
   208 
   230   expressions.  This essentially provides a simple mechanism for
   229   expressions.  This essentially provides a simple mechanism for
   231   syntactic macros; even heavier transformations may be written in ML
   230   syntactic macros; even heavier transformations may be written in ML
   232   \cite{isabelle-ref}.
   231   \cite{isabelle-ref}.
   233 
   232 
   234   \medskip A typical example of syntax translations is to decorate
   233   \medskip A typical example of syntax translations is to decorate
   235   relational expressions (i.e.\ set-membership of tuples) with nice
   234   relational expressions (set-membership of tuples) with symbolic
   236   symbolic notation, such as \isa{{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}\ {\isasymin}\ sim} versus \isa{x\ {\isasymapprox}\ y}.%
   235   notation, e.g.\ \isa{{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}\ {\isasymin}\ sim} versus \isa{x\ {\isasymapprox}\ y}.%
   237 \end{isamarkuptext}%
   236 \end{isamarkuptext}%
   238 \isamarkuptrue%
   237 \isamarkuptrue%
   239 \isacommand{consts}\isanewline
   238 \isacommand{consts}\isanewline
   240 \ \ sim\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set{\isachardoublequote}\isanewline
   239 \ \ sim\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set{\isachardoublequote}\isanewline
   241 \isanewline
   240 \isanewline
   294 
   293 
   295   \medskip The Isabelle document preparation system essentially acts
   294   \medskip The Isabelle document preparation system essentially acts
   296   as a front-end to {\LaTeX}.  After checking specifications and
   295   as a front-end to {\LaTeX}.  After checking specifications and
   297   proofs formally, the theory sources are turned into typesetting
   296   proofs formally, the theory sources are turned into typesetting
   298   instructions in a schematic manner.  This enables users to write
   297   instructions in a schematic manner.  This enables users to write
   299   authentic reports on theory developments with little effort, where
   298   authentic reports on theory developments with little effort --- many
   300   most consistency checks are handled by the system.
   299   technical consistency checks are handled by the system.
   301 
   300 
   302   Here is an example to illustrate the idea of Isabelle document
   301   Here is an example to illustrate the idea of Isabelle document
   303   preparation.
   302   preparation.%
   304 
   303 \end{isamarkuptext}%
   305   \bigskip The following datatype definition of \isa{{\isacharprime}a\ bintree}
   304 \isamarkuptrue%
   306   models binary trees with nodes being decorated by elements of type
   305 %
   307   \isa{{\isacharprime}a}.%
   306 \begin{quotation}
       
   307 %
       
   308 \begin{isamarkuptext}%
       
   309 The following datatype definition of \isa{{\isacharprime}a\ bintree} models
       
   310   binary trees with nodes being decorated by elements of type \isa{{\isacharprime}a}.%
   308 \end{isamarkuptext}%
   311 \end{isamarkuptext}%
   309 \isamarkuptrue%
   312 \isamarkuptrue%
   310 \isacommand{datatype}\ {\isacharprime}a\ bintree\ {\isacharequal}\isanewline
   313 \isacommand{datatype}\ {\isacharprime}a\ bintree\ {\isacharequal}\isanewline
   311 \ \ Leaf\ {\isacharbar}\ Branch\ {\isacharprime}a\ \ {\isachardoublequote}{\isacharprime}a\ bintree{\isachardoublequote}\ \ {\isachardoublequote}{\isacharprime}a\ bintree{\isachardoublequote}\isamarkupfalse%
   314 \ \ \ \ \ Leaf\ {\isacharbar}\ Branch\ {\isacharprime}a\ \ {\isachardoublequote}{\isacharprime}a\ bintree{\isachardoublequote}\ \ {\isachardoublequote}{\isacharprime}a\ bintree{\isachardoublequote}\isamarkupfalse%
   312 %
   315 %
   313 \begin{isamarkuptext}%
   316 \begin{isamarkuptext}%
   314 \noindent The datatype induction rule generated here is of the form
   317 \noindent The datatype induction rule generated here is of the form
   315   \begin{isabelle}%
   318   \begin{isabelle}%
   316 {\isasymlbrakk}P\ Leaf{\isacharsemicolon}\isanewline
   319 {\isasymlbrakk}P\ Leaf{\isacharsemicolon}\isanewline
   317 \isaindent{\ \ \ }{\isasymAnd}a\ bintree{\isadigit{1}}\ bintree{\isadigit{2}}{\isachardot}\isanewline
   320 \isaindent{\ \ \ }{\isasymAnd}a\ bintree{\isadigit{1}}\ bintree{\isadigit{2}}{\isachardot}\isanewline
   318 \isaindent{\ \ \ \ \ \ }{\isasymlbrakk}P\ bintree{\isadigit{1}}{\isacharsemicolon}\ P\ bintree{\isadigit{2}}{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}Branch\ a\ bintree{\isadigit{1}}\ bintree{\isadigit{2}}{\isacharparenright}{\isasymrbrakk}\isanewline
   321 \isaindent{\ \ \ \ \ \ }{\isasymlbrakk}P\ bintree{\isadigit{1}}{\isacharsemicolon}\ P\ bintree{\isadigit{2}}{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}Branch\ a\ bintree{\isadigit{1}}\ bintree{\isadigit{2}}{\isacharparenright}{\isasymrbrakk}\isanewline
   319 {\isasymLongrightarrow}\ P\ bintree%
   322 {\isasymLongrightarrow}\ P\ bintree%
   320 \end{isabelle}
   323 \end{isabelle}%
   321 
   324 \end{isamarkuptext}%
   322   \bigskip The above document output has been produced by the
   325 \isamarkuptrue%
   323   following theory specification:
   326 %
       
   327 \end{quotation}
       
   328 %
       
   329 \begin{isamarkuptext}%
       
   330 The above document output has been produced by the following theory
       
   331   specification:
   324 
   332 
   325   \begin{ttbox}
   333   \begin{ttbox}
   326   text {\ttlbrace}*
   334   text {\ttlbrace}*
   327     The following datatype definition of {\at}{\ttlbrace}text "'a bintree"{\ttrbrace}
   335     The following datatype definition of {\at}{\ttlbrace}text "'a bintree"{\ttrbrace}
   328     models binary trees with nodes being decorated by elements
   336     models binary trees with nodes being decorated by elements
   336     {\ttback}noindent The datatype induction rule generated here is
   344     {\ttback}noindent The datatype induction rule generated here is
   337     of the form {\at}{\ttlbrace}thm [display] bintree.induct [no_vars]{\ttrbrace}
   345     of the form {\at}{\ttlbrace}thm [display] bintree.induct [no_vars]{\ttrbrace}
   338   *{\ttrbrace}
   346   *{\ttrbrace}
   339   \end{ttbox}
   347   \end{ttbox}
   340 
   348 
   341   Here we have augmented the theory by formal comments (via
   349   \noindent Here we have augmented the theory by formal comments
   342   \isakeyword{text} blocks).  The informal parts may again refer to
   350   (using \isakeyword{text} blocks).  The informal parts may again
   343   formal entities by means of ``antiquotations'' (such as
   351   refer to formal entities by means of ``antiquotations'' (such as
   344   \texttt{\at}\verb,{text "'a bintree"}, or
   352   \texttt{\at}\verb,{text "'a bintree"}, or
   345   \texttt{\at}\verb,{typ 'a},; see also \S\ref{sec:doc-prep-text}.%
   353   \texttt{\at}\verb,{typ 'a},), see also \S\ref{sec:doc-prep-text}.%
   346 \end{isamarkuptext}%
   354 \end{isamarkuptext}%
   347 \isamarkuptrue%
   355 \isamarkuptrue%
   348 %
   356 %
   349 \isamarkupsubsection{Isabelle Sessions%
   357 \isamarkupsubsection{Isabelle Sessions%
   350 }
   358 }
   380   make sure that \texttt{pdflatex} is present; if all fails one may
   388   make sure that \texttt{pdflatex} is present; if all fails one may
   381   fall back on DVI output by changing \texttt{usedir} options in
   389   fall back on DVI output by changing \texttt{usedir} options in
   382   \texttt{IsaMakefile} \cite{isabelle-sys}.}
   390   \texttt{IsaMakefile} \cite{isabelle-sys}.}
   383 
   391 
   384   \medskip The detailed arrangement of the session sources is as
   392   \medskip The detailed arrangement of the session sources is as
   385   follows.  This may be ignored in the beginning, but some of these
   393   follows.
   386   files need to be edited in realistic applications.
       
   387 
   394 
   388   \begin{itemize}
   395   \begin{itemize}
   389 
   396 
   390   \item Directory \texttt{MySession} holds the required theory files
   397   \item Directory \texttt{MySession} holds the required theory files
   391   $T@1$\texttt{.thy}, \dots, $T@n$\texttt{.thy}.
   398   $T@1$\texttt{.thy}, \dots, $T@n$\texttt{.thy}.
   401 
   408 
   402   The latter file holds appropriate {\LaTeX} code to commence a
   409   The latter file holds appropriate {\LaTeX} code to commence a
   403   document (\verb,\documentclass, etc.), and to include the generated
   410   document (\verb,\documentclass, etc.), and to include the generated
   404   files $T@i$\texttt{.tex} for each theory.  Isabelle will generate a
   411   files $T@i$\texttt{.tex} for each theory.  Isabelle will generate a
   405   file \texttt{session.tex} holding {\LaTeX} commands to include all
   412   file \texttt{session.tex} holding {\LaTeX} commands to include all
   406   generated theory output files in topologically sorted order.  So
   413   generated theory output files in topologically sorted order, so
   407   \verb,\input{session}, in \texttt{root.tex} does the job in most
   414   \verb,\input{session}, in the body of \texttt{root.tex} does the job
   408   situations.
   415   in most situations.
   409 
   416 
   410   \item \texttt{IsaMakefile} holds appropriate dependencies and
   417   \item \texttt{IsaMakefile} holds appropriate dependencies and
   411   invocations of Isabelle tools to control the batch job.  In fact,
   418   invocations of Isabelle tools to control the batch job.  In fact,
   412   several sessions may be controlled by the same \texttt{IsaMakefile}.
   419   several sessions may be managed by the same \texttt{IsaMakefile}.
   413   See also \cite{isabelle-sys} for further details, especially on
   420   See also \cite{isabelle-sys} for further details, especially on
   414   \texttt{isatool usedir} and \texttt{isatool make}.
   421   \texttt{isatool usedir} and \texttt{isatool make}.
   415 
   422 
   416   \end{itemize}
   423   \end{itemize}
   417 
   424 
   428   symbols), and the final \texttt{pdfsetup} (provides sane defaults
   435   symbols), and the final \texttt{pdfsetup} (provides sane defaults
   429   for \texttt{hyperref}, including URL markup) --- all three are
   436   for \texttt{hyperref}, including URL markup) --- all three are
   430   distributed with Isabelle. Further packages may be required in
   437   distributed with Isabelle. Further packages may be required in
   431   particular applications, e.g.\ for unusual mathematical symbols.
   438   particular applications, e.g.\ for unusual mathematical symbols.
   432 
   439 
   433   \medskip Additional files for the {\LaTeX} stage may be put into the
   440   \medskip Any additional files for the {\LaTeX} stage go into the
   434   \texttt{MySession/document} directory, too.  In particular, adding
   441   \texttt{MySession/document} directory as well.  In particular,
   435   \texttt{root.bib} here (with that specific name) causes an automatic
   442   adding \texttt{root.bib} (with that specific name) causes an
   436   run of \texttt{bibtex} to process a bibliographic database; see also
   443   automatic run of \texttt{bibtex} to process a bibliographic
   437   \texttt{isatool document} covered in \cite{isabelle-sys}.
   444   database; see also \texttt{isatool document} in \cite{isabelle-sys}.
   438 
   445 
   439   \medskip Any failure of the document preparation phase in an
   446   \medskip Any failure of the document preparation phase in an
   440   Isabelle batch session leaves the generated sources in their target
   447   Isabelle batch session leaves the generated sources in their target
   441   location (as pointed out by the accompanied error message).  This
   448   location (as pointed out by the accompanied error message).  This
   442   enables users to trace {\LaTeX} problems with the target files at
   449   enables users to trace {\LaTeX} problems with the generated files at
   443   hand.%
   450   hand.%
   444 \end{isamarkuptext}%
   451 \end{isamarkuptext}%
   445 \isamarkuptrue%
   452 \isamarkuptrue%
   446 %
   453 %
   447 \isamarkupsubsection{Structure Markup%
   454 \isamarkupsubsection{Structure Markup%
   472   \end{tabular}
   479   \end{tabular}
   473 
   480 
   474   \medskip
   481   \medskip
   475 
   482 
   476   From the Isabelle perspective, each markup command takes a single
   483   From the Isabelle perspective, each markup command takes a single
   477   $text$ argument (delimited by \verb,",\dots\verb,", or
   484   $text$ argument (delimited by \verb,",~\isa{{\isasymdots}}~\verb,", or
   478   \verb,{,\verb,*,~\dots~\verb,*,\verb,},).  After stripping any
   485   \verb,{,\verb,*,~\isa{{\isasymdots}}~\verb,*,\verb,},).  After stripping any
   479   surrounding white space, the argument is passed to a {\LaTeX} macro
   486   surrounding white space, the argument is passed to a {\LaTeX} macro
   480   \verb,\isamarkupXYZ, for any command \isakeyword{XYZ}.  These macros
   487   \verb,\isamarkupXYZ, for any command \isakeyword{XYZ}.  These macros
   481   are defined in \verb,isabelle.sty, according to the meaning given in
   488   are defined in \verb,isabelle.sty, according to the meaning given in
   482   the rightmost column above.
   489   the rightmost column above.
   483 
   490 
   543 }
   550 }
   544 \isamarkuptrue%
   551 \isamarkuptrue%
   545 %
   552 %
   546 \begin{isamarkuptext}%
   553 \begin{isamarkuptext}%
   547 Isabelle \bfindex{source comments}, which are of the form
   554 Isabelle \bfindex{source comments}, which are of the form
   548   \verb,(,\verb,*,~\dots~\verb,*,\verb,),, essentially act like white
   555   \verb,(,\verb,*,~\isa{{\isasymdots}}~\verb,*,\verb,),, essentially act like
   549   space and do not really contribute to the content.  They mainly
   556   white space and do not really contribute to the content.  They
   550   serve technical purposes to mark certain oddities in the raw input
   557   mainly serve technical purposes to mark certain oddities in the raw
   551   text.  In contrast, \bfindex{formal comments} are portions of text
   558   input text.  In contrast, \bfindex{formal comments} are portions of
   552   that are associated with formal Isabelle/Isar commands
   559   text that are associated with formal Isabelle/Isar commands
   553   (\bfindex{marginal comments}), or as standalone paragraphs within a
   560   (\bfindex{marginal comments}), or as standalone paragraphs within a
   554   theory or proof context (\bfindex{text blocks}).
   561   theory or proof context (\bfindex{text blocks}).
   555 
   562 
   556   \medskip Marginal comments are part of each command's concrete
   563   \medskip Marginal comments are part of each command's concrete
   557   syntax \cite{isabelle-ref}; the common form is ``\verb,--,~$text$''
   564   syntax \cite{isabelle-ref}; the common form is ``\verb,--,~$text$''
   558   where $text$ is delimited by \verb,",\dots\verb,", or
   565   where $text$ is delimited by \verb,",\isa{{\isasymdots}}\verb,", or
   559   \verb,{,\verb,*,~\dots~\verb,*,\verb,}, as before.  Multiple
   566   \verb,{,\verb,*,~\isa{{\isasymdots}}~\verb,*,\verb,}, as before.  Multiple
   560   marginal comments may be given at the same time.  Here is a simple
   567   marginal comments may be given at the same time.  Here is a simple
   561   example:%
   568   example:%
   562 \end{isamarkuptext}%
   569 \end{isamarkuptext}%
   563 \isamarkuptrue%
   570 \isamarkuptrue%
   564 \isacommand{lemma}\ {\isachardoublequote}A\ {\isacharminus}{\isacharminus}{\isachargreater}\ A{\isachardoublequote}\isanewline
   571 \isacommand{lemma}\ {\isachardoublequote}A\ {\isacharminus}{\isacharminus}{\isachargreater}\ A{\isachardoublequote}\isanewline
   597   vertical space).  This behavior may be changed by redefining the
   604   vertical space).  This behavior may be changed by redefining the
   598   {\LaTeX} environments of \verb,isamarkuptext, or
   605   {\LaTeX} environments of \verb,isamarkuptext, or
   599   \verb,isamarkuptxt,, respectively (via \verb,\renewenvironment,) The
   606   \verb,isamarkuptxt,, respectively (via \verb,\renewenvironment,) The
   600   text style of the body is determined by \verb,\isastyletext, and
   607   text style of the body is determined by \verb,\isastyletext, and
   601   \verb,\isastyletxt,; the default setup uses a smaller font within
   608   \verb,\isastyletxt,; the default setup uses a smaller font within
   602   proofs.
   609   proofs.  This may be changed as follows:
       
   610 
       
   611 \begin{verbatim}
       
   612   \renewcommand{\isastyletxt}{\isastyletext}
       
   613 \end{verbatim}
   603 
   614 
   604   \medskip The $text$ part of each of the various markup commands
   615   \medskip The $text$ part of each of the various markup commands
   605   considered so far essentially inserts \emph{quoted material} into a
   616   considered so far essentially inserts \emph{quoted material} into a
   606   formal text, mainly for instruction of the reader.  An
   617   formal text, mainly for instruction of the reader.  An
   607   \bfindex{antiquotation} is again a formal object embedded into such
   618   \bfindex{antiquotation} is again a formal object embedded into such
   608   an informal portion.  The interpretation of antiquotations is
   619   an informal portion.  The interpretation of antiquotations is
   609   limited to some well-formedness checks, with the result being pretty
   620   limited to some well-formedness checks, with the result being pretty
   610   printed to the resulting document.  So quoted text blocks together
   621   printed to the resulting document.  Quoted text blocks together with
   611   with antiquotations provide very useful means to reference formal
   622   antiquotations provide very useful means to reference formal
   612   entities with good confidence in getting the technical details right
   623   entities, with good confidence in getting the technical details
   613   (especially syntax and types).
   624   right (especially syntax and types).
   614 
   625 
   615   The general syntax of antiquotations is as follows:
   626   The general syntax of antiquotations is as follows:
   616   \texttt{{\at}{\ttlbrace}$name$ $arguments${\ttrbrace}}, or
   627   \texttt{{\at}{\ttlbrace}$name$ $arguments${\ttrbrace}}, or
   617   \texttt{{\at}{\ttlbrace}$name$ [$options$] $arguments${\ttrbrace}}
   628   \texttt{{\at}{\ttlbrace}$name$ [$options$] $arguments${\ttrbrace}}
   618   for a comma-separated list of options consisting of a $name$ or
   629   for a comma-separated list of options consisting of a $name$ or
   656   \texttt{\at}\verb,{prop,~$\phi$\verb,}, & print proposition $\phi$ \\
   667   \texttt{\at}\verb,{prop,~$\phi$\verb,}, & print proposition $\phi$ \\
   657   \texttt{\at}\verb,{prop [display],~$\phi$\verb,}, & print large proposition $\phi$ (with linebreaks) \\
   668   \texttt{\at}\verb,{prop [display],~$\phi$\verb,}, & print large proposition $\phi$ (with linebreaks) \\
   658   \texttt{\at}\verb,{prop [source],~$\phi$\verb,}, & check proposition $\phi$, print its input \\
   669   \texttt{\at}\verb,{prop [source],~$\phi$\verb,}, & check proposition $\phi$, print its input \\
   659   \texttt{\at}\verb,{thm,~$a$\verb,}, & print fact $a$ \\
   670   \texttt{\at}\verb,{thm,~$a$\verb,}, & print fact $a$ \\
   660   \texttt{\at}\verb,{thm,~$a$~\verb,[no_vars]}, & print fact $a$, fixing schematic variables \\
   671   \texttt{\at}\verb,{thm,~$a$~\verb,[no_vars]}, & print fact $a$, fixing schematic variables \\
   661   \texttt{\at}\verb,{thm [source],~$a$\verb,}, & check availability of fact, print name $a$ \\
   672   \texttt{\at}\verb,{thm [source],~$a$\verb,}, & check availability of fact $a$, print its name \\
   662   \texttt{\at}\verb,{text,~$s$\verb,}, & print uninterpreted text $s$ \\
   673   \texttt{\at}\verb,{text,~$s$\verb,}, & print uninterpreted text $s$ \\
   663   \end{tabular}
   674   \end{tabular}
   664 
   675 
   665   \medskip
   676   \medskip
   666 
   677 
   700   redefining certain macros, say in \texttt{root.tex} of the document.
   711   redefining certain macros, say in \texttt{root.tex} of the document.
   701 
   712 
   702   \begin{enumerate}
   713   \begin{enumerate}
   703 
   714 
   704   \item 7-bit ASCII characters: letters \texttt{A\dots Z} and
   715   \item 7-bit ASCII characters: letters \texttt{A\dots Z} and
   705   \texttt{a\dots z} are output verbatim, digits are passed as an
   716   \texttt{a\dots z} are output directly, digits are passed as an
   706   argument to the \verb,\isadigit, macro, other characters are
   717   argument to the \verb,\isadigit, macro, other characters are
   707   replaced by specifically named macros of the form
   718   replaced by specifically named macros of the form
   708   \verb,\isacharXYZ,.
   719   \verb,\isacharXYZ,.
   709 
   720 
   710   \item Named symbols: \verb,\,\verb,<,$XYZ$\verb,>, become
   721   \item Named symbols: \verb,\,\verb,<,$XYZ$\verb,>, is turned into
   711   \verb,{\isasym,$XYZ$\verb,}, each (note the additional braces).
   722   \verb,{\isasym,$XYZ$\verb,},; note the additional braces.
   712 
   723 
   713   \item Named control symbols: \verb,{\isasym,$XYZ$\verb,}, become
   724   \item Named control symbols: \verb,\,\verb,<^,$XYZ$\verb,>, is
   714   \verb,\isactrl,$XYZ$ each; subsequent symbols may act as arguments
   725   turned into \verb,\isactrl,$XYZ$; subsequent symbols may act as
   715   if the corresponding macro is defined accordingly.
   726   arguments if the corresponding macro is defined accordingly.
   716 
   727 
   717   \end{enumerate}
   728   \end{enumerate}
   718 
   729 
   719   Users may occasionally wish to give new {\LaTeX} interpretations of
   730   Users may occasionally wish to give new {\LaTeX} interpretations of
   720   named symbols; this merely requires an appropriate definition of
   731   named symbols; this merely requires an appropriate definition of
   721   \verb,\,\verb,<,$XYZ$\verb,>, (see \texttt{isabelle.sty} for working
   732   \verb,\isasym,$XYZ$\verb,, for \verb,\,\verb,<,$XYZ$\verb,>, (see
   722   examples).  Control symbols are slightly more difficult to get
   733   \texttt{isabelle.sty} for working examples).  Control symbols are
   723   right, though.
   734   slightly more difficult to get right, though.
   724 
   735 
   725   \medskip The \verb,\isabellestyle, macro provides a high-level
   736   \medskip The \verb,\isabellestyle, macro provides a high-level
   726   interface to tune the general appearance of individual symbols.  For
   737   interface to tune the general appearance of individual symbols.  For
   727   example, \verb,\isabellestyle{it}, uses the italics text style to
   738   example, \verb,\isabellestyle{it}, uses the italics text style to
   728   mimic the general appearance of the {\LaTeX} math mode; double
   739   mimic the general appearance of the {\LaTeX} math mode; double
   740 By default, Isabelle's document system generates a {\LaTeX} source
   751 By default, Isabelle's document system generates a {\LaTeX} source
   741   file for each theory that gets loaded while running the session.
   752   file for each theory that gets loaded while running the session.
   742   The generated \texttt{session.tex} will include all of these in
   753   The generated \texttt{session.tex} will include all of these in
   743   order of appearance, which in turn gets included by the standard
   754   order of appearance, which in turn gets included by the standard
   744   \texttt{root.tex}.  Certainly one may change the order or suppress
   755   \texttt{root.tex}.  Certainly one may change the order or suppress
   745   unwanted theories by ignoring \texttt{session.tex} and include
   756   unwanted theories by ignoring \texttt{session.tex} and load
   746   individual files in \texttt{root.tex} by hand.  On the other hand,
   757   individual files directly in \texttt{root.tex}.  On the other hand,
   747   such an arrangement requires additional maintenance chores whenever
   758   such an arrangement requires additional maintenance whenever the
   748   the collection of theories changes.
   759   collection of theories changes.
   749 
   760 
   750   Alternatively, one may tune the theory loading process in
   761   Alternatively, one may tune the theory loading process in
   751   \texttt{ROOT.ML} itself: traversal of the theory dependency graph
   762   \texttt{ROOT.ML} itself: traversal of the theory dependency graph
   752   may be fine-tuned by adding \verb,use_thy, invocations, although
   763   may be fine-tuned by adding \verb,use_thy, invocations, although
   753   topological sorting still has to be observed.  Moreover, the ML
   764   topological sorting still has to be observed.  Moreover, the ML
   763   formal content in full.  In order to delimit \bfindex{ignored
   774   formal content in full.  In order to delimit \bfindex{ignored
   764   material} special source comments
   775   material} special source comments
   765   \verb,(,\verb,*,\verb,<,\verb,*,\verb,), and
   776   \verb,(,\verb,*,\verb,<,\verb,*,\verb,), and
   766   \verb,(,\verb,*,\verb,>,\verb,*,\verb,), may be included in the
   777   \verb,(,\verb,*,\verb,>,\verb,*,\verb,), may be included in the
   767   text.  Only the document preparation system is affected, the formal
   778   text.  Only the document preparation system is affected, the formal
   768   checking the theory is performed unchanged.
   779   checking of the theory is performed unchanged.
   769 
   780 
   770   In the following example we suppress the slightly formalistic
   781   In the subsequent example we suppress the slightly formalistic
   771   \isakeyword{theory} + \isakeyword{end} surroundings a theory.
   782   \isakeyword{theory} + \isakeyword{end} surroundings a theory.
   772 
   783 
   773   \medskip
   784   \medskip
   774 
   785 
   775   \begin{tabular}{l}
   786   \begin{tabular}{l}
   783   \end{tabular}
   794   \end{tabular}
   784 
   795 
   785   \medskip
   796   \medskip
   786 
   797 
   787   Text may be suppressed in a fine-grained manner.  We may even drop
   798   Text may be suppressed in a fine-grained manner.  We may even drop
   788   vital parts of a formal proof, pretending that things have been
   799   vital parts of a proof, pretending that things have been simpler
   789   simpler than in reality.  For example, the following ``fully
   800   than in reality.  For example, this ``fully automatic'' proof is
   790   automatic'' proof is actually a fake:%
   801   actually a fake:%
   791 \end{isamarkuptext}%
   802 \end{isamarkuptext}%
   792 \isamarkuptrue%
   803 \isamarkuptrue%
   793 \isacommand{lemma}\ {\isachardoublequote}x\ {\isasymnoteq}\ {\isacharparenleft}{\isadigit{0}}{\isacharcolon}{\isacharcolon}int{\isacharparenright}\ {\isasymLongrightarrow}\ {\isadigit{0}}\ {\isacharless}\ x\ {\isacharasterisk}\ x{\isachardoublequote}\isanewline
   804 \isacommand{lemma}\ {\isachardoublequote}x\ {\isasymnoteq}\ {\isacharparenleft}{\isadigit{0}}{\isacharcolon}{\isacharcolon}int{\isacharparenright}\ {\isasymLongrightarrow}\ {\isadigit{0}}\ {\isacharless}\ x\ {\isacharasterisk}\ x{\isachardoublequote}\isanewline
   794 \ \ \isamarkupfalse%
   805 \ \ \isamarkupfalse%
   795 \isacommand{by}\ {\isacharparenleft}auto{\isacharparenright}\isamarkupfalse%
   806 \isacommand{by}\ {\isacharparenleft}auto{\isacharparenright}\isamarkupfalse%
   801   by (auto(*<*)simp add: int_less_le(*>*))
   812   by (auto(*<*)simp add: int_less_le(*>*))
   802 \end{verbatim}
   813 \end{verbatim}
   803 %(*
   814 %(*
   804 
   815 
   805   \medskip Ignoring portions of printed text does demand some care by
   816   \medskip Ignoring portions of printed text does demand some care by
   806   the writer.  First of all, the writer is responsible not to
   817   the writer.  First of all, the user is responsible not to obfuscate
   807   obfuscate the underlying formal development in an unduly manner.  It
   818   the underlying theory development in an unduly manner.  It is fairly
   808   is fairly easy to invalidate the remaining visible text, e.g.\ by
   819   easy to invalidate the visible text, e.g.\ by referencing
   809   referencing questionable formal items (strange definitions,
   820   questionable formal items (strange definitions, arbitrary axioms
   810   arbitrary axioms etc.) that have been hidden from sight beforehand.
   821   etc.) that have been hidden from sight beforehand.
   811 
   822 
   812   Authentic reports of formal theories, say as part of a library,
   823   Authentic reports of Isabelle/Isar theories, say as part of a
   813   should refrain from suppressing parts of the text at all.  Other
   824   library, should refrain from suppressing parts of the text at all.
   814   users may need the full information for their own derivative work.
   825   Other users may need the full information for their own derivative
   815   If a particular formalization appears inadequate for general public
   826   work.  If a particular formalization appears inadequate for general
   816   coverage, it is often more appropriate to think of a better way in
   827   public coverage, it is often more appropriate to think of a better
   817   the first place.
   828   way in the first place.
   818 
   829 
   819   \medskip Some technical subtleties of the
   830   \medskip Some technical subtleties of the
   820   \verb,(,\verb,*,\verb,<,\verb,*,\verb,),~\verb,(,\verb,*,\verb,>,\verb,*,\verb,),
   831   \verb,(,\verb,*,\verb,<,\verb,*,\verb,),~\verb,(,\verb,*,\verb,>,\verb,*,\verb,),
   821   elements need to be kept in mind, too --- the system performs little
   832   elements need to be kept in mind, too --- the system performs little
   822   sanity checks here.  Arguments of markup commands and formal
   833   sanity checks here.  Arguments of markup commands and formal