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