| author | wenzelm | 
| Tue, 22 Apr 2014 23:57:17 +0200 | |
| changeset 56663 | 2d09b437c168 | 
| parent 56604 | 1b153b989860 | 
| child 57320 | 00f2c8d1aa0b | 
| permissions | -rw-r--r-- | 
| 48578 | 1 | theory Sessions | 
| 2 | imports Base | |
| 3 | begin | |
| 4 | ||
| 48584 | 5 | chapter {* Isabelle sessions and build management \label{ch:session} *}
 | 
| 6 | ||
| 7 | text {* An Isabelle \emph{session} consists of a collection of related
 | |
| 51055 | 8 | theories that may be associated with formal documents | 
| 9 |   (\chref{ch:present}).  There is also a notion of \emph{persistent
 | |
| 48604 | 10 | heap} image to capture the state of a session, similar to | 
| 11 | object-code in compiled programming languages. Thus the concept of | |
| 12 | session resembles that of a ``project'' in common IDE environments, | |
| 13 | but the specific name emphasizes the connection to interactive | |
| 14 | theorem proving: the session wraps-up the results of | |
| 15 | user-interaction with the prover in a persistent form. | |
| 48584 | 16 | |
| 48604 | 17 | Application sessions are built on a given parent session, which may | 
| 18 | be built recursively on other parents. Following this path in the | |
| 19 | hierarchy eventually leads to some major object-logic session like | |
| 20 |   @{text "HOL"}, which itself is based on @{text "Pure"} as the common
 | |
| 21 | root of all sessions. | |
| 48584 | 22 | |
| 48604 | 23 | Processing sessions may take considerable time. Isabelle build | 
| 24 | management helps to organize this efficiently. This includes | |
| 25 | support for parallel build jobs, in addition to the multithreaded | |
| 26 | theory and proof checking that is already provided by the prover | |
| 27 | process itself. *} | |
| 28 | ||
| 48578 | 29 | |
| 30 | section {* Session ROOT specifications \label{sec:session-root} *}
 | |
| 31 | ||
| 48579 | 32 | text {* Session specifications reside in files called @{verbatim ROOT}
 | 
| 33 | within certain directories, such as the home locations of registered | |
| 34 | Isabelle components or additional project directories given by the | |
| 35 | user. | |
| 36 | ||
| 37 | The ROOT file format follows the lexical conventions of the | |
| 38 |   \emph{outer syntax} of Isabelle/Isar, see also
 | |
| 39 |   \cite{isabelle-isar-ref}.  This defines common forms like
 | |
| 40 | identifiers, names, quoted strings, verbatim text, nested comments | |
| 51417 | 41 |   etc.  The grammar for @{syntax session_chapter} and @{syntax
 | 
| 42 | session_entry} is given as syntax diagram below; each ROOT file may | |
| 43 | contain multiple specifications like this. Chapters help to | |
| 44 |   organize browser info (\secref{sec:info}), but have no formal
 | |
| 45 |   meaning.  The default chapter is ``@{text "Unsorted"}''.
 | |
| 48579 | 46 | |
| 48582 | 47 |   Isabelle/jEdit (\secref{sec:tool-jedit}) includes a simple editing
 | 
| 51055 | 48 |   mode @{verbatim "isabelle-root"} for session ROOT files, which is
 | 
| 49 | enabled by default for any file of that name. | |
| 48579 | 50 | |
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
54705diff
changeset | 51 |   @{rail \<open>
 | 
| 51417 | 52 |     @{syntax_def session_chapter}: @'chapter' @{syntax name}
 | 
| 53 | ; | |
| 54 | ||
| 48579 | 55 |     @{syntax_def session_entry}: @'session' spec '=' (@{syntax name} '+')? body
 | 
| 56 | ; | |
| 56533 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 57 | body: description? options? (theories+) \<newline> files? (document_files*) | 
| 48579 | 58 | ; | 
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48737diff
changeset | 59 |     spec: @{syntax name} groups? dir?
 | 
| 48579 | 60 | ; | 
| 61 |     groups: '(' (@{syntax name} +) ')'
 | |
| 62 | ; | |
| 63 |     dir: @'in' @{syntax name}
 | |
| 64 | ; | |
| 65 |     description: @'description' @{syntax text}
 | |
| 66 | ; | |
| 67 | options: @'options' opts | |
| 68 | ; | |
| 48605 
e777363440d6
allow negative int values as well, according to real = int | float;
 wenzelm parents: 
48604diff
changeset | 69 |     opts: '[' ( (@{syntax name} '=' value | @{syntax name}) + ',' ) ']'
 | 
| 
e777363440d6
allow negative int values as well, according to real = int | float;
 wenzelm parents: 
48604diff
changeset | 70 | ; | 
| 
e777363440d6
allow negative int values as well, according to real = int | float;
 wenzelm parents: 
48604diff
changeset | 71 |     value: @{syntax name} | @{syntax real}
 | 
| 48579 | 72 | ; | 
| 48739 | 73 |     theories: @'theories' opts? ( @{syntax name} * )
 | 
| 48579 | 74 | ; | 
| 56533 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 75 |     files: @'files' (@{syntax name}+)
 | 
| 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 76 | ; | 
| 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 77 |     document_files: @'document_files' ('(' dir ')')? (@{syntax name}+)
 | 
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
54705diff
changeset | 78 | \<close>} | 
| 48579 | 79 | |
| 80 |   \begin{description}
 | |
| 81 | ||
| 82 |   \item \isakeyword{session}~@{text "A = B + body"} defines a new
 | |
| 83 |   session @{text "A"} based on parent session @{text "B"}, with its
 | |
| 84 |   content given in @{text body} (theories and auxiliary source files).
 | |
| 85 |   Note that a parent (like @{text "HOL"}) is mandatory in practical
 | |
| 86 | applications: only Isabelle/Pure can bootstrap itself from nothing. | |
| 87 | ||
| 88 | All such session specifications together describe a hierarchy (tree) | |
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48737diff
changeset | 89 | of sessions, with globally unique names. The new session name | 
| 51055 | 90 |   @{text "A"} should be sufficiently long and descriptive to stand on
 | 
| 91 | its own in a potentially large library. | |
| 48579 | 92 | |
| 93 |   \item \isakeyword{session}~@{text "A (groups)"} indicates a
 | |
| 94 | collection of groups where the new session is a member. Group names | |
| 95 | are uninterpreted and merely follow certain conventions. For | |
| 96 | example, the Isabelle distribution tags some important sessions by | |
| 48604 | 97 |   the group name called ``@{text "main"}''.  Other projects may invent
 | 
| 98 | their own conventions, but this requires some care to avoid clashes | |
| 99 | within this unchecked name space. | |
| 48579 | 100 | |
| 101 |   \item \isakeyword{session}~@{text "A"}~\isakeyword{in}~@{text "dir"}
 | |
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48737diff
changeset | 102 | specifies an explicit directory for this session; by default this is | 
| 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48737diff
changeset | 103 |   the current directory of the @{verbatim ROOT} file.
 | 
| 48579 | 104 | |
| 105 | All theories and auxiliary source files are located relatively to | |
| 106 | the session directory. The prover process is run within the same as | |
| 107 | its current working directory. | |
| 108 | ||
| 109 |   \item \isakeyword{description}~@{text "text"} is a free-form
 | |
| 110 | annotation for this session. | |
| 111 | ||
| 112 |   \item \isakeyword{options}~@{text "[x = a, y = b, z]"} defines
 | |
| 48604 | 113 |   separate options (\secref{sec:system-options}) that are used when
 | 
| 114 |   processing this session, but \emph{without} propagation to child
 | |
| 115 |   sessions.  Note that @{text "z"} abbreviates @{text "z = true"} for
 | |
| 116 | Boolean options. | |
| 48579 | 117 | |
| 118 |   \item \isakeyword{theories}~@{text "options names"} specifies a
 | |
| 119 | block of theories that are processed within an environment that is | |
| 48604 | 120 | augmented by the given options, in addition to the global session | 
| 121 |   options given before.  Any number of blocks of \isakeyword{theories}
 | |
| 122 | may be given. Options are only active for each | |
| 123 |   \isakeyword{theories} block separately.
 | |
| 48579 | 124 | |
| 125 |   \item \isakeyword{files}~@{text "files"} lists additional source
 | |
| 48604 | 126 | files that are involved in the processing of this session. This | 
| 127 | should cover anything outside the formal content of the theory | |
| 56533 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 128 | sources. In contrast, files that are loaded formally | 
| 51055 | 129 |   within a theory, e.g.\ via @{keyword "ML_file"}, need not be
 | 
| 130 | declared again. | |
| 48579 | 131 | |
| 56533 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 132 |   \item \isakeyword{document_files}~@{text "("}\isakeyword{in}~@{text
 | 
| 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 133 | "base_dir) files"} lists source files for document preparation, | 
| 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 134 |   typically @{verbatim ".tex"} and @{verbatim ".sty"} for {\LaTeX}.
 | 
| 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 135 | Only these explicitly given files are copied from the base directory | 
| 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 136 | to the document output directory, before formal document processing | 
| 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 137 |   is started (see also \secref{sec:tool-document}).  The local path
 | 
| 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 138 |   structure of the @{text files} is preserved, which allows to
 | 
| 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 139 |   reconstruct the original directory hierarchy of @{text "base_dir"}.
 | 
| 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 140 | |
| 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 141 |   \item \isakeyword{document_files}~@{text "files"} abbreviates
 | 
| 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 142 |   \isakeyword{document_files}~@{text "("}\isakeyword{in}~@{text
 | 
| 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 143 | "document) files"}, i.e.\ document sources are taken from the base | 
| 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 144 |   directory @{verbatim document} within the session root directory.
 | 
| 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 wenzelm parents: 
55112diff
changeset | 145 | |
| 48579 | 146 |   \end{description}
 | 
| 48580 | 147 | *} | 
| 48579 | 148 | |
| 51055 | 149 | |
| 48580 | 150 | subsubsection {* Examples *}
 | 
| 151 | ||
| 152 | text {* See @{file "~~/src/HOL/ROOT"} for a diversity of practically
 | |
| 51055 | 153 | relevant situations, although it uses relatively complex | 
| 154 |   quasi-hierarchic naming conventions like @{text "HOL\<dash>SPARK"},
 | |
| 155 |   @{text "HOL\<dash>SPARK\<dash>Examples"}.  An alternative is to use
 | |
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48737diff
changeset | 156 | unqualified names that are relatively long and descriptive, as in | 
| 54704 | 157 |   the Archive of Formal Proofs (@{url "http://afp.sourceforge.net"}), for
 | 
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48737diff
changeset | 158 | example. *} | 
| 48578 | 159 | |
| 160 | ||
| 161 | section {* System build options \label{sec:system-options} *}
 | |
| 162 | ||
| 48580 | 163 | text {* See @{file "~~/etc/options"} for the main defaults provided by
 | 
| 48582 | 164 |   the Isabelle distribution.  Isabelle/jEdit (\secref{sec:tool-jedit})
 | 
| 165 |   includes a simple editing mode @{verbatim "isabelle-options"} for
 | |
| 166 | this file-format. | |
| 48693 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 167 | |
| 56604 | 168 | The following options are particularly relevant to build Isabelle | 
| 51055 | 169 | sessions, in particular with document preparation | 
| 170 |   (\chref{ch:present}).
 | |
| 171 | ||
| 172 |   \begin{itemize}
 | |
| 173 | ||
| 174 |   \item @{system_option_def "browser_info"} controls output of HTML
 | |
| 175 |   browser info, see also \secref{sec:info}.
 | |
| 176 | ||
| 177 |   \item @{system_option_def "document"} specifies the document output
 | |
| 178 |   format, see @{tool document} option @{verbatim "-o"} in
 | |
| 179 |   \secref{sec:tool-document}.  In practice, the most relevant values
 | |
| 180 |   are @{verbatim "document=false"} or @{verbatim "document=pdf"}.
 | |
| 181 | ||
| 182 |   \item @{system_option_def "document_output"} specifies an
 | |
| 183 | alternative directory for generated output of the document | |
| 184 |   preparation system; the default is within the @{setting
 | |
| 185 | "ISABELLE_BROWSER_INFO"} hierarchy as explained in | |
| 186 |   \secref{sec:info}.  See also @{tool mkroot}, which generates a
 | |
| 187 | default configuration with output readily available to the author of | |
| 188 | the document. | |
| 189 | ||
| 190 |   \item @{system_option_def "document_variants"} specifies document
 | |
| 191 |   variants as a colon-separated list of @{text "name=tags"} entries,
 | |
| 192 |   corresponding to @{tool document} options @{verbatim "-n"} and
 | |
| 193 |   @{verbatim "-t"}.
 | |
| 194 | ||
| 195 |   For example, @{verbatim
 | |
| 196 | "document_variants=document:outline=/proof,/ML"} indicates two | |
| 197 |   documents: the one called @{verbatim document} with default tags,
 | |
| 198 |   and the other called @{verbatim outline} where proofs and ML
 | |
| 199 | sections are folded. | |
| 200 | ||
| 201 | Document variant names are just a matter of conventions. It is also | |
| 202 | possible to use different document variant names (without tags) for | |
| 203 | different document root entries, see also | |
| 204 |   \secref{sec:tool-document}.
 | |
| 205 | ||
| 206 |   \item @{system_option_def "document_graph"} tells whether the
 | |
| 207 | generated document files should include a theory graph (cf.\ | |
| 208 |   \secref{sec:browse} and \secref{sec:info}).  The resulting EPS or
 | |
| 209 |   PDF file can be included as graphics in {\LaTeX}.
 | |
| 210 | ||
| 211 | Note that this option is usually determined as static parameter of | |
| 212 |   some session (e.g.\ within its @{verbatim ROOT} file) and \emph{not}
 | |
| 213 |   given globally or on the command line of @{tool build}.
 | |
| 214 | ||
| 215 |   \item @{system_option_def "threads"} determines the number of worker
 | |
| 216 | threads for parallel checking of theories and proofs. The default | |
| 217 |   @{text "0"} means that a sensible maximum value is determined by the
 | |
| 218 | underlying hardware. For machines with many cores or with | |
| 219 | hyperthreading, this is often requires manual adjustment (on the | |
| 220 | command-line or within personal settings or preferences, not within | |
| 221 |   a session @{verbatim ROOT}).
 | |
| 222 | ||
| 223 |   \item @{system_option_def "condition"} specifies a comma-separated
 | |
| 224 | list of process environment variables (or Isabelle settings) that | |
| 225 | are required for the subsequent theories to be processed. | |
| 226 | Conditions are considered ``true'' if the corresponding environment | |
| 227 | value is defined and non-empty. | |
| 228 | ||
| 229 |   For example, the @{verbatim "condition=ISABELLE_FULL_TEST"} may be
 | |
| 230 | used to guard extraordinary theories, which are meant to be enabled | |
| 231 |   explicitly via some shell prefix @{verbatim "env ISABELLE_FULL_TEST=true"}
 | |
| 232 |   before invoking @{tool build}.
 | |
| 233 | ||
| 234 |   \item @{system_option_def "timeout"} specifies a real wall-clock
 | |
| 235 | timeout (in seconds) for the session as a whole. The timer is | |
| 236 | controlled outside the ML process by the JVM that runs | |
| 237 | Isabelle/Scala. Thus it is relatively reliable in canceling | |
| 238 | processes that get out of control, even if there is a deadlock | |
| 239 | without CPU time usage. | |
| 240 | ||
| 241 |   \end{itemize}
 | |
| 242 | ||
| 48693 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 243 |   The @{tool_def options} tool prints Isabelle system options.  Its
 | 
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 244 | command-line usage is: | 
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 245 | \begin{ttbox}
 | 
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 246 | Usage: isabelle options [OPTIONS] [MORE_OPTIONS ...] | 
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 247 | |
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 248 | Options are: | 
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 249 | -b include $ISABELLE_BUILD_OPTIONS | 
| 52735 | 250 | -g OPTION get value of OPTION | 
| 50531 
f841ac0cb757
clarified "isabelle options" command line, to make it more close to "isabelle components";
 wenzelm parents: 
50406diff
changeset | 251 | -l list options | 
| 48693 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 252 | -x FILE export to FILE in YXML format | 
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 253 | |
| 50531 
f841ac0cb757
clarified "isabelle options" command line, to make it more close to "isabelle components";
 wenzelm parents: 
50406diff
changeset | 254 | Report Isabelle system options, augmented by MORE_OPTIONS given as | 
| 48693 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 255 | arguments NAME=VAL or NAME. | 
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 256 | \end{ttbox}
 | 
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 257 | |
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 258 | The command line arguments provide additional system options of the | 
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 259 |   form @{text "name"}@{verbatim "="}@{text "value"} or @{text name}
 | 
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 260 | for Boolean options. | 
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 261 | |
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 262 |   Option @{verbatim "-b"} augments the implicit environment of system
 | 
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 263 |   options by the ones of @{setting ISABELLE_BUILD_OPTIONS}, cf.\
 | 
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 264 |   \secref{sec:tool-build}.
 | 
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 265 | |
| 52735 | 266 |   Option @{verbatim "-g"} prints the value of the given option.
 | 
| 54347 | 267 |   Option @{verbatim "-l"} lists all options with their declaration and
 | 
| 268 | current value. | |
| 52735 | 269 | |
| 48693 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 270 |   Option @{verbatim "-x"} specifies a file to export the result in
 | 
| 
ceeea46bdeba
"isabelle options" prints Isabelle system options;
 wenzelm parents: 
48684diff
changeset | 271 | YXML format, instead of printing it in human-readable form. | 
| 48580 | 272 | *} | 
| 48578 | 273 | |
| 274 | ||
| 275 | section {* Invoking the build process \label{sec:tool-build} *}
 | |
| 276 | ||
| 48602 | 277 | text {* The @{tool_def build} tool invokes the build process for
 | 
| 48578 | 278 | Isabelle sessions. It manages dependencies between sessions, | 
| 279 | related sources of theories and auxiliary files, and target heap | |
| 280 | images. Accordingly, it runs instances of the prover process with | |
| 281 | optional document preparation. Its command-line usage | |
| 282 |   is:\footnote{Isabelle/Scala provides the same functionality via
 | |
| 283 |   \texttt{isabelle.Build.build}.}
 | |
| 48602 | 284 | \begin{ttbox}
 | 
| 285 | Usage: isabelle build [OPTIONS] [SESSIONS ...] | |
| 48578 | 286 | |
| 287 | Options are: | |
| 48737 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 wenzelm parents: 
48693diff
changeset | 288 | -D DIR include session directory and select its sessions | 
| 49131 | 289 | -R operate on requirements of selected sessions | 
| 48578 | 290 | -a select all sessions | 
| 291 | -b build heap images | |
| 48595 | 292 | -c clean build | 
| 48737 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 wenzelm parents: 
48693diff
changeset | 293 | -d DIR include session directory | 
| 48578 | 294 | -g NAME select session group NAME | 
| 295 | -j INT maximum number of parallel jobs (default 1) | |
| 48903 | 296 | -l list session source files | 
| 48578 | 297 | -n no build -- test dependencies only | 
| 52056 | 298 | -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) | 
| 48578 | 299 | -s system build mode: produce output in ISABELLE_HOME | 
| 300 | -v verbose | |
| 301 | ||
| 302 | Build and manage Isabelle sessions, depending on implicit | |
| 303 | ISABELLE_BUILD_OPTIONS="..." | |
| 304 | ||
| 305 | ML_PLATFORM="..." | |
| 306 | ML_HOME="..." | |
| 307 | ML_SYSTEM="..." | |
| 308 | ML_OPTIONS="..." | |
| 309 | \end{ttbox}
 | |
| 310 | ||
| 311 | \medskip Isabelle sessions are defined via session ROOT files as | |
| 312 |   described in (\secref{sec:session-root}).  The totality of sessions
 | |
| 313 | is determined by collecting such specifications from all Isabelle | |
| 314 |   component directories (\secref{sec:components}), augmented by more
 | |
| 315 |   directories given via options @{verbatim "-d"}~@{text "DIR"} on the
 | |
| 316 | command line. Each such directory may contain a session | |
| 48650 | 317 |   \texttt{ROOT} file with several session specifications.
 | 
| 48578 | 318 | |
| 48684 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 wenzelm parents: 
48683diff
changeset | 319 | Any session root directory may refer recursively to further | 
| 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 wenzelm parents: 
48683diff
changeset | 320 | directories of the same kind, by listing them in a catalog file | 
| 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 wenzelm parents: 
48683diff
changeset | 321 |   @{verbatim "ROOTS"} line-by-line.  This helps to organize large
 | 
| 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 wenzelm parents: 
48683diff
changeset | 322 |   collections of session specifications, or to make @{verbatim "-d"}
 | 
| 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 wenzelm parents: 
48683diff
changeset | 323 |   command line options persistent (say within @{verbatim
 | 
| 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 wenzelm parents: 
48683diff
changeset | 324 | "$ISABELLE_HOME_USER/ROOTS"}). | 
| 
9170e10c651e
re-introduced ROOTS catalog files (cf. 47330b712f8f) which help to organize AFP or make -d options persistent;
 wenzelm parents: 
48683diff
changeset | 325 | |
| 48604 | 326 | \medskip The subset of sessions to be managed is determined via | 
| 48578 | 327 |   individual @{text "SESSIONS"} given as command-line arguments, or
 | 
| 48604 | 328 |   session groups that are given via one or more options @{verbatim
 | 
| 48578 | 329 |   "-g"}~@{text "NAME"}.  Option @{verbatim "-a"} selects all sessions.
 | 
| 48604 | 330 | The build tool takes session dependencies into account: the set of | 
| 331 | selected sessions is completed by including all ancestors. | |
| 48578 | 332 | |
| 49131 | 333 |   \medskip Option @{verbatim "-R"} reverses the selection in the sense
 | 
| 334 | that it refers to its requirements: all ancestor sessions excluding | |
| 335 | the original selection. This allows to prepare the stage for some | |
| 336 | build process with different options, before running the main build | |
| 337 |   itself (without option @{verbatim "-R"}).
 | |
| 338 | ||
| 48737 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 wenzelm parents: 
48693diff
changeset | 339 |   \medskip Option @{verbatim "-D"} is similar to @{verbatim "-d"}, but
 | 
| 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 wenzelm parents: 
48693diff
changeset | 340 | selects all sessions that are defined in the given directories. | 
| 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 wenzelm parents: 
48693diff
changeset | 341 | |
| 48604 | 342 | \medskip The build process depends on additional options | 
| 343 |   (\secref{sec:system-options}) that are passed to the prover
 | |
| 344 |   eventually.  The settings variable @{setting_ref
 | |
| 345 | ISABELLE_BUILD_OPTIONS} allows to provide additional defaults, e.g.\ | |
| 346 |   \texttt{ISABELLE_BUILD_OPTIONS="document=pdf threads=4"}. Moreover,
 | |
| 347 | the environment of system build options may be augmented on the | |
| 348 |   command line via @{verbatim "-o"}~@{text "name"}@{verbatim
 | |
| 349 |   "="}@{text "value"} or @{verbatim "-o"}~@{text "name"}, which
 | |
| 350 |   abbreviates @{verbatim "-o"}~@{text "name"}@{verbatim"=true"} for
 | |
| 351 |   Boolean options.  Multiple occurrences of @{verbatim "-o"} on the
 | |
| 352 | command-line are applied in the given order. | |
| 48578 | 353 | |
| 354 |   \medskip Option @{verbatim "-b"} ensures that heap images are
 | |
| 355 | produced for all selected sessions. By default, images are only | |
| 356 | saved for inner nodes of the hierarchy of sessions, as required for | |
| 357 | other sessions to continue later on. | |
| 358 | ||
| 48595 | 359 |   \medskip Option @{verbatim "-c"} cleans all descendants of the
 | 
| 360 | selected sessions before performing the specified build operation. | |
| 361 | ||
| 362 |   \medskip Option @{verbatim "-n"} omits the actual build process
 | |
| 363 | after the preparatory stage (including optional cleanup). Note that | |
| 364 | the return code always indicates the status of the set of selected | |
| 365 | sessions. | |
| 366 | ||
| 48578 | 367 |   \medskip Option @{verbatim "-j"} specifies the maximum number of
 | 
| 48604 | 368 | parallel build jobs (prover processes). Each prover process is | 
| 369 | subject to a separate limit of parallel worker threads, cf.\ system | |
| 370 |   option @{system_option_ref threads}.
 | |
| 48578 | 371 | |
| 372 |   \medskip Option @{verbatim "-s"} enables \emph{system mode}, which
 | |
| 373 | means that resulting heap images and log files are stored in | |
| 54705 | 374 |   @{file_unchecked "$ISABELLE_HOME/heaps"} instead of the default location
 | 
| 48578 | 375 |   @{setting ISABELLE_OUTPUT} (which is normally in @{setting
 | 
| 376 | ISABELLE_HOME_USER}, i.e.\ the user's home directory). | |
| 377 | ||
| 48903 | 378 |   \medskip Option @{verbatim "-v"} increases the general level of
 | 
| 379 |   verbosity.  Option @{verbatim "-l"} lists the source files that
 | |
| 380 | contribute to a session. | |
| 48578 | 381 | *} | 
| 382 | ||
| 383 | subsubsection {* Examples *}
 | |
| 384 | ||
| 385 | text {*
 | |
| 386 | Build a specific logic image: | |
| 387 | \begin{ttbox}
 | |
| 388 | isabelle build -b HOLCF | |
| 389 | \end{ttbox}
 | |
| 390 | ||
| 48604 | 391 | \smallskip Build the main group of logic images: | 
| 48578 | 392 | \begin{ttbox}
 | 
| 393 | isabelle build -b -g main | |
| 394 | \end{ttbox}
 | |
| 395 | ||
| 48595 | 396 | \smallskip Provide a general overview of the status of all Isabelle | 
| 397 | sessions, without building anything: | |
| 48578 | 398 | \begin{ttbox}
 | 
| 399 | isabelle build -a -n -v | |
| 400 | \end{ttbox}
 | |
| 401 | ||
| 48595 | 402 | \smallskip Build all sessions with HTML browser info and PDF | 
| 403 | document preparation: | |
| 48578 | 404 | \begin{ttbox}
 | 
| 405 | isabelle build -a -o browser_info -o document=pdf | |
| 406 | \end{ttbox}
 | |
| 407 | ||
| 48604 | 408 | \smallskip Build all sessions with a maximum of 8 parallel prover | 
| 409 | processes and 4 worker threads each (on a machine with many cores): | |
| 48578 | 410 | \begin{ttbox}
 | 
| 411 | isabelle build -a -j8 -o threads=4 | |
| 412 | \end{ttbox}
 | |
| 48595 | 413 | |
| 414 | \smallskip Build some session images with cleanup of their | |
| 415 | descendants, while retaining their ancestry: | |
| 416 | \begin{ttbox}
 | |
| 54445 | 417 | isabelle build -b -c HOL-Algebra HOL-Word | 
| 48595 | 418 | \end{ttbox}
 | 
| 419 | ||
| 420 | \smallskip Clean all sessions without building anything: | |
| 421 | \begin{ttbox}
 | |
| 422 | isabelle build -a -n -c | |
| 423 | \end{ttbox}
 | |
| 48737 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 wenzelm parents: 
48693diff
changeset | 424 | |
| 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 wenzelm parents: 
48693diff
changeset | 425 | \smallskip Build all sessions from some other directory hierarchy, | 
| 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 wenzelm parents: 
48693diff
changeset | 426 |   according to the settings variable @{verbatim "AFP"} that happens to
 | 
| 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 wenzelm parents: 
48693diff
changeset | 427 | be defined inside the Isabelle environment: | 
| 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 wenzelm parents: 
48693diff
changeset | 428 | \begin{ttbox}
 | 
| 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 wenzelm parents: 
48693diff
changeset | 429 | isabelle build -D '$AFP' | 
| 
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
 wenzelm parents: 
48693diff
changeset | 430 | \end{ttbox}
 | 
| 49131 | 431 | |
| 432 | \smallskip Inform about the status of all sessions required for AFP, | |
| 433 | without building anything yet: | |
| 434 | \begin{ttbox}
 | |
| 435 | isabelle build -D '$AFP' -R -v -n | |
| 436 | \end{ttbox}
 | |
| 48578 | 437 | *} | 
| 438 | ||
| 439 | end |