author | wenzelm |
Mon, 10 Aug 2015 11:20:16 +0200 | |
changeset 60871 | 9b26f3118e40 |
parent 58618 | 782f0b662cae |
child 61406 | eb2463fc4d7b |
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 |
||
58618 | 5 |
chapter \<open>Miscellaneous tools \label{ch:tools}\<close> |
28224 | 6 |
|
58618 | 7 |
text \<open> |
28224 | 8 |
Subsequently we describe various Isabelle related utilities, given |
9 |
in alphabetical order. |
|
58618 | 10 |
\<close> |
28224 | 11 |
|
12 |
||
58618 | 13 |
section \<open>Theory graph browser \label{sec:browse}\<close> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
14 |
|
58618 | 15 |
text \<open>The Isabelle graph browser is a general tool for visualizing |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
16 |
dependency graphs. Certain nodes of the graph (i.e.\ theories) can |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
17 |
be grouped together in ``directories'', whose contents may be |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
18 |
hidden, thus enabling the user to collapse irrelevant portions of |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
19 |
information. The browser is written in Java, it can be used both as |
58618 | 20 |
a stand-alone application and as an applet.\<close> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
21 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
22 |
|
58618 | 23 |
subsection \<open>Invoking the graph browser\<close> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
24 |
|
58618 | 25 |
text \<open>The stand-alone version of the graph browser is wrapped up as |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
26 |
@{tool_def browser}: |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
27 |
\begin{ttbox} |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
28 |
Usage: isabelle browser [OPTIONS] [GRAPHFILE] |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
29 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
30 |
Options are: |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
31 |
-b Admin/build only |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
32 |
-c cleanup -- remove GRAPHFILE after use |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
33 |
-o FILE output to FILE (ps, eps, pdf) |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
34 |
\end{ttbox} |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
35 |
When no file name is specified, the browser automatically changes to |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
36 |
the directory @{setting ISABELLE_BROWSER_INFO}. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
37 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
38 |
\medskip The @{verbatim "-b"} option indicates that this is for |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
39 |
administrative build only, i.e.\ no browser popup if no files are |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
40 |
given. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
41 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
42 |
The @{verbatim "-c"} option causes the input file to be removed |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
43 |
after use. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
44 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
45 |
The @{verbatim "-o"} option indicates batch-mode operation, with the |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
46 |
output written to the indicated file; note that @{verbatim pdf} |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
47 |
produces an @{verbatim eps} copy as well. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
48 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
49 |
\medskip The applet version of the browser is part of the standard |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
50 |
WWW theory presentation, see the link ``theory dependencies'' within |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
51 |
each session index. |
58618 | 52 |
\<close> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
53 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
54 |
|
58618 | 55 |
subsection \<open>Using the graph browser\<close> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
56 |
|
58618 | 57 |
text \<open> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
58 |
The browser's main window, which is shown in |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
59 |
\figref{fig:browserwindow}, consists of two sub-windows. In the |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
60 |
left sub-window, the directory tree is displayed. The graph itself |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
61 |
is displayed in the right sub-window. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
62 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
63 |
\begin{figure}[ht] |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
64 |
\includegraphics[width=\textwidth]{browser_screenshot} |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
65 |
\caption{\label{fig:browserwindow} Browser main window} |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
66 |
\end{figure} |
58618 | 67 |
\<close> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
68 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
69 |
|
58618 | 70 |
subsubsection \<open>The directory tree window\<close> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
71 |
|
58618 | 72 |
text \<open> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
73 |
We describe the usage of the directory browser and the meaning of |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
74 |
the different items in the browser window. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
75 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
76 |
\begin{itemize} |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
77 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
78 |
\item A red arrow before a directory name indicates that the |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
79 |
directory is currently ``folded'', i.e.~the nodes in this directory |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
80 |
are collapsed to one single node. In the right sub-window, the names |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
81 |
of nodes corresponding to folded directories are enclosed in square |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
82 |
brackets and displayed in red color. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
83 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
84 |
\item A green downward arrow before a directory name indicates that |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
85 |
the directory is currently ``unfolded''. It can be folded by |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
86 |
clicking on the directory name. Clicking on the name for a second |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
87 |
time unfolds the directory again. Alternatively, a directory can |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
88 |
also be unfolded by clicking on the corresponding node in the right |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
89 |
sub-window. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
90 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
91 |
\item Blue arrows stand before ordinary node names. When clicking on |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
92 |
such a name (i.e.\ that of a theory), the graph display window |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
93 |
focuses to the corresponding node. Double clicking invokes a text |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
94 |
viewer window in which the contents of the theory file are |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
95 |
displayed. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
96 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
97 |
\end{itemize} |
58618 | 98 |
\<close> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
99 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
100 |
|
58618 | 101 |
subsubsection \<open>The graph display window\<close> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
102 |
|
58618 | 103 |
text \<open> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
104 |
When pointing on an ordinary node, an upward and a downward arrow is |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
105 |
shown. Initially, both of these arrows are green. Clicking on the |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
106 |
upward or downward arrow collapses all predecessor or successor |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
107 |
nodes, respectively. The arrow's color then changes to red, |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
108 |
indicating that the predecessor or successor nodes are currently |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
109 |
collapsed. The node corresponding to the collapsed nodes has the |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
110 |
name ``@{verbatim "[....]"}''. To uncollapse the nodes again, simply |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
111 |
click on the red arrow or on the node with the name ``@{verbatim |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
112 |
"[....]"}''. Similar to the directory browser, the contents of |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
113 |
theory files can be displayed by double clicking on the |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
114 |
corresponding node. |
58618 | 115 |
\<close> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
116 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
117 |
|
58618 | 118 |
subsubsection \<open>The ``File'' menu\<close> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
119 |
|
58618 | 120 |
text \<open> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
121 |
Due to Java Applet security restrictions this menu is only available |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
122 |
in the full application version. The meaning of the menu items is as |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
123 |
follows: |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
124 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
125 |
\begin{description} |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
126 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
127 |
\item[Open \dots] Open a new graph file. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
128 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
129 |
\item[Export to PostScript] Outputs the current graph in Postscript |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
130 |
format, appropriately scaled to fit on one single sheet of A4 paper. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
131 |
The resulting file can be printed directly. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
132 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
133 |
\item[Export to EPS] Outputs the current graph in Encapsulated |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
134 |
Postscript format. The resulting file can be included in other |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
135 |
documents. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
136 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
137 |
\item[Quit] Quit the graph browser. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
138 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
139 |
\end{description} |
58618 | 140 |
\<close> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
141 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
142 |
|
58618 | 143 |
subsection \<open>Syntax of graph definition files\<close> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
144 |
|
58618 | 145 |
text \<open> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
146 |
A graph definition file has the following syntax: |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
147 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
148 |
\begin{center}\small |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
149 |
\begin{tabular}{rcl} |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
150 |
@{text graph} & @{text "="} & @{text "{ vertex"}~@{verbatim ";"}~@{text "}+"} \\ |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
151 |
@{text vertex} & @{text "="} & @{text "vertex_name vertex_ID dir_name ["}~@{verbatim "+"}~@{text "] path ["}~@{verbatim "<"}~@{text "|"}~@{verbatim ">"}~@{text "] { vertex_ID }*"} |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
152 |
\end{tabular} |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
153 |
\end{center} |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
154 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
155 |
The meaning of the items in a vertex description is as follows: |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
156 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
157 |
\begin{description} |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
158 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
159 |
\item[@{text vertex_name}] The name of the vertex. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
160 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
161 |
\item[@{text vertex_ID}] The vertex identifier. Note that there may |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
162 |
be several vertices with equal names, whereas identifiers must be |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
163 |
unique. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
164 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
165 |
\item[@{text dir_name}] The name of the ``directory'' the vertex |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
166 |
should be placed in. A ``@{verbatim "+"}'' sign after @{text |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
167 |
dir_name} indicates that the nodes in the directory are initially |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
168 |
visible. Directories are initially invisible by default. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
169 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
170 |
\item[@{text path}] The path of the corresponding theory file. This |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
171 |
is specified relatively to the path of the graph definition file. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
172 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
173 |
\item[List of successor/predecessor nodes] A ``@{verbatim "<"}'' |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
174 |
sign before the list means that successor nodes are listed, a |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
175 |
``@{verbatim ">"}'' sign means that predecessor nodes are listed. If |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
176 |
neither ``@{verbatim "<"}'' nor ``@{verbatim ">"}'' is found, the |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
177 |
browser assumes that successor nodes are listed. |
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
178 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
179 |
\end{description} |
58618 | 180 |
\<close> |
57330
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
181 |
|
d8a64a4cbfca
clarified role of old user interfaces as misc tools;
wenzelm
parents:
56439
diff
changeset
|
182 |
|
58618 | 183 |
section \<open>Resolving Isabelle components \label{sec:tool-components}\<close> |
48844 | 184 |
|
58618 | 185 |
text \<open> |
48844 | 186 |
The @{tool_def components} tool resolves Isabelle components: |
187 |
\begin{ttbox} |
|
188 |
Usage: isabelle components [OPTIONS] [COMPONENTS ...] |
|
189 |
||
190 |
Options are: |
|
50653
5c85f8b80b95
simplified quick start via "isabelle components -I";
wenzelm
parents:
50132
diff
changeset
|
191 |
-I init user settings |
48844 | 192 |
-R URL component repository |
193 |
(default $ISABELLE_COMPONENT_REPOSITORY) |
|
53435 | 194 |
-a resolve all missing components |
48844 | 195 |
-l list status |
196 |
||
197 |
Resolve Isabelle components via download and installation. |
|
198 |
COMPONENTS are identified via base name. |
|
199 |
||
200 |
ISABELLE_COMPONENT_REPOSITORY="http://isabelle.in.tum.de/components" |
|
201 |
\end{ttbox} |
|
202 |
||
203 |
Components are initialized as described in \secref{sec:components} |
|
204 |
in a permissive manner, which can mark components as ``missing''. |
|
205 |
This state is amended by letting @{tool "components"} download and |
|
206 |
unpack components that are published on the default component |
|
54703 | 207 |
repository @{url "http://isabelle.in.tum.de/components/"} in |
48844 | 208 |
particular. |
209 |
||
210 |
Option @{verbatim "-R"} specifies an alternative component |
|
211 |
repository. Note that @{verbatim "file:///"} URLs can be used for |
|
212 |
local directories. |
|
213 |
||
214 |
Option @{verbatim "-a"} selects all missing components to be |
|
53435 | 215 |
resolved. Explicit components may be named as command |
48844 | 216 |
line-arguments as well. Note that components are uniquely |
217 |
identified by their base name, while the installation takes place in |
|
218 |
the location that was specified in the attempt to initialize the |
|
219 |
component before. |
|
220 |
||
221 |
Option @{verbatim "-l"} lists the current state of available and |
|
222 |
missing components with their location (full name) within the |
|
50653
5c85f8b80b95
simplified quick start via "isabelle components -I";
wenzelm
parents:
50132
diff
changeset
|
223 |
file-system. |
5c85f8b80b95
simplified quick start via "isabelle components -I";
wenzelm
parents:
50132
diff
changeset
|
224 |
|
5c85f8b80b95
simplified quick start via "isabelle components -I";
wenzelm
parents:
50132
diff
changeset
|
225 |
Option @{verbatim "-I"} initializes the user settings file to |
5c85f8b80b95
simplified quick start via "isabelle components -I";
wenzelm
parents:
50132
diff
changeset
|
226 |
subscribe to the standard components specified in the Isabelle |
5c85f8b80b95
simplified quick start via "isabelle components -I";
wenzelm
parents:
50132
diff
changeset
|
227 |
repository clone --- this does not make any sense for regular |
5c85f8b80b95
simplified quick start via "isabelle components -I";
wenzelm
parents:
50132
diff
changeset
|
228 |
Isabelle releases. If the file already exists, it needs to be |
5c85f8b80b95
simplified quick start via "isabelle components -I";
wenzelm
parents:
50132
diff
changeset
|
229 |
edited manually according to the printed explanation. |
58618 | 230 |
\<close> |
48844 | 231 |
|
232 |
||
58618 | 233 |
section \<open>Raw ML console\<close> |
57439
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
234 |
|
58618 | 235 |
text \<open> |
57439
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
236 |
The @{tool_def console} tool runs the Isabelle process with raw ML console: |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
237 |
\begin{ttbox} |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
238 |
Usage: isabelle console [OPTIONS] |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
239 |
|
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
240 |
Options are: |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
241 |
-d DIR include session directory |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
242 |
-l NAME logic session name (default ISABELLE_LOGIC) |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
243 |
-m MODE add print mode for output |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
244 |
-n no build of session image on startup |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
245 |
-o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
246 |
-s system build mode for session image |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
247 |
|
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
248 |
Run Isabelle process with raw ML console and line editor |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
249 |
(default ISABELLE_LINE_EDITOR). |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
250 |
\end{ttbox} |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
251 |
|
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
252 |
The @{verbatim "-l"} option specifies the logic session name. By default, |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
253 |
its heap image is checked and built on demand, but the option @{verbatim |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
254 |
"-n"} skips that. |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
255 |
|
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
256 |
Options @{verbatim "-d"}, @{verbatim "-o"}, @{verbatim "-s"} are passed |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
257 |
directly to @{tool build} (\secref{sec:tool-build}). |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
258 |
|
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
259 |
Options @{verbatim "-m"}, @{verbatim "-o"} are passed directly to the |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
260 |
underlying Isabelle process (\secref{sec:isabelle-process}). |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
261 |
|
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
262 |
The Isabelle process is run through the line editor that is specified via |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
263 |
the settings variable @{setting ISABELLE_LINE_EDITOR} (e.g.\ |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
264 |
@{executable_def rlwrap} for GNU readline); the fall-back is to use plain |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
265 |
standard input/output. |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
266 |
|
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
267 |
Interaction works via the raw ML toplevel loop: this is neither |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
268 |
Isabelle/Isar nor Isabelle/ML within the usual formal context. Some useful |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
269 |
ML commands at this stage are @{ML cd}, @{ML pwd}, @{ML use}, @{ML use_thy}, |
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
270 |
@{ML use_thys}. |
58618 | 271 |
\<close> |
57439
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
272 |
|
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
57414
diff
changeset
|
273 |
|
58618 | 274 |
section \<open>Displaying documents \label{sec:tool-display}\<close> |
28224 | 275 |
|
58618 | 276 |
text \<open>The @{tool_def display} tool displays documents in DVI or PDF |
28224 | 277 |
format: |
278 |
\begin{ttbox} |
|
54683
cf48ddc266e5
clarified "isabelle display" and 'display_drafts': re-use file and program instance, open asynchronously via desktop environment;
wenzelm
parents:
53435
diff
changeset
|
279 |
Usage: isabelle display DOCUMENT |
28224 | 280 |
|
54683
cf48ddc266e5
clarified "isabelle display" and 'display_drafts': re-use file and program instance, open asynchronously via desktop environment;
wenzelm
parents:
53435
diff
changeset
|
281 |
Display DOCUMENT (in DVI or PDF format). |
28224 | 282 |
\end{ttbox} |
283 |
||
52444
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
284 |
\medskip The settings @{setting DVI_VIEWER} and @{setting |
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
285 |
PDF_VIEWER} determine the programs for viewing the corresponding |
54683
cf48ddc266e5
clarified "isabelle display" and 'display_drafts': re-use file and program instance, open asynchronously via desktop environment;
wenzelm
parents:
53435
diff
changeset
|
286 |
file formats. Normally this opens the document via the desktop |
cf48ddc266e5
clarified "isabelle display" and 'display_drafts': re-use file and program instance, open asynchronously via desktop environment;
wenzelm
parents:
53435
diff
changeset
|
287 |
environment, potentially in an asynchronous manner with re-use of |
cf48ddc266e5
clarified "isabelle display" and 'display_drafts': re-use file and program instance, open asynchronously via desktop environment;
wenzelm
parents:
53435
diff
changeset
|
288 |
previews views. |
58618 | 289 |
\<close> |
28224 | 290 |
|
291 |
||
58618 | 292 |
section \<open>Viewing documentation \label{sec:tool-doc}\<close> |
28224 | 293 |
|
58618 | 294 |
text \<open> |
52444
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
295 |
The @{tool_def doc} tool displays Isabelle documentation: |
28224 | 296 |
\begin{ttbox} |
52444
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
297 |
Usage: isabelle doc [DOC ...] |
28224 | 298 |
|
52444
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
299 |
View Isabelle documentation. |
28224 | 300 |
\end{ttbox} |
301 |
If called without arguments, it lists all available documents. Each |
|
302 |
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
|
303 |
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
|
304 |
display the corresponding document (see also |
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
wenzelm
parents:
52052
diff
changeset
|
305 |
\secref{sec:tool-display}). |
28224 | 306 |
|
307 |
\medskip The @{setting ISABELLE_DOCS} setting specifies the list of |
|
308 |
directories (separated by colons) to be scanned for documentations. |
|
58618 | 309 |
\<close> |
28224 | 310 |
|
311 |
||
58618 | 312 |
section \<open>Shell commands within the settings environment \label{sec:tool-env}\<close> |
47828 | 313 |
|
58618 | 314 |
text \<open>The @{tool_def env} tool is a direct wrapper for the standard |
48602 | 315 |
@{verbatim "/usr/bin/env"} command on POSIX systems, running within |
316 |
the Isabelle settings environment (\secref{sec:settings}). |
|
47828 | 317 |
|
318 |
The command-line arguments are that of the underlying version of |
|
319 |
@{verbatim env}. For example, the following invokes an instance of |
|
320 |
the GNU Bash shell within the Isabelle environment: |
|
321 |
\begin{alltt} |
|
322 |
isabelle env bash |
|
323 |
\end{alltt} |
|
58618 | 324 |
\<close> |
47828 | 325 |
|
326 |
||
58618 | 327 |
section \<open>Inspecting the settings environment \label{sec:tool-getenv}\<close> |
28224 | 328 |
|
58618 | 329 |
text \<open>The Isabelle settings environment --- as provided by the |
28224 | 330 |
site-default and user-specific settings files --- can be inspected |
48602 | 331 |
with the @{tool_def getenv} tool: |
28224 | 332 |
\begin{ttbox} |
48602 | 333 |
Usage: isabelle getenv [OPTIONS] [VARNAMES ...] |
28224 | 334 |
|
335 |
Options are: |
|
336 |
-a display complete environment |
|
337 |
-b print values only (doesn't work for -a) |
|
31497 | 338 |
-d FILE dump complete environment to FILE |
339 |
(null terminated entries) |
|
28224 | 340 |
|
341 |
Get value of VARNAMES from the Isabelle settings. |
|
342 |
\end{ttbox} |
|
343 |
||
344 |
With the @{verbatim "-a"} option, one may inspect the full process |
|
345 |
environment that Isabelle related programs are run in. This usually |
|
346 |
contains much more variables than are actually Isabelle settings. |
|
347 |
Normally, output is a list of lines of the form @{text |
|
348 |
name}@{verbatim "="}@{text value}. The @{verbatim "-b"} option |
|
349 |
causes only the values to be printed. |
|
31497 | 350 |
|
351 |
Option @{verbatim "-d"} produces a dump of the complete environment |
|
352 |
to the specified file. Entries are terminated by the ASCII null |
|
353 |
character, i.e.\ the C string terminator. |
|
58618 | 354 |
\<close> |
28224 | 355 |
|
356 |
||
58618 | 357 |
subsubsection \<open>Examples\<close> |
28224 | 358 |
|
58618 | 359 |
text \<open>Get the location of @{setting ISABELLE_HOME_USER} where |
48815 | 360 |
user-specific information is stored: |
28224 | 361 |
\begin{ttbox} |
48815 | 362 |
isabelle getenv ISABELLE_HOME_USER |
28224 | 363 |
\end{ttbox} |
364 |
||
48815 | 365 |
\medskip Get the value only of the same settings variable, which is |
366 |
particularly useful in shell scripts: |
|
28224 | 367 |
\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
|
368 |
isabelle getenv -b ISABELLE_OUTPUT |
28224 | 369 |
\end{ttbox} |
58618 | 370 |
\<close> |
28224 | 371 |
|
372 |
||
58618 | 373 |
section \<open>Installing standalone Isabelle executables \label{sec:tool-install}\<close> |
28224 | 374 |
|
58618 | 375 |
text \<open>By default, the main Isabelle binaries (@{executable |
48602 | 376 |
"isabelle"} etc.) are just run from their location within the |
377 |
distribution directory, probably indirectly by the shell through its |
|
378 |
@{setting PATH}. Other schemes of installation are supported by the |
|
379 |
@{tool_def install} tool: |
|
28224 | 380 |
\begin{ttbox} |
50132 | 381 |
Usage: isabelle install [OPTIONS] BINDIR |
28224 | 382 |
|
383 |
Options are: |
|
50132 | 384 |
-d DISTDIR refer to DISTDIR as Isabelle distribution |
28224 | 385 |
(default ISABELLE_HOME) |
386 |
||
50132 | 387 |
Install Isabelle executables with absolute references to the |
28224 | 388 |
distribution directory. |
389 |
\end{ttbox} |
|
390 |
||
391 |
The @{verbatim "-d"} option overrides the current Isabelle |
|
392 |
distribution directory as determined by @{setting ISABELLE_HOME}. |
|
393 |
||
50132 | 394 |
The @{text BINDIR} argument tells where executable wrapper scripts |
56439
95e2656b3b23
renamed "isabelle-process" to "isabelle_process", with shell function to avoid dynamic path lookups;
wenzelm
parents:
54703
diff
changeset
|
395 |
for @{executable "isabelle_process"} and @{executable isabelle} |
50132 | 396 |
should be placed, which is typically a directory in the shell's |
397 |
@{setting PATH}, such as @{verbatim "$HOME/bin"}. |
|
48815 | 398 |
|
50132 | 399 |
\medskip It is also possible to make symbolic links of the main |
400 |
Isabelle executables manually, but making separate copies outside |
|
58618 | 401 |
the Isabelle distribution directory will not work!\<close> |
28224 | 402 |
|
403 |
||
58618 | 404 |
section \<open>Creating instances of the Isabelle logo\<close> |
28224 | 405 |
|
58618 | 406 |
text \<open>The @{tool_def logo} tool creates instances of the generic |
49072
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
407 |
Isabelle logo as EPS and PDF, for inclusion in {\LaTeX} documents. |
28224 | 408 |
\begin{ttbox} |
49072
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
409 |
Usage: isabelle logo [OPTIONS] XYZ |
28224 | 410 |
|
49072
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
411 |
Create instance XYZ of the Isabelle logo (as EPS and PDF). |
28224 | 412 |
|
413 |
Options are: |
|
49072
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
414 |
-n NAME alternative output base name (default "isabelle_xyx") |
28224 | 415 |
-q quiet mode |
416 |
\end{ttbox} |
|
48936 | 417 |
|
49072
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
418 |
Option @{verbatim "-n"} specifies an altenative (base) name for the |
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
419 |
generated files. The default is @{verbatim "isabelle_"}@{text xyz} |
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
wenzelm
parents:
48985
diff
changeset
|
420 |
in lower-case. |
48936 | 421 |
|
422 |
Option @{verbatim "-q"} omits printing of the result file name. |
|
423 |
||
424 |
\medskip Implementors of Isabelle tools and applications are |
|
425 |
encouraged to make derived Isabelle logos for their own projects |
|
58618 | 426 |
using this template.\<close> |
28224 | 427 |
|
428 |
||
58618 | 429 |
section \<open>Output the version identifier of the Isabelle distribution\<close> |
28224 | 430 |
|
58618 | 431 |
text \<open> |
48602 | 432 |
The @{tool_def version} tool displays Isabelle version information: |
41511 | 433 |
\begin{ttbox} |
434 |
Usage: isabelle version [OPTIONS] |
|
435 |
||
436 |
Options are: |
|
437 |
-i short identification (derived from Mercurial id) |
|
438 |
||
439 |
Display Isabelle version information. |
|
440 |
\end{ttbox} |
|
441 |
||
442 |
\medskip The default is to output the full version string of the |
|
47827 | 443 |
Isabelle distribution, e.g.\ ``@{verbatim "Isabelle2012: May 2012"}. |
41511 | 444 |
|
445 |
The @{verbatim "-i"} option produces a short identification derived |
|
446 |
from the Mercurial id of the @{setting ISABELLE_HOME} directory. |
|
58618 | 447 |
\<close> |
28224 | 448 |
|
449 |
||
58618 | 450 |
section \<open>Convert XML to YXML\<close> |
28224 | 451 |
|
58618 | 452 |
text \<open> |
28224 | 453 |
The @{tool_def yxml} tool converts a standard XML document (stdin) |
454 |
to the much simpler and more efficient YXML format of Isabelle |
|
455 |
(stdout). The YXML format is defined as follows. |
|
456 |
||
457 |
\begin{enumerate} |
|
458 |
||
459 |
\item The encoding is always UTF-8. |
|
460 |
||
461 |
\item Body text is represented verbatim (no escaping, no special |
|
462 |
treatment of white space, no named entities, no CDATA chunks, no |
|
463 |
comments). |
|
464 |
||
465 |
\item Markup elements are represented via ASCII control characters |
|
466 |
@{text "\<^bold>X = 5"} and @{text "\<^bold>Y = 6"} as follows: |
|
467 |
||
468 |
\begin{tabular}{ll} |
|
469 |
XML & YXML \\\hline |
|
470 |
@{verbatim "<"}@{text "name attribute"}@{verbatim "="}@{text "value \<dots>"}@{verbatim ">"} & |
|
471 |
@{text "\<^bold>X\<^bold>Yname\<^bold>Yattribute"}@{verbatim "="}@{text "value\<dots>\<^bold>X"} \\ |
|
472 |
@{verbatim "</"}@{text name}@{verbatim ">"} & @{text "\<^bold>X\<^bold>Y\<^bold>X"} \\ |
|
473 |
\end{tabular} |
|
474 |
||
475 |
There is no special case for empty body text, i.e.\ @{verbatim |
|
476 |
"<foo/>"} is treated like @{verbatim "<foo></foo>"}. Also note that |
|
477 |
@{text "\<^bold>X"} and @{text "\<^bold>Y"} may never occur in |
|
478 |
well-formed XML documents. |
|
479 |
||
480 |
\end{enumerate} |
|
481 |
||
482 |
Parsing YXML is pretty straight-forward: split the text into chunks |
|
483 |
separated by @{text "\<^bold>X"}, then split each chunk into |
|
484 |
sub-chunks separated by @{text "\<^bold>Y"}. Markup chunks start |
|
485 |
with an empty sub-chunk, and a second empty sub-chunk indicates |
|
486 |
close of an element. Any other non-empty chunk consists of plain |
|
44799 | 487 |
text. For example, see @{file "~~/src/Pure/PIDE/yxml.ML"} or |
488 |
@{file "~~/src/Pure/PIDE/yxml.scala"}. |
|
28224 | 489 |
|
490 |
YXML documents may be detected quickly by checking that the first |
|
491 |
two characters are @{text "\<^bold>X\<^bold>Y"}. |
|
58618 | 492 |
\<close> |
28224 | 493 |
|
494 |
end |