| author | wenzelm | 
| Tue, 25 Feb 2014 20:46:09 +0100 | |
| changeset 55748 | 2e1398b484aa | 
| parent 54937 | ce4bf91331e7 | 
| child 56439 | 95e2656b3b23 | 
| 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: 
41955diff
changeset | 2 | imports Base | 
| 28215 | 3 | begin | 
| 4 | ||
| 5 | chapter {* The Isabelle system environment *}
 | |
| 6 | ||
| 47823 | 7 | text {* This manual describes Isabelle together with related tools and
 | 
| 8 | user interfaces as seen from a system oriented view. See also the | |
| 28916 
0a802cdda340
removed obsolete isabelle-interface executable and ISABELLE_INTERFACE setting;
 wenzelm parents: 
28914diff
changeset | 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
 | |
| 12 |   Manual}~\cite{isabelle-implementation} for the main concepts of the
 | |
| 13 | underlying implementation in Isabelle/ML. | |
| 28215 | 14 | |
| 28916 
0a802cdda340
removed obsolete isabelle-interface executable and ISABELLE_INTERFACE setting;
 wenzelm parents: 
28914diff
changeset | 15 | \medskip The Isabelle system environment provides the following | 
| 
0a802cdda340
removed obsolete isabelle-interface executable and ISABELLE_INTERFACE setting;
 wenzelm parents: 
28914diff
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: 
28914diff
changeset | 20 |   \item The \emph{Isabelle settings} mechanism provides process
 | 
| 
0a802cdda340
removed obsolete isabelle-interface executable and ISABELLE_INTERFACE setting;
 wenzelm parents: 
28914diff
changeset | 21 | environment variables to all Isabelle executables (including tools | 
| 
0a802cdda340
removed obsolete isabelle-interface executable and ISABELLE_INTERFACE setting;
 wenzelm parents: 
28914diff
changeset | 22 | and user interfaces). | 
| 28215 | 23 | |
| 47823 | 24 |   \item The raw \emph{Isabelle process} (@{executable_ref
 | 
| 28504 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
changeset | 25 | "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 | 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: 
28502diff
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: 
28502diff
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}
 | 
| 28215 | 36 | *} | 
| 37 | ||
| 38 | ||
| 39 | section {* Isabelle settings \label{sec:settings} *}
 | |
| 40 | ||
| 41 | text {*
 | |
| 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: 
40569diff
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. | |
| 28215 | 62 | *} | 
| 63 | ||
| 64 | ||
| 32323 | 65 | subsection {* Bootstrapping the environment \label{sec:boot} *}
 | 
| 28215 | 66 | |
| 32323 | 67 | text {* Isabelle executables need to be run within a proper settings
 | 
| 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: 
40569diff
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: 
40569diff
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: 
38253diff
changeset | 101 |   before --- usually to something like @{verbatim
 | 
| 47661 
012a887997f3
USER_HOME settings variable points to cross-platform user home directory;
 wenzelm parents: 
45028diff
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
 | 
| 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 | 
| 48602 | 136 |   the @{tool getenv} tool.  This might help to figure out the effect
 | 
| 137 | of complex settings scripts. *} | |
| 28215 | 138 | |
| 139 | ||
| 48602 | 140 | subsection {* Common variables *}
 | 
| 28215 | 141 | |
| 142 | text {*
 | |
| 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: 
45028diff
changeset | 157 | |
| 
012a887997f3
USER_HOME settings variable points to cross-platform user home directory;
 wenzelm parents: 
45028diff
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: 
45028diff
changeset | 159 | top-level Isabelle distribution directory. This is automatically | 
| 
012a887997f3
USER_HOME settings variable points to cross-platform user home directory;
 wenzelm parents: 
45028diff
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: 
45028diff
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: 
45028diff
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: 
45028diff
changeset | 167 |   Typically, the @{setting ISABELLE_HOME_USER} directory mimics
 | 
| 
012a887997f3
USER_HOME settings variable points to cross-platform user home directory;
 wenzelm parents: 
45028diff
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: 
45028diff
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: 
38253diff
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: 
33952diff
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: 
33952diff
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: 
33952diff
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: 
33952diff
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: 
33952diff
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: 
33952diff
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: 
33952diff
changeset | 185 | |
| 
cbb9ee265cdd
added ISABELLE_PLATFORM and ISABELLE_PLATFORM64 -- NB: ML and JVM may have a different idea;
 wenzelm parents: 
33952diff
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: 
33952diff
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: 
33952diff
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: 
33952diff
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
 | 
| 28215 | 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: 
28502diff
changeset | 197 | isabelle} executables, respectively. Thus other tools and scripts | 
| 40800 
330eb65c9469
Parse.liberal_name for document antiquotations and attributes;
 wenzelm parents: 
40569diff
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: 
40569diff
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 | |
| 51434 | 220 |   \item[@{setting_def ISABELLE_POLYML}@{text "\<^sup>*"}] is @{verbatim true}
 | 
| 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 | ||
| 252 |   \item[@{setting_def ISABELLE_LINE_EDITOR}] specifies the default
 | |
| 28238 | 253 |   line editor for the @{tool_ref tty} 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: 
28502diff
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: 
28502diff
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: 
50182diff
changeset | 266 | |
| 54683 
cf48ddc266e5
clarified "isabelle display" and 'display_drafts': re-use file and program instance, open asynchronously via desktop environment;
 wenzelm parents: 
52746diff
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: 
52746diff
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: 
52746diff
changeset | 269 | |
| 
cf48ddc266e5
clarified "isabelle display" and 'display_drafts': re-use file and program instance, open asynchronously via desktop environment;
 wenzelm parents: 
52746diff
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: 
52746diff
changeset | 271 |   for displaying @{verbatim dvi} files.
 | 
| 28215 | 272 | |
| 273 |   \item[@{setting_def ISABELLE_TMP_PREFIX}@{text "\<^sup>*"}] is the
 | |
| 28238 | 274 |   prefix from which any running @{executable "isabelle-process"}
 | 
| 275 | derives an individual directory for temporary files. The default is | |
| 54705 | 276 |   somewhere in @{file_unchecked "/tmp"}.
 | 
| 28215 | 277 | |
| 278 |   \end{description}
 | |
| 279 | *} | |
| 280 | ||
| 281 | ||
| 32323 | 282 | subsection {* Additional components \label{sec:components} *}
 | 
| 283 | ||
| 284 | text {* Any directory may be registered as an explicit \emph{Isabelle
 | |
| 285 | component}. The general layout conventions are that of the main | |
| 286 | Isabelle distribution itself, and the following two files (both | |
| 287 | optional) have a special meaning: | |
| 288 | ||
| 289 |   \begin{itemize}
 | |
| 290 | ||
| 291 |   \item @{verbatim "etc/settings"} holds additional settings that are
 | |
| 292 | initialized when bootstrapping the overall Isabelle environment, | |
| 293 |   cf.\ \secref{sec:boot}.  As usual, the content is interpreted as a
 | |
| 294 |   @{verbatim bash} script.  It may refer to the component's enclosing
 | |
| 295 |   directory via the @{verbatim "COMPONENT"} shell variable.
 | |
| 296 | ||
| 297 | For example, the following setting allows to refer to files within | |
| 298 | the component later on, without having to hardwire absolute paths: | |
| 299 | ||
| 48838 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 300 | \begin{ttbox}
 | 
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 301 | MY_COMPONENT_HOME="$COMPONENT" | 
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 302 | \end{ttbox}
 | 
| 32323 | 303 | |
| 304 | Components can also add to existing Isabelle settings such as | |
| 305 |   @{setting_def ISABELLE_TOOLS}, in order to provide
 | |
| 306 | component-specific tools that can be invoked by end-users. For | |
| 307 | example: | |
| 308 | ||
| 48838 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 309 | \begin{ttbox}
 | 
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 310 | ISABELLE_TOOLS="$ISABELLE_TOOLS:$COMPONENT/lib/Tools" | 
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 311 | \end{ttbox}
 | 
| 32323 | 312 | |
| 313 |   \item @{verbatim "etc/components"} holds a list of further
 | |
| 314 | sub-components of the same structure. The directory specifications | |
| 315 |   given here can be either absolute (with leading @{verbatim "/"}) or
 | |
| 316 | relative to the component's main directory. | |
| 317 | ||
| 318 |   \end{itemize}
 | |
| 319 | ||
| 320 |   The root of component initialization is @{setting ISABELLE_HOME}
 | |
| 321 | itself. After initializing all of its sub-components recursively, | |
| 322 |   @{setting ISABELLE_HOME_USER} is included in the same manner (if
 | |
| 40569 
ffcff7509a49
more explicit explanation of init_component shell function;
 wenzelm parents: 
40387diff
changeset | 323 | that directory exists). This allows to install private components | 
| 54705 | 324 |   via @{file_unchecked "$ISABELLE_HOME_USER/etc/components"}, although it is
 | 
| 40569 
ffcff7509a49
more explicit explanation of init_component shell function;
 wenzelm parents: 
40387diff
changeset | 325 | often more convenient to do that programmatically via the | 
| 
ffcff7509a49
more explicit explanation of init_component shell function;
 wenzelm parents: 
40387diff
changeset | 326 | \verb,init_component, shell function in the \verb,etc/settings, | 
| 
ffcff7509a49
more explicit explanation of init_component shell function;
 wenzelm parents: 
40387diff
changeset | 327 | script of \verb,$ISABELLE_HOME_USER, (or any other component | 
| 
ffcff7509a49
more explicit explanation of init_component shell function;
 wenzelm parents: 
40387diff
changeset | 328 | directory). For example: | 
| 48838 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 329 | \begin{ttbox}
 | 
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 330 | init_component "$HOME/screwdriver-2.0" | 
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 331 | \end{ttbox}
 | 
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 332 | |
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 333 | This is tolerant wrt.\ missing component directories, but might | 
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 334 | produce a warning. | 
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 335 | |
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 336 | \medskip More complex situations may be addressed by initializing | 
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 337 | components listed in a given catalog file, relatively to some base | 
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 338 | directory: | 
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 339 | |
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 340 | \begin{ttbox}
 | 
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 341 | init_components "$HOME/my_component_store" "some_catalog_file" | 
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 342 | \end{ttbox}
 | 
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 343 | |
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 344 | The component directories listed in the catalog file are treated as | 
| 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 345 | relative to the given base directory. | 
| 48844 | 346 | |
| 347 |   See also \secref{sec:tool-components} for some tool-support for
 | |
| 348 | resolving components that are formally initialized but not installed | |
| 349 | yet. | |
| 32323 | 350 | *} | 
| 351 | ||
| 352 | ||
| 28215 | 353 | section {* The raw Isabelle process *}
 | 
| 354 | ||
| 355 | text {*
 | |
| 28504 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
changeset | 356 |   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 | 357 | 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 | 358 | 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 | 359 | the actual heap file locations. Its usage is: | 
| 28215 | 360 | |
| 361 | \begin{ttbox}
 | |
| 28238 | 362 | Usage: isabelle-process [OPTIONS] [INPUT] [OUTPUT] | 
| 28215 | 363 | |
| 364 | Options are: | |
| 365 | -I startup Isar interaction mode | |
| 366 | -P startup Proof General interaction mode | |
| 367 | -S secure mode -- disallow critical operations | |
| 45028 | 368 | -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: 
36196diff
changeset | 369 | -W IN:OUT startup process wrapper, with input/output fifos | 
| 28215 | 370 | -e MLTEXT pass MLTEXT to the ML session | 
| 371 | -m MODE add print mode for output | |
| 52056 | 372 | -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) | 
| 28215 | 373 | -q non-interactive session | 
| 374 | -r open heap file read-only | |
| 375 | -w reset write permissions on OUTPUT | |
| 376 | ||
| 377 | INPUT (default "\$ISABELLE_LOGIC") and OUTPUT specify in/out heaps. | |
| 378 | These are either names to be searched in the Isabelle path, or | |
| 379 | actual file names (containing at least one /). | |
| 380 | If INPUT is "RAW_ML_SYSTEM", just start the bare bones ML system. | |
| 381 | \end{ttbox}
 | |
| 382 | ||
| 383 | Input files without path specifications are looked up in the | |
| 384 |   @{setting ISABELLE_PATH} setting, which may consist of multiple
 | |
| 385 | components separated by colons --- these are tried in the given | |
| 386 |   order with the value of @{setting ML_IDENTIFIER} appended
 | |
| 387 | internally. In a similar way, base names are relative to the | |
| 388 |   directory specified by @{setting ISABELLE_OUTPUT}.  In any case,
 | |
| 389 | actual file locations may also be given by including at least one | |
| 390 |   slash (@{verbatim "/"}) in the name (hint: use @{verbatim "./"} to
 | |
| 391 | refer to the current directory). | |
| 392 | *} | |
| 393 | ||
| 394 | ||
| 28223 | 395 | subsubsection {* Options *}
 | 
| 28215 | 396 | |
| 397 | text {*
 | |
| 398 | If the input heap file does not have write permission bits set, or | |
| 399 |   the @{verbatim "-r"} option is given explicitely, then the session
 | |
| 400 | started will be read-only. That is, the ML world cannot be | |
| 401 | committed back into the image file. Otherwise, a writable session | |
| 402 | enables commits into either the input file, or into another output | |
| 403 | heap file (if that is given as the second argument on the command | |
| 404 | line). | |
| 405 | ||
| 406 | The read-write state of sessions is determined at startup only, it | |
| 407 | cannot be changed intermediately. Also note that heap images may | |
| 408 | require considerable amounts of disk space (approximately | |
| 409 | 50--200~MB). Users are responsible for themselves to dispose their | |
| 410 | heap files when they are no longer needed. | |
| 411 | ||
| 412 |   \medskip The @{verbatim "-w"} option makes the output heap file
 | |
| 413 | read-only after terminating. Thus subsequent invocations cause the | |
| 414 | logic image to be read-only automatically. | |
| 415 | ||
| 416 |   \medskip Using the @{verbatim "-e"} option, arbitrary ML code may be
 | |
| 417 | passed to the Isabelle session from the command line. Multiple | |
| 418 |   @{verbatim "-e"}'s are evaluated in the given order. Strange things
 | |
| 419 | may happen when errorneous ML code is provided. Also make sure that | |
| 420 | the ML commands are terminated properly by semicolon. | |
| 421 | ||
| 422 |   \medskip The @{verbatim "-m"} option adds identifiers of print modes
 | |
| 423 | to be made active for this session. Typically, this is used by some | |
| 424 | user interface, e.g.\ to enable output of proper mathematical | |
| 425 | symbols. | |
| 426 | ||
| 427 | \medskip Isabelle normally enters an interactive top-level loop | |
| 428 |   (after processing the @{verbatim "-e"} texts). The @{verbatim "-q"}
 | |
| 429 | option inhibits interaction, thus providing a pure batch mode | |
| 430 | facility. | |
| 431 | ||
| 52061 | 432 |   \medskip Option @{verbatim "-o"} allows to override Isabelle system
 | 
| 52056 | 433 |   options for this process, see also \secref{sec:system-options}.
 | 
| 434 | ||
| 28215 | 435 |   \medskip The @{verbatim "-I"} option makes Isabelle enter Isar
 | 
| 436 | interaction mode on startup, instead of the primitive ML top-level. | |
| 437 |   The @{verbatim "-P"} option configures the top-level loop for
 | |
| 51932 | 438 | interaction with the Proof General user interface. | 
| 38253 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 wenzelm parents: 
36196diff
changeset | 439 | |
| 45028 | 440 |   \medskip The @{verbatim "-T"} or @{verbatim "-W"} option makes
 | 
| 49173 
fa01a202399c
eliminated potentially confusing terminology of Scala "layer";
 wenzelm parents: 
48985diff
changeset | 441 | Isabelle enter a special process wrapper for interaction via | 
| 
fa01a202399c
eliminated potentially confusing terminology of Scala "layer";
 wenzelm parents: 
48985diff
changeset | 442 |   Isabelle/Scala, see also @{file
 | 
| 45028 | 443 | "~~/src/Pure/System/isabelle_process.scala"}. The protocol between | 
| 444 | the ML and JVM process is private to the implementation. | |
| 28215 | 445 | |
| 446 |   \medskip The @{verbatim "-S"} option makes the Isabelle process more
 | |
| 447 | secure by disabling some critical operations, notably runtime | |
| 448 | compilation and evaluation of ML source code. | |
| 449 | *} | |
| 450 | ||
| 451 | ||
| 28223 | 452 | subsubsection {* Examples *}
 | 
| 28215 | 453 | |
| 454 | text {*
 | |
| 455 | Run an interactive session of the default object-logic (as specified | |
| 456 |   by the @{setting ISABELLE_LOGIC} setting) like this:
 | |
| 457 | \begin{ttbox}
 | |
| 28238 | 458 | isabelle-process | 
| 28215 | 459 | \end{ttbox}
 | 
| 460 | ||
| 461 |   Usually @{setting ISABELLE_LOGIC} refers to one of the standard
 | |
| 462 | logic images, which are read-only by default. A writable session | |
| 47823 | 463 |   --- based on @{verbatim HOL}, but output to @{verbatim Test} (in the
 | 
| 28238 | 464 |   directory specified by the @{setting ISABELLE_OUTPUT} setting) ---
 | 
| 28215 | 465 | may be invoked as follows: | 
| 466 | \begin{ttbox}
 | |
| 47823 | 467 | isabelle-process HOL Test | 
| 28215 | 468 | \end{ttbox}
 | 
| 469 | Ending this session normally (e.g.\ by typing control-D) dumps the | |
| 47823 | 470 |   whole ML system state into @{verbatim Test} (be prepared for more
 | 
| 471 | than 100\,MB): | |
| 28215 | 472 | |
| 47823 | 473 |   The @{verbatim Test} session may be continued later (still in
 | 
| 28215 | 474 | writable state) by: | 
| 475 | \begin{ttbox}
 | |
| 47823 | 476 | isabelle-process Test | 
| 28215 | 477 | \end{ttbox}
 | 
| 47823 | 478 |   A read-only @{verbatim Test} session may be started by:
 | 
| 28215 | 479 | \begin{ttbox}
 | 
| 47823 | 480 | isabelle-process -r Test | 
| 28215 | 481 | \end{ttbox}
 | 
| 482 | ||
| 28238 | 483 | \bigskip The next example demonstrates batch execution of Isabelle. | 
| 47823 | 484 |   We retrieve the @{verbatim Main} theory value from the theory loader
 | 
| 485 | within ML (observe the delicate quoting rules for the Bash shell | |
| 486 | vs.\ ML): | |
| 28215 | 487 | \begin{ttbox}
 | 
| 47823 | 488 | isabelle-process -e 'Thy_Info.get_theory "Main";' -q -r HOL | 
| 28215 | 489 | \end{ttbox}
 | 
| 490 | Note that the output text will be interspersed with additional junk | |
| 28238 | 491 |   messages by the ML runtime environment.  The @{verbatim "-W"} option
 | 
| 492 | allows to communicate with the Isabelle process via an external | |
| 493 | program in a more robust fashion. | |
| 494 | *} | |
| 495 | ||
| 496 | ||
| 48813 | 497 | section {* The Isabelle tool wrapper \label{sec:isabelle-tool} *}
 | 
| 28238 | 498 | |
| 499 | text {*
 | |
| 500 | 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 | 501 |   wrapper --- @{executable isabelle}:
 | 
| 28238 | 502 | |
| 503 | \begin{ttbox}
 | |
| 28504 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
changeset | 504 | Usage: isabelle TOOL [ARGS ...] | 
| 28238 | 505 | |
| 28506 | 506 | Start Isabelle tool NAME with ARGS; pass "-?" for tool specific help. | 
| 28238 | 507 | |
| 48858 | 508 | Available tools: | 
| 509 | \dots | |
| 28238 | 510 | \end{ttbox}
 | 
| 511 | ||
| 512 | In principle, Isabelle tools are ordinary executable scripts that | |
| 513 | are run within the Isabelle settings environment, see | |
| 514 |   \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 | 515 |   @{executable isabelle} from the directories listed in the @{setting
 | 
| 28238 | 516 | ISABELLE_TOOLS} setting. Do not try to call the scripts directly | 
| 517 | from the shell. Neither should you add the tool directories to your | |
| 518 | shell's search path! | |
| 519 | *} | |
| 520 | ||
| 521 | ||
| 522 | subsubsection {* Examples *}
 | |
| 523 | ||
| 48813 | 524 | text {* Show the list of available documentation of the Isabelle
 | 
| 525 | distribution: | |
| 28238 | 526 | |
| 527 | \begin{ttbox}
 | |
| 28504 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
28502diff
changeset | 528 | isabelle doc | 
| 28238 | 529 | \end{ttbox}
 | 
| 530 | ||
| 531 | View a certain document as follows: | |
| 532 | \begin{ttbox}
 | |
| 47823 | 533 | isabelle doc system | 
| 28238 | 534 | \end{ttbox}
 | 
| 535 | ||
| 48813 | 536 | Query the Isabelle settings environment: | 
| 28238 | 537 | \begin{ttbox}
 | 
| 48813 | 538 | isabelle getenv ISABELLE_HOME_USER | 
| 28238 | 539 | \end{ttbox}
 | 
| 28215 | 540 | *} | 
| 541 | ||
| 28223 | 542 | end |