| author | berghofe | 
| Tue, 19 Apr 2011 14:17:41 +0200 | |
| changeset 42416 | a8a9f4d79196 | 
| parent 41512 | 8445396e1e39 | 
| child 43564 | 9864182c6bad | 
| permissions | -rw-r--r-- | 
| 28224 | 1  | 
theory Misc  | 
2  | 
imports Pure  | 
|
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  | 
||
13  | 
section {* Displaying documents *}
 | 
|
14  | 
||
15  | 
text {*
 | 
|
16  | 
  The @{tool_def display} utility displays documents in DVI or PDF
 | 
|
17  | 
format:  | 
|
18  | 
\begin{ttbox}
 | 
|
19  | 
Usage: display [OPTIONS] FILE  | 
|
20  | 
||
21  | 
Options are:  | 
|
22  | 
-c cleanup -- remove FILE after use  | 
|
23  | 
||
24  | 
Display document FILE (in DVI format).  | 
|
25  | 
\end{ttbox}
 | 
|
26  | 
||
27  | 
  \medskip The @{verbatim "-c"} option causes the input file to be
 | 
|
28  | 
  removed after use.  The program for viewing @{verbatim dvi} files is
 | 
|
29  | 
  determined by the @{setting DVI_VIEWER} setting.
 | 
|
30  | 
*}  | 
|
31  | 
||
32  | 
||
33  | 
section {* Viewing documentation \label{sec:tool-doc} *}
 | 
|
34  | 
||
35  | 
text {*
 | 
|
36  | 
  The @{tool_def doc} utility displays online documentation:
 | 
|
37  | 
\begin{ttbox}
 | 
|
38  | 
Usage: doc [DOC]  | 
|
39  | 
||
40  | 
View Isabelle documentation DOC, or show list of available documents.  | 
|
41  | 
\end{ttbox}
 | 
|
42  | 
If called without arguments, it lists all available documents. Each  | 
|
43  | 
line starts with an identifier, followed by a short description. Any  | 
|
44  | 
of these identifiers may be specified as the first argument in order  | 
|
45  | 
to have the corresponding document displayed.  | 
|
46  | 
||
47  | 
  \medskip The @{setting ISABELLE_DOCS} setting specifies the list of
 | 
|
48  | 
directories (separated by colons) to be scanned for documentations.  | 
|
49  | 
  The program for viewing @{verbatim dvi} files is determined by the
 | 
|
50  | 
  @{setting DVI_VIEWER} setting.
 | 
|
51  | 
*}  | 
|
52  | 
||
53  | 
||
54  | 
section {* Getting logic images *}
 | 
|
55  | 
||
56  | 
text {*
 | 
|
57  | 
  The @{tool_def findlogics} utility traverses all directories
 | 
|
58  | 
  specified in @{setting ISABELLE_PATH}, looking for Isabelle logic
 | 
|
59  | 
images. Its usage is:  | 
|
60  | 
\begin{ttbox}
 | 
|
61  | 
Usage: findlogics  | 
|
62  | 
||
63  | 
Collect heap file names from ISABELLE_PATH.  | 
|
64  | 
\end{ttbox}
 | 
|
65  | 
||
66  | 
The base names of all files found on the path are printed --- sorted  | 
|
67  | 
  and with duplicates removed. Also note that lookup in @{setting
 | 
|
68  | 
  ISABELLE_PATH} includes the current values of @{setting ML_SYSTEM}
 | 
|
69  | 
  and @{setting ML_PLATFORM}. Thus switching to another ML compiler
 | 
|
70  | 
may change the set of logic images available.  | 
|
71  | 
*}  | 
|
72  | 
||
73  | 
||
74  | 
section {* Inspecting the settings environment \label{sec:tool-getenv} *}
 | 
|
75  | 
||
76  | 
text {*
 | 
|
77  | 
The Isabelle settings environment --- as provided by the  | 
|
78  | 
site-default and user-specific settings files --- can be inspected  | 
|
79  | 
  with the @{tool_def getenv} utility:
 | 
|
80  | 
\begin{ttbox}
 | 
|
81  | 
Usage: getenv [OPTIONS] [VARNAMES ...]  | 
|
82  | 
||
83  | 
Options are:  | 
|
84  | 
-a display complete environment  | 
|
85  | 
-b print values only (doesn't work for -a)  | 
|
| 31497 | 86  | 
-d FILE dump complete environment to FILE  | 
87  | 
(null terminated entries)  | 
|
| 28224 | 88  | 
|
89  | 
Get value of VARNAMES from the Isabelle settings.  | 
|
90  | 
\end{ttbox}
 | 
|
91  | 
||
92  | 
  With the @{verbatim "-a"} option, one may inspect the full process
 | 
|
93  | 
environment that Isabelle related programs are run in. This usually  | 
|
94  | 
contains much more variables than are actually Isabelle settings.  | 
|
95  | 
  Normally, output is a list of lines of the form @{text
 | 
|
96  | 
  name}@{verbatim "="}@{text value}. The @{verbatim "-b"} option
 | 
|
97  | 
causes only the values to be printed.  | 
|
| 31497 | 98  | 
|
99  | 
  Option @{verbatim "-d"} produces a dump of the complete environment
 | 
|
100  | 
to the specified file. Entries are terminated by the ASCII null  | 
|
101  | 
character, i.e.\ the C string terminator.  | 
|
| 28224 | 102  | 
*}  | 
103  | 
||
104  | 
||
105  | 
subsubsection {* Examples *}
 | 
|
106  | 
||
107  | 
text {*
 | 
|
108  | 
Get the ML system name and the location where the compiler binaries  | 
|
109  | 
are supposed to reside as follows:  | 
|
110  | 
\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
 | 
111  | 
isabelle getenv ML_SYSTEM ML_HOME  | 
| 28224 | 112  | 
{\out ML_SYSTEM=polyml}
 | 
113  | 
{\out ML_HOME=/usr/share/polyml/x86-linux}
 | 
|
114  | 
\end{ttbox}
 | 
|
115  | 
||
116  | 
The next one peeks at the output directory for Isabelle logic  | 
|
117  | 
images:  | 
|
118  | 
\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
 | 
119  | 
isabelle getenv -b ISABELLE_OUTPUT  | 
| 28224 | 120  | 
{\out /home/me/isabelle/heaps/polyml_x86-linux}
 | 
121  | 
\end{ttbox}
 | 
|
122  | 
  Here we have used the @{verbatim "-b"} option to suppress the
 | 
|
123  | 
  @{verbatim "ISABELLE_OUTPUT="} prefix.  The value above is what
 | 
|
124  | 
became of the following assignment in the default settings file:  | 
|
125  | 
\begin{ttbox}
 | 
|
126  | 
ISABELLE_OUTPUT="\$ISABELLE_HOME_USER/heaps"  | 
|
127  | 
\end{ttbox}
 | 
|
128  | 
||
129  | 
  Note how the @{setting ML_IDENTIFIER} value got appended
 | 
|
130  | 
automatically to each path component. This is a special feature of  | 
|
131  | 
  @{setting ISABELLE_OUTPUT}.
 | 
|
132  | 
*}  | 
|
133  | 
||
134  | 
||
135  | 
section {* Installing standalone Isabelle executables \label{sec:tool-install} *}
 | 
|
136  | 
||
137  | 
text {*
 | 
|
| 
28504
 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 
wenzelm 
parents: 
28253 
diff
changeset
 | 
138  | 
  By default, the main Isabelle binaries (@{executable "isabelle"}
 | 
| 
 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 
wenzelm 
parents: 
28253 
diff
changeset
 | 
139  | 
etc.) are just run from their location within the distribution  | 
| 
 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 
wenzelm 
parents: 
28253 
diff
changeset
 | 
140  | 
  directory, probably indirectly by the shell through its @{setting
 | 
| 
 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 
wenzelm 
parents: 
28253 
diff
changeset
 | 
141  | 
PATH}. Other schemes of installation are supported by the  | 
| 
 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 
wenzelm 
parents: 
28253 
diff
changeset
 | 
142  | 
  @{tool_def install} utility:
 | 
| 28224 | 143  | 
\begin{ttbox}
 | 
144  | 
Usage: install [OPTIONS]  | 
|
145  | 
||
146  | 
Options are:  | 
|
147  | 
-d DISTDIR use DISTDIR as Isabelle distribution  | 
|
148  | 
(default ISABELLE_HOME)  | 
|
149  | 
-p DIR install standalone binaries in DIR  | 
|
150  | 
||
151  | 
Install Isabelle executables with absolute references to the current  | 
|
152  | 
distribution directory.  | 
|
153  | 
\end{ttbox}
 | 
|
154  | 
||
155  | 
  The @{verbatim "-d"} option overrides the current Isabelle
 | 
|
156  | 
  distribution directory as determined by @{setting ISABELLE_HOME}.
 | 
|
157  | 
||
158  | 
  The @{verbatim "-p"} option installs executable wrapper scripts for
 | 
|
| 
28504
 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 
wenzelm 
parents: 
28253 
diff
changeset
 | 
159  | 
  @{executable "isabelle-process"}, @{executable isabelle},
 | 
| 28238 | 160  | 
  @{executable Isabelle}, containing proper absolute references to the
 | 
161  | 
  Isabelle distribution directory.  A typical @{verbatim DIR}
 | 
|
162  | 
specification would be some directory expected to be in the shell's  | 
|
163  | 
  @{setting PATH}, such as @{verbatim "/usr/local/bin"}.  It is
 | 
|
164  | 
important to note that a plain manual copy of the original Isabelle  | 
|
165  | 
executables does not work, since it disrupts the integrity of the  | 
|
166  | 
Isabelle distribution.  | 
|
| 28224 | 167  | 
*}  | 
168  | 
||
169  | 
||
170  | 
section {* Creating instances of the Isabelle logo *}
 | 
|
171  | 
||
172  | 
text {*
 | 
|
173  | 
  The @{tool_def logo} utility creates any instance of the generic
 | 
|
174  | 
Isabelle logo as an Encapsuled Postscript file (EPS):  | 
|
175  | 
\begin{ttbox}
 | 
|
176  | 
Usage: logo [OPTIONS] NAME  | 
|
177  | 
||
178  | 
Create instance NAME of the Isabelle logo (as EPS).  | 
|
179  | 
||
180  | 
Options are:  | 
|
181  | 
-o OUTFILE set output file (default determined from NAME)  | 
|
182  | 
-q quiet mode  | 
|
183  | 
\end{ttbox}
 | 
|
184  | 
You are encouraged to use this to create a derived logo for your  | 
|
| 
28504
 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 
wenzelm 
parents: 
28253 
diff
changeset
 | 
185  | 
  Isabelle project.  For example, @{verbatim isabelle} @{tool
 | 
| 28238 | 186  | 
  logo}~@{verbatim Bali} creates @{verbatim isabelle_bali.eps}.
 | 
| 28224 | 187  | 
*}  | 
188  | 
||
189  | 
||
190  | 
section {* Isabelle's version of make \label{sec:tool-make} *}
 | 
|
191  | 
||
192  | 
text {*
 | 
|
193  | 
  The Isabelle @{tool_def make} utility is a very simple wrapper for
 | 
|
194  | 
  ordinary Unix @{executable make}:
 | 
|
195  | 
\begin{ttbox}
 | 
|
196  | 
Usage: make [ARGS ...]  | 
|
197  | 
||
198  | 
Compile the logic in current directory using IsaMakefile.  | 
|
199  | 
ARGS are directly passed to the system make program.  | 
|
200  | 
\end{ttbox}
 | 
|
201  | 
||
202  | 
Note that the Isabelle settings environment is also active. Thus one  | 
|
203  | 
  may refer to its values within the @{verbatim IsaMakefile}, e.g.\
 | 
|
204  | 
  @{verbatim "$(ISABELLE_OUTPUT)"}. Furthermore, programs started from
 | 
|
205  | 
  the make file also inherit this environment.  Typically, @{verbatim
 | 
|
206  | 
  IsaMakefile}s defer the real work to the @{tool_ref usedir} utility.
 | 
|
207  | 
||
208  | 
  \medskip The basic @{verbatim IsaMakefile} convention is that the
 | 
|
209  | 
default target builds the actual logic, including its parents if  | 
|
210  | 
  appropriate.  The @{verbatim images} target is intended to build all
 | 
|
211  | 
  local logic images, while the @{verbatim test} target shall build
 | 
|
212  | 
  all related examples.  The @{verbatim all} target shall do
 | 
|
213  | 
  @{verbatim images} and @{verbatim test}.
 | 
|
214  | 
*}  | 
|
215  | 
||
216  | 
||
217  | 
subsubsection {* Examples *}
 | 
|
218  | 
||
219  | 
text {*
 | 
|
220  | 
  Refer to the @{verbatim IsaMakefile}s of the Isabelle distribution's
 | 
|
221  | 
object-logics as a model for your own developments. For example,  | 
|
| 
40800
 
330eb65c9469
Parse.liberal_name for document antiquotations and attributes;
 
wenzelm 
parents: 
32325 
diff
changeset
 | 
222  | 
  see @{file "~~/src/FOL/IsaMakefile"}.
 | 
| 28224 | 223  | 
*}  | 
224  | 
||
225  | 
||
226  | 
section {* Make all logics *}
 | 
|
227  | 
||
| 32325 | 228  | 
text {* The @{tool_def makeall} utility applies Isabelle make to any
 | 
229  | 
  Isabelle component (cf.\ \secref{sec:components}) that contains an
 | 
|
230  | 
  @{verbatim IsaMakefile}:
 | 
|
| 28224 | 231  | 
\begin{ttbox}
 | 
232  | 
Usage: makeall [ARGS ...]  | 
|
233  | 
||
| 32325 | 234  | 
Apply isabelle make to all components with IsaMakefile (passing ARGS).  | 
| 28224 | 235  | 
\end{ttbox}
 | 
236  | 
||
237  | 
  The arguments @{verbatim ARGS} are just passed verbatim to each
 | 
|
238  | 
  @{tool make} invocation.
 | 
|
239  | 
*}  | 
|
240  | 
||
241  | 
||
242  | 
section {* Printing documents *}
 | 
|
243  | 
||
244  | 
text {*
 | 
|
245  | 
  The @{tool_def print} utility prints documents:
 | 
|
246  | 
\begin{ttbox}
 | 
|
247  | 
Usage: print [OPTIONS] FILE  | 
|
248  | 
||
249  | 
Options are:  | 
|
250  | 
-c cleanup -- remove FILE after use  | 
|
251  | 
||
252  | 
Print document FILE.  | 
|
253  | 
\end{ttbox}
 | 
|
254  | 
||
255  | 
  The @{verbatim "-c"} option causes the input file to be removed
 | 
|
256  | 
  after use.  The printer spool command is determined by the @{setting
 | 
|
257  | 
PRINT_COMMAND} setting.  | 
|
258  | 
*}  | 
|
259  | 
||
260  | 
||
261  | 
section {* Remove awkward symbol names from theory sources *}
 | 
|
262  | 
||
263  | 
text {*
 | 
|
264  | 
  The @{tool_def unsymbolize} utility tunes Isabelle theory sources to
 | 
|
265  | 
improve readability for plain ASCII output (e.g.\ in email  | 
|
266  | 
  communication).  Most notably, @{tool unsymbolize} replaces awkward
 | 
|
267  | 
  arrow symbols such as @{verbatim "\\"}@{verbatim "<Longrightarrow>"}
 | 
|
268  | 
  by @{verbatim "==>"}.
 | 
|
269  | 
\begin{ttbox}
 | 
|
270  | 
Usage: unsymbolize [FILES|DIRS...]  | 
|
271  | 
||
272  | 
Recursively find .thy/.ML files, removing unreadable symbol names.  | 
|
273  | 
Note: this is an ad-hoc script; there is no systematic way to replace  | 
|
274  | 
symbols independently of the inner syntax of a theory!  | 
|
275  | 
||
276  | 
Renames old versions of FILES by appending "~~".  | 
|
277  | 
\end{ttbox}
 | 
|
278  | 
*}  | 
|
279  | 
||
280  | 
||
281  | 
section {* Output the version identifier of the Isabelle distribution *}
 | 
|
282  | 
||
283  | 
text {*
 | 
|
| 41511 | 284  | 
  The @{tool_def version} utility displays Isabelle version information:
 | 
285  | 
\begin{ttbox}
 | 
|
286  | 
Usage: isabelle version [OPTIONS]  | 
|
287  | 
||
288  | 
Options are:  | 
|
289  | 
-i short identification (derived from Mercurial id)  | 
|
290  | 
||
291  | 
Display Isabelle version information.  | 
|
292  | 
\end{ttbox}
 | 
|
293  | 
||
294  | 
\medskip The default is to output the full version string of the  | 
|
| 41512 | 295  | 
  Isabelle distribution, e.g.\ ``@{verbatim "Isabelle2011: January 2011"}.
 | 
| 41511 | 296  | 
|
297  | 
  The @{verbatim "-i"} option produces a short identification derived
 | 
|
298  | 
  from the Mercurial id of the @{setting ISABELLE_HOME} directory.
 | 
|
| 28224 | 299  | 
*}  | 
300  | 
||
301  | 
||
302  | 
section {* Convert XML to YXML *}
 | 
|
303  | 
||
304  | 
text {*
 | 
|
305  | 
  The @{tool_def yxml} tool converts a standard XML document (stdin)
 | 
|
306  | 
to the much simpler and more efficient YXML format of Isabelle  | 
|
307  | 
(stdout). The YXML format is defined as follows.  | 
|
308  | 
||
309  | 
  \begin{enumerate}
 | 
|
310  | 
||
311  | 
\item The encoding is always UTF-8.  | 
|
312  | 
||
313  | 
\item Body text is represented verbatim (no escaping, no special  | 
|
314  | 
treatment of white space, no named entities, no CDATA chunks, no  | 
|
315  | 
comments).  | 
|
316  | 
||
317  | 
\item Markup elements are represented via ASCII control characters  | 
|
318  | 
  @{text "\<^bold>X = 5"} and @{text "\<^bold>Y = 6"} as follows:
 | 
|
319  | 
||
320  | 
  \begin{tabular}{ll}
 | 
|
321  | 
XML & YXML \\\hline  | 
|
322  | 
    @{verbatim "<"}@{text "name attribute"}@{verbatim "="}@{text "value \<dots>"}@{verbatim ">"} &
 | 
|
323  | 
    @{text "\<^bold>X\<^bold>Yname\<^bold>Yattribute"}@{verbatim "="}@{text "value\<dots>\<^bold>X"} \\
 | 
|
324  | 
    @{verbatim "</"}@{text name}@{verbatim ">"} & @{text "\<^bold>X\<^bold>Y\<^bold>X"} \\
 | 
|
325  | 
  \end{tabular}
 | 
|
326  | 
||
327  | 
  There is no special case for empty body text, i.e.\ @{verbatim
 | 
|
328  | 
  "<foo/>"} is treated like @{verbatim "<foo></foo>"}.  Also note that
 | 
|
329  | 
  @{text "\<^bold>X"} and @{text "\<^bold>Y"} may never occur in
 | 
|
330  | 
well-formed XML documents.  | 
|
331  | 
||
332  | 
  \end{enumerate}
 | 
|
333  | 
||
334  | 
Parsing YXML is pretty straight-forward: split the text into chunks  | 
|
335  | 
  separated by @{text "\<^bold>X"}, then split each chunk into
 | 
|
336  | 
  sub-chunks separated by @{text "\<^bold>Y"}.  Markup chunks start
 | 
|
337  | 
with an empty sub-chunk, and a second empty sub-chunk indicates  | 
|
338  | 
close of an element. Any other non-empty chunk consists of plain  | 
|
| 
40800
 
330eb65c9469
Parse.liberal_name for document antiquotations and attributes;
 
wenzelm 
parents: 
32325 
diff
changeset
 | 
339  | 
  text.  For example, see @{file "~~/src/Pure/General/yxml.ML"} or
 | 
| 
 
330eb65c9469
Parse.liberal_name for document antiquotations and attributes;
 
wenzelm 
parents: 
32325 
diff
changeset
 | 
340  | 
  @{file "~~/src/Pure/General/yxml.scala"}.
 | 
| 28224 | 341  | 
|
342  | 
YXML documents may be detected quickly by checking that the first  | 
|
343  | 
  two characters are @{text "\<^bold>X\<^bold>Y"}.
 | 
|
344  | 
*}  | 
|
345  | 
||
346  | 
end  |