doc-src/IsarRef/Thy/document/Document_Preparation.tex
author wenzelm
Fri, 13 Mar 2009 15:50:05 +0100
changeset 30492 cb7e886e4b10
parent 30397 b6212ae21656
child 32893 dbef0e6438ec
permissions -rw-r--r--
get data from plain Proof.context;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
     1
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
     2
\begin{isabellebody}%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
     3
\def\isabellecontext{Document{\isacharunderscore}Preparation}%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
     4
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
     5
\isadelimtheory
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
     6
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
     7
\endisadelimtheory
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
     8
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
     9
\isatagtheory
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    10
\isacommand{theory}\isamarkupfalse%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    11
\ Document{\isacharunderscore}Preparation\isanewline
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    12
\isakeyword{imports}\ Main\isanewline
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    13
\isakeyword{begin}%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    14
\endisatagtheory
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    15
{\isafoldtheory}%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    16
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    17
\isadelimtheory
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    18
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    19
\endisadelimtheory
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    20
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    21
\isamarkupchapter{Document preparation \label{ch:document-prep}%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    22
}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    23
\isamarkuptrue%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    24
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    25
\begin{isamarkuptext}%
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    26
Isabelle/Isar provides a simple document preparation system
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    27
  based on regular {PDF-\LaTeX} technology, with full support for
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    28
  hyper-links and bookmarks.  Thus the results are well suited for WWW
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    29
  browsing and as printed copies.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    30
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    31
  \medskip Isabelle generates {\LaTeX} output while running a
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    32
  \emph{logic session} (see also \cite{isabelle-sys}).  Getting
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    33
  started with a working configuration for common situations is quite
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    34
  easy by using the Isabelle \verb|mkdir| and \verb|make|
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    35
  tools.  First invoke
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    36
\begin{ttbox}
28505
f98751bd715f updated generated file;
wenzelm
parents: 27881
diff changeset
    37
  isabelle mkdir Foo
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    38
\end{ttbox}
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    39
  to initialize a separate directory for session \verb|Foo| (it
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    40
  is safe to experiment, since \verb|isabelle mkdir| never
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    41
  overwrites existing files).  Ensure that \verb|Foo/ROOT.ML|
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    42
  holds ML commands to load all theories required for this session;
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    43
  furthermore \verb|Foo/document/root.tex| should include any
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    44
  special {\LaTeX} macro packages required for your document (the
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    45
  default is usually sufficient as a start).
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    46
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    47
  The session is controlled by a separate \verb|IsaMakefile|
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    48
  (with crude source dependencies by default).  This file is located
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    49
  one level up from the \verb|Foo| directory location.  Now
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    50
  invoke
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    51
\begin{ttbox}
28505
f98751bd715f updated generated file;
wenzelm
parents: 27881
diff changeset
    52
  isabelle make Foo
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    53
\end{ttbox}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    54
  to run the \verb|Foo| session, with browser information and
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    55
  document preparation enabled.  Unless any errors are reported by
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    56
  Isabelle or {\LaTeX}, the output will appear inside the directory
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    57
  defined by the \verb|ISABELLE_BROWSER_INFO| setting (as
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    58
  reported by the batch job in verbose mode).
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    59
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    60
  \medskip You may also consider to tune the \verb|usedir|
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    61
  options in \verb|IsaMakefile|, for example to switch the output
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    62
  format between \verb|pdf| and \verb|dvi|, or activate the
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    63
  \verb|-D| option to retain a second copy of the generated
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    64
  {\LaTeX} sources (for manual inspection or separate runs of
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    65
  \hyperlink{executable.latex}{\mbox{\isa{\isatt{latex}}}}).
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    66
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    67
  \medskip See \emph{The Isabelle System Manual} \cite{isabelle-sys}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    68
  for further details on Isabelle logic sessions and theory
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    69
  presentation.  The Isabelle/HOL tutorial \cite{isabelle-hol-book}
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    70
  also covers theory presentation to some extent.%
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    71
\end{isamarkuptext}%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    72
\isamarkuptrue%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    73
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    74
\isamarkupsection{Markup commands \label{sec:markup}%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    75
}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    76
\isamarkuptrue%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    77
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    78
\begin{isamarkuptext}%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    79
\begin{matharray}{rcl}
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    80
    \indexdef{}{command}{header}\hypertarget{command.header}{\hyperlink{command.header}{\mbox{\isa{\isacommand{header}}}}} & : & \isa{{\isachardoublequote}toplevel\ {\isasymrightarrow}\ toplevel{\isachardoublequote}} \\[0.5ex]
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    81
    \indexdef{}{command}{chapter}\hypertarget{command.chapter}{\hyperlink{command.chapter}{\mbox{\isa{\isacommand{chapter}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    82
    \indexdef{}{command}{section}\hypertarget{command.section}{\hyperlink{command.section}{\mbox{\isa{\isacommand{section}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    83
    \indexdef{}{command}{subsection}\hypertarget{command.subsection}{\hyperlink{command.subsection}{\mbox{\isa{\isacommand{subsection}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    84
    \indexdef{}{command}{subsubsection}\hypertarget{command.subsubsection}{\hyperlink{command.subsubsection}{\mbox{\isa{\isacommand{subsubsection}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    85
    \indexdef{}{command}{text}\hypertarget{command.text}{\hyperlink{command.text}{\mbox{\isa{\isacommand{text}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    86
    \indexdef{}{command}{text\_raw}\hypertarget{command.text-raw}{\hyperlink{command.text-raw}{\mbox{\isa{\isacommand{text{\isacharunderscore}raw}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\[0.5ex]
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    87
    \indexdef{}{command}{sect}\hypertarget{command.sect}{\hyperlink{command.sect}{\mbox{\isa{\isacommand{sect}}}}} & : & \isa{{\isachardoublequote}proof\ {\isasymrightarrow}\ proof{\isachardoublequote}} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    88
    \indexdef{}{command}{subsect}\hypertarget{command.subsect}{\hyperlink{command.subsect}{\mbox{\isa{\isacommand{subsect}}}}} & : & \isa{{\isachardoublequote}proof\ {\isasymrightarrow}\ proof{\isachardoublequote}} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    89
    \indexdef{}{command}{subsubsect}\hypertarget{command.subsubsect}{\hyperlink{command.subsubsect}{\mbox{\isa{\isacommand{subsubsect}}}}} & : & \isa{{\isachardoublequote}proof\ {\isasymrightarrow}\ proof{\isachardoublequote}} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    90
    \indexdef{}{command}{txt}\hypertarget{command.txt}{\hyperlink{command.txt}{\mbox{\isa{\isacommand{txt}}}}} & : & \isa{{\isachardoublequote}proof\ {\isasymrightarrow}\ proof{\isachardoublequote}} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    91
    \indexdef{}{command}{txt\_raw}\hypertarget{command.txt-raw}{\hyperlink{command.txt-raw}{\mbox{\isa{\isacommand{txt{\isacharunderscore}raw}}}}} & : & \isa{{\isachardoublequote}proof\ {\isasymrightarrow}\ proof{\isachardoublequote}} \\
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    92
  \end{matharray}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
    93
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    94
  Markup commands provide a structured way to insert text into the
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    95
  document generated from a theory.  Each markup command takes a
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    96
  single \hyperlink{syntax.text}{\mbox{\isa{text}}} argument, which is passed as argument to a
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    97
  corresponding {\LaTeX} macro.  The default macros provided by
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    98
  \hyperlink{file.~~/lib/texinputs/isabelle.sty}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}lib{\isacharslash}texinputs{\isacharslash}isabelle{\isachardot}sty}}}} can be redefined according
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
    99
  to the needs of the underlying document and {\LaTeX} styles.
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   100
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   101
  Note that formal comments (\secref{sec:comments}) are similar to
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   102
  markup commands, but have a different status within Isabelle/Isar
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   103
  syntax.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   104
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   105
  \begin{rail}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   106
    ('chapter' | 'section' | 'subsection' | 'subsubsection' | 'text') target? text
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   107
    ;
27052
5c48cecb981b updated generated file;
wenzelm
parents: 27042
diff changeset
   108
    ('header' | 'text\_raw' | 'sect' | 'subsect' | 'subsubsect' | 'txt' | 'txt\_raw') text
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   109
    ;
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   110
  \end{rail}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   111
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   112
  \begin{description}
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   113
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   114
  \item \hyperlink{command.header}{\mbox{\isa{\isacommand{header}}}} provides plain text markup just preceding
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   115
  the formal beginning of a theory.  The corresponding {\LaTeX} macro
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   116
  is \verb|\isamarkupheader|, which acts like \hyperlink{command.section}{\mbox{\isa{\isacommand{section}}}} by default.
27052
5c48cecb981b updated generated file;
wenzelm
parents: 27042
diff changeset
   117
  
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   118
  \item \hyperlink{command.chapter}{\mbox{\isa{\isacommand{chapter}}}}, \hyperlink{command.section}{\mbox{\isa{\isacommand{section}}}}, \hyperlink{command.subsection}{\mbox{\isa{\isacommand{subsection}}}},
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   119
  and \hyperlink{command.subsubsection}{\mbox{\isa{\isacommand{subsubsection}}}} mark chapter and section headings
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   120
  within the main theory body or local theory targets.  The
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   121
  corresponding {\LaTeX} macros are \verb|\isamarkupchapter|,
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   122
  \verb|\isamarkupsection|, \verb|\isamarkupsubsection| etc.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   123
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   124
  \item \hyperlink{command.sect}{\mbox{\isa{\isacommand{sect}}}}, \hyperlink{command.subsect}{\mbox{\isa{\isacommand{subsect}}}}, and \hyperlink{command.subsubsect}{\mbox{\isa{\isacommand{subsubsect}}}}
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   125
  mark section headings within proofs.  The corresponding {\LaTeX}
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   126
  macros are \verb|\isamarkupsect|, \verb|\isamarkupsubsect| etc.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   127
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   128
  \item \hyperlink{command.text}{\mbox{\isa{\isacommand{text}}}} and \hyperlink{command.txt}{\mbox{\isa{\isacommand{txt}}}} specify paragraphs of plain
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   129
  text.  This corresponds to a {\LaTeX} environment \verb|\begin{isamarkuptext}| \isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}} \verb|\end{isamarkuptext}| etc.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   130
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   131
  \item \hyperlink{command.text-raw}{\mbox{\isa{\isacommand{text{\isacharunderscore}raw}}}} and \hyperlink{command.txt-raw}{\mbox{\isa{\isacommand{txt{\isacharunderscore}raw}}}} insert {\LaTeX}
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   132
  source into the output, without additional markup.  Thus the full
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   133
  range of document manipulations becomes available, at the risk of
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   134
  messing up document output.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   135
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   136
  \end{description}
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   137
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   138
  Except for \hyperlink{command.text-raw}{\mbox{\isa{\isacommand{text{\isacharunderscore}raw}}}} and \hyperlink{command.txt-raw}{\mbox{\isa{\isacommand{txt{\isacharunderscore}raw}}}}, the text
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   139
  passed to any of the above markup commands may refer to formal
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   140
  entities via \emph{document antiquotations}, see also
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   141
  \secref{sec:antiq}.  These are interpreted in the present theory or
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   142
  proof context, or the named \isa{{\isachardoublequote}target{\isachardoublequote}}.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   143
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   144
  \medskip The proof markup commands closely resemble those for theory
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   145
  specifications, but have a different formal status and produce
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   146
  different {\LaTeX} macros.  The default definitions coincide for
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   147
  analogous commands such as \hyperlink{command.section}{\mbox{\isa{\isacommand{section}}}} and \hyperlink{command.sect}{\mbox{\isa{\isacommand{sect}}}}.%
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   148
\end{isamarkuptext}%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   149
\isamarkuptrue%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   150
%
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   151
\isamarkupsection{Document Antiquotations \label{sec:antiq}%
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   152
}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   153
\isamarkuptrue%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   154
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   155
\begin{isamarkuptext}%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   156
\begin{matharray}{rcl}
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   157
    \indexdef{}{antiquotation}{theory}\hypertarget{antiquotation.theory}{\hyperlink{antiquotation.theory}{\mbox{\isa{theory}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   158
    \indexdef{}{antiquotation}{thm}\hypertarget{antiquotation.thm}{\hyperlink{antiquotation.thm}{\mbox{\isa{thm}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   159
    \indexdef{}{antiquotation}{lemma}\hypertarget{antiquotation.lemma}{\hyperlink{antiquotation.lemma}{\mbox{\isa{lemma}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   160
    \indexdef{}{antiquotation}{prop}\hypertarget{antiquotation.prop}{\hyperlink{antiquotation.prop}{\mbox{\isa{prop}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   161
    \indexdef{}{antiquotation}{term}\hypertarget{antiquotation.term}{\hyperlink{antiquotation.term}{\mbox{\isa{term}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   162
    \indexdef{}{antiquotation}{const}\hypertarget{antiquotation.const}{\hyperlink{antiquotation.const}{\mbox{\isa{const}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   163
    \indexdef{}{antiquotation}{abbrev}\hypertarget{antiquotation.abbrev}{\hyperlink{antiquotation.abbrev}{\mbox{\isa{abbrev}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   164
    \indexdef{}{antiquotation}{typeof}\hypertarget{antiquotation.typeof}{\hyperlink{antiquotation.typeof}{\mbox{\isa{typeof}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   165
    \indexdef{}{antiquotation}{typ}\hypertarget{antiquotation.typ}{\hyperlink{antiquotation.typ}{\mbox{\isa{typ}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   166
    \indexdef{}{antiquotation}{thm\_style}\hypertarget{antiquotation.thm-style}{\hyperlink{antiquotation.thm-style}{\mbox{\isa{thm{\isacharunderscore}style}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   167
    \indexdef{}{antiquotation}{term\_style}\hypertarget{antiquotation.term-style}{\hyperlink{antiquotation.term-style}{\mbox{\isa{term{\isacharunderscore}style}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   168
    \indexdef{}{antiquotation}{text}\hypertarget{antiquotation.text}{\hyperlink{antiquotation.text}{\mbox{\isa{text}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   169
    \indexdef{}{antiquotation}{goals}\hypertarget{antiquotation.goals}{\hyperlink{antiquotation.goals}{\mbox{\isa{goals}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   170
    \indexdef{}{antiquotation}{subgoals}\hypertarget{antiquotation.subgoals}{\hyperlink{antiquotation.subgoals}{\mbox{\isa{subgoals}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   171
    \indexdef{}{antiquotation}{prf}\hypertarget{antiquotation.prf}{\hyperlink{antiquotation.prf}{\mbox{\isa{prf}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   172
    \indexdef{}{antiquotation}{full\_prf}\hypertarget{antiquotation.full-prf}{\hyperlink{antiquotation.full-prf}{\mbox{\isa{full{\isacharunderscore}prf}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   173
    \indexdef{}{antiquotation}{ML}\hypertarget{antiquotation.ML}{\hyperlink{antiquotation.ML}{\mbox{\isa{ML}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   174
    \indexdef{}{antiquotation}{ML\_type}\hypertarget{antiquotation.ML-type}{\hyperlink{antiquotation.ML-type}{\mbox{\isa{ML{\isacharunderscore}type}}}} & : & \isa{antiquotation} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   175
    \indexdef{}{antiquotation}{ML\_struct}\hypertarget{antiquotation.ML-struct}{\hyperlink{antiquotation.ML-struct}{\mbox{\isa{ML{\isacharunderscore}struct}}}} & : & \isa{antiquotation} \\
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   176
  \end{matharray}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   177
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   178
  The overall content of an Isabelle/Isar theory may alternate between
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   179
  formal and informal text.  The main body consists of formal
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   180
  specification and proof commands, interspersed with markup commands
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   181
  (\secref{sec:markup}) or document comments (\secref{sec:comments}).
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   182
  The argument of markup commands quotes informal text to be printed
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   183
  in the resulting document, but may again refer to formal entities
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   184
  via \emph{document antiquotations}.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   185
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   186
  For example, embedding of ``\isa{{\isacharat}{\isacharbraceleft}term\ {\isacharbrackleft}show{\isacharunderscore}types{\isacharbrackright}\ {\isachardoublequote}f\ x\ {\isacharequal}\ a\ {\isacharplus}\ x{\isachardoublequote}{\isacharbraceright}}''
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   187
  within a text block makes
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   188
  \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.
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   189
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   190
  Antiquotations usually spare the author tedious typing of logical
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   191
  entities in full detail.  Even more importantly, some degree of
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   192
  consistency-checking between the main body of formal text and its
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   193
  informal explanation is achieved, since terms and types appearing in
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   194
  antiquotations are checked within the current theory or proof
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   195
  context.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   196
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   197
  \begin{rail}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   198
    atsign lbrace antiquotation rbrace
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   199
    ;
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   200
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   201
    antiquotation:
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   202
      'theory' options name |
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   203
      'thm' options thmrefs |
27453
eecd9d84e41b added lemma antiquotation
haftmann
parents: 27054
diff changeset
   204
      'lemma' options prop 'by' method |
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   205
      'prop' options prop |
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   206
      'term' options term |
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   207
      'const' options term |
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   208
      'abbrev' options term |
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   209
      'typeof' options term |
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   210
      'typ' options type |
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   211
      'thm\_style' options name thmref |
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   212
      'term\_style' options name term |
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   213
      'text' options name |
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   214
      'goals' options |
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   215
      'subgoals' options |
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   216
      'prf' options thmrefs |
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   217
      'full\_prf' options thmrefs |
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   218
      'ML' options name |
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   219
      'ML\_type' options name |
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   220
      'ML\_struct' options name
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   221
    ;
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   222
    options: '[' (option * ',') ']'
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   223
    ;
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   224
    option: name | name '=' name
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   225
    ;
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   226
  \end{rail}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   227
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   228
  Note that the syntax of antiquotations may \emph{not} include source
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   229
  comments \verb|(*|~\isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}}~\verb|*)| nor verbatim
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   230
  text \verb|{|\verb|*|~\isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}}~\verb|*|\verb|}|.
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   231
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   232
  \begin{description}
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   233
  
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   234
  \item \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}theory\ A{\isacharbraceright}{\isachardoublequote}} prints the name \isa{{\isachardoublequote}A{\isachardoublequote}}, which is
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   235
  guaranteed to refer to a valid ancestor theory in the current
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   236
  context.
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   237
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   238
  \item \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}thm\ a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n{\isacharbraceright}{\isachardoublequote}} prints theorems \isa{{\isachardoublequote}a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n{\isachardoublequote}}.
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   239
  Full fact expressions are allowed here, including attributes
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   240
  (\secref{sec:syn-att}).
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   241
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   242
  \item \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}prop\ {\isasymphi}{\isacharbraceright}{\isachardoublequote}} prints a well-typed proposition \isa{{\isachardoublequote}{\isasymphi}{\isachardoublequote}}.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   243
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   244
  \item \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}lemma\ {\isasymphi}\ by\ m{\isacharbraceright}{\isachardoublequote}} proves a well-typed proposition
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   245
  \isa{{\isachardoublequote}{\isasymphi}{\isachardoublequote}} by method \isa{m} and prints the original \isa{{\isachardoublequote}{\isasymphi}{\isachardoublequote}}.
27453
eecd9d84e41b added lemma antiquotation
haftmann
parents: 27054
diff changeset
   246
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   247
  \item \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}term\ t{\isacharbraceright}{\isachardoublequote}} prints a well-typed term \isa{{\isachardoublequote}t{\isachardoublequote}}.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   248
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   249
  \item \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}const\ c{\isacharbraceright}{\isachardoublequote}} prints a logical or syntactic constant
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   250
  \isa{{\isachardoublequote}c{\isachardoublequote}}.
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   251
  
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   252
  \item \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}abbrev\ c\ x\isactrlsub {\isadigit{1}}\ {\isasymdots}\ x\isactrlsub n{\isacharbraceright}{\isachardoublequote}} prints a constant abbreviation
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   253
  \isa{{\isachardoublequote}c\ x\isactrlsub {\isadigit{1}}\ {\isasymdots}\ x\isactrlsub n\ {\isasymequiv}\ rhs{\isachardoublequote}} as defined in the current context.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   254
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   255
  \item \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}typeof\ t{\isacharbraceright}{\isachardoublequote}} prints the type of a well-typed term
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   256
  \isa{{\isachardoublequote}t{\isachardoublequote}}.
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   257
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   258
  \item \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}typ\ {\isasymtau}{\isacharbraceright}{\isachardoublequote}} prints a well-formed type \isa{{\isachardoublequote}{\isasymtau}{\isachardoublequote}}.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   259
  
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   260
  \item \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}thm{\isacharunderscore}style\ s\ a{\isacharbraceright}{\isachardoublequote}} prints theorem \isa{a},
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   261
  previously applying a style \isa{s} to it (see below).
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   262
  
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   263
  \item \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}term{\isacharunderscore}style\ s\ t{\isacharbraceright}{\isachardoublequote}} prints a well-typed term \isa{t}
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   264
  after applying a style \isa{s} to it (see below).
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   265
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   266
  \item \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}text\ s{\isacharbraceright}{\isachardoublequote}} prints uninterpreted source text \isa{s}.  This is particularly useful to print portions of text according
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   267
  to the Isabelle document style, without demanding well-formedness,
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   268
  e.g.\ small pieces of terms that should not be parsed or
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   269
  type-checked yet.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   270
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   271
  \item \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}goals{\isacharbraceright}{\isachardoublequote}} prints the current \emph{dynamic} goal
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   272
  state.  This is mainly for support of tactic-emulation scripts
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   273
  within Isar.  Presentation of goal states does not conform to the
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   274
  idea of human-readable proof documents!
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   275
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   276
  When explaining proofs in detail it is usually better to spell out
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   277
  the reasoning via proper Isar proof commands, instead of peeking at
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   278
  the internal machine configuration.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   279
  
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   280
  \item \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}subgoals{\isacharbraceright}{\isachardoublequote}} is similar to \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}goals{\isacharbraceright}{\isachardoublequote}}, but
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   281
  does not print the main goal.
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   282
  
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   283
  \item \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}prf\ a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n{\isacharbraceright}{\isachardoublequote}} prints the (compact) proof terms
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   284
  corresponding to the theorems \isa{{\isachardoublequote}a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n{\isachardoublequote}}. Note that this
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   285
  requires proof terms to be switched on for the current logic
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   286
  session.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   287
  
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   288
  \item \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}full{\isacharunderscore}prf\ a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n{\isacharbraceright}{\isachardoublequote}} is like \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}prf\ a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n{\isacharbraceright}{\isachardoublequote}}, but prints the full proof terms, i.e.\ also displays
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   289
  information omitted in the compact proof term, which is denoted by
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   290
  ``\isa{{\isacharunderscore}}'' placeholders there.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   291
  
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   292
  \item \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}ML\ s{\isacharbraceright}{\isachardoublequote}}, \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}ML{\isacharunderscore}type\ s{\isacharbraceright}{\isachardoublequote}}, and \isa{{\isachardoublequote}{\isacharat}{\isacharbraceleft}ML{\isacharunderscore}struct\ s{\isacharbraceright}{\isachardoublequote}} check text \isa{s} as ML value, type, and
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   293
  structure, respectively.  The source is printed verbatim.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   294
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   295
  \end{description}%
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   296
\end{isamarkuptext}%
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   297
\isamarkuptrue%
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   298
%
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   299
\isamarkupsubsubsection{Styled antiquotations%
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   300
}
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   301
\isamarkuptrue%
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   302
%
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   303
\begin{isamarkuptext}%
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   304
Some antiquotations like \isa{thm{\isacharunderscore}style} and \isa{term{\isacharunderscore}style} admit an extra \emph{style} specification to modify the
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   305
  printed result.  The following standard styles are available:
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   306
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   307
  \begin{description}
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   308
  
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   309
  \item \isa{lhs} extracts the first argument of any application
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   310
  form with at least two arguments --- typically meta-level or
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   311
  object-level equality, or any other binary relation.
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   312
  
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   313
  \item \isa{rhs} is like \isa{lhs}, but extracts the second
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   314
  argument.
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   315
  
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   316
  \item \isa{{\isachardoublequote}concl{\isachardoublequote}} extracts the conclusion \isa{C} from a rule
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   317
  in Horn-clause normal form \isa{{\isachardoublequote}A\isactrlsub {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isasymdots}\ A\isactrlsub n\ {\isasymLongrightarrow}\ C{\isachardoublequote}}.
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   318
  
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   319
  \item \isa{{\isachardoublequote}prem{\isadigit{1}}{\isachardoublequote}}, \dots, \isa{{\isachardoublequote}prem{\isadigit{9}}{\isachardoublequote}} extract premise number
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   320
  \isa{{\isachardoublequote}{\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isadigit{9}}{\isachardoublequote}}, respectively, from from a rule in Horn-clause
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   321
  normal form \isa{{\isachardoublequote}A\isactrlsub {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isasymdots}\ A\isactrlsub n\ {\isasymLongrightarrow}\ C{\isachardoublequote}}
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   322
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   323
  \end{description}%
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   324
\end{isamarkuptext}%
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   325
\isamarkuptrue%
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   326
%
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   327
\isamarkupsubsubsection{General options%
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   328
}
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   329
\isamarkuptrue%
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   330
%
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   331
\begin{isamarkuptext}%
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   332
The following options are available to tune the printed output
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   333
  of antiquotations.  Note that many of these coincide with global ML
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   334
  flags of the same names.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   335
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   336
  \begin{description}
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   337
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   338
  \item \indexdef{}{antiquotation option}{show\_types}\hypertarget{antiquotation option.show-types}{\hyperlink{antiquotation option.show-types}{\mbox{\isa{show{\isacharunderscore}types}}}}~\isa{{\isachardoublequote}{\isacharequal}\ bool{\isachardoublequote}} and
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   339
  \indexdef{}{antiquotation option}{show\_sorts}\hypertarget{antiquotation option.show-sorts}{\hyperlink{antiquotation option.show-sorts}{\mbox{\isa{show{\isacharunderscore}sorts}}}}~\isa{{\isachardoublequote}{\isacharequal}\ bool{\isachardoublequote}} control
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   340
  printing of explicit type and sort constraints.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   341
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   342
  \item \indexdef{}{antiquotation option}{show\_structs}\hypertarget{antiquotation option.show-structs}{\hyperlink{antiquotation option.show-structs}{\mbox{\isa{show{\isacharunderscore}structs}}}}~\isa{{\isachardoublequote}{\isacharequal}\ bool{\isachardoublequote}}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   343
  controls printing of implicit structures.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   344
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   345
  \item \indexdef{}{antiquotation option}{long\_names}\hypertarget{antiquotation option.long-names}{\hyperlink{antiquotation option.long-names}{\mbox{\isa{long{\isacharunderscore}names}}}}~\isa{{\isachardoublequote}{\isacharequal}\ bool{\isachardoublequote}} forces
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   346
  names of types and constants etc.\ to be printed in their fully
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   347
  qualified internal form.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   348
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   349
  \item \indexdef{}{antiquotation option}{short\_names}\hypertarget{antiquotation option.short-names}{\hyperlink{antiquotation option.short-names}{\mbox{\isa{short{\isacharunderscore}names}}}}~\isa{{\isachardoublequote}{\isacharequal}\ bool{\isachardoublequote}}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   350
  forces names of types and constants etc.\ to be printed unqualified.
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   351
  Note that internalizing the output again in the current context may
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   352
  well yield a different result.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   353
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   354
  \item \indexdef{}{antiquotation option}{unique\_names}\hypertarget{antiquotation option.unique-names}{\hyperlink{antiquotation option.unique-names}{\mbox{\isa{unique{\isacharunderscore}names}}}}~\isa{{\isachardoublequote}{\isacharequal}\ bool{\isachardoublequote}}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   355
  determines whether the printed version of qualified names should be
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   356
  made sufficiently long to avoid overlap with names declared further
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   357
  back.  Set to \isa{false} for more concise output.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   358
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   359
  \item \indexdef{}{antiquotation option}{eta\_contract}\hypertarget{antiquotation option.eta-contract}{\hyperlink{antiquotation option.eta-contract}{\mbox{\isa{eta{\isacharunderscore}contract}}}}~\isa{{\isachardoublequote}{\isacharequal}\ bool{\isachardoublequote}}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   360
  prints terms in \isa{{\isasymeta}}-contracted form.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   361
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   362
  \item \indexdef{}{antiquotation option}{display}\hypertarget{antiquotation option.display}{\hyperlink{antiquotation option.display}{\mbox{\isa{display}}}}~\isa{{\isachardoublequote}{\isacharequal}\ bool{\isachardoublequote}} indicates
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   363
  if the text is to be output as multi-line ``display material'',
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   364
  rather than a small piece of text without line breaks (which is the
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   365
  default).
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   366
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   367
  In this mode the embedded entities are printed in the same style as
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   368
  the main theory text.
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   369
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   370
  \item \indexdef{}{antiquotation option}{break}\hypertarget{antiquotation option.break}{\hyperlink{antiquotation option.break}{\mbox{\isa{break}}}}~\isa{{\isachardoublequote}{\isacharequal}\ bool{\isachardoublequote}} controls
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   371
  line breaks in non-display material.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   372
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   373
  \item \indexdef{}{antiquotation option}{quotes}\hypertarget{antiquotation option.quotes}{\hyperlink{antiquotation option.quotes}{\mbox{\isa{quotes}}}}~\isa{{\isachardoublequote}{\isacharequal}\ bool{\isachardoublequote}} indicates
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   374
  if the output should be enclosed in double quotes.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   375
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   376
  \item \indexdef{}{antiquotation option}{mode}\hypertarget{antiquotation option.mode}{\hyperlink{antiquotation option.mode}{\mbox{\isa{mode}}}}~\isa{{\isachardoublequote}{\isacharequal}\ name{\isachardoublequote}} adds \isa{name} to the print mode to be used for presentation.  Note that the
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   377
  standard setup for {\LaTeX} output is already present by default,
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   378
  including the modes \isa{latex} and \isa{xsymbols}.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   379
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   380
  \item \indexdef{}{antiquotation option}{margin}\hypertarget{antiquotation option.margin}{\hyperlink{antiquotation option.margin}{\mbox{\isa{margin}}}}~\isa{{\isachardoublequote}{\isacharequal}\ nat{\isachardoublequote}} and
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   381
  \indexdef{}{antiquotation option}{indent}\hypertarget{antiquotation option.indent}{\hyperlink{antiquotation option.indent}{\mbox{\isa{indent}}}}~\isa{{\isachardoublequote}{\isacharequal}\ nat{\isachardoublequote}} change the margin
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   382
  or indentation for pretty printing of display material.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   383
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   384
  \item \indexdef{}{antiquotation option}{goals\_limit}\hypertarget{antiquotation option.goals-limit}{\hyperlink{antiquotation option.goals-limit}{\mbox{\isa{goals{\isacharunderscore}limit}}}}~\isa{{\isachardoublequote}{\isacharequal}\ nat{\isachardoublequote}}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   385
  determines the maximum number of goals to be printed (for goal-based
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   386
  antiquotation).
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   387
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   388
  \item \indexdef{}{antiquotation option}{source}\hypertarget{antiquotation option.source}{\hyperlink{antiquotation option.source}{\mbox{\isa{source}}}}~\isa{{\isachardoublequote}{\isacharequal}\ bool{\isachardoublequote}} prints the
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   389
  original source text of the antiquotation arguments, rather than its
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   390
  internal representation.  Note that formal checking of
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   391
  \hyperlink{antiquotation.thm}{\mbox{\isa{thm}}}, \hyperlink{antiquotation.term}{\mbox{\isa{term}}}, etc. is still
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   392
  enabled; use the \hyperlink{antiquotation.text}{\mbox{\isa{text}}} antiquotation for unchecked
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   393
  output.
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   394
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   395
  Regular \isa{{\isachardoublequote}term{\isachardoublequote}} and \isa{{\isachardoublequote}typ{\isachardoublequote}} antiquotations with \isa{{\isachardoublequote}source\ {\isacharequal}\ false{\isachardoublequote}} involve a full round-trip from the original source
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   396
  to an internalized logical entity back to a source form, according
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   397
  to the syntax of the current context.  Thus the printed output is
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   398
  not under direct control of the author, it may even fluctuate a bit
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   399
  as the underlying theory is changed later on.
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   400
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   401
  In contrast, \indexdef{}{antiquotation option}{source}\hypertarget{antiquotation option.source}{\hyperlink{antiquotation option.source}{\mbox{\isa{source}}}}~\isa{{\isachardoublequote}{\isacharequal}\ true{\isachardoublequote}}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   402
  admits direct printing of the given source text, with the desirable
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   403
  well-formedness check in the background, but without modification of
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30172
diff changeset
   404
  the printed text.
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   405
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   406
  \end{description}
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   407
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   408
  For boolean flags, ``\isa{{\isachardoublequote}name\ {\isacharequal}\ true{\isachardoublequote}}'' may be abbreviated as
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   409
  ``\isa{name}''.  All of the above flags are disabled by default,
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   410
  unless changed from ML, say in the \verb|ROOT.ML| of the
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   411
  logic session.%
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   412
\end{isamarkuptext}%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   413
\isamarkuptrue%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   414
%
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   415
\isamarkupsection{Markup via command tags \label{sec:tags}%
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   416
}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   417
\isamarkuptrue%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   418
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   419
\begin{isamarkuptext}%
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   420
Each Isabelle/Isar command may be decorated by additional
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   421
  presentation tags, to indicate some modification in the way it is
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   422
  printed in the document.
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   423
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   424
  \indexouternonterm{tags}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   425
  \begin{rail}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   426
    tags: ( tag * )
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   427
    ;
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   428
    tag: '\%' (ident | string)
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   429
  \end{rail}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   430
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   431
  Some tags are pre-declared for certain classes of commands, serving
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   432
  as default markup if no tags are given in the text:
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   433
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   434
  \medskip
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   435
  \begin{tabular}{ll}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   436
    \isa{{\isachardoublequote}theory{\isachardoublequote}} & theory begin/end \\
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   437
    \isa{{\isachardoublequote}proof{\isachardoublequote}} & all proof commands \\
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   438
    \isa{{\isachardoublequote}ML{\isachardoublequote}} & all commands involving ML code \\
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   439
  \end{tabular}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   440
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   441
  \medskip The Isabelle document preparation system
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   442
  \cite{isabelle-sys} allows tagged command regions to be presented
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   443
  specifically, e.g.\ to fold proof texts, or drop parts of the text
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   444
  completely.
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   445
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   446
  For example ``\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{{\isachardoublequote}{\isacharpercent}invisible\ auto{\isachardoublequote}}'' causes
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   447
  that piece of proof to be treated as \isa{invisible} instead of
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   448
  \isa{{\isachardoublequote}proof{\isachardoublequote}} (the default), which may be shown or hidden
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   449
  depending on the document setup.  In contrast, ``\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{{\isachardoublequote}{\isacharpercent}visible\ auto{\isachardoublequote}}'' forces this text to be shown
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   450
  invariably.
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   451
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   452
  Explicit tag specifications within a proof apply to all subsequent
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   453
  commands of the same level of nesting.  For example, ``\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isachardoublequote}{\isacharpercent}visible\ {\isasymdots}{\isachardoublequote}}~\hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}'' forces the whole
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   454
  sub-proof to be typeset as \isa{visible} (unless some of its parts
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   455
  are tagged differently).
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   456
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   457
  \medskip Command tags merely produce certain markup environments for
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   458
  type-setting.  The meaning of these is determined by {\LaTeX}
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   459
  macros, as defined in \hyperlink{file.~~/lib/texinputs/isabelle.sty}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}lib{\isacharslash}texinputs{\isacharslash}isabelle{\isachardot}sty}}}} or
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   460
  by the document author.  The Isabelle document preparation tools
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   461
  also provide some high-level options to specify the meaning of
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   462
  arbitrary tags to ``keep'', ``drop'', or ``fold'' the corresponding
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   463
  parts of the text.  Logic sessions may also specify ``document
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   464
  versions'', where given tags are interpreted in some particular way.
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   465
  Again see \cite{isabelle-sys} for further details.%
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   466
\end{isamarkuptext}%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   467
\isamarkuptrue%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   468
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   469
\isamarkupsection{Draft presentation%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   470
}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   471
\isamarkuptrue%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   472
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   473
\begin{isamarkuptext}%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   474
\begin{matharray}{rcl}
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   475
    \indexdef{}{command}{display\_drafts}\hypertarget{command.display-drafts}{\hyperlink{command.display-drafts}{\mbox{\isa{\isacommand{display{\isacharunderscore}drafts}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}any\ {\isasymrightarrow}{\isachardoublequote}} \\
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   476
    \indexdef{}{command}{print\_drafts}\hypertarget{command.print-drafts}{\hyperlink{command.print-drafts}{\mbox{\isa{\isacommand{print{\isacharunderscore}drafts}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}any\ {\isasymrightarrow}{\isachardoublequote}} \\
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   477
  \end{matharray}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   478
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   479
  \begin{rail}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   480
    ('display\_drafts' | 'print\_drafts') (name +)
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   481
    ;
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   482
  \end{rail}
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   483
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   484
  \begin{description}
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   485
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   486
  \item \hyperlink{command.display-drafts}{\mbox{\isa{\isacommand{display{\isacharunderscore}drafts}}}}~\isa{paths} and \hyperlink{command.print-drafts}{\mbox{\isa{\isacommand{print{\isacharunderscore}drafts}}}}~\isa{paths} perform simple output of a given list
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   487
  of raw source files.  Only those symbols that do not require
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   488
  additional {\LaTeX} packages are displayed properly, everything else
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   489
  is left verbatim.
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   490
28788
ff9d8a8932e4 updated generated files;
wenzelm
parents: 28505
diff changeset
   491
  \end{description}%
27042
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   492
\end{isamarkuptext}%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   493
\isamarkuptrue%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   494
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   495
\isadelimtheory
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   496
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   497
\endisadelimtheory
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   498
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   499
\isatagtheory
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   500
\isacommand{end}\isamarkupfalse%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   501
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   502
\endisatagtheory
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   503
{\isafoldtheory}%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   504
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   505
\isadelimtheory
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   506
%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   507
\endisadelimtheory
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   508
\isanewline
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   509
\end{isabellebody}%
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   510
%%% Local Variables:
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   511
%%% mode: latex
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   512
%%% TeX-master: "root"
8fcf19f2168b updated generated file;
wenzelm
parents:
diff changeset
   513
%%% End: