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