| author | wenzelm | 
| Sat, 10 Oct 2020 21:45:58 +0200 | |
| changeset 72428 | b7351ffe0dbc | 
| parent 72309 | 564012e31db1 | 
| child 72574 | d892f6d66402 | 
| permissions | -rw-r--r-- | 
| 61656 | 1 | (*:maxLineLen=78:*) | 
| 61575 | 2 | |
| 28221 | 3 | theory Presentation | 
| 43564 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42009diff
changeset | 4 | imports Base | 
| 28221 | 5 | begin | 
| 6 | ||
| 58618 | 7 | chapter \<open>Presenting theories \label{ch:present}\<close>
 | 
| 28221 | 8 | |
| 61575 | 9 | text \<open> | 
| 10 | Isabelle provides several ways to present the outcome of formal | |
| 11 | developments, including WWW-based browsable libraries or actual printable | |
| 12 | documents. Presentation is centered around the concept of \<^emph>\<open>sessions\<close> | |
| 13 |   (\chref{ch:session}). The global session structure is that of a tree, with
 | |
| 14 | Isabelle Pure at its root, further object-logics derived (e.g.\ HOLCF from | |
| 15 | HOL, and HOL from Pure), and application sessions further on in the | |
| 16 | hierarchy. | |
| 28221 | 17 | |
| 61575 | 18 |   The tools @{tool_ref mkroot} and @{tool_ref build} provide the primary means
 | 
| 19 | for managing Isabelle sessions, including proper setup for presentation; | |
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62013diff
changeset | 20 |   @{tool build} tells the Isabelle process to run any additional stages
 | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62013diff
changeset | 21 |   required for document preparation, notably the @{tool_ref document} and
 | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62013diff
changeset | 22 |   @{tool_ref latex}. The complete tool chain for managing batch-mode Isabelle
 | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62013diff
changeset | 23 |   sessions is illustrated in \figref{fig:session-tools}.
 | 
| 28221 | 24 | |
| 25 |   \begin{figure}[htbp]
 | |
| 26 |   \begin{center}
 | |
| 27 |   \begin{tabular}{lp{0.6\textwidth}}
 | |
| 28 | ||
| 48814 
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
 wenzelm parents: 
48722diff
changeset | 29 |       @{tool_ref mkroot} & invoked once by the user to initialize the
 | 
| 61503 | 30 | session \<^verbatim>\<open>ROOT\<close> with optional \<^verbatim>\<open>document\<close> | 
| 48814 
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
 wenzelm parents: 
48722diff
changeset | 31 | directory; \\ | 
| 28221 | 32 | |
| 48814 
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
 wenzelm parents: 
48722diff
changeset | 33 |       @{tool_ref build} & invoked repeatedly by the user to keep
 | 
| 
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
 wenzelm parents: 
48722diff
changeset | 34 | session output up-to-date (HTML, documents etc.); \\ | 
| 28221 | 35 | |
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62013diff
changeset | 36 |       @{tool_ref process} & run through @{tool_ref build}; \\
 | 
| 28221 | 37 | |
| 48602 | 38 |       @{tool_ref document} & run by the Isabelle process if document
 | 
| 39 | preparation is enabled; \\ | |
| 28221 | 40 | |
| 48602 | 41 |       @{tool_ref latex} & universal {\LaTeX} tool wrapper invoked
 | 
| 42 |       multiple times by @{tool_ref document}; also useful for manual
 | |
| 43 | experiments; \\ | |
| 28221 | 44 | |
| 45 |   \end{tabular}
 | |
| 46 |   \caption{The tool chain of Isabelle session presentation} \label{fig:session-tools}
 | |
| 47 |   \end{center}
 | |
| 48 |   \end{figure}
 | |
| 58618 | 49 | \<close> | 
| 28221 | 50 | |
| 51 | ||
| 62013 | 52 | section \<open>Generating HTML browser information \label{sec:info}\<close>
 | 
| 28221 | 53 | |
| 58618 | 54 | text \<open> | 
| 61575 | 55 | As a side-effect of building sessions, Isabelle is able to generate theory | 
| 56 | browsing information, including HTML documents that show the theory sources | |
| 57 | and the relationship with its ancestors and descendants. Besides the HTML | |
| 58 | file that is generated for every theory, Isabelle stores links to all | |
| 59 | theories of a session in an index file. As a second hierarchy, groups of | |
| 60 | sessions are organized as \<^emph>\<open>chapters\<close>, with a separate index. Note that the | |
| 61 | implicit tree structure of the session build hierarchy is \<^emph>\<open>not\<close> relevant | |
| 51417 | 62 | for the presentation. | 
| 28221 | 63 | |
| 61406 | 64 | \<^medskip> | 
| 62013 | 65 | To generate theory browsing information for an existing session, just invoke | 
| 66 |   @{tool build} with suitable options:
 | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 67 |   @{verbatim [display] \<open>isabelle build -o browser_info -v -c FOL\<close>}
 | 
| 28221 | 68 | |
| 61575 | 69 | The presentation output will appear in \<^verbatim>\<open>$ISABELLE_BROWSER_INFO/FOL/FOL\<close> as | 
| 70 | reported by the above verbose invocation of the build process. | |
| 28221 | 71 | |
| 67219 | 72 | Many Isabelle sessions (such as \<^session>\<open>HOL-Library\<close> in | 
| 73 | \<^dir>\<open>~~/src/HOL/Library\<close>) also provide printable documents in PDF. These are | |
| 74 | prepared automatically as well if enabled like this: | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 75 |   @{verbatim [display] \<open>isabelle build -o browser_info -o document=pdf -v -c HOL-Library\<close>}
 | 
| 48814 
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
 wenzelm parents: 
48722diff
changeset | 76 | |
| 61575 | 77 | Enabling both browser info and document preparation simultaneously causes an | 
| 78 | appropriate ``document'' link to be included in the HTML index. Documents | |
| 79 | may be generated independently of browser information as well, see | |
| 80 |   \secref{sec:tool-document} for further details.
 | |
| 28221 | 81 | |
| 61406 | 82 | \<^bigskip> | 
| 61575 | 83 | The theory browsing information is stored in a sub-directory directory | 
| 84 |   determined by the @{setting_ref ISABELLE_BROWSER_INFO} setting plus a prefix
 | |
| 85 | corresponding to the session chapter and identifier. In order to present | |
| 86 | Isabelle applications on the web, the corresponding subdirectory from | |
| 87 |   @{setting ISABELLE_BROWSER_INFO} can be put on a WWW server.
 | |
| 88 | \<close> | |
| 28221 | 89 | |
| 61406 | 90 | |
| 58618 | 91 | section \<open>Preparing session root directories \label{sec:tool-mkroot}\<close>
 | 
| 28221 | 92 | |
| 61575 | 93 | text \<open> | 
| 94 |   The @{tool_def mkroot} tool configures a given directory as session root,
 | |
| 95 | with some \<^verbatim>\<open>ROOT\<close> file and optional document source directory. Its usage is: | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 96 |   @{verbatim [display]
 | 
| 67043 | 97 | \<open>Usage: isabelle mkroot [OPTIONS] [DIRECTORY] | 
| 28221 | 98 | |
| 99 | Options are: | |
| 67043 | 100 | -A LATEX provide author in LaTeX notation (default: user name) | 
| 67069 
f11486d31586
init Mercurial repository for the generated session files;
 wenzelm parents: 
67043diff
changeset | 101 | -I init Mercurial repository and add generated files | 
| 67043 | 102 | -T LATEX provide title in LaTeX notation (default: session name) | 
| 103 | -n NAME alternative session name (default: directory base name) | |
| 28221 | 104 | |
| 67043 | 105 | Prepare session root directory (default: current directory). | 
| 106 | \<close>} | |
| 28221 | 107 | |
| 61575 | 108 | The results are placed in the given directory \<open>dir\<close>, which refers to the | 
| 109 |   current directory by default. The @{tool mkroot} tool is conservative in the
 | |
| 110 | sense that it does not overwrite existing files or directories. Earlier | |
| 111 | attempts to generate a session root need to be deleted manually. | |
| 28221 | 112 | |
| 67042 
677cab7c2b85
adapted to changed ROOT syntax (see 13857f49d215);
 wenzelm parents: 
63680diff
changeset | 113 | The generated session template will be accompanied by a formal document, | 
| 67043 | 114 |   with \<open>DIRECTORY\<close>\<^verbatim>\<open>/document/root.tex\<close> as its {\LaTeX} entry point (see also
 | 
| 67042 
677cab7c2b85
adapted to changed ROOT syntax (see 13857f49d215);
 wenzelm parents: 
63680diff
changeset | 115 |   \chref{ch:present}).
 | 
| 28221 | 116 | |
| 67043 | 117 | Options \<^verbatim>\<open>-T\<close> and \<^verbatim>\<open>-A\<close> specify the document title and author explicitly, | 
| 118 |   using {\LaTeX} source notation.
 | |
| 119 | ||
| 67069 
f11486d31586
init Mercurial repository for the generated session files;
 wenzelm parents: 
67043diff
changeset | 120 | Option \<^verbatim>\<open>-I\<close> initializes a Mercurial repository in the target directory, and | 
| 
f11486d31586
init Mercurial repository for the generated session files;
 wenzelm parents: 
67043diff
changeset | 121 | adds all generated files (without commit). | 
| 
f11486d31586
init Mercurial repository for the generated session files;
 wenzelm parents: 
67043diff
changeset | 122 | |
| 67043 | 123 | Option \<^verbatim>\<open>-n\<close> specifies an alternative session name; otherwise the base name | 
| 124 | of the given directory is used. | |
| 28221 | 125 | |
| 61406 | 126 | \<^medskip> | 
| 61575 | 127 |   The implicit Isabelle settings variable @{setting ISABELLE_LOGIC} specifies
 | 
| 67043 | 128 | the parent session. | 
| 61503 | 129 | \<close> | 
| 28221 | 130 | |
| 131 | ||
| 58618 | 132 | subsubsection \<open>Examples\<close> | 
| 28221 | 133 | |
| 61575 | 134 | text \<open> | 
| 67042 
677cab7c2b85
adapted to changed ROOT syntax (see 13857f49d215);
 wenzelm parents: 
63680diff
changeset | 135 | Produce session \<^verbatim>\<open>Test\<close> within a separate directory of the same name: | 
| 
677cab7c2b85
adapted to changed ROOT syntax (see 13857f49d215);
 wenzelm parents: 
63680diff
changeset | 136 |   @{verbatim [display] \<open>isabelle mkroot Test && isabelle build -D Test\<close>}
 | 
| 28221 | 137 | |
| 61406 | 138 | \<^medskip> | 
| 61575 | 139 | Upgrade the current directory into a session ROOT with document preparation, | 
| 140 | and build it: | |
| 67042 
677cab7c2b85
adapted to changed ROOT syntax (see 13857f49d215);
 wenzelm parents: 
63680diff
changeset | 141 |   @{verbatim [display] \<open>isabelle mkroot && isabelle build -D .\<close>}
 | 
| 58618 | 142 | \<close> | 
| 28221 | 143 | |
| 144 | ||
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 145 | section \<open>Preparing Isabelle session documents \label{sec:tool-document}\<close>
 | 
| 28221 | 146 | |
| 61575 | 147 | text \<open> | 
| 148 |   The @{tool_def document} tool prepares logic session documents, processing
 | |
| 149 | the sources as provided by the user and generated by Isabelle. Its usage is: | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 150 |   @{verbatim [display]
 | 
| 67176 
13b5c3ff1954
re-implemented "isabelle document" in Isabelle/Scala, include latex_errors here;
 wenzelm parents: 
67151diff
changeset | 151 | \<open>Usage: isabelle document [OPTIONS] | 
| 28221 | 152 | |
| 153 | Options are: | |
| 67176 
13b5c3ff1954
re-implemented "isabelle document" in Isabelle/Scala, include latex_errors here;
 wenzelm parents: 
67151diff
changeset | 154 | -d DIR document output directory (default "output/document") | 
| 
13b5c3ff1954
re-implemented "isabelle document" in Isabelle/Scala, include latex_errors here;
 wenzelm parents: 
67151diff
changeset | 155 | -n NAME document name (default "document") | 
| 
13b5c3ff1954
re-implemented "isabelle document" in Isabelle/Scala, include latex_errors here;
 wenzelm parents: 
67151diff
changeset | 156 | -t TAGS markup for tagged regions | 
| 28221 | 157 | |
| 67176 
13b5c3ff1954
re-implemented "isabelle document" in Isabelle/Scala, include latex_errors here;
 wenzelm parents: 
67151diff
changeset | 158 | Prepare the theory session document in document directory, producing the | 
| 
13b5c3ff1954
re-implemented "isabelle document" in Isabelle/Scala, include latex_errors here;
 wenzelm parents: 
67151diff
changeset | 159 | specified output format. | 
| 
13b5c3ff1954
re-implemented "isabelle document" in Isabelle/Scala, include latex_errors here;
 wenzelm parents: 
67151diff
changeset | 160 | \<close>} | 
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 161 | |
| 48814 
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
 wenzelm parents: 
48722diff
changeset | 162 | This tool is usually run automatically as part of the Isabelle build | 
| 
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
 wenzelm parents: 
48722diff
changeset | 163 | process, provided document preparation has been enabled via suitable | 
| 61575 | 164 | options. It may be manually invoked on the generated browser information | 
| 165 | document output as well, e.g.\ in case of errors encountered in the batch | |
| 166 | run. | |
| 28221 | 167 | |
| 61406 | 168 | \<^medskip> | 
| 67351 | 169 | Option \<^verbatim>\<open>-d\<close> specifies an alternative document output directory. The default | 
| 67176 
13b5c3ff1954
re-implemented "isabelle document" in Isabelle/Scala, include latex_errors here;
 wenzelm parents: 
67151diff
changeset | 170 | is \<^verbatim>\<open>output/document\<close> (derived from the document name). Note that the result | 
| 
13b5c3ff1954
re-implemented "isabelle document" in Isabelle/Scala, include latex_errors here;
 wenzelm parents: 
67151diff
changeset | 171 | will appear in the parent of this directory. | 
| 
13b5c3ff1954
re-implemented "isabelle document" in Isabelle/Scala, include latex_errors here;
 wenzelm parents: 
67151diff
changeset | 172 | |
| 
13b5c3ff1954
re-implemented "isabelle document" in Isabelle/Scala, include latex_errors here;
 wenzelm parents: 
67151diff
changeset | 173 | \<^medskip> | 
| 72309 
564012e31db1
discontinued obsolete DVI document format and related settings/tools;
 wenzelm parents: 
67399diff
changeset | 174 | Option \<^verbatim>\<open>-n\<close> specifies the resulting document file name, the default is | 
| 
564012e31db1
discontinued obsolete DVI document format and related settings/tools;
 wenzelm parents: 
67399diff
changeset | 175 | ``\<^verbatim>\<open>document\<close>'' (leading to ``\<^verbatim>\<open>document.pdf\<close>''). | 
| 28221 | 176 | |
| 61406 | 177 | \<^medskip> | 
| 61575 | 178 |   The \<^verbatim>\<open>-t\<close> option tells {\LaTeX} how to interpret tagged Isabelle command
 | 
| 179 | regions. Tags are specified as a comma separated list of modifier/name | |
| 180 | pairs: ``\<^verbatim>\<open>+\<close>\<open>foo\<close>'' (or just ``\<open>foo\<close>'') means to keep, ``\<^verbatim>\<open>-\<close>\<open>foo\<close>'' to | |
| 181 | drop, and ``\<^verbatim>\<open>/\<close>\<open>foo\<close>'' to fold text tagged as \<open>foo\<close>. The builtin default is | |
| 182 | equivalent to the tag specification | |
| 67151 | 183 | ``\<^verbatim>\<open>+document,+theory,+proof,+ML,+visible,-invisible,+important,+unimportant\<close>''; | 
| 184 |   see also the {\LaTeX} macros \<^verbatim>\<open>\isakeeptag\<close>, \<^verbatim>\<open>\isadroptag\<close>, and
 | |
| 185 | \<^verbatim>\<open>\isafoldtag\<close>, in \<^file>\<open>~~/lib/texinputs/isabelle.sty\<close>. | |
| 28221 | 186 | |
| 61406 | 187 | \<^medskip> | 
| 61575 | 188 | Document preparation requires a \<^verbatim>\<open>document\<close> directory within the session | 
| 189 | sources. This directory is supposed to contain all the files needed to | |
| 190 | produce the final document --- apart from the actual theories which are | |
| 191 | generated by Isabelle. | |
| 28221 | 192 | |
| 61406 | 193 | \<^medskip> | 
| 61575 | 194 |   For most practical purposes, @{tool document} is smart enough to create any
 | 
| 195 | of the specified output formats, taking \<^verbatim>\<open>root.tex\<close> supplied by the user as | |
| 196 |   a starting point. This even includes multiple runs of {\LaTeX} to
 | |
| 197 | accommodate references and bibliographies (the latter assumes \<^verbatim>\<open>root.bib\<close> | |
| 198 | within the same directory). | |
| 28221 | 199 | |
| 61575 | 200 | In more complex situations, a separate \<^verbatim>\<open>build\<close> script for the document | 
| 201 | sources may be given. It is invoked with command-line arguments for the | |
| 202 | document format and the document variant name. The script needs to produce | |
| 203 | corresponding output files, e.g.\ \<^verbatim>\<open>root.pdf\<close> for target format \<^verbatim>\<open>pdf\<close> (and | |
| 204 |   default variants). The main work can be again delegated to @{tool latex},
 | |
| 205 |   but it is also possible to harvest generated {\LaTeX} sources and copy them
 | |
| 206 | elsewhere. | |
| 28221 | 207 | |
| 61406 | 208 | \<^medskip> | 
| 61575 | 209 | When running the session, Isabelle copies the content of the original | 
| 210 |   \<^verbatim>\<open>document\<close> directory into its proper place within @{setting
 | |
| 211 | ISABELLE_BROWSER_INFO}, according to the session path and document variant. | |
| 212 |   Then, for any processed theory \<open>A\<close> some {\LaTeX} source is generated and put
 | |
| 213 | there as \<open>A\<close>\<^verbatim>\<open>.tex\<close>. Furthermore, a list of all generated theory files is | |
| 214 |   put into \<^verbatim>\<open>session.tex\<close>. Typically, the root {\LaTeX} file provided by the
 | |
| 215 | user would include \<^verbatim>\<open>session.tex\<close> to get a document containing all the | |
| 216 | theories. | |
| 28221 | 217 | |
| 63680 | 218 |   The {\LaTeX} versions of the theories require some macros defined in
 | 
| 219 |   \<^file>\<open>~~/lib/texinputs/isabelle.sty\<close>. Doing \<^verbatim>\<open>\usepackage{isabelle}\<close> in
 | |
| 61575 | 220 |   \<^verbatim>\<open>root.tex\<close> should be fine; the underlying @{tool latex} already includes an
 | 
| 221 |   appropriate path specification for {\TeX} inputs.
 | |
| 28221 | 222 | |
| 61575 | 223 | If the text contains any references to Isabelle symbols (such as \<^verbatim>\<open>\<forall>\<close>) then | 
| 224 | \<^verbatim>\<open>isabellesym.sty\<close> should be included as well. This package contains a | |
| 225 |   standard set of {\LaTeX} macro definitions \<^verbatim>\<open>\isasym\<close>\<open>foo\<close> corresponding to
 | |
| 226 |   \<^verbatim>\<open>\\<close>\<^verbatim>\<open><\<close>\<open>foo\<close>\<^verbatim>\<open>>\<close>, see @{cite "isabelle-implementation"} for a complete list
 | |
| 227 | of predefined Isabelle symbols. Users may invent further symbols as well, | |
| 63680 | 228 |   just by providing {\LaTeX} macros in a similar fashion as in
 | 
| 229 | \<^file>\<open>~~/lib/texinputs/isabellesym.sty\<close> of the Isabelle distribution. | |
| 28221 | 230 | |
| 72309 
564012e31db1
discontinued obsolete DVI document format and related settings/tools;
 wenzelm parents: 
67399diff
changeset | 231 | For proper setup of PDF documents (with hyperlinks and bookmarks), | 
| 63680 | 232 | we recommend to include \<^file>\<open>~~/lib/texinputs/pdfsetup.sty\<close> as well. | 
| 28221 | 233 | |
| 61406 | 234 | \<^medskip> | 
| 67185 | 235 |   As a final step of Isabelle document preparation, @{tool document} is run on
 | 
| 236 | the generated \<^verbatim>\<open>output/document\<close> directory. Thus the actual output document | |
| 237 | is built and installed in its proper place. The generated sources are | |
| 238 |   deleted after successful run of {\LaTeX} and friends.
 | |
| 48814 
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
 wenzelm parents: 
48722diff
changeset | 239 | |
| 
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
 wenzelm parents: 
48722diff
changeset | 240 | Some care is needed if the document output location is configured | 
| 
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
 wenzelm parents: 
48722diff
changeset | 241 | differently, say within a directory whose content is still required | 
| 
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
 wenzelm parents: 
48722diff
changeset | 242 | afterwards! | 
| 58618 | 243 | \<close> | 
| 28221 | 244 | |
| 245 | ||
| 58618 | 246 | section \<open>Running {\LaTeX} within the Isabelle environment
 | 
| 247 |   \label{sec:tool-latex}\<close>
 | |
| 28221 | 248 | |
| 61575 | 249 | text \<open> | 
| 250 |   The @{tool_def latex} tool provides the basic interface for Isabelle
 | |
| 251 | document preparation. Its usage is: | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 252 |   @{verbatim [display]
 | 
| 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 253 | \<open>Usage: isabelle latex [OPTIONS] [FILE] | 
| 28221 | 254 | |
| 255 | Options are: | |
| 72309 
564012e31db1
discontinued obsolete DVI document format and related settings/tools;
 wenzelm parents: 
67399diff
changeset | 256 | -o FORMAT specify output format: pdf (default), bbl, idx, sty | 
| 28221 | 257 | |
| 258 | Run LaTeX (and related tools) on FILE (default root.tex), | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 259 | producing the specified output format.\<close>} | 
| 28221 | 260 | |
| 61575 | 261 |   Appropriate {\LaTeX}-related programs are run on the input file, according
 | 
| 72309 
564012e31db1
discontinued obsolete DVI document format and related settings/tools;
 wenzelm parents: 
67399diff
changeset | 262 |   to the given output format: @{executable pdflatex}, @{executable bibtex}
 | 
| 
564012e31db1
discontinued obsolete DVI document format and related settings/tools;
 wenzelm parents: 
67399diff
changeset | 263 |   (for \<^verbatim>\<open>bbl\<close>), and @{executable makeindex} (for \<^verbatim>\<open>idx\<close>). The actual commands
 | 
| 
564012e31db1
discontinued obsolete DVI document format and related settings/tools;
 wenzelm parents: 
67399diff
changeset | 264 |   are determined from the settings environment (@{setting ISABELLE_PDFLATEX}
 | 
| 
564012e31db1
discontinued obsolete DVI document format and related settings/tools;
 wenzelm parents: 
67399diff
changeset | 265 | etc.). | 
| 28221 | 266 | |
| 61575 | 267 | The \<^verbatim>\<open>sty\<close> output format causes the Isabelle style files to be updated from | 
| 268 | the distribution. This is useful in special situations where the document | |
| 269 | sources are to be processed another time by separate tools. | |
| 58618 | 270 | \<close> | 
| 28221 | 271 | |
| 272 | ||
| 58618 | 273 | subsubsection \<open>Examples\<close> | 
| 28221 | 274 | |
| 61575 | 275 | text \<open> | 
| 276 |   Invoking @{tool latex} by hand may be occasionally useful when debugging
 | |
| 277 | failed attempts of the automatic document preparation stage of batch-mode | |
| 278 | Isabelle. The abortive process leaves the sources at a certain place within | |
| 279 |   @{setting ISABELLE_BROWSER_INFO}, see the runtime error message for details.
 | |
| 280 |   This enables users to inspect {\LaTeX} runs in further detail, e.g.\ like
 | |
| 281 | this: | |
| 48602 | 282 | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 283 |   @{verbatim [display]
 | 
| 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 284 | \<open>cd "$(isabelle getenv -b ISABELLE_BROWSER_INFO)/Unsorted/Test/document" | 
| 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 285 | isabelle latex -o pdf\<close>} | 
| 58618 | 286 | \<close> | 
| 28221 | 287 | |
| 67399 | 288 | end |