author | blanchet |
Tue, 10 Sep 2013 15:56:51 +0200 | |
changeset 53500 | 53b9326196fe |
parent 53435 | 2220f0fb5581 |
child 54683 | cf48ddc266e5 |
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) |
|
53435 | 24 |
-a resolve all missing components |
48844 | 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 |
|
53435 | 45 |
resolved. Explicit components may be named as command |
48844 | 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 |
||
52444
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
63 |
section {* Displaying documents \label{sec:tool-display} *} |
28224 | 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 |
||
52444
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
73 |
Display document FILE (in DVI or PDF format). |
28224 | 74 |
\end{ttbox} |
75 |
||
76 |
\medskip The @{verbatim "-c"} option causes the input file to be |
|
52444
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
77 |
removed after use. |
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
78 |
|
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
79 |
\medskip The settings @{setting DVI_VIEWER} and @{setting |
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
80 |
PDF_VIEWER} determine the programs for viewing the corresponding |
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
81 |
file formats. |
28224 | 82 |
*} |
83 |
||
84 |
||
85 |
section {* Viewing documentation \label{sec:tool-doc} *} |
|
86 |
||
87 |
text {* |
|
52444
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
88 |
The @{tool_def doc} tool displays Isabelle documentation: |
28224 | 89 |
\begin{ttbox} |
52444
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
90 |
Usage: isabelle doc [DOC ...] |
28224 | 91 |
|
52444
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
92 |
View Isabelle documentation. |
28224 | 93 |
\end{ttbox} |
94 |
If called without arguments, it lists all available documents. Each |
|
95 |
line starts with an identifier, followed by a short description. Any |
|
52444
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
96 |
of these identifiers may be specified as arguments, in order to |
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
97 |
display the corresponding document (see also |
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
98 |
\secref{sec:tool-display}). |
28224 | 99 |
|
100 |
\medskip The @{setting ISABELLE_DOCS} setting specifies the list of |
|
101 |
directories (separated by colons) to be scanned for documentations. |
|
102 |
*} |
|
103 |
||
104 |
||
47828 | 105 |
section {* Shell commands within the settings environment \label{sec:tool-env} *} |
106 |
||
48602 | 107 |
text {* The @{tool_def env} tool is a direct wrapper for the standard |
108 |
@{verbatim "/usr/bin/env"} command on POSIX systems, running within |
|
109 |
the Isabelle settings environment (\secref{sec:settings}). |
|
47828 | 110 |
|
111 |
The command-line arguments are that of the underlying version of |
|
112 |
@{verbatim env}. For example, the following invokes an instance of |
|
113 |
the GNU Bash shell within the Isabelle environment: |
|
114 |
\begin{alltt} |
|
115 |
isabelle env bash |
|
116 |
\end{alltt} |
|
117 |
*} |
|
118 |
||
119 |
||
28224 | 120 |
section {* Getting logic images *} |
121 |
||
48602 | 122 |
text {* The @{tool_def findlogics} tool traverses all directories |
28224 | 123 |
specified in @{setting ISABELLE_PATH}, looking for Isabelle logic |
124 |
images. Its usage is: |
|
125 |
\begin{ttbox} |
|
48577 | 126 |
Usage: isabelle findlogics |
28224 | 127 |
|
128 |
Collect heap file names from ISABELLE_PATH. |
|
129 |
\end{ttbox} |
|
130 |
||
131 |
The base names of all files found on the path are printed --- sorted |
|
132 |
and with duplicates removed. Also note that lookup in @{setting |
|
133 |
ISABELLE_PATH} includes the current values of @{setting ML_SYSTEM} |
|
134 |
and @{setting ML_PLATFORM}. Thus switching to another ML compiler |
|
135 |
may change the set of logic images available. |
|
136 |
*} |
|
137 |
||
138 |
||
139 |
section {* Inspecting the settings environment \label{sec:tool-getenv} *} |
|
140 |
||
48602 | 141 |
text {* The Isabelle settings environment --- as provided by the |
28224 | 142 |
site-default and user-specific settings files --- can be inspected |
48602 | 143 |
with the @{tool_def getenv} tool: |
28224 | 144 |
\begin{ttbox} |
48602 | 145 |
Usage: isabelle getenv [OPTIONS] [VARNAMES ...] |
28224 | 146 |
|
147 |
Options are: |
|
148 |
-a display complete environment |
|
149 |
-b print values only (doesn't work for -a) |
|
31497 | 150 |
-d FILE dump complete environment to FILE |
151 |
(null terminated entries) |
|
28224 | 152 |
|
153 |
Get value of VARNAMES from the Isabelle settings. |
|
154 |
\end{ttbox} |
|
155 |
||
156 |
With the @{verbatim "-a"} option, one may inspect the full process |
|
157 |
environment that Isabelle related programs are run in. This usually |
|
158 |
contains much more variables than are actually Isabelle settings. |
|
159 |
Normally, output is a list of lines of the form @{text |
|
160 |
name}@{verbatim "="}@{text value}. The @{verbatim "-b"} option |
|
161 |
causes only the values to be printed. |
|
31497 | 162 |
|
163 |
Option @{verbatim "-d"} produces a dump of the complete environment |
|
164 |
to the specified file. Entries are terminated by the ASCII null |
|
165 |
character, i.e.\ the C string terminator. |
|
28224 | 166 |
*} |
167 |
||
168 |
||
169 |
subsubsection {* Examples *} |
|
170 |
||
48815 | 171 |
text {* Get the location of @{setting ISABELLE_HOME_USER} where |
172 |
user-specific information is stored: |
|
28224 | 173 |
\begin{ttbox} |
48815 | 174 |
isabelle getenv ISABELLE_HOME_USER |
28224 | 175 |
\end{ttbox} |
176 |
||
48815 | 177 |
\medskip Get the value only of the same settings variable, which is |
178 |
particularly useful in shell scripts: |
|
28224 | 179 |
\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
|
180 |
isabelle getenv -b ISABELLE_OUTPUT |
28224 | 181 |
\end{ttbox} |
182 |
*} |
|
183 |
||
184 |
||
185 |
section {* Installing standalone Isabelle executables \label{sec:tool-install} *} |
|
186 |
||
48602 | 187 |
text {* By default, the main Isabelle binaries (@{executable |
188 |
"isabelle"} etc.) are just run from their location within the |
|
189 |
distribution directory, probably indirectly by the shell through its |
|
190 |
@{setting PATH}. Other schemes of installation are supported by the |
|
191 |
@{tool_def install} tool: |
|
28224 | 192 |
\begin{ttbox} |
50132 | 193 |
Usage: isabelle install [OPTIONS] BINDIR |
28224 | 194 |
|
195 |
Options are: |
|
50132 | 196 |
-d DISTDIR refer to DISTDIR as Isabelle distribution |
28224 | 197 |
(default ISABELLE_HOME) |
198 |
||
50132 | 199 |
Install Isabelle executables with absolute references to the |
28224 | 200 |
distribution directory. |
201 |
\end{ttbox} |
|
202 |
||
203 |
The @{verbatim "-d"} option overrides the current Isabelle |
|
204 |
distribution directory as determined by @{setting ISABELLE_HOME}. |
|
205 |
||
50132 | 206 |
The @{text BINDIR} argument tells where executable wrapper scripts |
207 |
for @{executable "isabelle-process"} and @{executable isabelle} |
|
208 |
should be placed, which is typically a directory in the shell's |
|
209 |
@{setting PATH}, such as @{verbatim "$HOME/bin"}. |
|
48815 | 210 |
|
50132 | 211 |
\medskip It is also possible to make symbolic links of the main |
212 |
Isabelle executables manually, but making separate copies outside |
|
213 |
the Isabelle distribution directory will not work! *} |
|
28224 | 214 |
|
215 |
||
216 |
section {* Creating instances of the Isabelle logo *} |
|
217 |
||
49072
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
218 |
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
|
219 |
Isabelle logo as EPS and PDF, for inclusion in {\LaTeX} documents. |
28224 | 220 |
\begin{ttbox} |
49072
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
221 |
Usage: isabelle logo [OPTIONS] XYZ |
28224 | 222 |
|
49072
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
223 |
Create instance XYZ of the Isabelle logo (as EPS and PDF). |
28224 | 224 |
|
225 |
Options are: |
|
49072
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
226 |
-n NAME alternative output base name (default "isabelle_xyx") |
28224 | 227 |
-q quiet mode |
228 |
\end{ttbox} |
|
48936 | 229 |
|
49072
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
230 |
Option @{verbatim "-n"} specifies an altenative (base) name for the |
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
231 |
generated files. The default is @{verbatim "isabelle_"}@{text xyz} |
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
232 |
in lower-case. |
48936 | 233 |
|
234 |
Option @{verbatim "-q"} omits printing of the result file name. |
|
235 |
||
236 |
\medskip Implementors of Isabelle tools and applications are |
|
237 |
encouraged to make derived Isabelle logos for their own projects |
|
238 |
using this template. *} |
|
28224 | 239 |
|
240 |
||
241 |
section {* Remove awkward symbol names from theory sources *} |
|
242 |
||
243 |
text {* |
|
48602 | 244 |
The @{tool_def unsymbolize} tool tunes Isabelle theory sources to |
28224 | 245 |
improve readability for plain ASCII output (e.g.\ in email |
246 |
communication). Most notably, @{tool unsymbolize} replaces awkward |
|
247 |
arrow symbols such as @{verbatim "\\"}@{verbatim "<Longrightarrow>"} |
|
248 |
by @{verbatim "==>"}. |
|
249 |
\begin{ttbox} |
|
48602 | 250 |
Usage: isabelle unsymbolize [FILES|DIRS...] |
28224 | 251 |
|
252 |
Recursively find .thy/.ML files, removing unreadable symbol names. |
|
253 |
Note: this is an ad-hoc script; there is no systematic way to replace |
|
254 |
symbols independently of the inner syntax of a theory! |
|
255 |
||
256 |
Renames old versions of FILES by appending "~~". |
|
257 |
\end{ttbox} |
|
258 |
*} |
|
259 |
||
260 |
||
261 |
section {* Output the version identifier of the Isabelle distribution *} |
|
262 |
||
263 |
text {* |
|
48602 | 264 |
The @{tool_def version} tool displays Isabelle version information: |
41511 | 265 |
\begin{ttbox} |
266 |
Usage: isabelle version [OPTIONS] |
|
267 |
||
268 |
Options are: |
|
269 |
-i short identification (derived from Mercurial id) |
|
270 |
||
271 |
Display Isabelle version information. |
|
272 |
\end{ttbox} |
|
273 |
||
274 |
\medskip The default is to output the full version string of the |
|
47827 | 275 |
Isabelle distribution, e.g.\ ``@{verbatim "Isabelle2012: May 2012"}. |
41511 | 276 |
|
277 |
The @{verbatim "-i"} option produces a short identification derived |
|
278 |
from the Mercurial id of the @{setting ISABELLE_HOME} directory. |
|
28224 | 279 |
*} |
280 |
||
281 |
||
282 |
section {* Convert XML to YXML *} |
|
283 |
||
284 |
text {* |
|
285 |
The @{tool_def yxml} tool converts a standard XML document (stdin) |
|
286 |
to the much simpler and more efficient YXML format of Isabelle |
|
287 |
(stdout). The YXML format is defined as follows. |
|
288 |
||
289 |
\begin{enumerate} |
|
290 |
||
291 |
\item The encoding is always UTF-8. |
|
292 |
||
293 |
\item Body text is represented verbatim (no escaping, no special |
|
294 |
treatment of white space, no named entities, no CDATA chunks, no |
|
295 |
comments). |
|
296 |
||
297 |
\item Markup elements are represented via ASCII control characters |
|
298 |
@{text "\<^bold>X = 5"} and @{text "\<^bold>Y = 6"} as follows: |
|
299 |
||
300 |
\begin{tabular}{ll} |
|
301 |
XML & YXML \\\hline |
|
302 |
@{verbatim "<"}@{text "name attribute"}@{verbatim "="}@{text "value \<dots>"}@{verbatim ">"} & |
|
303 |
@{text "\<^bold>X\<^bold>Yname\<^bold>Yattribute"}@{verbatim "="}@{text "value\<dots>\<^bold>X"} \\ |
|
304 |
@{verbatim "</"}@{text name}@{verbatim ">"} & @{text "\<^bold>X\<^bold>Y\<^bold>X"} \\ |
|
305 |
\end{tabular} |
|
306 |
||
307 |
There is no special case for empty body text, i.e.\ @{verbatim |
|
308 |
"<foo/>"} is treated like @{verbatim "<foo></foo>"}. Also note that |
|
309 |
@{text "\<^bold>X"} and @{text "\<^bold>Y"} may never occur in |
|
310 |
well-formed XML documents. |
|
311 |
||
312 |
\end{enumerate} |
|
313 |
||
314 |
Parsing YXML is pretty straight-forward: split the text into chunks |
|
315 |
separated by @{text "\<^bold>X"}, then split each chunk into |
|
316 |
sub-chunks separated by @{text "\<^bold>Y"}. Markup chunks start |
|
317 |
with an empty sub-chunk, and a second empty sub-chunk indicates |
|
318 |
close of an element. Any other non-empty chunk consists of plain |
|
44799 | 319 |
text. For example, see @{file "~~/src/Pure/PIDE/yxml.ML"} or |
320 |
@{file "~~/src/Pure/PIDE/yxml.scala"}. |
|
28224 | 321 |
|
322 |
YXML documents may be detected quickly by checking that the first |
|
323 |
two characters are @{text "\<^bold>X\<^bold>Y"}. |
|
324 |
*} |
|
325 |
||
326 |
end |