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