author | haftmann |
Wed, 14 Jul 2010 16:45:30 +0200 | |
changeset 37828 | 9e1758c7ff06 |
parent 35587 | f037aa6699c3 |
child 40387 | e4c9e0dad473 |
permissions | -rw-r--r-- |
28221 | 1 |
theory Presentation |
2 |
imports Pure |
|
3 |
begin |
|
4 |
||
5 |
chapter {* Presenting theories \label{ch:present} *} |
|
6 |
||
7 |
text {* |
|
8 |
Isabelle provides several ways to present the outcome of formal |
|
9 |
developments, including WWW-based browsable libraries or actual |
|
10 |
printable documents. Presentation is centered around the concept of |
|
11 |
\emph{logic sessions}. The global session structure is that of a |
|
12 |
tree, with Isabelle Pure at its root, further object-logics derived |
|
13 |
(e.g.\ HOLCF from HOL, and HOL from Pure), and application sessions |
|
14 |
in leaf positions (usually without a separate image). |
|
15 |
||
16 |
The Isabelle tools @{tool_ref mkdir} and @{tool_ref make} provide |
|
17 |
the primary means for managing Isabelle sessions, including proper |
|
18 |
setup for presentation. Here the @{tool_ref usedir} tool takes care |
|
19 |
to let @{executable_ref "isabelle-process"} process run any |
|
20 |
additional stages required for document preparation, notably the |
|
21 |
tools @{tool_ref document} and @{tool_ref latex}. The complete tool |
|
22 |
chain for managing batch-mode Isabelle sessions is illustrated in |
|
23 |
\figref{fig:session-tools}. |
|
24 |
||
25 |
\begin{figure}[htbp] |
|
26 |
\begin{center} |
|
27 |
\begin{tabular}{lp{0.6\textwidth}} |
|
28 |
||
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
29 |
@{verbatim isabelle} @{tool_ref mkdir} & invoked once by the user |
28238 | 30 |
to create the initial source setup (common @{verbatim |
31 |
IsaMakefile} plus a single session directory); \\ |
|
28221 | 32 |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
33 |
@{verbatim isabelle} @{tool make} & invoked repeatedly by the |
28238 | 34 |
user to keep session output up-to-date (HTML, documents etc.); \\ |
35 |
||
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
36 |
@{verbatim isabelle} @{tool usedir} & part of the standard |
28238 | 37 |
@{verbatim IsaMakefile} entry of a session; \\ |
28221 | 38 |
|
28238 | 39 |
@{executable "isabelle-process"} & run through @{verbatim |
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
40 |
isabelle} @{tool_ref usedir}; \\ |
28221 | 41 |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
42 |
@{verbatim isabelle} @{tool_ref document} & run by the Isabelle |
28238 | 43 |
process if document preparation is enabled; \\ |
28221 | 44 |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
45 |
@{verbatim isabelle} @{tool_ref latex} & universal {\LaTeX} tool |
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
46 |
wrapper invoked multiple times by @{verbatim isabelle} @{tool_ref |
28238 | 47 |
document}; also useful for manual experiments; \\ |
28221 | 48 |
|
49 |
\end{tabular} |
|
50 |
\caption{The tool chain of Isabelle session presentation} \label{fig:session-tools} |
|
51 |
\end{center} |
|
52 |
\end{figure} |
|
53 |
*} |
|
54 |
||
55 |
||
56 |
section {* Generating theory browser information \label{sec:info} *} |
|
57 |
||
58 |
text {* |
|
59 |
\index{theory browsing information|bold} |
|
60 |
||
61 |
As a side-effect of running a logic sessions, Isabelle is able to |
|
62 |
generate theory browsing information, including HTML documents that |
|
63 |
show a theory's definition, the theorems proved in its ML file and |
|
64 |
the relationship with its ancestors and descendants. Besides the |
|
65 |
HTML file that is generated for every theory, Isabelle stores links |
|
66 |
to all theories in an index file. These indexes are linked with |
|
67 |
other indexes to represent the overall tree structure of logic |
|
68 |
sessions. |
|
69 |
||
70 |
Isabelle also generates graph files that represent the theory |
|
71 |
hierarchy of a logic. There is a graph browser Java applet embedded |
|
72 |
in the generated HTML pages, and also a stand-alone application that |
|
73 |
allows browsing theory graphs without having to start a WWW client |
|
74 |
first. The latter version also includes features such as generating |
|
75 |
Postscript files, which are not available in the applet version. |
|
76 |
See \secref{sec:browse} for further information. |
|
77 |
||
78 |
\medskip |
|
79 |
||
80 |
The easiest way to let Isabelle generate theory browsing information |
|
81 |
for existing sessions is to append ``@{verbatim "-i true"}'' to the |
|
82 |
@{setting_ref ISABELLE_USEDIR_OPTIONS} before invoking @{verbatim |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
83 |
isabelle} @{tool make} (or @{"file" "$ISABELLE_HOME/build"}). For |
28221 | 84 |
example, add something like this to your Isabelle settings file |
85 |
||
86 |
\begin{ttbox} |
|
87 |
ISABELLE_USEDIR_OPTIONS="-i true" |
|
88 |
\end{ttbox} |
|
89 |
||
28238 | 90 |
and then change into the @{"file" "~~/src/FOL"} directory and run |
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
91 |
@{verbatim isabelle} @{tool make}, or even @{verbatim isabelle} @{tool |
28238 | 92 |
make}~@{verbatim all}. The presentation output will appear in |
93 |
@{verbatim "ISABELLE_BROWSER_INFO/FOL"}, which usually refers to |
|
28914
f993cbffc42a
default for ISABELLE_HOME_USER is now ~/.isabelle instead of ~/isabelle;
wenzelm
parents:
28838
diff
changeset
|
94 |
@{verbatim "~/.isabelle/browser_info/FOL"}. Note that option |
28221 | 95 |
@{verbatim "-v true"} will make the internal runs of @{tool usedir} |
96 |
more explicit about such details. |
|
97 |
||
28238 | 98 |
Many standard Isabelle sessions (such as @{"file" "~~/src/HOL/ex"}) |
99 |
also provide actual printable documents. These are prepared |
|
28221 | 100 |
automatically as well if enabled like this, using the @{verbatim |
101 |
"-d"} option |
|
102 |
\begin{ttbox} |
|
103 |
ISABELLE_USEDIR_OPTIONS="-i true -d dvi" |
|
104 |
\end{ttbox} |
|
105 |
Enabling options @{verbatim "-i"} and @{verbatim "-d"} |
|
28225 | 106 |
simultaneously as shown above causes an appropriate ``document'' |
28221 | 107 |
link to be included in the HTML index. Documents (or raw document |
108 |
sources) may be generated independently of browser information as |
|
109 |
well, see \secref{sec:tool-document} for further details. |
|
110 |
||
111 |
\bigskip The theory browsing information is stored in a |
|
112 |
sub-directory directory determined by the @{setting_ref |
|
113 |
ISABELLE_BROWSER_INFO} setting plus a prefix corresponding to the |
|
114 |
session identifier (according to the tree structure of sub-sessions |
|
115 |
by default). A complete WWW view of all standard object-logics and |
|
28225 | 116 |
examples of the Isabelle distribution is available at the usual |
117 |
Isabelle sites: |
|
28221 | 118 |
\begin{center}\small |
119 |
\begin{tabular}{l} |
|
28225 | 120 |
\url{http://isabelle.in.tum.de/dist/library/} \\ |
121 |
\url{http://www.cl.cam.ac.uk/research/hvg/Isabelle/dist/library/} \\ |
|
122 |
\url{http://mirror.cse.unsw.edu.au/pub/isabelle/dist/library/} \\ |
|
28221 | 123 |
\end{tabular} |
124 |
\end{center} |
|
125 |
||
126 |
\medskip In order to present your own theories on the web, simply |
|
127 |
copy the corresponding subdirectory from @{setting |
|
128 |
ISABELLE_BROWSER_INFO} to your WWW server, having generated browser |
|
129 |
info like this: |
|
130 |
\begin{ttbox} |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
131 |
isabelle usedir -i true HOL Foo |
28221 | 132 |
\end{ttbox} |
133 |
||
134 |
This assumes that directory @{verbatim Foo} contains some @{verbatim |
|
135 |
ROOT.ML} file to load all your theories, and HOL is your parent |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
136 |
logic image (@{verbatim isabelle} @{tool_ref mkdir} assists in |
28221 | 137 |
setting up Isabelle session directories. Theory browser information |
138 |
for HOL should have been generated already beforehand. |
|
139 |
Alternatively, one may specify an external link to an existing body |
|
140 |
of HTML data by giving @{tool usedir} a @{verbatim "-P"} option like |
|
141 |
this: |
|
142 |
\begin{ttbox} |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
143 |
isabelle usedir -i true -P http://isabelle.in.tum.de/library/ HOL Foo |
28221 | 144 |
\end{ttbox} |
145 |
||
146 |
\medskip For production use, the @{tool usedir} tool is usually |
|
147 |
invoked in an appropriate @{verbatim IsaMakefile}, via the Isabelle |
|
148 |
@{tool make} tool. There is a separate @{tool mkdir} tool to |
|
149 |
provide easy setup of all this, with only minimal manual editing |
|
150 |
required. |
|
151 |
\begin{ttbox} |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
152 |
isabelle mkdir HOL Foo && isabelle make |
28221 | 153 |
\end{ttbox} |
154 |
See \secref{sec:tool-mkdir} for more information on preparing |
|
155 |
Isabelle session directories, including the setup for documents. |
|
156 |
*} |
|
157 |
||
158 |
||
159 |
section {* Browsing theory graphs \label{sec:browse} *} |
|
160 |
||
161 |
text {* |
|
162 |
\index{theory graph browser|bold} |
|
163 |
||
164 |
The Isabelle graph browser is a general tool for visualizing |
|
165 |
dependency graphs. Certain nodes of the graph (i.e.~theories) can |
|
166 |
be grouped together in ``directories'', whose contents may be |
|
167 |
hidden, thus enabling the user to collapse irrelevant portions of |
|
168 |
information. The browser is written in Java, it can be used both as |
|
169 |
a stand-alone application and as an applet. Note that the option |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
170 |
@{verbatim "-g"} of @{verbatim isabelle} @{tool_ref usedir} creates |
28221 | 171 |
graph presentations in batch mode for inclusion in session |
172 |
documents. |
|
173 |
*} |
|
174 |
||
175 |
||
176 |
subsection {* Invoking the graph browser *} |
|
177 |
||
178 |
text {* |
|
179 |
The stand-alone version of the graph browser is wrapped up as an |
|
180 |
Isabelle tool called @{tool_def browser}: |
|
181 |
||
182 |
\begin{ttbox} |
|
183 |
Usage: browser [OPTIONS] [GRAPHFILE] |
|
184 |
||
185 |
Options are: |
|
35587 | 186 |
-b Admin/build only |
28221 | 187 |
-c cleanup -- remove GRAPHFILE after use |
188 |
-o FILE output to FILE (ps, eps, pdf) |
|
189 |
\end{ttbox} |
|
190 |
When no filename is specified, the browser automatically changes to |
|
191 |
the directory @{setting ISABELLE_BROWSER_INFO}. |
|
192 |
||
35587 | 193 |
\medskip The @{verbatim "-b"} option indicates that this is for |
194 |
administrative build only, i.e.\ no browser popup if no files are |
|
195 |
given. |
|
196 |
||
197 |
The @{verbatim "-c"} option causes the input file to be removed |
|
198 |
after use. |
|
28221 | 199 |
|
200 |
The @{verbatim "-o"} option indicates batch-mode operation, with the |
|
201 |
output written to the indicated file; note that @{verbatim pdf} |
|
202 |
produces an @{verbatim eps} copy as well. |
|
203 |
||
204 |
\medskip The applet version of the browser is part of the standard |
|
205 |
WWW theory presentation, see the link ``theory dependencies'' within |
|
206 |
each session index. |
|
207 |
*} |
|
208 |
||
209 |
||
210 |
subsection {* Using the graph browser *} |
|
211 |
||
212 |
text {* |
|
213 |
The browser's main window, which is shown in |
|
214 |
\figref{fig:browserwindow}, consists of two sub-windows. In the |
|
215 |
left sub-window, the directory tree is displayed. The graph itself |
|
216 |
is displayed in the right sub-window. |
|
217 |
||
218 |
\begin{figure}[ht] |
|
219 |
\includegraphics[width=\textwidth]{browser_screenshot} |
|
220 |
\caption{\label{fig:browserwindow} Browser main window} |
|
221 |
\end{figure} |
|
222 |
*} |
|
223 |
||
224 |
||
225 |
subsubsection {* The directory tree window *} |
|
226 |
||
227 |
text {* |
|
228 |
We describe the usage of the directory browser and the meaning of |
|
229 |
the different items in the browser window. |
|
230 |
||
231 |
\begin{itemize} |
|
232 |
||
233 |
\item A red arrow before a directory name indicates that the |
|
234 |
directory is currently ``folded'', i.e.~the nodes in this directory |
|
235 |
are collapsed to one single node. In the right sub-window, the names |
|
236 |
of nodes corresponding to folded directories are enclosed in square |
|
237 |
brackets and displayed in red color. |
|
238 |
||
239 |
\item A green downward arrow before a directory name indicates that |
|
240 |
the directory is currently ``unfolded''. It can be folded by |
|
241 |
clicking on the directory name. Clicking on the name for a second |
|
242 |
time unfolds the directory again. Alternatively, a directory can |
|
243 |
also be unfolded by clicking on the corresponding node in the right |
|
244 |
sub-window. |
|
245 |
||
246 |
\item Blue arrows stand before ordinary node names. When clicking on |
|
247 |
such a name (i.e.\ that of a theory), the graph display window |
|
248 |
focuses to the corresponding node. Double clicking invokes a text |
|
249 |
viewer window in which the contents of the theory file are |
|
250 |
displayed. |
|
251 |
||
252 |
\end{itemize} |
|
253 |
*} |
|
254 |
||
255 |
||
256 |
subsubsection {* The graph display window *} |
|
257 |
||
258 |
text {* |
|
259 |
When pointing on an ordinary node, an upward and a downward arrow is |
|
260 |
shown. Initially, both of these arrows are green. Clicking on the |
|
261 |
upward or downward arrow collapses all predecessor or successor |
|
262 |
nodes, respectively. The arrow's color then changes to red, |
|
263 |
indicating that the predecessor or successor nodes are currently |
|
264 |
collapsed. The node corresponding to the collapsed nodes has the |
|
265 |
name ``@{verbatim "[....]"}''. To uncollapse the nodes again, simply |
|
266 |
click on the red arrow or on the node with the name ``@{verbatim |
|
267 |
"[....]"}''. Similar to the directory browser, the contents of |
|
268 |
theory files can be displayed by double clicking on the |
|
269 |
corresponding node. |
|
270 |
*} |
|
271 |
||
272 |
||
273 |
subsubsection {* The ``File'' menu *} |
|
274 |
||
275 |
text {* |
|
276 |
Due to Java Applet security restrictions this menu is only available |
|
277 |
in the full application version. The meaning of the menu items is as |
|
278 |
follows: |
|
279 |
||
280 |
\begin{description} |
|
281 |
||
282 |
\item[Open \dots] Open a new graph file. |
|
283 |
||
284 |
\item[Export to PostScript] Outputs the current graph in Postscript |
|
285 |
format, appropriately scaled to fit on one single sheet of A4 paper. |
|
286 |
The resulting file can be printed directly. |
|
287 |
||
288 |
\item[Export to EPS] Outputs the current graph in Encapsulated |
|
289 |
Postscript format. The resulting file can be included in other |
|
290 |
documents. |
|
291 |
||
292 |
\item[Quit] Quit the graph browser. |
|
293 |
||
294 |
\end{description} |
|
295 |
*} |
|
296 |
||
297 |
||
298 |
subsection {* Syntax of graph definition files *} |
|
299 |
||
300 |
text {* |
|
301 |
A graph definition file has the following syntax: |
|
302 |
||
28225 | 303 |
\begin{center}\small |
28221 | 304 |
\begin{tabular}{rcl} |
31688 | 305 |
@{text graph} & @{text "="} & @{text "{ vertex"}~@{verbatim ";"}~@{text "}+"} \\ |
306 |
@{text vertex} & @{text "="} & @{text "vertex_name vertex_ID dir_name ["}~@{verbatim "+"}~@{text "] path ["}~@{verbatim "<"}~@{text "|"}~@{verbatim ">"}~@{text "] { vertex_ID }*"} |
|
28221 | 307 |
\end{tabular} |
28225 | 308 |
\end{center} |
28221 | 309 |
|
310 |
The meaning of the items in a vertex description is as follows: |
|
311 |
||
312 |
\begin{description} |
|
313 |
||
314 |
\item[@{text vertex_name}] The name of the vertex. |
|
315 |
||
316 |
\item[@{text vertex_ID}] The vertex identifier. Note that there may |
|
317 |
be several vertices with equal names, whereas identifiers must be |
|
318 |
unique. |
|
319 |
||
320 |
\item[@{text dir_name}] The name of the ``directory'' the vertex |
|
321 |
should be placed in. A ``@{verbatim "+"}'' sign after @{text |
|
322 |
dir_name} indicates that the nodes in the directory are initially |
|
323 |
visible. Directories are initially invisible by default. |
|
324 |
||
325 |
\item[@{text path}] The path of the corresponding theory file. This |
|
326 |
is specified relatively to the path of the graph definition file. |
|
327 |
||
328 |
\item[List of successor/predecessor nodes] A ``@{verbatim "<"}'' |
|
329 |
sign before the list means that successor nodes are listed, a |
|
330 |
``@{verbatim ">"}'' sign means that predecessor nodes are listed. If |
|
331 |
neither ``@{verbatim "<"}'' nor ``@{verbatim ">"}'' is found, the |
|
332 |
browser assumes that successor nodes are listed. |
|
333 |
||
334 |
\end{description} |
|
335 |
*} |
|
336 |
||
337 |
||
338 |
section {* Creating Isabelle session directories |
|
339 |
\label{sec:tool-mkdir} *} |
|
340 |
||
341 |
text {* |
|
342 |
The @{tool_def mkdir} utility prepares Isabelle session source |
|
343 |
directories, including a sensible default setup of @{verbatim |
|
344 |
IsaMakefile}, @{verbatim ROOT.ML}, and a @{verbatim document} |
|
345 |
directory with a minimal @{verbatim root.tex} that is sufficient to |
|
346 |
print all theories of the session (in the order of appearance); see |
|
347 |
\secref{sec:tool-document} for further information on Isabelle |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
348 |
document preparation. The usage of @{verbatim isabelle} @{tool |
28238 | 349 |
mkdir} is: |
28221 | 350 |
|
351 |
\begin{ttbox} |
|
352 |
Usage: mkdir [OPTIONS] [LOGIC] NAME |
|
353 |
||
354 |
Options are: |
|
355 |
-I FILE alternative IsaMakefile output |
|
356 |
-P include parent logic target |
|
357 |
-b setup build mode (session outputs heap image) |
|
358 |
-q quiet mode |
|
359 |
||
360 |
Prepare session directory, including IsaMakefile and document source, |
|
361 |
with parent LOGIC (default ISABELLE_LOGIC=\$ISABELLE_LOGIC) |
|
362 |
\end{ttbox} |
|
363 |
||
364 |
The @{tool mkdir} tool is conservative in the sense that any |
|
365 |
existing @{verbatim IsaMakefile} etc.\ is left unchanged. Thus it |
|
366 |
is safe to invoke it multiple times, although later runs may not |
|
367 |
have the desired effect. |
|
368 |
||
369 |
Note that @{tool mkdir} is unable to change @{verbatim IsaMakefile} |
|
370 |
incrementally --- manual changes are required for multiple |
|
371 |
sub-sessions. On order to get an initial working session, the only |
|
372 |
editing needed is to add appropriate @{ML use_thy} calls to the |
|
373 |
generated @{verbatim ROOT.ML} file. |
|
374 |
*} |
|
375 |
||
376 |
||
377 |
subsubsection {* Options *} |
|
378 |
||
379 |
text {* |
|
380 |
The @{verbatim "-I"} option specifies an alternative to @{verbatim |
|
381 |
IsaMakefile} for dependencies. Note that ``@{verbatim "-"}'' refers |
|
382 |
to \emph{stdout}, i.e.\ ``@{verbatim "-I-"}'' provides an easy way |
|
383 |
to peek at @{tool mkdir}'s idea of @{tool make} setup required for |
|
384 |
some particular of Isabelle session. |
|
385 |
||
386 |
\medskip The @{verbatim "-P"} option includes a target for the |
|
387 |
parent @{verbatim LOGIC} session in the generated @{verbatim |
|
388 |
IsaMakefile}. The corresponding sources are assumed to be located |
|
389 |
within the Isabelle distribution. |
|
390 |
||
391 |
\medskip The @{verbatim "-b"} option sets up the current directory |
|
392 |
as the base for a new session that provides an actual logic image, |
|
393 |
as opposed to one that only runs several theories based on an |
|
394 |
existing image. Note that in the latter case, everything except |
|
395 |
@{verbatim IsaMakefile} would be placed into a separate directory |
|
396 |
@{verbatim NAME}, rather than the current one. See |
|
397 |
\secref{sec:tool-usedir} for further information on \emph{build |
|
398 |
mode} vs.\ \emph{example mode} of the @{tool usedir} utility. |
|
399 |
||
400 |
\medskip The @{verbatim "-q"} option enables quiet mode, suppressing |
|
401 |
further notes on how to proceed. |
|
402 |
*} |
|
403 |
||
404 |
||
405 |
subsubsection {* Examples *} |
|
406 |
||
407 |
text {* |
|
408 |
The standard setup of a single ``example session'' based on the |
|
409 |
default logic, with proper document generation is generated like |
|
410 |
this: |
|
411 |
\begin{ttbox} |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
412 |
isabelle mkdir Foo && isabelle make |
28221 | 413 |
\end{ttbox} |
414 |
||
415 |
\noindent The theory sources should be put into the @{verbatim Foo} |
|
416 |
directory, and its @{verbatim ROOT.ML} should be edited to load all |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
417 |
required theories. Invoking @{verbatim isabelle} @{tool make} again |
28238 | 418 |
would run the whole session, generating browser information and the |
28221 | 419 |
document automatically. The @{verbatim IsaMakefile} is typically |
420 |
tuned manually later, e.g.\ adding source dependencies, or changing |
|
421 |
the options passed to @{tool usedir}. |
|
422 |
||
423 |
\medskip Large projects may demand further sessions, potentially |
|
424 |
with separate logic images being created. This usually requires |
|
425 |
manual editing of the generated @{verbatim IsaMakefile}, which is |
|
426 |
meant to cover all of the sub-session directories at the same time |
|
427 |
(this is the deeper reasong why @{verbatim IsaMakefile} is not made |
|
28238 | 428 |
part of the initial session directory created by @{verbatim |
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
429 |
isabelle} @{tool mkdir}). See @{"file" "~~/src/HOL/IsaMakefile"} for |
28238 | 430 |
a full-blown example. |
28221 | 431 |
*} |
432 |
||
433 |
||
434 |
section {* Running Isabelle sessions \label{sec:tool-usedir} *} |
|
435 |
||
436 |
text {* |
|
437 |
The @{tool_def usedir} utility builds object-logic images, or runs |
|
438 |
example sessions based on existing logics. Its usage is: |
|
439 |
\begin{ttbox} |
|
440 |
||
441 |
Usage: usedir [OPTIONS] LOGIC NAME |
|
442 |
||
443 |
Options are: |
|
444 |
-C BOOL copy existing document directory to -D PATH (default true) |
|
445 |
-D PATH dump generated document sources into PATH |
|
446 |
-M MAX multithreading: maximum number of worker threads (default 1) |
|
447 |
-P PATH set path for remote theory browsing information |
|
448 |
-T LEVEL multithreading: trace level (default 0) |
|
449 |
-V VERSION declare alternative document VERSION |
|
450 |
-b build mode (output heap image, using current dir) |
|
451 |
-d FORMAT build document as FORMAT (default false) |
|
452 |
-f NAME use ML file NAME (default ROOT.ML) |
|
453 |
-g BOOL generate session graph image for document (default false) |
|
454 |
-i BOOL generate theory browser information (default false) |
|
455 |
-m MODE add print mode for output |
|
32061
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
wenzelm
parents:
31688
diff
changeset
|
456 |
-p LEVEL set level of detail for proof objects (default 0) |
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
wenzelm
parents:
31688
diff
changeset
|
457 |
-q LEVEL set level of parallel proof checking (default 1) |
28221 | 458 |
-r reset session path |
459 |
-s NAME override session NAME |
|
31688 | 460 |
-t BOOL internal session timing (default false) |
28221 | 461 |
-v BOOL be verbose (default false) |
462 |
||
463 |
Build object-logic or run examples. Also creates browsing |
|
464 |
information (HTML etc.) according to settings. |
|
465 |
||
466 |
ISABELLE_USEDIR_OPTIONS= |
|
29435
a5f84ac14609
added parallel_proofs flag (default true, cf. usedir option -Q), which can be disabled in low-memory situations;
wenzelm
parents:
28914
diff
changeset
|
467 |
|
a5f84ac14609
added parallel_proofs flag (default true, cf. usedir option -Q), which can be disabled in low-memory situations;
wenzelm
parents:
28914
diff
changeset
|
468 |
ML_PLATFORM=x86-linux |
a5f84ac14609
added parallel_proofs flag (default true, cf. usedir option -Q), which can be disabled in low-memory situations;
wenzelm
parents:
28914
diff
changeset
|
469 |
ML_HOME=/usr/local/polyml-5.2.1/x86-linux |
a5f84ac14609
added parallel_proofs flag (default true, cf. usedir option -Q), which can be disabled in low-memory situations;
wenzelm
parents:
28914
diff
changeset
|
470 |
ML_SYSTEM=polyml-5.2.1 |
a5f84ac14609
added parallel_proofs flag (default true, cf. usedir option -Q), which can be disabled in low-memory situations;
wenzelm
parents:
28914
diff
changeset
|
471 |
ML_OPTIONS=-H 500 |
28221 | 472 |
\end{ttbox} |
473 |
||
474 |
Note that the value of the @{setting_ref ISABELLE_USEDIR_OPTIONS} |
|
475 |
setting is implicitly prefixed to \emph{any} @{tool usedir} |
|
476 |
call. Since the @{verbatim IsaMakefile}s of all object-logics |
|
28238 | 477 |
distributed with Isabelle just invoke @{tool usedir} for the real |
28221 | 478 |
work, one may control compilation options globally via above |
479 |
variable. In particular, generation of \rmindex{HTML} browsing |
|
480 |
information and document preparation is controlled here. |
|
481 |
*} |
|
482 |
||
483 |
||
484 |
subsubsection {* Options *} |
|
485 |
||
486 |
text {* |
|
487 |
Basically, there are two different modes of operation: \emph{build |
|
488 |
mode} (enabled through the @{verbatim "-b"} option) and |
|
489 |
\emph{example mode} (default). |
|
490 |
||
491 |
Calling @{tool usedir} with @{verbatim "-b"} runs @{executable |
|
492 |
"isabelle-process"} with input image @{verbatim LOGIC} and output to |
|
493 |
@{verbatim NAME}, as provided on the command line. This will be a |
|
494 |
batch session, running @{verbatim ROOT.ML} from the current |
|
495 |
directory and then quitting. It is assumed that @{verbatim ROOT.ML} |
|
496 |
contains all ML commands required to build the logic. |
|
497 |
||
28238 | 498 |
In example mode, @{tool usedir} runs a read-only session of |
28221 | 499 |
@{verbatim LOGIC} and automatically runs @{verbatim ROOT.ML} from |
500 |
within directory @{verbatim NAME}. It assumes that this file |
|
501 |
contains appropriate ML commands to run the desired examples. |
|
502 |
||
503 |
\medskip The @{verbatim "-i"} option controls theory browser data |
|
504 |
generation. It may be explicitly turned on or off --- as usual, the |
|
505 |
last occurrence of @{verbatim "-i"} on the command line wins. |
|
506 |
||
507 |
The @{verbatim "-P"} option specifies a path (or actual URL) to be |
|
508 |
prefixed to any \emph{non-local} reference of existing theories. |
|
509 |
Thus user sessions may easily link to existing Isabelle libraries |
|
510 |
already present on the WWW. |
|
511 |
||
512 |
The @{verbatim "-m"} options specifies additional print modes to be |
|
513 |
activated temporarily while the session is processed. |
|
514 |
||
515 |
\medskip The @{verbatim "-d"} option controls document preparation. |
|
516 |
Valid arguments are @{verbatim false} (do not prepare any document; |
|
517 |
this is default), or any of @{verbatim dvi}, @{verbatim dvi.gz}, |
|
518 |
@{verbatim ps}, @{verbatim ps.gz}, @{verbatim pdf}. The logic |
|
519 |
session has to provide a properly setup @{verbatim document} |
|
520 |
directory. See \secref{sec:tool-document} and |
|
521 |
\secref{sec:tool-latex} for more details. |
|
522 |
||
523 |
\medskip The @{verbatim "-V"} option declares alternative document |
|
524 |
versions, consisting of name/tags pairs (cf.\ options @{verbatim |
|
525 |
"-n"} and @{verbatim "-t"} of the @{tool_ref document} tool). The |
|
526 |
standard document is equivalent to ``@{verbatim |
|
527 |
"document=theory,proof,ML"}'', which means that all theory begin/end |
|
528 |
commands, proof body texts, and ML code will be presented |
|
529 |
faithfully. An alternative version ``@{verbatim |
|
530 |
"outline=/proof/ML"}'' would fold proof and ML parts, replacing the |
|
531 |
original text by a short place-holder. The form ``@{text |
|
532 |
name}@{verbatim "=-"},'' means to remove document @{text name} from |
|
533 |
the list of versions to be processed. Any number of @{verbatim |
|
534 |
"-V"} options may be given; later declarations have precedence over |
|
535 |
earlier ones. |
|
536 |
||
537 |
\medskip The @{verbatim "-g"} option produces images of the theory |
|
538 |
dependency graph (cf.\ \secref{sec:browse}) for inclusion in the |
|
539 |
generated document, both as @{verbatim session_graph.eps} and |
|
540 |
@{verbatim session_graph.pdf} at the same time. To include this in |
|
541 |
the final {\LaTeX} document one could say @{verbatim |
|
542 |
"\\includegraphics{session_graph}"} in @{verbatim |
|
543 |
"document/root.tex"} (omitting the file-name extension enables |
|
544 |
{\LaTeX} to select to correct version, either for the DVI or PDF |
|
545 |
output path). |
|
546 |
||
547 |
\medskip The @{verbatim "-D"} option causes the generated document |
|
548 |
sources to be dumped at location @{verbatim PATH}; this path is |
|
549 |
relative to the session's main directory. If the @{verbatim "-C"} |
|
550 |
option is true, this will include a copy of an existing @{verbatim |
|
551 |
document} directory as provided by the user. For example, |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
552 |
@{verbatim isabelle} @{tool usedir}~@{verbatim "-D generated HOL |
28238 | 553 |
Foo"} produces a complete set of document sources at @{verbatim |
554 |
"Foo/generated"}. Subsequent invocation of @{verbatim |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
555 |
isabelle} @{tool document}~@{verbatim "Foo/generated"} (see also |
28238 | 556 |
\secref{sec:tool-document}) will process the final result |
557 |
independently of an Isabelle job. This decoupled mode of operation |
|
558 |
facilitates debugging of serious {\LaTeX} errors, for example. |
|
28221 | 559 |
|
560 |
\medskip The @{verbatim "-p"} option determines the level of detail |
|
561 |
for internal proof objects, see also the \emph{Isabelle Reference |
|
562 |
Manual}~\cite{isabelle-ref}. |
|
563 |
||
32061
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
wenzelm
parents:
31688
diff
changeset
|
564 |
\medskip The @{verbatim "-q"} option specifies the level of parallel |
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
wenzelm
parents:
31688
diff
changeset
|
565 |
proof checking: @{verbatim 0} no proofs, @{verbatim 1} toplevel |
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
wenzelm
parents:
31688
diff
changeset
|
566 |
proofs (default), @{verbatim 2} toplevel and nested Isar proofs. |
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
wenzelm
parents:
31688
diff
changeset
|
567 |
The resulting speedup may vary, depending on the number of worker |
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
wenzelm
parents:
31688
diff
changeset
|
568 |
threads, granularity of proofs, and whether proof terms are enabled. |
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
wenzelm
parents:
31688
diff
changeset
|
569 |
|
31688 | 570 |
\medskip The @{verbatim "-t"} option produces a more detailed |
571 |
internal timing report of the session. |
|
572 |
||
28221 | 573 |
\medskip The @{verbatim "-v"} option causes additional information |
574 |
to be printed while running the session, notably the location of |
|
575 |
prepared documents. |
|
576 |
||
577 |
\medskip The @{verbatim "-M"} option specifies the maximum number of |
|
578 |
parallel threads used for processing independent tasks when checking |
|
579 |
theory sources (multithreading only works on suitable ML platforms). |
|
28238 | 580 |
The special value of @{verbatim 0} or @{verbatim max} refers to the |
581 |
number of actual CPU cores of the underlying machine, which is a |
|
582 |
good starting point for optimal performance tuning. The @{verbatim |
|
583 |
"-T"} option determines the level of detail in tracing output |
|
584 |
concerning the internal locking and scheduling in multithreaded |
|
585 |
operation. This may be helpful in isolating performance |
|
586 |
bottle-necks, e.g.\ due to excessive wait states when locking |
|
587 |
critical code sections. |
|
28221 | 588 |
|
589 |
\medskip Any @{tool usedir} session is named by some \emph{session |
|
590 |
identifier}. These accumulate, documenting the way sessions depend |
|
591 |
on others. For example, consider @{verbatim "Pure/FOL/ex"}, which |
|
592 |
refers to the examples of FOL, which in turn is built upon Pure. |
|
593 |
||
594 |
The current session's identifier is by default just the base name of |
|
595 |
the @{verbatim LOGIC} argument (in build mode), or of the @{verbatim |
|
596 |
NAME} argument (in example mode). This may be overridden explicitly |
|
597 |
via the @{verbatim "-s"} option. |
|
598 |
*} |
|
599 |
||
600 |
||
601 |
subsubsection {* Examples *} |
|
602 |
||
603 |
text {* |
|
604 |
Refer to the @{verbatim IsaMakefile}s of the Isabelle distribution's |
|
605 |
object-logics as a model for your own developments. For example, |
|
28238 | 606 |
see @{"file" "~~/src/FOL/IsaMakefile"}. The Isabelle @{tool_ref |
28221 | 607 |
mkdir} tool creates @{verbatim IsaMakefile}s with proper invocation |
608 |
of @{tool usedir} as well. |
|
609 |
*} |
|
610 |
||
611 |
||
612 |
section {* Preparing Isabelle session documents |
|
613 |
\label{sec:tool-document} *} |
|
614 |
||
615 |
text {* |
|
616 |
The @{tool_def document} utility prepares logic session documents, |
|
617 |
processing the sources both as provided by the user and generated by |
|
618 |
Isabelle. Its usage is: |
|
619 |
\begin{ttbox} |
|
620 |
Usage: document [OPTIONS] [DIR] |
|
621 |
||
622 |
Options are: |
|
623 |
-c cleanup -- be aggressive in removing old stuff |
|
624 |
-n NAME specify document name (default 'document') |
|
625 |
-o FORMAT specify output format: dvi (default), dvi.gz, ps, |
|
626 |
ps.gz, pdf |
|
627 |
-t TAGS specify tagged region markup |
|
628 |
||
629 |
Prepare the theory session document in DIR (default 'document') |
|
630 |
producing the specified output format. |
|
631 |
\end{ttbox} |
|
632 |
This tool is usually run automatically as part of the corresponding |
|
633 |
Isabelle batch process, provided document preparation has been |
|
634 |
enabled (cf.\ the @{verbatim "-d"} option of the @{tool_ref usedir} |
|
635 |
tool). It may be manually invoked on the generated browser |
|
636 |
information document output as well, e.g.\ in case of errors |
|
637 |
encountered in the batch run. |
|
638 |
||
639 |
\medskip The @{verbatim "-c"} option tells the @{tool document} tool |
|
640 |
to dispose the document sources after successful operation. This is |
|
641 |
the right thing to do for sources generated by an Isabelle process, |
|
642 |
but take care of your files in manual document preparation! |
|
643 |
||
644 |
\medskip The @{verbatim "-n"} and @{verbatim "-o"} option specify |
|
645 |
the final output file name and format, the default is ``@{verbatim |
|
646 |
document.dvi}''. Note that the result will appear in the parent of |
|
647 |
the target @{verbatim DIR}. |
|
648 |
||
649 |
\medskip The @{verbatim "-t"} option tells {\LaTeX} how to interpret |
|
650 |
tagged Isabelle command regions. Tags are specified as a comma |
|
651 |
separated list of modifier/name pairs: ``@{verbatim "+"}@{text |
|
652 |
foo}'' (or just ``@{text foo}'') means to keep, ``@{verbatim |
|
653 |
"-"}@{text foo}'' to drop, and ``@{verbatim "/"}@{text foo}'' to |
|
654 |
fold text tagged as @{text foo}. The builtin default is equivalent |
|
655 |
to the tag specification ``@{verbatim |
|
30113
5ea17e90b08a
isabelle document: adapted (postulated) defaults for tags to actual isabelle.sty;
wenzelm
parents:
29435
diff
changeset
|
656 |
"+theory,+proof,+ML,+visible,-invisible"}''; see also the {\LaTeX} |
28221 | 657 |
macros @{verbatim "\\isakeeptag"}, @{verbatim "\\isadroptag"}, and |
28238 | 658 |
@{verbatim "\\isafoldtag"}, in @{"file" |
659 |
"~~/lib/texinputs/isabelle.sty"}. |
|
28221 | 660 |
|
661 |
\medskip Document preparation requires a properly setup ``@{verbatim |
|
662 |
document}'' directory within the logic session sources. This |
|
663 |
directory is supposed to contain all the files needed to produce the |
|
664 |
final document --- apart from the actual theories which are |
|
665 |
generated by Isabelle. |
|
666 |
||
667 |
\medskip For most practical purposes, the @{tool document} tool is |
|
668 |
smart enough to create any of the specified output formats, taking |
|
669 |
@{verbatim root.tex} supplied by the user as a starting point. This |
|
670 |
even includes multiple runs of {\LaTeX} to accommodate references |
|
671 |
and bibliographies (the latter assumes @{verbatim root.bib} within |
|
672 |
the same directory). |
|
673 |
||
674 |
In more complex situations, a separate @{verbatim IsaMakefile} for |
|
675 |
the document sources may be given instead. This should provide |
|
676 |
targets for any admissible document format; these have to produce |
|
677 |
corresponding output files named after @{verbatim root} as well, |
|
678 |
e.g.\ @{verbatim root.dvi} for target format @{verbatim dvi}. |
|
679 |
||
680 |
\medskip When running the session, Isabelle copies the original |
|
681 |
@{verbatim document} directory into its proper place within |
|
28238 | 682 |
@{setting ISABELLE_BROWSER_INFO} according to the session path. |
28221 | 683 |
Then, for any processed theory @{text A} some {\LaTeX} source is |
684 |
generated and put there as @{text A}@{verbatim ".tex"}. |
|
685 |
Furthermore, a list of all generated theory files is put into |
|
686 |
@{verbatim session.tex}. Typically, the root {\LaTeX} file provided |
|
687 |
by the user would include @{verbatim session.tex} to get a document |
|
688 |
containing all the theories. |
|
689 |
||
690 |
The {\LaTeX} versions of the theories require some macros defined in |
|
28238 | 691 |
@{"file" "~~/lib/texinputs/isabelle.sty"}. Doing @{verbatim |
692 |
"\\usepackage{isabelle}"} in @{verbatim root.tex} should be fine; |
|
693 |
the underlying Isabelle @{tool latex} tool already includes an |
|
694 |
appropriate path specification for {\TeX} inputs. |
|
28221 | 695 |
|
696 |
If the text contains any references to Isabelle symbols (such as |
|
697 |
@{verbatim "\\"}@{verbatim "<forall>"}) then @{verbatim |
|
28238 | 698 |
"isabellesym.sty"} should be included as well. This package |
699 |
contains a standard set of {\LaTeX} macro definitions @{verbatim |
|
28221 | 700 |
"\\isasym"}@{text foo} corresponding to @{verbatim "\\"}@{verbatim |
28838
d5db6dfcb34a
moved table of standard Isabelle symbols to isar-ref manual;
wenzelm
parents:
28504
diff
changeset
|
701 |
"<"}@{text foo}@{verbatim ">"}, see \cite{isabelle-implementation} for a |
d5db6dfcb34a
moved table of standard Isabelle symbols to isar-ref manual;
wenzelm
parents:
28504
diff
changeset
|
702 |
complete list of predefined Isabelle symbols. Users may invent |
28221 | 703 |
further symbols as well, just by providing {\LaTeX} macros in a |
28238 | 704 |
similar fashion as in @{"file" "~~/lib/texinputs/isabellesym.sty"} of |
705 |
the distribution. |
|
28221 | 706 |
|
707 |
For proper setup of DVI and PDF documents (with hyperlinks and |
|
28238 | 708 |
bookmarks), we recommend to include @{"file" |
709 |
"~~/lib/texinputs/pdfsetup.sty"} as well. |
|
28221 | 710 |
|
711 |
\medskip As a final step of document preparation within Isabelle, |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
712 |
@{verbatim isabelle} @{tool document}~@{verbatim "-c"} is run on the |
28238 | 713 |
resulting @{verbatim document} directory. Thus the actual output |
714 |
document is built and installed in its proper place (as linked by |
|
715 |
the session's @{verbatim index.html} if option @{verbatim "-i"} of |
|
716 |
@{tool_ref usedir} has been enabled, cf.\ \secref{sec:info}). The |
|
717 |
generated sources are deleted after successful run of {\LaTeX} and |
|
718 |
friends. Note that a separate copy of the sources may be retained |
|
719 |
by passing an option @{verbatim "-D"} to the @{tool usedir} utility |
|
720 |
when running the session. |
|
28221 | 721 |
*} |
722 |
||
723 |
||
724 |
section {* Running {\LaTeX} within the Isabelle environment |
|
725 |
\label{sec:tool-latex} *} |
|
726 |
||
727 |
text {* |
|
728 |
The @{tool_def latex} utility provides the basic interface for |
|
729 |
Isabelle document preparation. Its usage is: |
|
730 |
\begin{ttbox} |
|
731 |
Usage: latex [OPTIONS] [FILE] |
|
732 |
||
733 |
Options are: |
|
734 |
-o FORMAT specify output format: dvi (default), dvi.gz, ps, |
|
735 |
ps.gz, pdf, bbl, idx, sty, syms |
|
736 |
||
737 |
Run LaTeX (and related tools) on FILE (default root.tex), |
|
738 |
producing the specified output format. |
|
739 |
\end{ttbox} |
|
740 |
||
741 |
Appropriate {\LaTeX}-related programs are run on the input file, |
|
742 |
according to the given output format: @{executable latex}, |
|
743 |
@{executable pdflatex}, @{executable dvips}, @{executable bibtex} |
|
744 |
(for @{verbatim bbl}), and @{executable makeindex} (for @{verbatim |
|
745 |
idx}). The actual commands are determined from the settings |
|
746 |
environment (@{setting ISABELLE_LATEX} etc.). |
|
747 |
||
748 |
The @{verbatim sty} output format causes the Isabelle style files to |
|
749 |
be updated from the distribution. This is useful in special |
|
750 |
situations where the document sources are to be processed another |
|
751 |
time by separate tools (cf.\ option @{verbatim "-D"} of the @{tool |
|
752 |
usedir} utility). |
|
753 |
||
754 |
The @{verbatim syms} output is for internal use; it generates lists |
|
755 |
of symbols that are available without loading additional {\LaTeX} |
|
756 |
packages. |
|
757 |
*} |
|
758 |
||
759 |
||
760 |
subsubsection {* Examples *} |
|
761 |
||
762 |
text {* |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
763 |
Invoking @{verbatim isabelle} @{tool latex} by hand may be |
28238 | 764 |
occasionally useful when debugging failed attempts of the automatic |
765 |
document preparation stage of batch-mode Isabelle. The abortive |
|
766 |
process leaves the sources at a certain place within @{setting |
|
28221 | 767 |
ISABELLE_BROWSER_INFO}, see the runtime error message for details. |
768 |
This enables users to inspect {\LaTeX} runs in further detail, e.g.\ |
|
769 |
like this: |
|
770 |
\begin{ttbox} |
|
28914
f993cbffc42a
default for ISABELLE_HOME_USER is now ~/.isabelle instead of ~/isabelle;
wenzelm
parents:
28838
diff
changeset
|
771 |
cd ~/.isabelle/browser_info/HOL/Test/document |
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28238
diff
changeset
|
772 |
isabelle latex -o pdf |
28221 | 773 |
\end{ttbox} |
774 |
*} |
|
775 |
||
776 |
end |