author | blanchet |
Fri, 11 Jan 2013 16:30:56 +0100 | |
changeset 50826 | 18ace05656cf |
parent 50653 | 5c85f8b80b95 |
child 52052 | 892061142ba6 |
permissions | -rw-r--r-- |
28224 | 1 |
theory Misc |
43564
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
41512
diff
changeset
|
2 |
imports Base |
28224 | 3 |
begin |
4 |
||
5 |
chapter {* Miscellaneous tools \label{ch:tools} *} |
|
6 |
||
7 |
text {* |
|
8 |
Subsequently we describe various Isabelle related utilities, given |
|
9 |
in alphabetical order. |
|
10 |
*} |
|
11 |
||
12 |
||
48844 | 13 |
section {* Resolving Isabelle components \label{sec:tool-components} *} |
14 |
||
15 |
text {* |
|
16 |
The @{tool_def components} tool resolves Isabelle components: |
|
17 |
\begin{ttbox} |
|
18 |
Usage: isabelle components [OPTIONS] [COMPONENTS ...] |
|
19 |
||
20 |
Options are: |
|
50653
5c85f8b80b95
simplified quick start via "isabelle components -I";
wenzelm
parents:
50132
diff
changeset
|
21 |
-I init user settings |
48844 | 22 |
-R URL component repository |
23 |
(default $ISABELLE_COMPONENT_REPOSITORY) |
|
24 |
-a all missing components |
|
25 |
-l list status |
|
26 |
||
27 |
Resolve Isabelle components via download and installation. |
|
28 |
COMPONENTS are identified via base name. |
|
29 |
||
30 |
ISABELLE_COMPONENT_REPOSITORY="http://isabelle.in.tum.de/components" |
|
31 |
\end{ttbox} |
|
32 |
||
33 |
Components are initialized as described in \secref{sec:components} |
|
34 |
in a permissive manner, which can mark components as ``missing''. |
|
35 |
This state is amended by letting @{tool "components"} download and |
|
36 |
unpack components that are published on the default component |
|
37 |
repository \url{http://isabelle.in.tum.de/components/} in |
|
38 |
particular. |
|
39 |
||
40 |
Option @{verbatim "-R"} specifies an alternative component |
|
41 |
repository. Note that @{verbatim "file:///"} URLs can be used for |
|
42 |
local directories. |
|
43 |
||
44 |
Option @{verbatim "-a"} selects all missing components to be |
|
45 |
installed. Explicit components may be named as command |
|
46 |
line-arguments as well. Note that components are uniquely |
|
47 |
identified by their base name, while the installation takes place in |
|
48 |
the location that was specified in the attempt to initialize the |
|
49 |
component before. |
|
50 |
||
51 |
Option @{verbatim "-l"} lists the current state of available and |
|
52 |
missing components with their location (full name) within the |
|
50653
5c85f8b80b95
simplified quick start via "isabelle components -I";
wenzelm
parents:
50132
diff
changeset
|
53 |
file-system. |
5c85f8b80b95
simplified quick start via "isabelle components -I";
wenzelm
parents:
50132
diff
changeset
|
54 |
|
5c85f8b80b95
simplified quick start via "isabelle components -I";
wenzelm
parents:
50132
diff
changeset
|
55 |
Option @{verbatim "-I"} initializes the user settings file to |
5c85f8b80b95
simplified quick start via "isabelle components -I";
wenzelm
parents:
50132
diff
changeset
|
56 |
subscribe to the standard components specified in the Isabelle |
5c85f8b80b95
simplified quick start via "isabelle components -I";
wenzelm
parents:
50132
diff
changeset
|
57 |
repository clone --- this does not make any sense for regular |
5c85f8b80b95
simplified quick start via "isabelle components -I";
wenzelm
parents:
50132
diff
changeset
|
58 |
Isabelle releases. If the file already exists, it needs to be |
5c85f8b80b95
simplified quick start via "isabelle components -I";
wenzelm
parents:
50132
diff
changeset
|
59 |
edited manually according to the printed explanation. |
5c85f8b80b95
simplified quick start via "isabelle components -I";
wenzelm
parents:
50132
diff
changeset
|
60 |
*} |
48844 | 61 |
|
62 |
||
28224 | 63 |
section {* Displaying documents *} |
64 |
||
48602 | 65 |
text {* The @{tool_def display} tool displays documents in DVI or PDF |
28224 | 66 |
format: |
67 |
\begin{ttbox} |
|
48602 | 68 |
Usage: isabelle display [OPTIONS] FILE |
28224 | 69 |
|
70 |
Options are: |
|
71 |
-c cleanup -- remove FILE after use |
|
72 |
||
73 |
Display document FILE (in DVI format). |
|
74 |
\end{ttbox} |
|
75 |
||
76 |
\medskip The @{verbatim "-c"} option causes the input file to be |
|
77 |
removed after use. The program for viewing @{verbatim dvi} files is |
|
78 |
determined by the @{setting DVI_VIEWER} setting. |
|
79 |
*} |
|
80 |
||
81 |
||
82 |
section {* Viewing documentation \label{sec:tool-doc} *} |
|
83 |
||
84 |
text {* |
|
48602 | 85 |
The @{tool_def doc} tool displays online documentation: |
28224 | 86 |
\begin{ttbox} |
48602 | 87 |
Usage: isabelle doc [DOC] |
28224 | 88 |
|
89 |
View Isabelle documentation DOC, or show list of available documents. |
|
90 |
\end{ttbox} |
|
91 |
If called without arguments, it lists all available documents. Each |
|
92 |
line starts with an identifier, followed by a short description. Any |
|
93 |
of these identifiers may be specified as the first argument in order |
|
94 |
to have the corresponding document displayed. |
|
95 |
||
96 |
\medskip The @{setting ISABELLE_DOCS} setting specifies the list of |
|
97 |
directories (separated by colons) to be scanned for documentations. |
|
98 |
The program for viewing @{verbatim dvi} files is determined by the |
|
99 |
@{setting DVI_VIEWER} setting. |
|
100 |
*} |
|
101 |
||
102 |
||
47828 | 103 |
section {* Shell commands within the settings environment \label{sec:tool-env} *} |
104 |
||
48602 | 105 |
text {* The @{tool_def env} tool is a direct wrapper for the standard |
106 |
@{verbatim "/usr/bin/env"} command on POSIX systems, running within |
|
107 |
the Isabelle settings environment (\secref{sec:settings}). |
|
47828 | 108 |
|
109 |
The command-line arguments are that of the underlying version of |
|
110 |
@{verbatim env}. For example, the following invokes an instance of |
|
111 |
the GNU Bash shell within the Isabelle environment: |
|
112 |
\begin{alltt} |
|
113 |
isabelle env bash |
|
114 |
\end{alltt} |
|
115 |
*} |
|
116 |
||
117 |
||
28224 | 118 |
section {* Getting logic images *} |
119 |
||
48602 | 120 |
text {* The @{tool_def findlogics} tool traverses all directories |
28224 | 121 |
specified in @{setting ISABELLE_PATH}, looking for Isabelle logic |
122 |
images. Its usage is: |
|
123 |
\begin{ttbox} |
|
48577 | 124 |
Usage: isabelle findlogics |
28224 | 125 |
|
126 |
Collect heap file names from ISABELLE_PATH. |
|
127 |
\end{ttbox} |
|
128 |
||
129 |
The base names of all files found on the path are printed --- sorted |
|
130 |
and with duplicates removed. Also note that lookup in @{setting |
|
131 |
ISABELLE_PATH} includes the current values of @{setting ML_SYSTEM} |
|
132 |
and @{setting ML_PLATFORM}. Thus switching to another ML compiler |
|
133 |
may change the set of logic images available. |
|
134 |
*} |
|
135 |
||
136 |
||
137 |
section {* Inspecting the settings environment \label{sec:tool-getenv} *} |
|
138 |
||
48602 | 139 |
text {* The Isabelle settings environment --- as provided by the |
28224 | 140 |
site-default and user-specific settings files --- can be inspected |
48602 | 141 |
with the @{tool_def getenv} tool: |
28224 | 142 |
\begin{ttbox} |
48602 | 143 |
Usage: isabelle getenv [OPTIONS] [VARNAMES ...] |
28224 | 144 |
|
145 |
Options are: |
|
146 |
-a display complete environment |
|
147 |
-b print values only (doesn't work for -a) |
|
31497 | 148 |
-d FILE dump complete environment to FILE |
149 |
(null terminated entries) |
|
28224 | 150 |
|
151 |
Get value of VARNAMES from the Isabelle settings. |
|
152 |
\end{ttbox} |
|
153 |
||
154 |
With the @{verbatim "-a"} option, one may inspect the full process |
|
155 |
environment that Isabelle related programs are run in. This usually |
|
156 |
contains much more variables than are actually Isabelle settings. |
|
157 |
Normally, output is a list of lines of the form @{text |
|
158 |
name}@{verbatim "="}@{text value}. The @{verbatim "-b"} option |
|
159 |
causes only the values to be printed. |
|
31497 | 160 |
|
161 |
Option @{verbatim "-d"} produces a dump of the complete environment |
|
162 |
to the specified file. Entries are terminated by the ASCII null |
|
163 |
character, i.e.\ the C string terminator. |
|
28224 | 164 |
*} |
165 |
||
166 |
||
167 |
subsubsection {* Examples *} |
|
168 |
||
48815 | 169 |
text {* Get the location of @{setting ISABELLE_HOME_USER} where |
170 |
user-specific information is stored: |
|
28224 | 171 |
\begin{ttbox} |
48815 | 172 |
isabelle getenv ISABELLE_HOME_USER |
28224 | 173 |
\end{ttbox} |
174 |
||
48815 | 175 |
\medskip Get the value only of the same settings variable, which is |
176 |
particularly useful in shell scripts: |
|
28224 | 177 |
\begin{ttbox} |
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28253
diff
changeset
|
178 |
isabelle getenv -b ISABELLE_OUTPUT |
28224 | 179 |
\end{ttbox} |
180 |
*} |
|
181 |
||
182 |
||
183 |
section {* Installing standalone Isabelle executables \label{sec:tool-install} *} |
|
184 |
||
48602 | 185 |
text {* By default, the main Isabelle binaries (@{executable |
186 |
"isabelle"} etc.) are just run from their location within the |
|
187 |
distribution directory, probably indirectly by the shell through its |
|
188 |
@{setting PATH}. Other schemes of installation are supported by the |
|
189 |
@{tool_def install} tool: |
|
28224 | 190 |
\begin{ttbox} |
50132 | 191 |
Usage: isabelle install [OPTIONS] BINDIR |
28224 | 192 |
|
193 |
Options are: |
|
50132 | 194 |
-d DISTDIR refer to DISTDIR as Isabelle distribution |
28224 | 195 |
(default ISABELLE_HOME) |
196 |
||
50132 | 197 |
Install Isabelle executables with absolute references to the |
28224 | 198 |
distribution directory. |
199 |
\end{ttbox} |
|
200 |
||
201 |
The @{verbatim "-d"} option overrides the current Isabelle |
|
202 |
distribution directory as determined by @{setting ISABELLE_HOME}. |
|
203 |
||
50132 | 204 |
The @{text BINDIR} argument tells where executable wrapper scripts |
205 |
for @{executable "isabelle-process"} and @{executable isabelle} |
|
206 |
should be placed, which is typically a directory in the shell's |
|
207 |
@{setting PATH}, such as @{verbatim "$HOME/bin"}. |
|
48815 | 208 |
|
50132 | 209 |
\medskip It is also possible to make symbolic links of the main |
210 |
Isabelle executables manually, but making separate copies outside |
|
211 |
the Isabelle distribution directory will not work! *} |
|
28224 | 212 |
|
213 |
||
214 |
section {* Creating instances of the Isabelle logo *} |
|
215 |
||
49072
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
216 |
text {* The @{tool_def logo} tool creates instances of the generic |
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
217 |
Isabelle logo as EPS and PDF, for inclusion in {\LaTeX} documents. |
28224 | 218 |
\begin{ttbox} |
49072
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
219 |
Usage: isabelle logo [OPTIONS] XYZ |
28224 | 220 |
|
49072
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
221 |
Create instance XYZ of the Isabelle logo (as EPS and PDF). |
28224 | 222 |
|
223 |
Options are: |
|
49072
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
224 |
-n NAME alternative output base name (default "isabelle_xyx") |
28224 | 225 |
-q quiet mode |
226 |
\end{ttbox} |
|
48936 | 227 |
|
49072
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
228 |
Option @{verbatim "-n"} specifies an altenative (base) name for the |
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
229 |
generated files. The default is @{verbatim "isabelle_"}@{text xyz} |
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
230 |
in lower-case. |
48936 | 231 |
|
232 |
Option @{verbatim "-q"} omits printing of the result file name. |
|
233 |
||
234 |
\medskip Implementors of Isabelle tools and applications are |
|
235 |
encouraged to make derived Isabelle logos for their own projects |
|
236 |
using this template. *} |
|
28224 | 237 |
|
238 |
||
48814
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
239 |
section {* Isabelle wrapper for make \label{sec:tool-make} *} |
28224 | 240 |
|
48814
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
241 |
text {* The old @{tool_def make} tool is a very simple wrapper for |
28224 | 242 |
ordinary Unix @{executable make}: |
243 |
\begin{ttbox} |
|
48602 | 244 |
Usage: isabelle make [ARGS ...] |
28224 | 245 |
|
246 |
Compile the logic in current directory using IsaMakefile. |
|
247 |
ARGS are directly passed to the system make program. |
|
248 |
\end{ttbox} |
|
249 |
||
250 |
Note that the Isabelle settings environment is also active. Thus one |
|
251 |
may refer to its values within the @{verbatim IsaMakefile}, e.g.\ |
|
48814
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
252 |
@{verbatim "$(ISABELLE_HOME)"}. Furthermore, programs started from |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
253 |
the make file also inherit this environment. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
254 |
*} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
255 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
256 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
257 |
section {* Creating Isabelle session directories |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
258 |
\label{sec:tool-mkdir} *} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
259 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
260 |
text {* The old @{tool_def mkdir} tool prepares Isabelle session |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
261 |
source directories, including a sensible default setup of @{verbatim |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
262 |
IsaMakefile}, @{verbatim ROOT.ML}, and a @{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:
48736
diff
changeset
|
263 |
directory with a minimal @{verbatim root.tex} that is sufficient to |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
264 |
print all theories of the session (in the order of appearance); see |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
265 |
\secref{sec:tool-document} for further information on Isabelle |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
266 |
document preparation. The usage of @{tool mkdir} is: |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
267 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
268 |
\begin{ttbox} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
269 |
Usage: isabelle mkdir [OPTIONS] [LOGIC] NAME |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
270 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
271 |
Options are: |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
272 |
-I FILE alternative IsaMakefile output |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
273 |
-P include parent logic target |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
274 |
-b setup build mode (session outputs heap image) |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
275 |
-q quiet mode |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
276 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
277 |
Prepare session directory, including IsaMakefile and document source, |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
278 |
with parent LOGIC (default ISABELLE_LOGIC=\$ISABELLE_LOGIC) |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
279 |
\end{ttbox} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
280 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
281 |
The @{tool mkdir} tool is conservative in the sense that any |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
282 |
existing @{verbatim IsaMakefile} etc.\ is left unchanged. Thus it |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
283 |
is safe to invoke it multiple times, although later runs may not |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
284 |
have the desired effect. |
28224 | 285 |
|
48814
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
286 |
Note that @{tool mkdir} is unable to change @{verbatim IsaMakefile} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
287 |
incrementally --- manual changes are required for multiple |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
288 |
sub-sessions. On order to get an initial working session, the only |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
289 |
editing needed is to add appropriate @{ML use_thy} calls to the |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
290 |
generated @{verbatim ROOT.ML} file. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
291 |
*} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
292 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
293 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
294 |
subsubsection {* Options *} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
295 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
296 |
text {* |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
297 |
The @{verbatim "-I"} option specifies an alternative to @{verbatim |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
298 |
IsaMakefile} for dependencies. Note that ``@{verbatim "-"}'' refers |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
299 |
to \emph{stdout}, i.e.\ ``@{verbatim "-I-"}'' provides an easy way |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
300 |
to peek at @{tool mkdir}'s idea of @{tool make} setup required for |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
301 |
some particular of Isabelle session. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
302 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
303 |
\medskip The @{verbatim "-P"} option includes a target for the |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
304 |
parent @{verbatim LOGIC} session in the generated @{verbatim |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
305 |
IsaMakefile}. The corresponding sources are assumed to be located |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
306 |
within the Isabelle distribution. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
307 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
308 |
\medskip The @{verbatim "-b"} option sets up the current directory |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
309 |
as the base for a new session that provides an actual logic image, |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
310 |
as opposed to one that only runs several theories based on an |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
311 |
existing image. Note that in the latter case, everything except |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
312 |
@{verbatim IsaMakefile} would be placed into a separate directory |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
313 |
@{verbatim NAME}, rather than the current one. See |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
314 |
\secref{sec:tool-usedir} for further information on \emph{build |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
315 |
mode} vs.\ \emph{example mode} of @{tool usedir}. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
316 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
317 |
\medskip The @{verbatim "-q"} option enables quiet mode, suppressing |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
318 |
further notes on how to proceed. |
28224 | 319 |
*} |
320 |
||
321 |
||
322 |
section {* Printing documents *} |
|
323 |
||
324 |
text {* |
|
48602 | 325 |
The @{tool_def print} tool prints documents: |
28224 | 326 |
\begin{ttbox} |
48602 | 327 |
Usage: isabelle print [OPTIONS] FILE |
28224 | 328 |
|
329 |
Options are: |
|
330 |
-c cleanup -- remove FILE after use |
|
331 |
||
332 |
Print document FILE. |
|
333 |
\end{ttbox} |
|
334 |
||
335 |
The @{verbatim "-c"} option causes the input file to be removed |
|
336 |
after use. The printer spool command is determined by the @{setting |
|
337 |
PRINT_COMMAND} setting. |
|
338 |
*} |
|
339 |
||
340 |
||
341 |
section {* Remove awkward symbol names from theory sources *} |
|
342 |
||
343 |
text {* |
|
48602 | 344 |
The @{tool_def unsymbolize} tool tunes Isabelle theory sources to |
28224 | 345 |
improve readability for plain ASCII output (e.g.\ in email |
346 |
communication). Most notably, @{tool unsymbolize} replaces awkward |
|
347 |
arrow symbols such as @{verbatim "\\"}@{verbatim "<Longrightarrow>"} |
|
348 |
by @{verbatim "==>"}. |
|
349 |
\begin{ttbox} |
|
48602 | 350 |
Usage: isabelle unsymbolize [FILES|DIRS...] |
28224 | 351 |
|
352 |
Recursively find .thy/.ML files, removing unreadable symbol names. |
|
353 |
Note: this is an ad-hoc script; there is no systematic way to replace |
|
354 |
symbols independently of the inner syntax of a theory! |
|
355 |
||
356 |
Renames old versions of FILES by appending "~~". |
|
357 |
\end{ttbox} |
|
358 |
*} |
|
359 |
||
360 |
||
48814
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
361 |
section {* Running Isabelle sessions \label{sec:tool-usedir} *} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
362 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
363 |
text {* The old @{tool_def usedir} tool builds object-logic images, or |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
364 |
runs example sessions based on existing logics. Its usage is: |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
365 |
\begin{ttbox} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
366 |
Usage: isabelle usedir [OPTIONS] LOGIC NAME |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
367 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
368 |
Options are: |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
369 |
-C BOOL copy existing document directory to -D PATH (default true) |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
370 |
-D PATH dump generated document sources into PATH |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
371 |
-M MAX multithreading: maximum number of worker threads (default 1) |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
372 |
-P PATH set path for remote theory browsing information |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
373 |
-Q INT set threshold for sub-proof parallelization (default 50) |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
374 |
-T LEVEL multithreading: trace level (default 0) |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
375 |
-V VARIANT declare alternative document VARIANT |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
376 |
-b build mode (output heap image, using current dir) |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
377 |
-d FORMAT build document as FORMAT (default false) |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
378 |
-f NAME use ML file NAME (default ROOT.ML) |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
379 |
-g BOOL generate session graph image for document (default false) |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
380 |
-i BOOL generate theory browser information (default false) |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
381 |
-m MODE add print mode for output |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
382 |
-p LEVEL set level of detail for proof objects (default 0) |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
383 |
-q LEVEL set level of parallel proof checking (default 1) |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
384 |
-r reset session path |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
385 |
-s NAME override session NAME |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
386 |
-t BOOL internal session timing (default false) |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
387 |
-v BOOL be verbose (default false) |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
388 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
389 |
Build object-logic or run examples. Also creates browsing |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
390 |
information (HTML etc.) according to settings. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
391 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
392 |
ISABELLE_USEDIR_OPTIONS=... |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
393 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
394 |
ML_PLATFORM=... |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
395 |
ML_HOME=... |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
396 |
ML_SYSTEM=... |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
397 |
ML_OPTIONS=... |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
398 |
\end{ttbox} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
399 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
400 |
Note that the value of the @{setting_ref ISABELLE_USEDIR_OPTIONS} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
401 |
setting is implicitly prefixed to \emph{any} @{tool usedir} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
402 |
call. Since the @{verbatim IsaMakefile}s of all object-logics |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
403 |
distributed with Isabelle just invoke @{tool usedir} for the real |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
404 |
work, one may control compilation options globally via above |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
405 |
variable. In particular, generation of \rmindex{HTML} browsing |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
406 |
information and document preparation is controlled here. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
407 |
*} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
408 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
409 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
410 |
subsubsection {* Options *} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
411 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
412 |
text {* |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
413 |
Basically, there are two different modes of operation: \emph{build |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
414 |
mode} (enabled through the @{verbatim "-b"} option) and |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
415 |
\emph{example mode} (default). |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
416 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
417 |
Calling @{tool usedir} with @{verbatim "-b"} runs @{executable |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
418 |
"isabelle-process"} with input image @{verbatim LOGIC} and output to |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
419 |
@{verbatim NAME}, as provided on the command line. This will be a |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
420 |
batch session, running @{verbatim ROOT.ML} from the current |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
421 |
directory and then quitting. It is assumed that @{verbatim ROOT.ML} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
422 |
contains all ML commands required to build the logic. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
423 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
424 |
In example mode, @{tool usedir} runs a read-only session of |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
425 |
@{verbatim LOGIC} and automatically runs @{verbatim ROOT.ML} from |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
426 |
within directory @{verbatim NAME}. It assumes that this file |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
427 |
contains appropriate ML commands to run the desired examples. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
428 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
429 |
\medskip The @{verbatim "-i"} option controls theory browser data |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
430 |
generation. It may be explicitly turned on or off --- as usual, the |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
431 |
last occurrence of @{verbatim "-i"} on the command line wins. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
432 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
433 |
The @{verbatim "-P"} option specifies a path (or actual URL) to be |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
434 |
prefixed to any \emph{non-local} reference of existing theories. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
435 |
Thus user sessions may easily link to existing Isabelle libraries |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
436 |
already present on the WWW. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
437 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
438 |
The @{verbatim "-m"} options specifies additional print modes to be |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
439 |
activated temporarily while the session is processed. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
440 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
441 |
\medskip The @{verbatim "-d"} option controls 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:
48736
diff
changeset
|
442 |
Valid arguments are @{verbatim false} (do not prepare any document; |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
443 |
this is default), or any of @{verbatim dvi}, @{verbatim dvi.gz}, |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
444 |
@{verbatim ps}, @{verbatim ps.gz}, @{verbatim pdf}. The logic |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
445 |
session has to provide a properly setup @{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:
48736
diff
changeset
|
446 |
directory. See \secref{sec:tool-document} and |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
447 |
\secref{sec:tool-latex} for more details. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
448 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
449 |
\medskip The @{verbatim "-V"} option declares alternative document |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
450 |
variants, consisting of name/tags pairs (cf.\ options @{verbatim |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
451 |
"-n"} and @{verbatim "-t"} of @{tool_ref document}). The standard |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
452 |
document is equivalent to ``@{verbatim |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
453 |
"document=theory,proof,ML"}'', which means that all theory begin/end |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
454 |
commands, proof body texts, and ML code will be presented |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
455 |
faithfully. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
456 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
457 |
An alternative variant ``@{verbatim "outline=/proof/ML"}'' would |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
458 |
fold proof and ML parts, replacing the original text by a short |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
459 |
place-holder. The form ``@{text name}@{verbatim "=-"},'' means to |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
460 |
remove document @{text name} from the list of variants to be |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
461 |
processed. Any number of @{verbatim "-V"} options may be given; |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
462 |
later declarations have precedence over earlier ones. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
463 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
464 |
Some document variant @{text name} may use an alternative {\LaTeX} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
465 |
entry point called @{verbatim "document/root_"}@{text |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
466 |
"name"}@{verbatim ".tex"} if that file exists; otherwise the common |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
467 |
@{verbatim "document/root.tex"} is used. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
468 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
469 |
\medskip The @{verbatim "-g"} option produces images of the theory |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
470 |
dependency graph (cf.\ \secref{sec:browse}) for inclusion 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:
48736
diff
changeset
|
471 |
generated document, both as @{verbatim session_graph.eps} and |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
472 |
@{verbatim session_graph.pdf} at the same time. To include this in |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
473 |
the final {\LaTeX} document one could say @{verbatim |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
474 |
"\\includegraphics{session_graph}"} in @{verbatim |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
475 |
"document/root.tex"} (omitting the file-name extension enables |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
476 |
{\LaTeX} to select to correct version, either for the DVI or PDF |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
477 |
output path). |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
478 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
479 |
\medskip The @{verbatim "-D"} option causes the generated document |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
480 |
sources to be dumped at location @{verbatim PATH}; this path is |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
481 |
relative to the session's main directory. If the @{verbatim "-C"} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
482 |
option is true, this will include a copy of an existing @{verbatim |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
483 |
document} directory as provided by the user. For example, @{tool |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
484 |
usedir}~@{verbatim "-D generated HOL Foo"} produces a complete set |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
485 |
of document sources at @{verbatim "Foo/generated"}. Subsequent |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
486 |
invocation of @{tool document}~@{verbatim "Foo/generated"} (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:
48736
diff
changeset
|
487 |
\secref{sec:tool-document}) will process the final result |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
488 |
independently of an Isabelle job. This decoupled mode of operation |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
489 |
facilitates debugging of serious {\LaTeX} errors, for example. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
490 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
491 |
\medskip The @{verbatim "-p"} option determines the level of detail |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
492 |
for internal proof objects, see also the \emph{Isabelle Reference |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
493 |
Manual}~\cite{isabelle-ref}. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
494 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
495 |
\medskip The @{verbatim "-q"} option specifies the level of parallel |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
496 |
proof checking: @{verbatim 0} no proofs, @{verbatim 1} toplevel |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
497 |
proofs (default), @{verbatim 2} toplevel and nested Isar proofs. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
498 |
The option @{verbatim "-Q"} specifies a threshold for @{verbatim |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
499 |
"-q2"}: nested proofs are only parallelized when the current number |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
500 |
of forked proofs falls below the given value (default 50), |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
501 |
multiplied by the number of worker threads (see option @{verbatim |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
502 |
"-M"}). |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
503 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
504 |
\medskip The @{verbatim "-t"} option produces a more detailed |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
505 |
internal timing report of the session. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
506 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
507 |
\medskip The @{verbatim "-v"} option causes additional information |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
508 |
to be printed while running the session, notably the location of |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
509 |
prepared documents. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
510 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
511 |
\medskip The @{verbatim "-M"} option specifies the maximum number of |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
512 |
parallel worker threads used for processing independent tasks when |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
513 |
checking theory sources (multithreading only works on suitable ML |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
514 |
platforms). The special value of @{verbatim 0} or @{verbatim max} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
515 |
refers to the number of actual CPU cores of the underlying machine, |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
516 |
which is a good starting point for optimal performance tuning. The |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
517 |
@{verbatim "-T"} option determines the level of detail in tracing |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
518 |
output concerning the internal locking and scheduling in |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
519 |
multithreaded operation. This may be helpful in isolating |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
520 |
performance bottle-necks, e.g.\ due to excessive wait states when |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
521 |
locking critical code sections. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
522 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
523 |
\medskip Any @{tool usedir} session is named by some \emph{session |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
524 |
identifier}. These accumulate, documenting the way sessions depend |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
525 |
on others. For example, consider @{verbatim "Pure/FOL/ex"}, which |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
526 |
refers to the examples of FOL, which in turn is built upon Pure. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
527 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
528 |
The current session's identifier is by default just the base name of |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
529 |
the @{verbatim LOGIC} argument (in build mode), or of the @{verbatim |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
530 |
NAME} argument (in example mode). This may be overridden explicitly |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
531 |
via the @{verbatim "-s"} option. |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
532 |
*} |
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
533 |
|
d488a5f25bf6
some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents:
48736
diff
changeset
|
534 |
|
28224 | 535 |
section {* Output the version identifier of the Isabelle distribution *} |
536 |
||
537 |
text {* |
|
48602 | 538 |
The @{tool_def version} tool displays Isabelle version information: |
41511 | 539 |
\begin{ttbox} |
540 |
Usage: isabelle version [OPTIONS] |
|
541 |
||
542 |
Options are: |
|
543 |
-i short identification (derived from Mercurial id) |
|
544 |
||
545 |
Display Isabelle version information. |
|
546 |
\end{ttbox} |
|
547 |
||
548 |
\medskip The default is to output the full version string of the |
|
47827 | 549 |
Isabelle distribution, e.g.\ ``@{verbatim "Isabelle2012: May 2012"}. |
41511 | 550 |
|
551 |
The @{verbatim "-i"} option produces a short identification derived |
|
552 |
from the Mercurial id of the @{setting ISABELLE_HOME} directory. |
|
28224 | 553 |
*} |
554 |
||
555 |
||
556 |
section {* Convert XML to YXML *} |
|
557 |
||
558 |
text {* |
|
559 |
The @{tool_def yxml} tool converts a standard XML document (stdin) |
|
560 |
to the much simpler and more efficient YXML format of Isabelle |
|
561 |
(stdout). The YXML format is defined as follows. |
|
562 |
||
563 |
\begin{enumerate} |
|
564 |
||
565 |
\item The encoding is always UTF-8. |
|
566 |
||
567 |
\item Body text is represented verbatim (no escaping, no special |
|
568 |
treatment of white space, no named entities, no CDATA chunks, no |
|
569 |
comments). |
|
570 |
||
571 |
\item Markup elements are represented via ASCII control characters |
|
572 |
@{text "\<^bold>X = 5"} and @{text "\<^bold>Y = 6"} as follows: |
|
573 |
||
574 |
\begin{tabular}{ll} |
|
575 |
XML & YXML \\\hline |
|
576 |
@{verbatim "<"}@{text "name attribute"}@{verbatim "="}@{text "value \<dots>"}@{verbatim ">"} & |
|
577 |
@{text "\<^bold>X\<^bold>Yname\<^bold>Yattribute"}@{verbatim "="}@{text "value\<dots>\<^bold>X"} \\ |
|
578 |
@{verbatim "</"}@{text name}@{verbatim ">"} & @{text "\<^bold>X\<^bold>Y\<^bold>X"} \\ |
|
579 |
\end{tabular} |
|
580 |
||
581 |
There is no special case for empty body text, i.e.\ @{verbatim |
|
582 |
"<foo/>"} is treated like @{verbatim "<foo></foo>"}. Also note that |
|
583 |
@{text "\<^bold>X"} and @{text "\<^bold>Y"} may never occur in |
|
584 |
well-formed XML documents. |
|
585 |
||
586 |
\end{enumerate} |
|
587 |
||
588 |
Parsing YXML is pretty straight-forward: split the text into chunks |
|
589 |
separated by @{text "\<^bold>X"}, then split each chunk into |
|
590 |
sub-chunks separated by @{text "\<^bold>Y"}. Markup chunks start |
|
591 |
with an empty sub-chunk, and a second empty sub-chunk indicates |
|
592 |
close of an element. Any other non-empty chunk consists of plain |
|
44799 | 593 |
text. For example, see @{file "~~/src/Pure/PIDE/yxml.ML"} or |
594 |
@{file "~~/src/Pure/PIDE/yxml.scala"}. |
|
28224 | 595 |
|
596 |
YXML documents may be detected quickly by checking that the first |
|
597 |
two characters are @{text "\<^bold>X\<^bold>Y"}. |
|
598 |
*} |
|
599 |
||
600 |
end |