| author | bulwahn | 
| Thu, 21 Oct 2010 19:13:09 +0200 | |
| changeset 40051 | b6acda4d1c29 | 
| parent 38253 | 3d4e521014f7 | 
| child 40387 | e4c9e0dad473 | 
| permissions | -rw-r--r-- | 
| 28215 | 1 | theory Basics | 
| 2 | imports Pure | |
| 3 | begin | |
| 4 | ||
| 5 | chapter {* The Isabelle system environment *}
 | |
| 6 | ||
| 7 | text {*
 | |
| 8 | This manual describes Isabelle together with related tools and user | |
| 28916 
0a802cdda340
removed obsolete isabelle-interface executable and ISABELLE_INTERFACE setting;
 wenzelm parents: 
28914diff
changeset | 9 | interfaces as seen from a system oriented view. See also the | 
| 
0a802cdda340
removed obsolete isabelle-interface executable and ISABELLE_INTERFACE setting;
 wenzelm parents: 
28914diff
changeset | 10 |   \emph{Isabelle/Isar Reference Manual}~\cite{isabelle-isar-ref} for
 | 
| 
0a802cdda340
removed obsolete isabelle-interface executable and ISABELLE_INTERFACE setting;
 wenzelm parents: 
28914diff
changeset | 11 | the actual Isabelle input language and related concepts. | 
| 28215 | 12 | |
| 28916 
0a802cdda340
removed obsolete isabelle-interface executable and ISABELLE_INTERFACE setting;
 wenzelm parents: 
28914diff
changeset | 13 | \medskip The Isabelle system environment provides the following | 
| 
0a802cdda340
removed obsolete isabelle-interface executable and ISABELLE_INTERFACE setting;
 wenzelm parents: 
28914diff
changeset | 14 | basic infrastructure to integrate tools smoothly. | 
| 28215 | 15 | |
| 28238 | 16 |   \begin{enumerate}
 | 
| 28215 | 17 | |
| 28916 
0a802cdda340
removed obsolete isabelle-interface executable and ISABELLE_INTERFACE setting;
 wenzelm parents: 
28914diff
changeset | 18 |   \item The \emph{Isabelle settings} mechanism provides process
 | 
| 
0a802cdda340
removed obsolete isabelle-interface executable and ISABELLE_INTERFACE setting;
 wenzelm parents: 
28914diff
changeset | 19 | environment variables to all Isabelle executables (including tools | 
| 
0a802cdda340
removed obsolete isabelle-interface executable and ISABELLE_INTERFACE setting;
 wenzelm parents: 
28914diff
changeset | 20 | and user interfaces). | 
| 28215 | 21 | |
| 28504 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
changeset | 22 |   \item The \emph{raw Isabelle process} (@{executable_ref
 | 
| 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
changeset | 23 | "isabelle-process"}) runs logic sessions either interactively or in | 
| 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
changeset | 24 | 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: 
28502diff
changeset | 25 | 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: 
28502diff
changeset | 26 | care about the low-level process. | 
| 28215 | 27 | |
| 28504 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
changeset | 28 |   \item The \emph{Isabelle tools wrapper} (@{executable_ref isabelle})
 | 
| 28238 | 29 | provides a generic startup environment Isabelle related utilities, | 
| 30 | user interfaces etc. Such tools automatically benefit from the | |
| 31 | settings mechanism. | |
| 28215 | 32 | |
| 28238 | 33 |   \end{enumerate}
 | 
| 28215 | 34 | *} | 
| 35 | ||
| 36 | ||
| 37 | section {* Isabelle settings \label{sec:settings} *}
 | |
| 38 | ||
| 39 | text {*
 | |
| 40 |   The Isabelle system heavily depends on the \emph{settings
 | |
| 41 |   mechanism}\indexbold{settings}.  Essentially, this is a statically
 | |
| 42 |   scoped collection of environment variables, such as @{setting
 | |
| 43 |   ISABELLE_HOME}, @{setting ML_SYSTEM}, @{setting ML_HOME}.  These
 | |
| 44 |   variables are \emph{not} intended to be set directly from the shell,
 | |
| 45 | though. Isabelle employs a somewhat more sophisticated scheme of | |
| 46 |   \emph{settings files} --- one for site-wide defaults, another for
 | |
| 47 | additional user-specific modifications. With all configuration | |
| 48 | variables in at most two places, this scheme is more maintainable | |
| 49 | and user-friendly than global shell environment variables. | |
| 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. | |
| 28238 | 56 |   Occasionally, users would still want to put the @{"file"
 | 
| 57 | "$ISABELLE_HOME/bin"} directory into their shell's search path, but | |
| 58 | this is not required. | |
| 28215 | 59 | *} | 
| 60 | ||
| 61 | ||
| 32323 | 62 | subsection {* Bootstrapping the environment \label{sec:boot} *}
 | 
| 28215 | 63 | |
| 32323 | 64 | text {* Isabelle executables need to be run within a proper settings
 | 
| 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 | |
| 71 |   \begin{enumerate}
 | |
| 72 | ||
| 73 |   \item The special variable @{setting_def ISABELLE_HOME} is
 | |
| 74 | determined automatically from the location of the binary that has | |
| 75 | been run. | |
| 76 | ||
| 77 |   You should not try to set @{setting ISABELLE_HOME} manually. Also
 | |
| 78 | note that the Isabelle executables either have to be run from their | |
| 79 | original location in the distribution directory, or via the | |
| 28238 | 80 |   executable objects created by the @{tool install} utility.  Symbolic
 | 
| 81 |   links are admissible, but a plain copy of the @{"file"
 | |
| 82 | "$ISABELLE_HOME/bin"} files will not work! | |
| 83 | ||
| 32324 | 84 |   \item The file @{"file" "$ISABELLE_HOME/etc/settings"} is run as a
 | 
| 28238 | 85 |   @{executable_ref bash} shell script with the auto-export option for
 | 
| 86 | variables enabled. | |
| 28215 | 87 | |
| 88 | This file holds a rather long list of shell variable assigments, | |
| 89 | thus providing the site-wide default settings. The Isabelle | |
| 90 | distribution already contains a global settings file with sensible | |
| 91 | defaults for most variables. When installing the system, only a few | |
| 92 |   of these may have to be adapted (probably @{setting ML_SYSTEM}
 | |
| 93 | etc.). | |
| 94 | ||
| 28285 | 95 |   \item The file @{verbatim "$ISABELLE_HOME_USER/etc/settings"} (if it
 | 
| 28215 | 96 | exists) is run in the same way as the site default settings. Note | 
| 97 |   that the variable @{setting ISABELLE_HOME_USER} has already been set
 | |
| 28914 
f993cbffc42a
default for ISABELLE_HOME_USER is now ~/.isabelle instead of ~/isabelle;
 wenzelm parents: 
28506diff
changeset | 98 |   before --- usually to @{verbatim "~/.isabelle"}.
 | 
| 28215 | 99 | |
| 100 | Thus individual users may override the site-wide defaults. See also | |
| 28238 | 101 |   file @{"file" "$ISABELLE_HOME/etc/user-settings.sample"} in the
 | 
| 102 | distribution. Typically, a user settings file would contain only a | |
| 103 | few lines, just the assigments that are really changed. One should | |
| 104 |   definitely \emph{not} start with a full copy the basic @{"file"
 | |
| 28215 | 105 | "$ISABELLE_HOME/etc/settings"}. This could cause very annoying | 
| 106 | maintainance problems later, when the Isabelle installation is | |
| 107 | updated or changed otherwise. | |
| 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
 | 
| 28504 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
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 | 
| 136 |   the Isabelle tool @{tool getenv}.  This might help to figure out the
 | |
| 137 | effect of complex settings scripts. | |
| 28215 | 138 | *} | 
| 139 | ||
| 140 | ||
| 141 | subsection{* Common variables *}
 | |
| 142 | ||
| 143 | text {*
 | |
| 144 | This is a reference of common Isabelle settings variables. Note that | |
| 145 | the list is somewhat open-ended. Third-party utilities or interfaces | |
| 146 | may add their own selection. Variables that are special in some | |
| 147 |   sense are marked with @{text "\<^sup>*"}.
 | |
| 148 | ||
| 149 |   \begin{description}
 | |
| 150 | ||
| 151 |   \item[@{setting_def ISABELLE_HOME}@{text "\<^sup>*"}] is the
 | |
| 152 | location of the top-level Isabelle distribution directory. This is | |
| 153 | automatically determined from the Isabelle executable that has been | |
| 154 |   invoked.  Do not attempt to set @{setting ISABELLE_HOME} yourself
 | |
| 28238 | 155 | from the shell! | 
| 28215 | 156 | |
| 157 |   \item[@{setting_def ISABELLE_HOME_USER}] is the user-specific
 | |
| 158 |   counterpart of @{setting ISABELLE_HOME}. The default value is
 | |
| 28914 
f993cbffc42a
default for ISABELLE_HOME_USER is now ~/.isabelle instead of ~/isabelle;
 wenzelm parents: 
28506diff
changeset | 159 |   @{verbatim "~/.isabelle"}, under rare circumstances this may be
 | 
| 28215 | 160 |   changed in the global setting file.  Typically, the @{setting
 | 
| 161 |   ISABELLE_HOME_USER} directory mimics @{setting ISABELLE_HOME} to
 | |
| 162 | some extend. In particular, site-wide defaults may be overridden by | |
| 28285 | 163 |   a private @{verbatim "$ISABELLE_HOME_USER/etc/settings"}.
 | 
| 28215 | 164 | |
| 36196 
cbb9ee265cdd
added ISABELLE_PLATFORM and ISABELLE_PLATFORM64 -- NB: ML and JVM may have a different idea;
 wenzelm parents: 
33952diff
changeset | 165 |   \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: 
33952diff
changeset | 166 | 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: 
33952diff
changeset | 167 | 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: 
33952diff
changeset | 168 | 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: 
33952diff
changeset | 169 | 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: 
33952diff
changeset | 170 | which binaries are actually run. | 
| 
cbb9ee265cdd
added ISABELLE_PLATFORM and ISABELLE_PLATFORM64 -- NB: ML and JVM may have a different idea;
 wenzelm parents: 
33952diff
changeset | 171 | |
| 
cbb9ee265cdd
added ISABELLE_PLATFORM and ISABELLE_PLATFORM64 -- NB: ML and JVM may have a different idea;
 wenzelm parents: 
33952diff
changeset | 172 |   \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: 
33952diff
changeset | 173 |   @{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: 
33952diff
changeset | 174 | on a platform that supports this; the value is empty for 32 bit. | 
| 
cbb9ee265cdd
added ISABELLE_PLATFORM and ISABELLE_PLATFORM64 -- NB: ML and JVM may have a different idea;
 wenzelm parents: 
33952diff
changeset | 175 | |
| 28502 | 176 |   \item[@{setting_def ISABELLE_PROCESS}@{text "\<^sup>*"}, @{setting
 | 
| 28500 | 177 |   ISABELLE_TOOL}@{text "\<^sup>*"}] are automatically set to the full path
 | 
| 28215 | 178 |   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: 
28502diff
changeset | 179 | isabelle} executables, respectively. Thus other tools and scripts | 
| 28238 | 180 |   need not assume that the @{"file" "$ISABELLE_HOME/bin"} directory is
 | 
| 181 | on the current search path of the shell. | |
| 28215 | 182 | |
| 183 |   \item[@{setting_def ISABELLE_IDENTIFIER}@{text "\<^sup>*"}] refers
 | |
| 184 |   to the name of this Isabelle distribution, e.g.\ ``@{verbatim
 | |
| 185 | Isabelle2008}''. | |
| 186 | ||
| 187 |   \item[@{setting_def ML_SYSTEM}, @{setting_def ML_HOME},
 | |
| 188 |   @{setting_def ML_OPTIONS}, @{setting_def ML_PLATFORM}, @{setting_def
 | |
| 189 |   ML_IDENTIFIER}@{text "\<^sup>*"}] specify the underlying ML system
 | |
| 190 | to be used for Isabelle. There is only a fixed set of admissable | |
| 28238 | 191 |   @{setting ML_SYSTEM} names (see the @{"file"
 | 
| 192 | "$ISABELLE_HOME/etc/settings"} file of the distribution). | |
| 28215 | 193 | |
| 194 |   The actual compiler binary will be run from the directory @{setting
 | |
| 195 |   ML_HOME}, with @{setting ML_OPTIONS} as first arguments on the
 | |
| 196 |   command line.  The optional @{setting ML_PLATFORM} may specify the
 | |
| 197 | binary format of ML heap images, which is useful for cross-platform | |
| 198 |   installations.  The value of @{setting ML_IDENTIFIER} is
 | |
| 199 |   automatically obtained by composing the values of @{setting
 | |
| 200 |   ML_SYSTEM}, @{setting ML_PLATFORM} and the Isabelle version values.
 | |
| 201 | ||
| 202 |   \item[@{setting_def ISABELLE_PATH}] is a list of directories
 | |
| 203 | (separated by colons) where Isabelle logic images may reside. When | |
| 204 |   looking up heaps files, the value of @{setting ML_IDENTIFIER} is
 | |
| 205 | appended to each component internally. | |
| 206 | ||
| 207 |   \item[@{setting_def ISABELLE_OUTPUT}@{text "\<^sup>*"}] is a
 | |
| 208 | directory where output heap files should be stored by default. The | |
| 209 | ML system and Isabelle version identifier is appended here, too. | |
| 210 | ||
| 211 |   \item[@{setting_def ISABELLE_BROWSER_INFO}] is the directory where
 | |
| 212 | theory browser information (HTML text, graph data, and printable | |
| 213 |   documents) is stored (see also \secref{sec:info}).  The default
 | |
| 214 |   value is @{verbatim "$ISABELLE_HOME_USER/browser_info"}.
 | |
| 215 | ||
| 216 |   \item[@{setting_def ISABELLE_LOGIC}] specifies the default logic to
 | |
| 217 | load if none is given explicitely by the user. The default value is | |
| 218 |   @{verbatim HOL}.
 | |
| 219 | ||
| 220 |   \item[@{setting_def ISABELLE_LINE_EDITOR}] specifies the default
 | |
| 28238 | 221 |   line editor for the @{tool_ref tty} interface.
 | 
| 28215 | 222 | |
| 223 |   \item[@{setting_def ISABELLE_USEDIR_OPTIONS}] is implicitly prefixed
 | |
| 28238 | 224 |   to the command line of any @{tool_ref usedir} invocation. This
 | 
| 225 |   typically contains compilation options for object-logics --- @{tool
 | |
| 226 | usedir} is the basic utility for managing logic sessions (cf.\ the | |
| 227 |   @{verbatim IsaMakefile}s in the distribution).
 | |
| 28215 | 228 | |
| 229 |   \item[@{setting_def ISABELLE_FILE_IDENT}] specifies a shell command
 | |
| 230 | for producing a source file identification, based on the actual | |
| 231 | content instead of the full physical path and date stamp (which is | |
| 232 | the default). A typical identification would produce a ``digest'' of | |
| 233 | the text, using a cryptographic hash function like SHA-1, for | |
| 234 | example. | |
| 235 | ||
| 236 |   \item[@{setting_def ISABELLE_LATEX}, @{setting_def
 | |
| 237 |   ISABELLE_PDFLATEX}, @{setting_def ISABELLE_BIBTEX}, @{setting_def
 | |
| 238 |   ISABELLE_DVIPS}] refer to {\LaTeX} related tools for Isabelle
 | |
| 239 |   document preparation (see also \secref{sec:tool-latex}).
 | |
| 240 | ||
| 241 |   \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: 
28502diff
changeset | 242 |   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: 
28502diff
changeset | 243 |   utility programs (see also \secref{sec:isabelle-tool}).
 | 
| 28215 | 244 | |
| 245 |   \item[@{setting_def ISABELLE_DOCS}] is a colon separated list of
 | |
| 246 | directories with documentation files. | |
| 247 | ||
| 248 |   \item[@{setting_def ISABELLE_DOC_FORMAT}] specifies the preferred
 | |
| 249 |   document format, typically @{verbatim dvi} or @{verbatim pdf}.
 | |
| 250 | ||
| 251 |   \item[@{setting_def DVI_VIEWER}] specifies the command to be used
 | |
| 252 |   for displaying @{verbatim dvi} files.
 | |
| 253 | ||
| 254 |   \item[@{setting_def PDF_VIEWER}] specifies the command to be used
 | |
| 255 |   for displaying @{verbatim pdf} files.
 | |
| 256 | ||
| 257 |   \item[@{setting_def PRINT_COMMAND}] specifies the standard printer
 | |
| 258 |   spool command, which is expected to accept @{verbatim ps} files.
 | |
| 259 | ||
| 260 |   \item[@{setting_def ISABELLE_TMP_PREFIX}@{text "\<^sup>*"}] is the
 | |
| 28238 | 261 |   prefix from which any running @{executable "isabelle-process"}
 | 
| 262 | derives an individual directory for temporary files. The default is | |
| 28215 | 263 |   somewhere in @{verbatim "/tmp"}.
 | 
| 264 | ||
| 265 |   \end{description}
 | |
| 266 | *} | |
| 267 | ||
| 268 | ||
| 32323 | 269 | subsection {* Additional components \label{sec:components} *}
 | 
| 270 | ||
| 271 | text {* Any directory may be registered as an explicit \emph{Isabelle
 | |
| 272 | component}. The general layout conventions are that of the main | |
| 273 | Isabelle distribution itself, and the following two files (both | |
| 274 | optional) have a special meaning: | |
| 275 | ||
| 276 |   \begin{itemize}
 | |
| 277 | ||
| 278 |   \item @{verbatim "etc/settings"} holds additional settings that are
 | |
| 279 | initialized when bootstrapping the overall Isabelle environment, | |
| 280 |   cf.\ \secref{sec:boot}.  As usual, the content is interpreted as a
 | |
| 281 |   @{verbatim bash} script.  It may refer to the component's enclosing
 | |
| 282 |   directory via the @{verbatim "COMPONENT"} shell variable.
 | |
| 283 | ||
| 284 | For example, the following setting allows to refer to files within | |
| 285 | the component later on, without having to hardwire absolute paths: | |
| 286 | ||
| 287 |   \begin{ttbox}
 | |
| 288 | MY_COMPONENT_HOME="$COMPONENT" | |
| 289 |   \end{ttbox}
 | |
| 290 | ||
| 291 | Components can also add to existing Isabelle settings such as | |
| 292 |   @{setting_def ISABELLE_TOOLS}, in order to provide
 | |
| 293 | component-specific tools that can be invoked by end-users. For | |
| 294 | example: | |
| 295 | ||
| 296 |   \begin{ttbox}
 | |
| 297 | ISABELLE_TOOLS="$ISABELLE_TOOLS:$COMPONENT/lib/Tools" | |
| 298 |   \end{ttbox}
 | |
| 299 | ||
| 300 |   \item @{verbatim "etc/components"} holds a list of further
 | |
| 301 | sub-components of the same structure. The directory specifications | |
| 302 |   given here can be either absolute (with leading @{verbatim "/"}) or
 | |
| 303 | relative to the component's main directory. | |
| 304 | ||
| 305 |   \end{itemize}
 | |
| 306 | ||
| 307 |   The root of component initialization is @{setting ISABELLE_HOME}
 | |
| 308 | itself. After initializing all of its sub-components recursively, | |
| 309 |   @{setting ISABELLE_HOME_USER} is included in the same manner (if
 | |
| 310 | that directory exists). Thus users can easily add private | |
| 311 |   components to @{verbatim "$ISABELLE_HOME_USER/etc/components"}.
 | |
| 33952 | 312 | |
| 313 | It is also possible to initialize components programmatically via | |
| 314 | the \verb,init_component, shell function, say within the | |
| 315 | \verb,settings, script of another component. | |
| 32323 | 316 | *} | 
| 317 | ||
| 318 | ||
| 28215 | 319 | section {* The raw Isabelle process *}
 | 
| 320 | ||
| 321 | text {*
 | |
| 28504 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
changeset | 322 |   The @{executable_def "isabelle-process"} executable runs bare-bones
 | 
| 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
changeset | 323 | 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: 
28502diff
changeset | 324 | 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: 
28502diff
changeset | 325 | the actual heap file locations. Its usage is: | 
| 28215 | 326 | |
| 327 | \begin{ttbox}
 | |
| 28238 | 328 | Usage: isabelle-process [OPTIONS] [INPUT] [OUTPUT] | 
| 28215 | 329 | |
| 330 | Options are: | |
| 331 | -I startup Isar interaction mode | |
| 332 | -P startup Proof General interaction mode | |
| 333 | -S secure mode -- disallow critical operations | |
| 38253 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 wenzelm parents: 
36196diff
changeset | 334 | -W IN:OUT startup process wrapper, with input/output fifos | 
| 28215 | 335 | -X startup PGIP interaction mode | 
| 336 | -e MLTEXT pass MLTEXT to the ML session | |
| 337 | -f pass 'Session.finish();' to the ML session | |
| 338 | -m MODE add print mode for output | |
| 339 | -q non-interactive session | |
| 340 | -r open heap file read-only | |
| 341 | -u pass 'use"ROOT.ML";' to the ML session | |
| 342 | -w reset write permissions on OUTPUT | |
| 343 | ||
| 344 | INPUT (default "\$ISABELLE_LOGIC") and OUTPUT specify in/out heaps. | |
| 345 | These are either names to be searched in the Isabelle path, or | |
| 346 | actual file names (containing at least one /). | |
| 347 | If INPUT is "RAW_ML_SYSTEM", just start the bare bones ML system. | |
| 348 | \end{ttbox}
 | |
| 349 | ||
| 350 | Input files without path specifications are looked up in the | |
| 351 |   @{setting ISABELLE_PATH} setting, which may consist of multiple
 | |
| 352 | components separated by colons --- these are tried in the given | |
| 353 |   order with the value of @{setting ML_IDENTIFIER} appended
 | |
| 354 | internally. In a similar way, base names are relative to the | |
| 355 |   directory specified by @{setting ISABELLE_OUTPUT}.  In any case,
 | |
| 356 | actual file locations may also be given by including at least one | |
| 357 |   slash (@{verbatim "/"}) in the name (hint: use @{verbatim "./"} to
 | |
| 358 | refer to the current directory). | |
| 359 | *} | |
| 360 | ||
| 361 | ||
| 28223 | 362 | subsubsection {* Options *}
 | 
| 28215 | 363 | |
| 364 | text {*
 | |
| 365 | If the input heap file does not have write permission bits set, or | |
| 366 |   the @{verbatim "-r"} option is given explicitely, then the session
 | |
| 367 | started will be read-only. That is, the ML world cannot be | |
| 368 | committed back into the image file. Otherwise, a writable session | |
| 369 | enables commits into either the input file, or into another output | |
| 370 | heap file (if that is given as the second argument on the command | |
| 371 | line). | |
| 372 | ||
| 373 | The read-write state of sessions is determined at startup only, it | |
| 374 | cannot be changed intermediately. Also note that heap images may | |
| 375 | require considerable amounts of disk space (approximately | |
| 376 | 50--200~MB). Users are responsible for themselves to dispose their | |
| 377 | heap files when they are no longer needed. | |
| 378 | ||
| 379 |   \medskip The @{verbatim "-w"} option makes the output heap file
 | |
| 380 | read-only after terminating. Thus subsequent invocations cause the | |
| 381 | logic image to be read-only automatically. | |
| 382 | ||
| 383 |   \medskip Using the @{verbatim "-e"} option, arbitrary ML code may be
 | |
| 384 | passed to the Isabelle session from the command line. Multiple | |
| 385 |   @{verbatim "-e"}'s are evaluated in the given order. Strange things
 | |
| 386 | may happen when errorneous ML code is provided. Also make sure that | |
| 387 | the ML commands are terminated properly by semicolon. | |
| 388 | ||
| 389 |   \medskip The @{verbatim "-u"} option is a shortcut for @{verbatim
 | |
| 390 |   "-e"} passing ``@{verbatim "use \"ROOT.ML\";"}'' to the ML session.
 | |
| 28223 | 391 |   The @{verbatim "-f"} option passes ``@{verbatim
 | 
| 392 | "Session.finish();"}'', which is intended mainly for administrative | |
| 393 | purposes. | |
| 28215 | 394 | |
| 395 |   \medskip The @{verbatim "-m"} option adds identifiers of print modes
 | |
| 396 | to be made active for this session. Typically, this is used by some | |
| 397 | user interface, e.g.\ to enable output of proper mathematical | |
| 398 | symbols. | |
| 399 | ||
| 400 | \medskip Isabelle normally enters an interactive top-level loop | |
| 401 |   (after processing the @{verbatim "-e"} texts). The @{verbatim "-q"}
 | |
| 402 | option inhibits interaction, thus providing a pure batch mode | |
| 403 | facility. | |
| 404 | ||
| 405 |   \medskip The @{verbatim "-I"} option makes Isabelle enter Isar
 | |
| 406 | interaction mode on startup, instead of the primitive ML top-level. | |
| 407 |   The @{verbatim "-P"} option configures the top-level loop for
 | |
| 408 | interaction with the Proof General user interface, and the | |
| 38253 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 wenzelm parents: 
36196diff
changeset | 409 |   @{verbatim "-X"} option enables XML-based PGIP communication.
 | 
| 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 wenzelm parents: 
36196diff
changeset | 410 | |
| 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 wenzelm parents: 
36196diff
changeset | 411 |   \medskip The @{verbatim "-W"} option makes Isabelle enter a special
 | 
| 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 wenzelm parents: 
36196diff
changeset | 412 | process wrapper for interaction via the Isabelle/Scala layer, see | 
| 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 wenzelm parents: 
36196diff
changeset | 413 |   also @{"file" "~~/src/Pure/System/isabelle_process.scala"}.  The
 | 
| 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 wenzelm parents: 
36196diff
changeset | 414 | protocol between the ML and JVM process is private to the | 
| 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 wenzelm parents: 
36196diff
changeset | 415 | implementation. | 
| 28215 | 416 | |
| 417 |   \medskip The @{verbatim "-S"} option makes the Isabelle process more
 | |
| 418 | secure by disabling some critical operations, notably runtime | |
| 419 | compilation and evaluation of ML source code. | |
| 420 | *} | |
| 421 | ||
| 422 | ||
| 28223 | 423 | subsubsection {* Examples *}
 | 
| 28215 | 424 | |
| 425 | text {*
 | |
| 426 | Run an interactive session of the default object-logic (as specified | |
| 427 |   by the @{setting ISABELLE_LOGIC} setting) like this:
 | |
| 428 | \begin{ttbox}
 | |
| 28238 | 429 | isabelle-process | 
| 28215 | 430 | \end{ttbox}
 | 
| 431 | ||
| 432 |   Usually @{setting ISABELLE_LOGIC} refers to one of the standard
 | |
| 433 | logic images, which are read-only by default. A writable session | |
| 434 |   --- based on @{verbatim FOL}, but output to @{verbatim Foo} (in the
 | |
| 28238 | 435 |   directory specified by the @{setting ISABELLE_OUTPUT} setting) ---
 | 
| 28215 | 436 | may be invoked as follows: | 
| 437 | \begin{ttbox}
 | |
| 28238 | 438 | isabelle-process FOL Foo | 
| 28215 | 439 | \end{ttbox}
 | 
| 440 | Ending this session normally (e.g.\ by typing control-D) dumps the | |
| 441 |   whole ML system state into @{verbatim Foo}. Be prepared for several
 | |
| 442 | tens of megabytes. | |
| 443 | ||
| 444 |   The @{verbatim Foo} session may be continued later (still in
 | |
| 445 | writable state) by: | |
| 446 | \begin{ttbox}
 | |
| 28238 | 447 | isabelle-process Foo | 
| 28215 | 448 | \end{ttbox}
 | 
| 449 |   A read-only @{verbatim Foo} session may be started by:
 | |
| 450 | \begin{ttbox}
 | |
| 28238 | 451 | isabelle-process -r Foo | 
| 28215 | 452 | \end{ttbox}
 | 
| 453 | ||
| 454 | \medskip Note that manual session management like this does | |
| 455 |   \emph{not} provide proper setup for theory presentation.  This would
 | |
| 28238 | 456 |   require the @{tool usedir} utility.
 | 
| 28215 | 457 | |
| 28238 | 458 | \bigskip The next example demonstrates batch execution of Isabelle. | 
| 459 |   We retrieve the @{verbatim FOL} theory value from the theory loader
 | |
| 460 | within ML: | |
| 28215 | 461 | \begin{ttbox}
 | 
| 28238 | 462 | isabelle-process -e 'theory "FOL";' -q -r FOL | 
| 28215 | 463 | \end{ttbox}
 | 
| 464 | Note that the output text will be interspersed with additional junk | |
| 28238 | 465 |   messages by the ML runtime environment.  The @{verbatim "-W"} option
 | 
| 466 | allows to communicate with the Isabelle process via an external | |
| 467 | program in a more robust fashion. | |
| 468 | *} | |
| 469 | ||
| 470 | ||
| 28504 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
changeset | 471 | section {* The Isabelle tools wrapper \label{sec:isabelle-tool} *}
 | 
| 28238 | 472 | |
| 473 | text {*
 | |
| 474 | 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: 
28502diff
changeset | 475 |   wrapper --- @{executable isabelle}:
 | 
| 28238 | 476 | |
| 477 | \begin{ttbox}
 | |
| 28504 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
changeset | 478 | Usage: isabelle TOOL [ARGS ...] | 
| 28238 | 479 | |
| 28506 | 480 | Start Isabelle tool NAME with ARGS; pass "-?" for tool specific help. | 
| 28238 | 481 | |
| 482 | Available tools are: | |
| 483 | ||
| 484 | browser - Isabelle graph browser | |
| 485 | \dots | |
| 486 | \end{ttbox}
 | |
| 487 | ||
| 488 | In principle, Isabelle tools are ordinary executable scripts that | |
| 489 | are run within the Isabelle settings environment, see | |
| 490 |   \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: 
28502diff
changeset | 491 |   @{executable isabelle} from the directories listed in the @{setting
 | 
| 28238 | 492 | ISABELLE_TOOLS} setting. Do not try to call the scripts directly | 
| 493 | from the shell. Neither should you add the tool directories to your | |
| 494 | shell's search path! | |
| 495 | *} | |
| 496 | ||
| 497 | ||
| 498 | subsubsection {* Examples *}
 | |
| 499 | ||
| 500 | text {*
 | |
| 501 | Show the list of available documentation of the current Isabelle | |
| 502 | installation like this: | |
| 503 | ||
| 504 | \begin{ttbox}
 | |
| 28504 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
changeset | 505 | isabelle doc | 
| 28238 | 506 | \end{ttbox}
 | 
| 507 | ||
| 508 | View a certain document as follows: | |
| 509 | \begin{ttbox}
 | |
| 28504 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
changeset | 510 | isabelle doc isar-ref | 
| 28238 | 511 | \end{ttbox}
 | 
| 512 | ||
| 513 | Create an Isabelle session derived from HOL (see also | |
| 514 |   \secref{sec:tool-mkdir} and \secref{sec:tool-make}):
 | |
| 515 | \begin{ttbox}
 | |
| 28504 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
changeset | 516 | isabelle mkdir HOL Test && isabelle make | 
| 28238 | 517 | \end{ttbox}
 | 
| 28504 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
changeset | 518 |   Note that @{verbatim "isabelle mkdir"} is usually only invoked once;
 | 
| 28238 | 519 | existing sessions (including document output etc.) are then updated | 
| 28504 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
changeset | 520 |   by @{verbatim "isabelle make"} alone.
 | 
| 28215 | 521 | *} | 
| 522 | ||
| 28223 | 523 | end |