author | wenzelm |
Wed, 04 Nov 2015 18:32:47 +0100 | |
changeset 61572 | ddb3ac3fef45 |
parent 61503 | 28e788ca2c5d |
child 61575 | f18f6e51e901 |
permissions | -rw-r--r-- |
28215 | 1 |
theory Basics |
43564
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
41955
diff
changeset
|
2 |
imports Base |
28215 | 3 |
begin |
4 |
||
58618 | 5 |
chapter \<open>The Isabelle system environment\<close> |
28215 | 6 |
|
58618 | 7 |
text \<open>This manual describes Isabelle together with related tools and |
47823 | 8 |
user interfaces as seen from a system oriented view. See also the |
61477 | 9 |
\<^emph>\<open>Isabelle/Isar Reference Manual\<close> @{cite "isabelle-isar-ref"} for |
47823 | 10 |
the actual Isabelle input language and related concepts, and |
61477 | 11 |
\<^emph>\<open>The Isabelle/Isar Implementation |
12 |
Manual\<close> @{cite "isabelle-implementation"} for the main concepts of the |
|
47823 | 13 |
underlying implementation in Isabelle/ML. |
28215 | 14 |
|
61406 | 15 |
\<^medskip> |
16 |
The Isabelle system environment provides the following |
|
28916
0a802cdda340
removed obsolete isabelle-interface executable and ISABELLE_INTERFACE setting;
wenzelm
parents:
28914
diff
changeset
|
17 |
basic infrastructure to integrate tools smoothly. |
28215 | 18 |
|
61477 | 19 |
\<^enum> The \<^emph>\<open>Isabelle settings\<close> mechanism provides process |
28916
0a802cdda340
removed obsolete isabelle-interface executable and ISABELLE_INTERFACE setting;
wenzelm
parents:
28914
diff
changeset
|
20 |
environment variables to all Isabelle executables (including tools |
0a802cdda340
removed obsolete isabelle-interface executable and ISABELLE_INTERFACE setting;
wenzelm
parents:
28914
diff
changeset
|
21 |
and user interfaces). |
28215 | 22 |
|
61477 | 23 |
\<^enum> The raw \<^emph>\<open>Isabelle process\<close> (@{executable_ref |
56439
95e2656b3b23
renamed "isabelle-process" to "isabelle_process", with shell function to avoid dynamic path lookups;
wenzelm
parents:
54937
diff
changeset
|
24 |
"isabelle_process"}) runs logic sessions either interactively or in |
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28502
diff
changeset
|
25 |
batch mode. In particular, this view abstracts over the invocation |
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28502
diff
changeset
|
26 |
of the actual ML system to be used. Regular users rarely need to |
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28502
diff
changeset
|
27 |
care about the low-level process. |
28215 | 28 |
|
61477 | 29 |
\<^enum> The main \<^emph>\<open>Isabelle tool wrapper\<close> (@{executable_ref |
47823 | 30 |
isabelle}) provides a generic startup environment Isabelle related |
31 |
utilities, user interfaces etc. Such tools automatically benefit |
|
32 |
from the settings mechanism. |
|
58618 | 33 |
\<close> |
28215 | 34 |
|
35 |
||
58618 | 36 |
section \<open>Isabelle settings \label{sec:settings}\<close> |
28215 | 37 |
|
58618 | 38 |
text \<open> |
61477 | 39 |
The Isabelle system heavily depends on the \<^emph>\<open>settings |
40 |
mechanism\<close>\indexbold{settings}. Essentially, this is a statically |
|
28215 | 41 |
scoped collection of environment variables, such as @{setting |
42 |
ISABELLE_HOME}, @{setting ML_SYSTEM}, @{setting ML_HOME}. These |
|
61477 | 43 |
variables are \<^emph>\<open>not\<close> intended to be set directly from the shell, |
28215 | 44 |
though. Isabelle employs a somewhat more sophisticated scheme of |
61477 | 45 |
\<^emph>\<open>settings files\<close> --- one for site-wide defaults, another for |
28215 | 46 |
additional user-specific modifications. With all configuration |
47823 | 47 |
variables in clearly defined places, this scheme is more |
48 |
maintainable and user-friendly than global shell environment |
|
49 |
variables. |
|
28215 | 50 |
|
51 |
In particular, we avoid the typical situation where prospective |
|
52 |
users of a software package are told to put several things into |
|
53 |
their shell startup scripts, before being able to actually run the |
|
54 |
program. Isabelle requires none such administrative chores of its |
|
55 |
end-users --- the executables can be invoked straight away. |
|
40800
330eb65c9469
Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents:
40569
diff
changeset
|
56 |
Occasionally, users would still want to put the @{file |
28238 | 57 |
"$ISABELLE_HOME/bin"} directory into their shell's search path, but |
58 |
this is not required. |
|
58618 | 59 |
\<close> |
28215 | 60 |
|
61 |
||
58618 | 62 |
subsection \<open>Bootstrapping the environment \label{sec:boot}\<close> |
28215 | 63 |
|
58618 | 64 |
text \<open>Isabelle executables need to be run within a proper settings |
32323 | 65 |
environment. This is bootstrapped as described below, on the first |
66 |
invocation of one of the outer wrapper scripts (such as |
|
67 |
@{executable_ref isabelle}). This happens only once for each |
|
68 |
process tree, i.e.\ the environment is passed to subprocesses |
|
69 |
according to regular Unix conventions. |
|
28215 | 70 |
|
61406 | 71 |
\<^enum> The special variable @{setting_def ISABELLE_HOME} is |
28215 | 72 |
determined automatically from the location of the binary that has |
73 |
been run. |
|
74 |
||
75 |
You should not try to set @{setting ISABELLE_HOME} manually. Also |
|
76 |
note that the Isabelle executables either have to be run from their |
|
77 |
original location in the distribution directory, or via the |
|
48602 | 78 |
executable objects created by the @{tool install} tool. Symbolic |
40800
330eb65c9469
Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents:
40569
diff
changeset
|
79 |
links are admissible, but a plain copy of the @{file |
28238 | 80 |
"$ISABELLE_HOME/bin"} files will not work! |
81 |
||
61406 | 82 |
\<^enum> The file @{file "$ISABELLE_HOME/etc/settings"} is run as a |
28238 | 83 |
@{executable_ref bash} shell script with the auto-export option for |
84 |
variables enabled. |
|
28215 | 85 |
|
86 |
This file holds a rather long list of shell variable assigments, |
|
87 |
thus providing the site-wide default settings. The Isabelle |
|
88 |
distribution already contains a global settings file with sensible |
|
89 |
defaults for most variables. When installing the system, only a few |
|
90 |
of these may have to be adapted (probably @{setting ML_SYSTEM} |
|
91 |
etc.). |
|
92 |
||
61406 | 93 |
\<^enum> The file @{file_unchecked "$ISABELLE_HOME_USER/etc/settings"} (if it |
28215 | 94 |
exists) is run in the same way as the site default settings. Note |
95 |
that the variable @{setting ISABELLE_HOME_USER} has already been set |
|
61503 | 96 |
before --- usually to something like \<^verbatim>\<open>$USER_HOME/.isabelle/IsabelleXXXX\<close>. |
28215 | 97 |
|
54937 | 98 |
Thus individual users may override the site-wide defaults. |
99 |
Typically, a user settings file contains only a few lines, with some |
|
100 |
assignments that are actually changed. Never copy the central |
|
101 |
@{file "$ISABELLE_HOME/etc/settings"} file! |
|
61458 | 102 |
|
28215 | 103 |
|
28238 | 104 |
Since settings files are regular GNU @{executable_def bash} scripts, |
61503 | 105 |
one may use complex shell commands, such as \<^verbatim>\<open>if\<close> or |
106 |
\<^verbatim>\<open>case\<close> statements to set variables depending on the |
|
28215 | 107 |
system architecture or other environment variables. Such advanced |
108 |
features should be added only with great care, though. In |
|
109 |
particular, external environment references should be kept at a |
|
110 |
minimum. |
|
111 |
||
61406 | 112 |
\<^medskip> |
113 |
A few variables are somewhat special: |
|
28215 | 114 |
|
61406 | 115 |
\<^item> @{setting_def ISABELLE_PROCESS} and @{setting_def ISABELLE_TOOL} are set |
28215 | 116 |
automatically to the absolute path names of the @{executable |
56439
95e2656b3b23
renamed "isabelle-process" to "isabelle_process", with shell function to avoid dynamic path lookups;
wenzelm
parents:
54937
diff
changeset
|
117 |
"isabelle_process"} and @{executable isabelle} executables, |
28215 | 118 |
respectively. |
119 |
||
61406 | 120 |
\<^item> @{setting_ref ISABELLE_OUTPUT} will have the identifiers of |
28215 | 121 |
the Isabelle distribution (cf.\ @{setting ISABELLE_IDENTIFIER}) and |
122 |
the ML system (cf.\ @{setting ML_IDENTIFIER}) appended automatically |
|
123 |
to its value. |
|
124 |
||
125 |
||
61406 | 126 |
\<^medskip> |
127 |
Note that the settings environment may be inspected with |
|
48602 | 128 |
the @{tool getenv} tool. This might help to figure out the effect |
58618 | 129 |
of complex settings scripts.\<close> |
28215 | 130 |
|
131 |
||
58618 | 132 |
subsection \<open>Common variables\<close> |
28215 | 133 |
|
58618 | 134 |
text \<open> |
28215 | 135 |
This is a reference of common Isabelle settings variables. Note that |
136 |
the list is somewhat open-ended. Third-party utilities or interfaces |
|
137 |
may add their own selection. Variables that are special in some |
|
61493 | 138 |
sense are marked with \<open>\<^sup>*\<close>. |
28215 | 139 |
|
61493 | 140 |
\<^descr>[@{setting_def USER_HOME}\<open>\<^sup>*\<close>] Is the cross-platform |
47823 | 141 |
user home directory. On Unix systems this is usually the same as |
142 |
@{setting HOME}, but on Windows it is the regular home directory of |
|
143 |
the user, not the one of within the Cygwin root |
|
61572 | 144 |
file-system.\<^footnote>\<open>Cygwin itself offers another choice whether |
61412 | 145 |
its HOME should point to the @{file_unchecked "/home"} directory tree or the |
61572 | 146 |
Windows user home.\<close> |
47661
012a887997f3
USER_HOME settings variable points to cross-platform user home directory;
wenzelm
parents:
45028
diff
changeset
|
147 |
|
61493 | 148 |
\<^descr>[@{setting_def ISABELLE_HOME}\<open>\<^sup>*\<close>] is the location of the |
47661
012a887997f3
USER_HOME settings variable points to cross-platform user home directory;
wenzelm
parents:
45028
diff
changeset
|
149 |
top-level Isabelle distribution directory. This is automatically |
012a887997f3
USER_HOME settings variable points to cross-platform user home directory;
wenzelm
parents:
45028
diff
changeset
|
150 |
determined from the Isabelle executable that has been invoked. Do |
012a887997f3
USER_HOME settings variable points to cross-platform user home directory;
wenzelm
parents:
45028
diff
changeset
|
151 |
not attempt to set @{setting ISABELLE_HOME} yourself from the shell! |
28215 | 152 |
|
61439 | 153 |
\<^descr>[@{setting_def ISABELLE_HOME_USER}] is the user-specific |
28215 | 154 |
counterpart of @{setting ISABELLE_HOME}. The default value is |
54705 | 155 |
relative to @{file_unchecked "$USER_HOME/.isabelle"}, under rare |
47661
012a887997f3
USER_HOME settings variable points to cross-platform user home directory;
wenzelm
parents:
45028
diff
changeset
|
156 |
circumstances this may be changed in the global setting file. |
012a887997f3
USER_HOME settings variable points to cross-platform user home directory;
wenzelm
parents:
45028
diff
changeset
|
157 |
Typically, the @{setting ISABELLE_HOME_USER} directory mimics |
012a887997f3
USER_HOME settings variable points to cross-platform user home directory;
wenzelm
parents:
45028
diff
changeset
|
158 |
@{setting ISABELLE_HOME} to some extend. In particular, site-wide |
61503 | 159 |
defaults may be overridden by a private |
160 |
\<^verbatim>\<open>$ISABELLE_HOME_USER/etc/settings\<close>. |
|
50182 | 161 |
|
61493 | 162 |
\<^descr>[@{setting_def ISABELLE_PLATFORM_FAMILY}\<open>\<^sup>*\<close>] is |
61503 | 163 |
automatically set to the general platform family: \<^verbatim>\<open>linux\<close>, |
164 |
\<^verbatim>\<open>macos\<close>, \<^verbatim>\<open>windows\<close>. Note that |
|
50182 | 165 |
platform-dependent tools usually need to refer to the more specific |
166 |
identification according to @{setting ISABELLE_PLATFORM}, @{setting |
|
167 |
ISABELLE_PLATFORM32}, @{setting ISABELLE_PLATFORM64}. |
|
168 |
||
61493 | 169 |
\<^descr>[@{setting_def ISABELLE_PLATFORM}\<open>\<^sup>*\<close>] is automatically |
36196
cbb9ee265cdd
added ISABELLE_PLATFORM and ISABELLE_PLATFORM64 -- NB: ML and JVM may have a different idea;
wenzelm
parents:
33952
diff
changeset
|
170 |
set to a symbolic identifier for the underlying hardware and |
cbb9ee265cdd
added ISABELLE_PLATFORM and ISABELLE_PLATFORM64 -- NB: ML and JVM may have a different idea;
wenzelm
parents:
33952
diff
changeset
|
171 |
operating system. The Isabelle platform identification always |
cbb9ee265cdd
added ISABELLE_PLATFORM and ISABELLE_PLATFORM64 -- NB: ML and JVM may have a different idea;
wenzelm
parents:
33952
diff
changeset
|
172 |
refers to the 32 bit variant, even this is a 64 bit machine. Note |
cbb9ee265cdd
added ISABELLE_PLATFORM and ISABELLE_PLATFORM64 -- NB: ML and JVM may have a different idea;
wenzelm
parents:
33952
diff
changeset
|
173 |
that the ML or Java runtime may have a different idea, depending on |
cbb9ee265cdd
added ISABELLE_PLATFORM and ISABELLE_PLATFORM64 -- NB: ML and JVM may have a different idea;
wenzelm
parents:
33952
diff
changeset
|
174 |
which binaries are actually run. |
cbb9ee265cdd
added ISABELLE_PLATFORM and ISABELLE_PLATFORM64 -- NB: ML and JVM may have a different idea;
wenzelm
parents:
33952
diff
changeset
|
175 |
|
61493 | 176 |
\<^descr>[@{setting_def ISABELLE_PLATFORM64}\<open>\<^sup>*\<close>] is similar to |
36196
cbb9ee265cdd
added ISABELLE_PLATFORM and ISABELLE_PLATFORM64 -- NB: ML and JVM may have a different idea;
wenzelm
parents:
33952
diff
changeset
|
177 |
@{setting ISABELLE_PLATFORM} but refers to the proper 64 bit variant |
cbb9ee265cdd
added ISABELLE_PLATFORM and ISABELLE_PLATFORM64 -- NB: ML and JVM may have a different idea;
wenzelm
parents:
33952
diff
changeset
|
178 |
on a platform that supports this; the value is empty for 32 bit. |
47823 | 179 |
Note that the following bash expression (including the quotes) |
180 |
prefers the 64 bit platform, if that is available: |
|
181 |
||
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
182 |
@{verbatim [display] \<open>"${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM}"\<close>} |
36196
cbb9ee265cdd
added ISABELLE_PLATFORM and ISABELLE_PLATFORM64 -- NB: ML and JVM may have a different idea;
wenzelm
parents:
33952
diff
changeset
|
183 |
|
61493 | 184 |
\<^descr>[@{setting_def ISABELLE_PROCESS}\<open>\<^sup>*\<close>, @{setting |
185 |
ISABELLE_TOOL}\<open>\<^sup>*\<close>] are automatically set to the full path |
|
56439
95e2656b3b23
renamed "isabelle-process" to "isabelle_process", with shell function to avoid dynamic path lookups;
wenzelm
parents:
54937
diff
changeset
|
186 |
names of the @{executable "isabelle_process"} and @{executable |
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28502
diff
changeset
|
187 |
isabelle} executables, respectively. Thus other tools and scripts |
40800
330eb65c9469
Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents:
40569
diff
changeset
|
188 |
need not assume that the @{file "$ISABELLE_HOME/bin"} directory is |
28238 | 189 |
on the current search path of the shell. |
28215 | 190 |
|
61493 | 191 |
\<^descr>[@{setting_def ISABELLE_IDENTIFIER}\<open>\<^sup>*\<close>] refers |
61503 | 192 |
to the name of this Isabelle distribution, e.g.\ ``\<^verbatim>\<open>Isabelle2012\<close>''. |
28215 | 193 |
|
61439 | 194 |
\<^descr>[@{setting_def ML_SYSTEM}, @{setting_def ML_HOME}, |
28215 | 195 |
@{setting_def ML_OPTIONS}, @{setting_def ML_PLATFORM}, @{setting_def |
61493 | 196 |
ML_IDENTIFIER}\<open>\<^sup>*\<close>] specify the underlying ML system |
28215 | 197 |
to be used for Isabelle. There is only a fixed set of admissable |
40800
330eb65c9469
Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents:
40569
diff
changeset
|
198 |
@{setting ML_SYSTEM} names (see the @{file |
28238 | 199 |
"$ISABELLE_HOME/etc/settings"} file of the distribution). |
28215 | 200 |
|
201 |
The actual compiler binary will be run from the directory @{setting |
|
202 |
ML_HOME}, with @{setting ML_OPTIONS} as first arguments on the |
|
203 |
command line. The optional @{setting ML_PLATFORM} may specify the |
|
204 |
binary format of ML heap images, which is useful for cross-platform |
|
205 |
installations. The value of @{setting ML_IDENTIFIER} is |
|
206 |
automatically obtained by composing the values of @{setting |
|
207 |
ML_SYSTEM}, @{setting ML_PLATFORM} and the Isabelle version values. |
|
47823 | 208 |
|
61503 | 209 |
\<^descr>[@{setting_def ML_SYSTEM_POLYML}\<open>\<^sup>*\<close>] is \<^verbatim>\<open>true\<close> |
51434 | 210 |
for @{setting ML_SYSTEM} values derived from Poly/ML, as opposed to |
211 |
SML/NJ where it is empty. This is particularly useful with the |
|
212 |
build option @{system_option condition} |
|
213 |
(\secref{sec:system-options}) to restrict big sessions to something |
|
214 |
that SML/NJ can still handle. |
|
215 |
||
61439 | 216 |
\<^descr>[@{setting_def ISABELLE_JDK_HOME}] needs to point to a full JDK |
61503 | 217 |
(Java Development Kit) installation with \<^verbatim>\<open>javac\<close> and |
218 |
\<^verbatim>\<open>jar\<close> executables. This is essential for Isabelle/Scala |
|
47823 | 219 |
and other JVM-based tools to work properly. Note that conventional |
61503 | 220 |
\<^verbatim>\<open>JAVA_HOME\<close> usually points to the JRE (Java Runtime |
47823 | 221 |
Environment), not JDK. |
28215 | 222 |
|
61439 | 223 |
\<^descr>[@{setting_def ISABELLE_PATH}] is a list of directories |
28215 | 224 |
(separated by colons) where Isabelle logic images may reside. When |
225 |
looking up heaps files, the value of @{setting ML_IDENTIFIER} is |
|
226 |
appended to each component internally. |
|
227 |
||
61493 | 228 |
\<^descr>[@{setting_def ISABELLE_OUTPUT}\<open>\<^sup>*\<close>] is a |
28215 | 229 |
directory where output heap files should be stored by default. The |
230 |
ML system and Isabelle version identifier is appended here, too. |
|
231 |
||
61439 | 232 |
\<^descr>[@{setting_def ISABELLE_BROWSER_INFO}] is the directory where |
28215 | 233 |
theory browser information (HTML text, graph data, and printable |
234 |
documents) is stored (see also \secref{sec:info}). The default |
|
54705 | 235 |
value is @{file_unchecked "$ISABELLE_HOME_USER/browser_info"}. |
28215 | 236 |
|
61439 | 237 |
\<^descr>[@{setting_def ISABELLE_LOGIC}] specifies the default logic to |
28215 | 238 |
load if none is given explicitely by the user. The default value is |
61503 | 239 |
\<^verbatim>\<open>HOL\<close>. |
28215 | 240 |
|
61439 | 241 |
\<^descr>[@{setting_def ISABELLE_LINE_EDITOR}] specifies the |
57439
0e41f26a0250
"isabelle tty" is superseded by "isabelle console";
wenzelm
parents:
56604
diff
changeset
|
242 |
line editor for the @{tool_ref console} interface. |
28215 | 243 |
|
61439 | 244 |
\<^descr>[@{setting_def ISABELLE_LATEX}, @{setting_def |
52746 | 245 |
ISABELLE_PDFLATEX}, @{setting_def ISABELLE_BIBTEX}] refer to {\LaTeX} |
246 |
related tools for Isabelle document preparation (see also |
|
247 |
\secref{sec:tool-latex}). |
|
28215 | 248 |
|
61439 | 249 |
\<^descr>[@{setting_def ISABELLE_TOOLS}] is a colon separated list of |
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28502
diff
changeset
|
250 |
directories that are scanned by @{executable isabelle} for external |
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28502
diff
changeset
|
251 |
utility programs (see also \secref{sec:isabelle-tool}). |
28215 | 252 |
|
61439 | 253 |
\<^descr>[@{setting_def ISABELLE_DOCS}] is a colon separated list of |
28215 | 254 |
directories with documentation files. |
50197
b385d134926d
eval PDF_VIEWER/DVI_VIEWER command line, which allows additional quotes for program name, for example;
wenzelm
parents:
50182
diff
changeset
|
255 |
|
61439 | 256 |
\<^descr>[@{setting_def PDF_VIEWER}] specifies the program to be used |
61503 | 257 |
for displaying \<^verbatim>\<open>pdf\<close> files. |
54683
cf48ddc266e5
clarified "isabelle display" and 'display_drafts': re-use file and program instance, open asynchronously via desktop environment;
wenzelm
parents:
52746
diff
changeset
|
258 |
|
61439 | 259 |
\<^descr>[@{setting_def DVI_VIEWER}] specifies the program to be used |
61503 | 260 |
for displaying \<^verbatim>\<open>dvi\<close> files. |
28215 | 261 |
|
61493 | 262 |
\<^descr>[@{setting_def ISABELLE_TMP_PREFIX}\<open>\<^sup>*\<close>] is the |
56439
95e2656b3b23
renamed "isabelle-process" to "isabelle_process", with shell function to avoid dynamic path lookups;
wenzelm
parents:
54937
diff
changeset
|
263 |
prefix from which any running @{executable "isabelle_process"} |
58639
1df53737c59b
prefer Unix standard-conformant $TMPDIR over hard-wired /tmp;
wenzelm
parents:
58618
diff
changeset
|
264 |
derives an individual directory for temporary files. |
58618 | 265 |
\<close> |
28215 | 266 |
|
267 |
||
58618 | 268 |
subsection \<open>Additional components \label{sec:components}\<close> |
32323 | 269 |
|
61477 | 270 |
text \<open>Any directory may be registered as an explicit \<^emph>\<open>Isabelle |
271 |
component\<close>. The general layout conventions are that of the main |
|
32323 | 272 |
Isabelle distribution itself, and the following two files (both |
273 |
optional) have a special meaning: |
|
274 |
||
61503 | 275 |
\<^item> \<^verbatim>\<open>etc/settings\<close> holds additional settings that are |
32323 | 276 |
initialized when bootstrapping the overall Isabelle environment, |
277 |
cf.\ \secref{sec:boot}. As usual, the content is interpreted as a |
|
61503 | 278 |
\<^verbatim>\<open>bash\<close> script. It may refer to the component's enclosing |
279 |
directory via the \<^verbatim>\<open>COMPONENT\<close> shell variable. |
|
32323 | 280 |
|
281 |
For example, the following setting allows to refer to files within |
|
282 |
the component later on, without having to hardwire absolute paths: |
|
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
283 |
@{verbatim [display] \<open>MY_COMPONENT_HOME="$COMPONENT"\<close>} |
32323 | 284 |
|
285 |
Components can also add to existing Isabelle settings such as |
|
286 |
@{setting_def ISABELLE_TOOLS}, in order to provide |
|
287 |
component-specific tools that can be invoked by end-users. For |
|
288 |
example: |
|
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
289 |
@{verbatim [display] \<open>ISABELLE_TOOLS="$ISABELLE_TOOLS:$COMPONENT/lib/Tools"\<close>} |
32323 | 290 |
|
61503 | 291 |
\<^item> \<^verbatim>\<open>etc/components\<close> holds a list of further |
32323 | 292 |
sub-components of the same structure. The directory specifications |
61503 | 293 |
given here can be either absolute (with leading \<^verbatim>\<open>/\<close>) or |
32323 | 294 |
relative to the component's main directory. |
295 |
||
296 |
||
297 |
The root of component initialization is @{setting ISABELLE_HOME} |
|
298 |
itself. After initializing all of its sub-components recursively, |
|
299 |
@{setting ISABELLE_HOME_USER} is included in the same manner (if |
|
40569
ffcff7509a49
more explicit explanation of init_component shell function;
wenzelm
parents:
40387
diff
changeset
|
300 |
that directory exists). This allows to install private components |
54705 | 301 |
via @{file_unchecked "$ISABELLE_HOME_USER/etc/components"}, although it is |
40569
ffcff7509a49
more explicit explanation of init_component shell function;
wenzelm
parents:
40387
diff
changeset
|
302 |
often more convenient to do that programmatically via the |
61503 | 303 |
\<^verbatim>\<open>init_component\<close> shell function in the \<^verbatim>\<open>etc/settings\<close> |
304 |
script of \<^verbatim>\<open>$ISABELLE_HOME_USER\<close> (or any other component |
|
40569
ffcff7509a49
more explicit explanation of init_component shell function;
wenzelm
parents:
40387
diff
changeset
|
305 |
directory). For example: |
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
306 |
@{verbatim [display] \<open>init_component "$HOME/screwdriver-2.0"\<close>} |
48838
623ba165d059
direct support for component forests via init_components;
wenzelm
parents:
48813
diff
changeset
|
307 |
|
623ba165d059
direct support for component forests via init_components;
wenzelm
parents:
48813
diff
changeset
|
308 |
This is tolerant wrt.\ missing component directories, but might |
623ba165d059
direct support for component forests via init_components;
wenzelm
parents:
48813
diff
changeset
|
309 |
produce a warning. |
623ba165d059
direct support for component forests via init_components;
wenzelm
parents:
48813
diff
changeset
|
310 |
|
61406 | 311 |
\<^medskip> |
312 |
More complex situations may be addressed by initializing |
|
48838
623ba165d059
direct support for component forests via init_components;
wenzelm
parents:
48813
diff
changeset
|
313 |
components listed in a given catalog file, relatively to some base |
623ba165d059
direct support for component forests via init_components;
wenzelm
parents:
48813
diff
changeset
|
314 |
directory: |
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
315 |
@{verbatim [display] \<open>init_components "$HOME/my_component_store" "some_catalog_file"\<close>} |
48838
623ba165d059
direct support for component forests via init_components;
wenzelm
parents:
48813
diff
changeset
|
316 |
|
623ba165d059
direct support for component forests via init_components;
wenzelm
parents:
48813
diff
changeset
|
317 |
The component directories listed in the catalog file are treated as |
623ba165d059
direct support for component forests via init_components;
wenzelm
parents:
48813
diff
changeset
|
318 |
relative to the given base directory. |
48844 | 319 |
|
320 |
See also \secref{sec:tool-components} for some tool-support for |
|
321 |
resolving components that are formally initialized but not installed |
|
322 |
yet. |
|
58618 | 323 |
\<close> |
32323 | 324 |
|
325 |
||
58618 | 326 |
section \<open>The raw Isabelle process \label{sec:isabelle-process}\<close> |
28215 | 327 |
|
58618 | 328 |
text \<open> |
56439
95e2656b3b23
renamed "isabelle-process" to "isabelle_process", with shell function to avoid dynamic path lookups;
wenzelm
parents:
54937
diff
changeset
|
329 |
The @{executable_def "isabelle_process"} executable runs bare-bones |
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28502
diff
changeset
|
330 |
Isabelle logic sessions --- either interactively or in batch mode. |
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28502
diff
changeset
|
331 |
It provides an abstraction over the underlying ML system, and over |
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28502
diff
changeset
|
332 |
the actual heap file locations. Its usage is: |
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
333 |
@{verbatim [display] |
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
334 |
\<open>Usage: isabelle_process [OPTIONS] [INPUT] [OUTPUT] |
28215 | 335 |
|
336 |
Options are: |
|
57581
74bbe9317aa4
provide explicit options file -- avoid multiple Scala/JVM invocation;
wenzelm
parents:
57439
diff
changeset
|
337 |
-O system options from given YXML file |
59350
acba5d6fdb2f
discontinued fifo channel, always use portable socket;
wenzelm
parents:
58846
diff
changeset
|
338 |
-P SOCKET startup process wrapper via TCP socket |
28215 | 339 |
-S secure mode -- disallow critical operations |
340 |
-e MLTEXT pass MLTEXT to the ML session |
|
341 |
-m MODE add print mode for output |
|
52056 | 342 |
-o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) |
28215 | 343 |
-q non-interactive session |
344 |
-r open heap file read-only |
|
345 |
-w reset write permissions on OUTPUT |
|
346 |
||
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
347 |
INPUT (default "$ISABELLE_LOGIC") and OUTPUT specify in/out heaps. |
28215 | 348 |
These are either names to be searched in the Isabelle path, or |
349 |
actual file names (containing at least one /). |
|
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
350 |
If INPUT is "RAW_ML_SYSTEM", just start the bare bones ML system.\<close>} |
28215 | 351 |
|
352 |
Input files without path specifications are looked up in the |
|
353 |
@{setting ISABELLE_PATH} setting, which may consist of multiple |
|
354 |
components separated by colons --- these are tried in the given |
|
355 |
order with the value of @{setting ML_IDENTIFIER} appended |
|
356 |
internally. In a similar way, base names are relative to the |
|
357 |
directory specified by @{setting ISABELLE_OUTPUT}. In any case, |
|
358 |
actual file locations may also be given by including at least one |
|
61503 | 359 |
slash (\<^verbatim>\<open>/\<close>) in the name (hint: use \<^verbatim>\<open>./\<close> to |
28215 | 360 |
refer to the current directory). |
58618 | 361 |
\<close> |
28215 | 362 |
|
363 |
||
58618 | 364 |
subsubsection \<open>Options\<close> |
28215 | 365 |
|
58618 | 366 |
text \<open> |
28215 | 367 |
If the input heap file does not have write permission bits set, or |
61503 | 368 |
the \<^verbatim>\<open>-r\<close> option is given explicitly, then the session |
28215 | 369 |
started will be read-only. That is, the ML world cannot be |
370 |
committed back into the image file. Otherwise, a writable session |
|
371 |
enables commits into either the input file, or into another output |
|
372 |
heap file (if that is given as the second argument on the command |
|
373 |
line). |
|
374 |
||
375 |
The read-write state of sessions is determined at startup only, it |
|
376 |
cannot be changed intermediately. Also note that heap images may |
|
59904 | 377 |
require considerable amounts of disk space (hundreds of MB or some GB). |
378 |
Users are responsible for themselves to dispose their heap files |
|
379 |
when they are no longer needed. |
|
28215 | 380 |
|
61406 | 381 |
\<^medskip> |
61503 | 382 |
The \<^verbatim>\<open>-w\<close> option makes the output heap file |
28215 | 383 |
read-only after terminating. Thus subsequent invocations cause the |
384 |
logic image to be read-only automatically. |
|
385 |
||
61406 | 386 |
\<^medskip> |
61503 | 387 |
Using the \<^verbatim>\<open>-e\<close> option, arbitrary ML code may be |
28215 | 388 |
passed to the Isabelle session from the command line. Multiple |
61503 | 389 |
\<^verbatim>\<open>-e\<close> options are evaluated in the given order. Strange things |
56604 | 390 |
may happen when erroneous ML code is provided. Also make sure that |
28215 | 391 |
the ML commands are terminated properly by semicolon. |
392 |
||
61406 | 393 |
\<^medskip> |
61503 | 394 |
The \<^verbatim>\<open>-m\<close> option adds identifiers of print modes |
28215 | 395 |
to be made active for this session. Typically, this is used by some |
396 |
user interface, e.g.\ to enable output of proper mathematical |
|
397 |
symbols. |
|
398 |
||
61406 | 399 |
\<^medskip> |
400 |
Isabelle normally enters an interactive top-level loop |
|
61503 | 401 |
(after processing the \<^verbatim>\<open>-e\<close> texts). The \<^verbatim>\<open>-q\<close> |
28215 | 402 |
option inhibits interaction, thus providing a pure batch mode |
403 |
facility. |
|
404 |
||
61406 | 405 |
\<^medskip> |
61503 | 406 |
Option \<^verbatim>\<open>-o\<close> allows to override Isabelle system |
52056 | 407 |
options for this process, see also \secref{sec:system-options}. |
61503 | 408 |
Alternatively, option \<^verbatim>\<open>-O\<close> specifies the full environment of |
57581
74bbe9317aa4
provide explicit options file -- avoid multiple Scala/JVM invocation;
wenzelm
parents:
57439
diff
changeset
|
409 |
system options as a file in YXML format (according to the Isabelle/Scala |
61503 | 410 |
function \<^verbatim>\<open>isabelle.Options.encode\<close>). |
52056 | 411 |
|
61406 | 412 |
\<^medskip> |
61503 | 413 |
The \<^verbatim>\<open>-P\<close> option starts the Isabelle process wrapper |
59350
acba5d6fdb2f
discontinued fifo channel, always use portable socket;
wenzelm
parents:
58846
diff
changeset
|
414 |
for Isabelle/Scala, with a private protocol running over the specified TCP |
acba5d6fdb2f
discontinued fifo channel, always use portable socket;
wenzelm
parents:
58846
diff
changeset
|
415 |
socket. Isabelle/ML and Isabelle/Scala provide various programming |
acba5d6fdb2f
discontinued fifo channel, always use portable socket;
wenzelm
parents:
58846
diff
changeset
|
416 |
interfaces to invoke protocol functions over untyped strings and XML |
acba5d6fdb2f
discontinued fifo channel, always use portable socket;
wenzelm
parents:
58846
diff
changeset
|
417 |
trees. |
28215 | 418 |
|
61406 | 419 |
\<^medskip> |
61503 | 420 |
The \<^verbatim>\<open>-S\<close> option makes the Isabelle process more |
28215 | 421 |
secure by disabling some critical operations, notably runtime |
422 |
compilation and evaluation of ML source code. |
|
58618 | 423 |
\<close> |
28215 | 424 |
|
425 |
||
58618 | 426 |
subsubsection \<open>Examples\<close> |
28215 | 427 |
|
58618 | 428 |
text \<open> |
28215 | 429 |
Run an interactive session of the default object-logic (as specified |
430 |
by the @{setting ISABELLE_LOGIC} setting) like this: |
|
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
431 |
@{verbatim [display] \<open>isabelle_process\<close>} |
28215 | 432 |
|
433 |
Usually @{setting ISABELLE_LOGIC} refers to one of the standard |
|
434 |
logic images, which are read-only by default. A writable session |
|
61503 | 435 |
--- based on \<^verbatim>\<open>HOL\<close>, but output to \<^verbatim>\<open>Test\<close> (in the |
28238 | 436 |
directory specified by the @{setting ISABELLE_OUTPUT} setting) --- |
28215 | 437 |
may be invoked as follows: |
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
438 |
@{verbatim [display] \<open>isabelle_process HOL Test\<close>} |
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
439 |
|
28215 | 440 |
Ending this session normally (e.g.\ by typing control-D) dumps the |
61503 | 441 |
whole ML system state into \<^verbatim>\<open>Test\<close> (be prepared for more |
47823 | 442 |
than 100\,MB): |
28215 | 443 |
|
61503 | 444 |
The \<^verbatim>\<open>Test\<close> session may be continued later (still in |
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
445 |
writable state) by: @{verbatim [display] \<open>isabelle_process Test\<close>} |
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
446 |
|
61503 | 447 |
A read-only \<^verbatim>\<open>Test\<close> session may be started by: |
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
448 |
@{verbatim [display] \<open>isabelle_process -r Test\<close>} |
28215 | 449 |
|
61406 | 450 |
\<^bigskip> |
451 |
The next example demonstrates batch execution of Isabelle. |
|
61503 | 452 |
We retrieve the \<^verbatim>\<open>Main\<close> theory value from the theory loader |
47823 | 453 |
within ML (observe the delicate quoting rules for the Bash shell |
454 |
vs.\ ML): |
|
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
455 |
@{verbatim [display] \<open>isabelle_process -e 'Thy_Info.get_theory "Main";' -q -r HOL\<close>} |
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
456 |
|
28215 | 457 |
Note that the output text will be interspersed with additional junk |
61503 | 458 |
messages by the ML runtime environment. The \<^verbatim>\<open>-W\<close> option |
28238 | 459 |
allows to communicate with the Isabelle process via an external |
460 |
program in a more robust fashion. |
|
58618 | 461 |
\<close> |
28238 | 462 |
|
463 |
||
58618 | 464 |
section \<open>The Isabelle tool wrapper \label{sec:isabelle-tool}\<close> |
28238 | 465 |
|
58618 | 466 |
text \<open> |
28238 | 467 |
All Isabelle related tools and interfaces are called via a common |
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28502
diff
changeset
|
468 |
wrapper --- @{executable isabelle}: |
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
469 |
@{verbatim [display] |
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
470 |
\<open>Usage: isabelle TOOL [ARGS ...] |
28238 | 471 |
|
28506 | 472 |
Start Isabelle tool NAME with ARGS; pass "-?" for tool specific help. |
28238 | 473 |
|
48858 | 474 |
Available tools: |
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
475 |
...\<close>} |
28238 | 476 |
|
477 |
In principle, Isabelle tools are ordinary executable scripts that |
|
478 |
are run within the Isabelle settings environment, see |
|
479 |
\secref{sec:settings}. The set of available tools is collected by |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
28502
diff
changeset
|
480 |
@{executable isabelle} from the directories listed in the @{setting |
28238 | 481 |
ISABELLE_TOOLS} setting. Do not try to call the scripts directly |
482 |
from the shell. Neither should you add the tool directories to your |
|
483 |
shell's search path! |
|
58618 | 484 |
\<close> |
28238 | 485 |
|
486 |
||
58618 | 487 |
subsubsection \<open>Examples\<close> |
28238 | 488 |
|
58618 | 489 |
text \<open>Show the list of available documentation of the Isabelle |
48813 | 490 |
distribution: |
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
491 |
@{verbatim [display] \<open>isabelle doc\<close>} |
28238 | 492 |
|
493 |
View a certain document as follows: |
|
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
494 |
@{verbatim [display] \<open>isabelle doc system\<close>} |
28238 | 495 |
|
48813 | 496 |
Query the Isabelle settings environment: |
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
497 |
@{verbatim [display] \<open>isabelle getenv ISABELLE_HOME_USER\<close>} |
58618 | 498 |
\<close> |
28215 | 499 |
|
28223 | 500 |
end |