| author | wenzelm | 
| Fri, 28 Apr 2017 18:24:58 +0200 | |
| changeset 65617 | 823bbc467dfa | 
| parent 64509 | 80aaa4ff7fed | 
| child 66732 | e566fb4d43d4 | 
| child 66785 | 6fbd7fc824a9 | 
| permissions | -rw-r--r-- | 
| 62643 | 1 | (*:maxLineLen=78:*) | 
| 61575 | 2 | |
| 62640 | 3 | theory Environment | 
| 43564 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
41955diff
changeset | 4 | imports Base | 
| 28215 | 5 | begin | 
| 6 | ||
| 58618 | 7 | chapter \<open>The Isabelle system environment\<close> | 
| 28215 | 8 | |
| 61575 | 9 | text \<open> | 
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 10 | This manual describes Isabelle together with related tools as seen from a | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 11 |   system oriented view. See also the \<^emph>\<open>Isabelle/Isar Reference Manual\<close> @{cite
 | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 12 | "isabelle-isar-ref"} for the actual Isabelle input language and related | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 13 |   concepts, and \<^emph>\<open>The Isabelle/Isar Implementation Manual\<close> @{cite
 | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 14 | "isabelle-implementation"} for the main concepts of the underlying | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 15 | implementation in Isabelle/ML. | 
| 58618 | 16 | \<close> | 
| 28215 | 17 | |
| 18 | ||
| 58618 | 19 | section \<open>Isabelle settings \label{sec:settings}\<close>
 | 
| 28215 | 20 | |
| 58618 | 21 | text \<open> | 
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 22 | Isabelle executables may depend on the \<^emph>\<open>Isabelle settings\<close> within the | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 23 | process environment. This is a statically scoped collection of environment | 
| 62013 | 24 |   variables, such as @{setting ISABELLE_HOME}, @{setting ML_SYSTEM}, @{setting
 | 
| 25 | ML_HOME}. These variables are \<^emph>\<open>not\<close> intended to be set directly from the | |
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 26 | shell, but are provided by Isabelle \<^emph>\<open>components\<close> their \<^emph>\<open>settings files\<close> as | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 27 | explained below. | 
| 58618 | 28 | \<close> | 
| 28215 | 29 | |
| 30 | ||
| 58618 | 31 | subsection \<open>Bootstrapping the environment \label{sec:boot}\<close>
 | 
| 28215 | 32 | |
| 61575 | 33 | text \<open> | 
| 34 | Isabelle executables need to be run within a proper settings environment. | |
| 35 | This is bootstrapped as described below, on the first invocation of one of | |
| 36 |   the outer wrapper scripts (such as @{executable_ref isabelle}). This happens
 | |
| 37 | only once for each process tree, i.e.\ the environment is passed to | |
| 38 | subprocesses according to regular Unix conventions. | |
| 39 | ||
| 40 |     \<^enum> The special variable @{setting_def ISABELLE_HOME} is determined
 | |
| 41 | automatically from the location of the binary that has been run. | |
| 28215 | 42 | |
| 61575 | 43 |     You should not try to set @{setting ISABELLE_HOME} manually. Also note
 | 
| 44 | that the Isabelle executables either have to be run from their original | |
| 45 | location in the distribution directory, or via the executable objects | |
| 46 |     created by the @{tool install} tool. Symbolic links are admissible, but a
 | |
| 63680 | 47 | plain copy of the \<^dir>\<open>$ISABELLE_HOME/bin\<close> files will not work! | 
| 61575 | 48 | |
| 63680 | 49 |     \<^enum> The file \<^file>\<open>$ISABELLE_HOME/etc/settings\<close> is run as a @{executable_ref
 | 
| 50 | bash} shell script with the auto-export option for variables enabled. | |
| 28238 | 51 | |
| 61575 | 52 | This file holds a rather long list of shell variable assignments, thus | 
| 53 | providing the site-wide default settings. The Isabelle distribution | |
| 54 | already contains a global settings file with sensible defaults for most | |
| 55 | variables. When installing the system, only a few of these may have to be | |
| 56 |     adapted (probably @{setting ML_SYSTEM} etc.).
 | |
| 57 | ||
| 63669 | 58 |     \<^enum> The file @{path "$ISABELLE_HOME_USER/etc/settings"} (if it
 | 
| 61575 | 59 | exists) is run in the same way as the site default settings. Note that the | 
| 60 |     variable @{setting ISABELLE_HOME_USER} has already been set before ---
 | |
| 61 | usually to something like \<^verbatim>\<open>$USER_HOME/.isabelle/IsabelleXXXX\<close>. | |
| 61458 | 62 | |
| 61575 | 63 | Thus individual users may override the site-wide defaults. Typically, a | 
| 64 | user settings file contains only a few lines, with some assignments that | |
| 63680 | 65 | are actually changed. Never copy the central | 
| 66 | \<^file>\<open>$ISABELLE_HOME/etc/settings\<close> file! | |
| 28215 | 67 | |
| 61575 | 68 |   Since settings files are regular GNU @{executable_def bash} scripts, one may
 | 
| 69 | use complex shell commands, such as \<^verbatim>\<open>if\<close> or \<^verbatim>\<open>case\<close> statements to set | |
| 70 | variables depending on the system architecture or other environment | |
| 71 | variables. Such advanced features should be added only with great care, | |
| 72 | though. In particular, external environment references should be kept at a | |
| 28215 | 73 | minimum. | 
| 74 | ||
| 61406 | 75 | \<^medskip> | 
| 76 | A few variables are somewhat special: | |
| 28215 | 77 | |
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 78 |     \<^item> @{setting_def ISABELLE_TOOL} is set automatically to the absolute path
 | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 79 |     name of the @{executable isabelle} executables.
 | 
| 28215 | 80 | |
| 61575 | 81 |     \<^item> @{setting_ref ISABELLE_OUTPUT} will have the identifiers of the Isabelle
 | 
| 82 |     distribution (cf.\ @{setting ISABELLE_IDENTIFIER}) and the ML system (cf.\
 | |
| 83 |     @{setting ML_IDENTIFIER}) appended automatically to its value.
 | |
| 28215 | 84 | |
| 61406 | 85 | \<^medskip> | 
| 61575 | 86 |   Note that the settings environment may be inspected with the @{tool getenv}
 | 
| 87 | tool. This might help to figure out the effect of complex settings scripts. | |
| 88 | \<close> | |
| 28215 | 89 | |
| 90 | ||
| 58618 | 91 | subsection \<open>Common variables\<close> | 
| 28215 | 92 | |
| 58618 | 93 | text \<open> | 
| 61575 | 94 | This is a reference of common Isabelle settings variables. Note that the | 
| 95 | list is somewhat open-ended. Third-party utilities or interfaces may add | |
| 96 | their own selection. Variables that are special in some sense are marked | |
| 97 | with \<open>\<^sup>*\<close>. | |
| 28215 | 98 | |
| 61575 | 99 |   \<^descr>[@{setting_def USER_HOME}\<open>\<^sup>*\<close>] Is the cross-platform user home directory.
 | 
| 100 |   On Unix systems this is usually the same as @{setting HOME}, but on Windows
 | |
| 101 | it is the regular home directory of the user, not the one of within the | |
| 102 | Cygwin root file-system.\<^footnote>\<open>Cygwin itself offers another choice whether its | |
| 63669 | 103 |   HOME should point to the @{path "/home"} directory tree or the Windows user
 | 
| 104 | home.\<close> | |
| 47661 
012a887997f3
USER_HOME settings variable points to cross-platform user home directory;
 wenzelm parents: 
45028diff
changeset | 105 | |
| 61575 | 106 |   \<^descr>[@{setting_def ISABELLE_HOME}\<open>\<^sup>*\<close>] is the location of the top-level
 | 
| 107 | Isabelle distribution directory. This is automatically determined from the | |
| 108 |   Isabelle executable that has been invoked. Do not attempt to set @{setting
 | |
| 109 | ISABELLE_HOME} yourself from the shell! | |
| 50182 | 110 | |
| 61575 | 111 |   \<^descr>[@{setting_def ISABELLE_HOME_USER}] is the user-specific counterpart of
 | 
| 63669 | 112 |   @{setting ISABELLE_HOME}. The default value is relative to @{path
 | 
| 61575 | 113 | "$USER_HOME/.isabelle"}, under rare circumstances this may be changed in the | 
| 114 |   global setting file. Typically, the @{setting ISABELLE_HOME_USER} directory
 | |
| 115 |   mimics @{setting ISABELLE_HOME} to some extend. In particular, site-wide
 | |
| 116 | defaults may be overridden by a private \<^verbatim>\<open>$ISABELLE_HOME_USER/etc/settings\<close>. | |
| 117 | ||
| 118 |   \<^descr>[@{setting_def ISABELLE_PLATFORM_FAMILY}\<open>\<^sup>*\<close>] is automatically set to the
 | |
| 119 | general platform family: \<^verbatim>\<open>linux\<close>, \<^verbatim>\<open>macos\<close>, \<^verbatim>\<open>windows\<close>. Note that | |
| 50182 | 120 | platform-dependent tools usually need to refer to the more specific | 
| 121 |   identification according to @{setting ISABELLE_PLATFORM}, @{setting
 | |
| 122 |   ISABELLE_PLATFORM32}, @{setting ISABELLE_PLATFORM64}.
 | |
| 123 | ||
| 61575 | 124 |   \<^descr>[@{setting_def ISABELLE_PLATFORM}\<open>\<^sup>*\<close>] is automatically set to a symbolic
 | 
| 125 | identifier for the underlying hardware and operating system. The Isabelle | |
| 126 | platform identification always refers to the 32 bit variant, even this is a | |
| 127 | 64 bit machine. Note that the ML or Java runtime may have a different idea, | |
| 128 | depending on which binaries are actually run. | |
| 36196 
cbb9ee265cdd
added ISABELLE_PLATFORM and ISABELLE_PLATFORM64 -- NB: ML and JVM may have a different idea;
 wenzelm parents: 
33952diff
changeset | 129 | |
| 61575 | 130 |   \<^descr>[@{setting_def ISABELLE_PLATFORM64}\<open>\<^sup>*\<close>] is similar to @{setting
 | 
| 131 | ISABELLE_PLATFORM} but refers to the proper 64 bit variant on a platform | |
| 132 | that supports this; the value is empty for 32 bit. Note that the following | |
| 133 | bash expression (including the quotes) prefers the 64 bit platform, if that | |
| 134 | is available: | |
| 47823 | 135 | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 136 |   @{verbatim [display] \<open>"${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM}"\<close>}
 | 
| 36196 
cbb9ee265cdd
added ISABELLE_PLATFORM and ISABELLE_PLATFORM64 -- NB: ML and JVM may have a different idea;
 wenzelm parents: 
33952diff
changeset | 137 | |
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 138 |   \<^descr>[@{setting ISABELLE_TOOL}\<open>\<^sup>*\<close>] is automatically set to the full path name
 | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 139 |   of the @{executable isabelle} executable. Thus other tools and scripts need
 | 
| 63680 | 140 | not assume that the \<^dir>\<open>$ISABELLE_HOME/bin\<close> directory is on the current | 
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 141 | search path of the shell. | 
| 28215 | 142 | |
| 61575 | 143 |   \<^descr>[@{setting_def ISABELLE_IDENTIFIER}\<open>\<^sup>*\<close>] refers to the name of this
 | 
| 144 | Isabelle distribution, e.g.\ ``\<^verbatim>\<open>Isabelle2012\<close>''. | |
| 145 | ||
| 146 |   \<^descr>[@{setting_def ML_SYSTEM}, @{setting_def ML_HOME}, @{setting_def
 | |
| 147 |   ML_OPTIONS}, @{setting_def ML_PLATFORM}, @{setting_def ML_IDENTIFIER}\<open>\<^sup>*\<close>]
 | |
| 148 | specify the underlying ML system to be used for Isabelle. There is only a | |
| 63680 | 149 |   fixed set of admissable @{setting ML_SYSTEM} names (see the
 | 
| 150 | \<^file>\<open>$ISABELLE_HOME/etc/settings\<close> file of the distribution). | |
| 61575 | 151 | |
| 28215 | 152 |   The actual compiler binary will be run from the directory @{setting
 | 
| 61575 | 153 |   ML_HOME}, with @{setting ML_OPTIONS} as first arguments on the command line.
 | 
| 154 |   The optional @{setting ML_PLATFORM} may specify the binary format of ML heap
 | |
| 155 | images, which is useful for cross-platform installations. The value of | |
| 156 |   @{setting ML_IDENTIFIER} is automatically obtained by composing the values
 | |
| 157 |   of @{setting ML_SYSTEM}, @{setting ML_PLATFORM} and the Isabelle version
 | |
| 158 | values. | |
| 47823 | 159 | |
| 61575 | 160 |   \<^descr>[@{setting_def ISABELLE_JDK_HOME}] needs to point to a full JDK (Java
 | 
| 161 | Development Kit) installation with \<^verbatim>\<open>javac\<close> and \<^verbatim>\<open>jar\<close> executables. This is | |
| 162 | essential for Isabelle/Scala and other JVM-based tools to work properly. | |
| 163 | Note that conventional \<^verbatim>\<open>JAVA_HOME\<close> usually points to the JRE (Java Runtime | |
| 47823 | 164 | Environment), not JDK. | 
| 61575 | 165 | |
| 166 |   \<^descr>[@{setting_def ISABELLE_PATH}] is a list of directories (separated by
 | |
| 167 | colons) where Isabelle logic images may reside. When looking up heaps files, | |
| 168 |   the value of @{setting ML_IDENTIFIER} is appended to each component
 | |
| 169 | internally. | |
| 170 | ||
| 171 |   \<^descr>[@{setting_def ISABELLE_OUTPUT}\<open>\<^sup>*\<close>] is a directory where output heap files
 | |
| 172 | should be stored by default. The ML system and Isabelle version identifier | |
| 173 | is appended here, too. | |
| 174 | ||
| 175 |   \<^descr>[@{setting_def ISABELLE_BROWSER_INFO}] is the directory where theory
 | |
| 62013 | 176 |   browser information is stored as HTML and PDF (see also \secref{sec:info}).
 | 
| 63669 | 177 |   The default value is @{path "$ISABELLE_HOME_USER/browser_info"}.
 | 
| 61575 | 178 | |
| 179 |   \<^descr>[@{setting_def ISABELLE_LOGIC}] specifies the default logic to load if none
 | |
| 180 | is given explicitely by the user. The default value is \<^verbatim>\<open>HOL\<close>. | |
| 181 | ||
| 62562 
905a5db3932d
back to external line editor, due to problems of JLine with multithreading of in vs. out;
 wenzelm parents: 
62559diff
changeset | 182 |   \<^descr>[@{setting_def ISABELLE_LINE_EDITOR}] specifies the line editor for the
 | 
| 
905a5db3932d
back to external line editor, due to problems of JLine with multithreading of in vs. out;
 wenzelm parents: 
62559diff
changeset | 183 |   @{tool_ref console} interface.
 | 
| 
905a5db3932d
back to external line editor, due to problems of JLine with multithreading of in vs. out;
 wenzelm parents: 
62559diff
changeset | 184 | |
| 61575 | 185 |   \<^descr>[@{setting_def ISABELLE_LATEX}, @{setting_def ISABELLE_PDFLATEX},
 | 
| 186 |   @{setting_def ISABELLE_BIBTEX}] refer to {\LaTeX} related tools for Isabelle
 | |
| 187 |   document preparation (see also \secref{sec:tool-latex}).
 | |
| 188 | ||
| 189 |   \<^descr>[@{setting_def ISABELLE_TOOLS}] is a colon separated list of directories
 | |
| 190 |   that are scanned by @{executable isabelle} for external utility programs
 | |
| 191 |   (see also \secref{sec:isabelle-tool}).
 | |
| 50197 
b385d134926d
eval PDF_VIEWER/DVI_VIEWER command line, which allows additional quotes for program name, for example;
 wenzelm parents: 
50182diff
changeset | 192 | |
| 61575 | 193 |   \<^descr>[@{setting_def ISABELLE_DOCS}] is a colon separated list of directories
 | 
| 194 | with documentation files. | |
| 195 | ||
| 196 |   \<^descr>[@{setting_def PDF_VIEWER}] specifies the program to be used for displaying
 | |
| 197 | \<^verbatim>\<open>pdf\<close> files. | |
| 54683 
cf48ddc266e5
clarified "isabelle display" and 'display_drafts': re-use file and program instance, open asynchronously via desktop environment;
 wenzelm parents: 
52746diff
changeset | 198 | |
| 61575 | 199 |   \<^descr>[@{setting_def DVI_VIEWER}] specifies the program to be used for displaying
 | 
| 200 | \<^verbatim>\<open>dvi\<close> files. | |
| 201 | ||
| 202 |   \<^descr>[@{setting_def ISABELLE_TMP_PREFIX}\<open>\<^sup>*\<close>] is the prefix from which any
 | |
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 203 | running Isabelle ML process derives an individual directory for temporary | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 204 | files. | 
| 58618 | 205 | \<close> | 
| 28215 | 206 | |
| 207 | ||
| 58618 | 208 | subsection \<open>Additional components \label{sec:components}\<close>
 | 
| 32323 | 209 | |
| 61575 | 210 | text \<open> | 
| 211 | Any directory may be registered as an explicit \<^emph>\<open>Isabelle component\<close>. The | |
| 212 | general layout conventions are that of the main Isabelle distribution | |
| 213 | itself, and the following two files (both optional) have a special meaning: | |
| 32323 | 214 | |
| 61575 | 215 | \<^item> \<^verbatim>\<open>etc/settings\<close> holds additional settings that are initialized when | 
| 216 |     bootstrapping the overall Isabelle environment, cf.\ \secref{sec:boot}. As
 | |
| 64509 | 217 | usual, the content is interpreted as a GNU bash script. It may refer to | 
| 218 | the component's enclosing directory via the \<^verbatim>\<open>COMPONENT\<close> shell variable. | |
| 32323 | 219 | |
| 61575 | 220 | For example, the following setting allows to refer to files within the | 
| 221 | component later on, without having to hardwire absolute paths: | |
| 222 |     @{verbatim [display] \<open>MY_COMPONENT_HOME="$COMPONENT"\<close>}
 | |
| 32323 | 223 | |
| 61575 | 224 | Components can also add to existing Isabelle settings such as | 
| 225 |     @{setting_def ISABELLE_TOOLS}, in order to provide component-specific
 | |
| 226 | tools that can be invoked by end-users. For example: | |
| 227 |     @{verbatim [display] \<open>ISABELLE_TOOLS="$ISABELLE_TOOLS:$COMPONENT/lib/Tools"\<close>}
 | |
| 32323 | 228 | |
| 61575 | 229 | \<^item> \<^verbatim>\<open>etc/components\<close> holds a list of further sub-components of the same | 
| 230 | structure. The directory specifications given here can be either absolute | |
| 231 | (with leading \<^verbatim>\<open>/\<close>) or relative to the component's main directory. | |
| 32323 | 232 | |
| 61575 | 233 |   The root of component initialization is @{setting ISABELLE_HOME} itself.
 | 
| 234 |   After initializing all of its sub-components recursively, @{setting
 | |
| 235 | ISABELLE_HOME_USER} is included in the same manner (if that directory | |
| 63669 | 236 |   exists). This allows to install private components via @{path
 | 
| 61575 | 237 | "$ISABELLE_HOME_USER/etc/components"}, although it is often more convenient | 
| 238 | to do that programmatically via the \<^verbatim>\<open>init_component\<close> shell function in the | |
| 239 | \<^verbatim>\<open>etc/settings\<close> script of \<^verbatim>\<open>$ISABELLE_HOME_USER\<close> (or any other component | |
| 63669 | 240 |   directory). For example: @{verbatim [display] \<open>init_component
 | 
| 241 | "$HOME/screwdriver-2.0"\<close>} | |
| 48838 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 242 | |
| 61575 | 243 | This is tolerant wrt.\ missing component directories, but might produce a | 
| 244 | warning. | |
| 48838 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 245 | |
| 61406 | 246 | \<^medskip> | 
| 61575 | 247 | More complex situations may be addressed by initializing components listed | 
| 248 | in a given catalog file, relatively to some base directory: | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 249 |   @{verbatim [display] \<open>init_components "$HOME/my_component_store" "some_catalog_file"\<close>}
 | 
| 48838 
623ba165d059
direct support for component forests via init_components;
 wenzelm parents: 
48813diff
changeset | 250 | |
| 61575 | 251 | The component directories listed in the catalog file are treated as relative | 
| 252 | to the given base directory. | |
| 48844 | 253 | |
| 61575 | 254 |   See also \secref{sec:tool-components} for some tool-support for resolving
 | 
| 255 | components that are formally initialized but not installed yet. | |
| 58618 | 256 | \<close> | 
| 32323 | 257 | |
| 258 | ||
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 259 | section \<open>The Isabelle tool wrapper \label{sec:isabelle-tool}\<close>
 | 
| 28215 | 260 | |
| 58618 | 261 | text \<open> | 
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 262 | The main \<^emph>\<open>Isabelle tool wrapper\<close> provides a generic startup environment for | 
| 64509 | 263 | Isabelle-related utilities, user interfaces, add-on applications etc. Such | 
| 264 | tools automatically benefit from the settings mechanism | |
| 265 |   (\secref{sec:settings}). Moreover, this is the standard way to invoke
 | |
| 266 | Isabelle/Scala functionality as a separate operating-system process. | |
| 267 | Isabelle command-line tools are run uniformly via a common wrapper --- | |
| 268 |   @{executable_ref isabelle}:
 | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 269 |   @{verbatim [display]
 | 
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 270 | \<open>Usage: isabelle TOOL [ARGS ...] | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 271 | |
| 62829 
4141c2a8458b
clarified Isabelle tool wrapper: bash, Scala, no perl, no ML;
 wenzelm parents: 
62677diff
changeset | 272 | Start Isabelle TOOL with ARGS; pass "-?" for tool-specific help. | 
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 273 | |
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 274 | Available tools: | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 275 | ...\<close>} | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 276 | |
| 64509 | 277 | Tools may be implemented in Isabelle/Scala or as stand-alone executables | 
| 278 |   (usually as GNU bash scripts). In the invocation of ``@{executable
 | |
| 279 | isabelle}~\<open>tool\<close>'', the named \<open>tool\<close> is resolved as follows (and in the | |
| 280 | given order). | |
| 281 | ||
| 282 |     \<^enum> An external tool found on the directories listed in the @{setting
 | |
| 283 | ISABELLE_TOOLS} settings variable (colon-separated list in standard POSIX | |
| 284 | notation). | |
| 285 | ||
| 286 | \<^enum> If a file ``\<open>tool\<close>\<^verbatim>\<open>.scala\<close>'' is found, the source needs to define | |
| 287 | some object that extends the class \<^verbatim>\<open>Isabelle_Tool.Body\<close>. The Scala | |
| 288 | compiler is invoked on the spot (which may take some time), and the body | |
| 289 | function is run with the command-line arguments as \<^verbatim>\<open>List[String]\<close>. | |
| 290 | ||
| 291 | \<^enum> If an executable file ``\<open>tool\<close>'' is found, it is invoked as | |
| 292 | stand-alone program with the command-line arguments provided as \<^verbatim>\<open>argv\<close> | |
| 293 | array. | |
| 294 | ||
| 295 | \<^enum> An internal tool that is registered in \<^verbatim>\<open>Isabelle_Tool.internal_tools\<close> | |
| 296 | within the Isabelle/Scala namespace of \<^verbatim>\<open>Pure.jar\<close>. This is the preferred | |
| 297 | entry-point for high-end tools implemented in Isabelle/Scala --- compiled | |
| 298 | once when the Isabelle distribution is built. These tools are provided by | |
| 299 | Isabelle/Pure and cannot be augmented in user-space. | |
| 300 | ||
| 301 | There are also some administrative tools that are available from a bare | |
| 302 | repository clone of Isabelle, but not in regular distributions. | |
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 303 | \<close> | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 304 | |
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 305 | |
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 306 | subsubsection \<open>Examples\<close> | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 307 | |
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 308 | text \<open> | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 309 | Show the list of available documentation of the Isabelle distribution: | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 310 |   @{verbatim [display] \<open>isabelle doc\<close>}
 | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 311 | |
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 312 | View a certain document as follows: | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 313 |   @{verbatim [display] \<open>isabelle doc system\<close>}
 | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 314 | |
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 315 | Query the Isabelle settings environment: | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 316 |   @{verbatim [display] \<open>isabelle getenv ISABELLE_HOME_USER\<close>}
 | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 317 | \<close> | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 318 | |
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 319 | |
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 320 | section \<open>The raw Isabelle ML process\<close> | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 321 | |
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 322 | subsection \<open>Batch mode \label{sec:tool-process}\<close>
 | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 323 | |
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 324 | text \<open> | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 325 |   The @{tool_def process} tool runs the raw ML process in batch mode:
 | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 326 |   @{verbatim [display]
 | 
| 62634 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 wenzelm parents: 
62588diff
changeset | 327 | \<open>Usage: isabelle process [OPTIONS] | 
| 28215 | 328 | |
| 329 | Options are: | |
| 62677 | 330 | -T THEORY load theory | 
| 62639 | 331 | -d DIR include session directory | 
| 62506 | 332 | -e ML_EXPR evaluate ML expression on startup | 
| 333 | -f ML_FILE evaluate ML file on startup | |
| 62634 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 wenzelm parents: 
62588diff
changeset | 334 | -l NAME logic session name (default ISABELLE_LOGIC="HOL") | 
| 28215 | 335 | -m MODE add print mode for output | 
| 52056 | 336 | -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) | 
| 28215 | 337 | |
| 62634 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 wenzelm parents: 
62588diff
changeset | 338 | Run the raw Isabelle ML process in batch mode.\<close>} | 
| 28215 | 339 | |
| 62643 | 340 | \<^medskip> | 
| 62506 | 341 | Options \<^verbatim>\<open>-e\<close> and \<^verbatim>\<open>-f\<close> allow to evaluate ML code, before the ML process is | 
| 342 | started. The source is either given literally or taken from a file. Multiple | |
| 343 | \<^verbatim>\<open>-e\<close> and \<^verbatim>\<open>-f\<close> options are evaluated in the given order. Errors lead to | |
| 344 | premature exit of the ML process with return code 1. | |
| 28215 | 345 | |
| 61406 | 346 | \<^medskip> | 
| 62677 | 347 | Option \<^verbatim>\<open>-T\<close> loads a specified theory file. This is a wrapper for \<^verbatim>\<open>-e\<close> with | 
| 348 |   a suitable @{ML use_thy} invocation.
 | |
| 349 | ||
| 350 | \<^medskip> | |
| 62639 | 351 | Option \<^verbatim>\<open>-l\<close> specifies the logic session name. Option \<^verbatim>\<open>-d\<close> specifies | 
| 352 |   additional directories for session roots, see also \secref{sec:tool-build}.
 | |
| 62634 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 wenzelm parents: 
62588diff
changeset | 353 | |
| 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 wenzelm parents: 
62588diff
changeset | 354 | \<^medskip> | 
| 61575 | 355 | The \<^verbatim>\<open>-m\<close> option adds identifiers of print modes to be made active for this | 
| 62573 | 356 | session. For example, \<^verbatim>\<open>-m ASCII\<close> prefers ASCII replacement syntax over | 
| 357 | mathematical Isabelle symbols. | |
| 28215 | 358 | |
| 61406 | 359 | \<^medskip> | 
| 61575 | 360 | Option \<^verbatim>\<open>-o\<close> allows to override Isabelle system options for this process, | 
| 62573 | 361 |   see also \secref{sec:system-options}.
 | 
| 58618 | 362 | \<close> | 
| 28215 | 363 | |
| 364 | ||
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 365 | subsubsection \<open>Example\<close> | 
| 28215 | 366 | |
| 58618 | 367 | text \<open> | 
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 368 | The subsequent example retrieves the \<^verbatim>\<open>Main\<close> theory value from the theory | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 369 | loader within ML: | 
| 62634 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 wenzelm parents: 
62588diff
changeset | 370 |   @{verbatim [display] \<open>isabelle process -e 'Thy_Info.get_theory "Main"'\<close>}
 | 
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 371 | |
| 64509 | 372 | Observe the delicate quoting rules for the GNU bash shell vs.\ ML. The | 
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 373 | Isabelle/ML and Scala libraries provide functions for that, but here we need | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 374 | to do it manually. | 
| 58618 | 375 | \<close> | 
| 28238 | 376 | |
| 377 | ||
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 378 | subsection \<open>Interactive mode\<close> | 
| 28238 | 379 | |
| 58618 | 380 | text \<open> | 
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 381 |   The @{tool_def console} tool runs the raw ML process with interactive
 | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 382 | console and line editor: | 
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 383 |   @{verbatim [display]
 | 
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 384 | \<open>Usage: isabelle console [OPTIONS] | 
| 28238 | 385 | |
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 386 | Options are: | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 387 | -d DIR include session directory | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 388 | -l NAME logic session name (default ISABELLE_LOGIC) | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 389 | -m MODE add print mode for output | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 390 | -n no build of session image on startup | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 391 | -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) | 
| 62643 | 392 | -r bootstrap from raw Poly/ML | 
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 393 | -s system build mode for session image | 
| 28238 | 394 | |
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 395 | Build a logic session image and run the raw Isabelle ML process | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 396 | in interactive mode, with line editor ISABELLE_LINE_EDITOR.\<close>} | 
| 28238 | 397 | |
| 62643 | 398 | \<^medskip> | 
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 399 | Option \<^verbatim>\<open>-l\<close> specifies the logic session name. By default, its heap image is | 
| 62643 | 400 | checked and built on demand, but the option \<^verbatim>\<open>-n\<close> skips that. | 
| 28238 | 401 | |
| 62643 | 402 | Option \<^verbatim>\<open>-r\<close> indicates a bootstrap from the raw Poly/ML system, which is | 
| 403 | relevant for Isabelle/Pure development. | |
| 404 | ||
| 405 | \<^medskip> | |
| 62639 | 406 |   Options \<^verbatim>\<open>-d\<close>, \<^verbatim>\<open>-m\<close>, \<^verbatim>\<open>-o\<close> have the same meaning as for @{tool process}
 | 
| 407 |   (\secref{sec:tool-process}).
 | |
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 408 | |
| 62639 | 409 |   Option \<^verbatim>\<open>-s\<close> has the same meaning as for @{tool build}
 | 
| 410 |   (\secref{sec:tool-build}).
 | |
| 28238 | 411 | |
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 412 | \<^medskip> | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 413 | The Isabelle/ML process is run through the line editor that is specified via | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 414 |   the settings variable @{setting ISABELLE_LINE_EDITOR} (e.g.\
 | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 415 |   @{executable_def rlwrap} for GNU readline); the fall-back is to use plain
 | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 416 | standard input/output. | 
| 28238 | 417 | |
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 418 | The user is connected to the raw ML toplevel loop: this is neither | 
| 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62576diff
changeset | 419 | Isabelle/Isar nor Isabelle/ML within the usual formal context. The most | 
| 62847 | 420 |   relevant ML commands at this stage are @{ML use} (for ML files) and
 | 
| 421 |   @{ML use_thy} (for theory files).
 | |
| 58618 | 422 | \<close> | 
| 28215 | 423 | |
| 62451 | 424 | |
| 63995 | 425 | section \<open>The raw Isabelle Java process \label{sec:isabelle-java}\<close>
 | 
| 426 | ||
| 427 | text \<open> | |
| 428 |   The @{executable_ref isabelle_java} executable allows to run a Java process
 | |
| 429 | within the name space of Java and Scala components that are bundled with | |
| 430 | Isabelle, but \<^emph>\<open>without\<close> the Isabelle settings environment | |
| 431 |   (\secref{sec:settings}).
 | |
| 432 | ||
| 433 | After such a JVM cold-start, the Isabelle environment can be accessed via | |
| 434 | \<^verbatim>\<open>Isabelle_System.getenv\<close> as usual, but the underlying process environment | |
| 435 | remains clean. This is e.g.\ relevant when invoking other processes that | |
| 436 | should remain separate from the current Isabelle installation. | |
| 437 | ||
| 438 | \<^medskip> | |
| 439 | Note that under normal circumstances, Isabelle command-line tools are run | |
| 440 |   \<^emph>\<open>within\<close> the settings environment, as provided by the @{executable
 | |
| 441 |   isabelle} wrapper (\secref{sec:isabelle-tool} and \secref{sec:tool-java}).
 | |
| 442 | \<close> | |
| 443 | ||
| 444 | ||
| 445 | subsubsection \<open>Example\<close> | |
| 446 | ||
| 447 | text \<open> | |
| 448 | The subsequent example creates a raw Java process on the command-line and | |
| 449 | invokes the main Isabelle application entry point: | |
| 450 |   @{verbatim [display] \<open>isabelle_java isabelle.Main\<close>}
 | |
| 451 | \<close> | |
| 452 | ||
| 453 | ||
| 62451 | 454 | section \<open>YXML versus XML\<close> | 
| 455 | ||
| 456 | text \<open> | |
| 457 | Isabelle tools often use YXML, which is a simple and efficient syntax for | |
| 458 | untyped XML trees. The YXML format is defined as follows. | |
| 459 | ||
| 460 | \<^enum> The encoding is always UTF-8. | |
| 461 | ||
| 462 | \<^enum> Body text is represented verbatim (no escaping, no special treatment of | |
| 463 | white space, no named entities, no CDATA chunks, no comments). | |
| 464 | ||
| 465 | \<^enum> Markup elements are represented via ASCII control characters \<open>\<^bold>X = 5\<close> | |
| 466 | and \<open>\<^bold>Y = 6\<close> as follows: | |
| 467 | ||
| 468 |     \begin{tabular}{ll}
 | |
| 469 | XML & YXML \\\hline | |
| 470 | \<^verbatim>\<open><\<close>\<open>name attribute\<close>\<^verbatim>\<open>=\<close>\<open>value \<dots>\<close>\<^verbatim>\<open>>\<close> & | |
| 471 | \<open>\<^bold>X\<^bold>Yname\<^bold>Yattribute\<close>\<^verbatim>\<open>=\<close>\<open>value\<dots>\<^bold>X\<close> \\ | |
| 472 | \<^verbatim>\<open></\<close>\<open>name\<close>\<^verbatim>\<open>>\<close> & \<open>\<^bold>X\<^bold>Y\<^bold>X\<close> \\ | |
| 473 |     \end{tabular}
 | |
| 474 | ||
| 475 | There is no special case for empty body text, i.e.\ \<^verbatim>\<open><foo/>\<close> is treated | |
| 476 | like \<^verbatim>\<open><foo></foo>\<close>. Also note that \<open>\<^bold>X\<close> and \<open>\<^bold>Y\<close> may never occur in | |
| 477 | well-formed XML documents. | |
| 478 | ||
| 479 | Parsing YXML is pretty straight-forward: split the text into chunks | |
| 480 | separated by \<open>\<^bold>X\<close>, then split each chunk into sub-chunks separated by \<open>\<^bold>Y\<close>. | |
| 481 | Markup chunks start with an empty sub-chunk, and a second empty sub-chunk | |
| 482 | indicates close of an element. Any other non-empty chunk consists of plain | |
| 63680 | 483 | text. For example, see \<^file>\<open>~~/src/Pure/PIDE/yxml.ML\<close> or | 
| 484 | \<^file>\<open>~~/src/Pure/PIDE/yxml.scala\<close>. | |
| 62451 | 485 | |
| 486 | YXML documents may be detected quickly by checking that the first two | |
| 487 | characters are \<open>\<^bold>X\<^bold>Y\<close>. | |
| 488 | \<close> | |
| 489 | ||
| 28223 | 490 | end |