| author | wenzelm | 
| Tue, 31 Jul 2012 16:23:20 +0200 | |
| changeset 48616 | be8002ee43d8 | 
| parent 48605 | e777363440d6 | 
| child 48650 | 47330b712f8f | 
| 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
 | |
| 48604 | 8 | theories that may be associated with formal documents (see also | 
| 9 |   \chref{ch:present}).  There is also a notion of \emph{persistent
 | |
| 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 | |
| 41 |   etc.  The grammar for a single @{syntax session_entry} is given as
 | |
| 42 | syntax diagram below; each ROOT file may contain multiple session | |
| 43 | specifications like this. | |
| 44 | ||
| 48582 | 45 |   Isabelle/jEdit (\secref{sec:tool-jedit}) includes a simple editing
 | 
| 46 |   mode @{verbatim "isabelle-root"} for session ROOT files.
 | |
| 48579 | 47 | |
| 48 |   @{rail "
 | |
| 49 |     @{syntax_def session_entry}: @'session' spec '=' (@{syntax name} '+')? body
 | |
| 50 | ; | |
| 51 | body: description? options? ( theories * ) files? | |
| 52 | ; | |
| 53 |     spec: @{syntax name} '!'? groups? dir?
 | |
| 54 | ; | |
| 55 |     groups: '(' (@{syntax name} +) ')'
 | |
| 56 | ; | |
| 57 |     dir: @'in' @{syntax name}
 | |
| 58 | ; | |
| 59 |     description: @'description' @{syntax text}
 | |
| 60 | ; | |
| 61 | options: @'options' opts | |
| 62 | ; | |
| 48605 
e777363440d6
allow negative int values as well, according to real = int | float;
 wenzelm parents: 
48604diff
changeset | 63 |     opts: '[' ( (@{syntax name} '=' value | @{syntax name}) + ',' ) ']'
 | 
| 
e777363440d6
allow negative int values as well, according to real = int | float;
 wenzelm parents: 
48604diff
changeset | 64 | ; | 
| 
e777363440d6
allow negative int values as well, according to real = int | float;
 wenzelm parents: 
48604diff
changeset | 65 |     value: @{syntax name} | @{syntax real}
 | 
| 48579 | 66 | ; | 
| 67 |     theories: @'theories' opts? ( @{syntax name} + )
 | |
| 68 | ; | |
| 69 |     files: @'files' ( @{syntax name} + )
 | |
| 70 | "} | |
| 71 | ||
| 72 |   \begin{description}
 | |
| 73 | ||
| 74 |   \item \isakeyword{session}~@{text "A = B + body"} defines a new
 | |
| 75 |   session @{text "A"} based on parent session @{text "B"}, with its
 | |
| 76 |   content given in @{text body} (theories and auxiliary source files).
 | |
| 77 |   Note that a parent (like @{text "HOL"}) is mandatory in practical
 | |
| 78 | applications: only Isabelle/Pure can bootstrap itself from nothing. | |
| 79 | ||
| 80 | All such session specifications together describe a hierarchy (tree) | |
| 81 | of sessions, with globally unique names. By default, names are | |
| 82 |   derived from parent ones by concatenation, i.e.\ @{text "B\<dash>A"}
 | |
| 83 |   above.  Cumulatively, this leads to session paths of the form @{text
 | |
| 84 | "X\<dash>Y\<dash>Z\<dash>W"}. Note that in the specification, | |
| 85 |   @{text B} is already such a fully-qualified name, while @{text "A"}
 | |
| 86 | is the new base name. | |
| 87 | ||
| 88 |   \item \isakeyword{session}~@{text "A! = B"} indicates a fresh start
 | |
| 89 |   in the naming scheme: the session is called just @{text "A"} instead
 | |
| 90 |   of @{text "B\<dash>A"}.  Here the name @{text "A"} should be
 | |
| 91 | sufficiently long to stand on its own in a potentially large | |
| 92 | library. | |
| 93 | ||
| 94 |   \item \isakeyword{session}~@{text "A (groups)"} indicates a
 | |
| 95 | collection of groups where the new session is a member. Group names | |
| 96 | are uninterpreted and merely follow certain conventions. For | |
| 97 | example, the Isabelle distribution tags some important sessions by | |
| 48604 | 98 |   the group name called ``@{text "main"}''.  Other projects may invent
 | 
| 99 | their own conventions, but this requires some care to avoid clashes | |
| 100 | within this unchecked name space. | |
| 48579 | 101 | |
| 102 |   \item \isakeyword{session}~@{text "A"}~\isakeyword{in}~@{text "dir"}
 | |
| 103 | specifies an explicit directory for this session. By default, | |
| 104 |   \isakeyword{session}~@{text "A"} abbreviates
 | |
| 105 |   \isakeyword{session}~@{text "A"}~\isakeyword{in}~@{text "A"}.  This
 | |
| 106 | accommodates the common scheme where some base directory contains | |
| 107 | several sessions in sub-directories of corresponding names. Another | |
| 108 |   common scheme is \isakeyword{session}~@{text
 | |
| 109 |   "A"}~\isakeyword{in}~@{verbatim "\".\""} to refer to the current
 | |
| 110 | directory of the ROOT file. | |
| 111 | ||
| 112 | All theories and auxiliary source files are located relatively to | |
| 113 | the session directory. The prover process is run within the same as | |
| 114 | its current working directory. | |
| 115 | ||
| 116 |   \item \isakeyword{description}~@{text "text"} is a free-form
 | |
| 117 | annotation for this session. | |
| 118 | ||
| 119 |   \item \isakeyword{options}~@{text "[x = a, y = b, z]"} defines
 | |
| 48604 | 120 |   separate options (\secref{sec:system-options}) that are used when
 | 
| 121 |   processing this session, but \emph{without} propagation to child
 | |
| 122 |   sessions.  Note that @{text "z"} abbreviates @{text "z = true"} for
 | |
| 123 | Boolean options. | |
| 48579 | 124 | |
| 125 |   \item \isakeyword{theories}~@{text "options names"} specifies a
 | |
| 126 | block of theories that are processed within an environment that is | |
| 48604 | 127 | augmented by the given options, in addition to the global session | 
| 128 |   options given before.  Any number of blocks of \isakeyword{theories}
 | |
| 129 | may be given. Options are only active for each | |
| 130 |   \isakeyword{theories} block separately.
 | |
| 48579 | 131 | |
| 132 |   \item \isakeyword{files}~@{text "files"} lists additional source
 | |
| 48604 | 133 | files that are involved in the processing of this session. This | 
| 134 | should cover anything outside the formal content of the theory | |
| 48579 | 135 |   sources, say some auxiliary {\TeX} files that are required for
 | 
| 136 | document processing. In contrast, files that are specified in | |
| 137 |   formal theory headers as @{keyword "uses"} need not be declared
 | |
| 138 | again. | |
| 139 | ||
| 140 |   \end{description}
 | |
| 48580 | 141 | *} | 
| 48579 | 142 | |
| 48580 | 143 | subsubsection {* Examples *}
 | 
| 144 | ||
| 145 | text {* See @{file "~~/src/HOL/ROOT"} for a diversity of practically
 | |
| 146 | relevant situations. *} | |
| 48578 | 147 | |
| 148 | ||
| 149 | section {* System build options \label{sec:system-options} *}
 | |
| 150 | ||
| 48580 | 151 | text {* See @{file "~~/etc/options"} for the main defaults provided by
 | 
| 48582 | 152 |   the Isabelle distribution.  Isabelle/jEdit (\secref{sec:tool-jedit})
 | 
| 153 |   includes a simple editing mode @{verbatim "isabelle-options"} for
 | |
| 154 | this file-format. | |
| 48580 | 155 | *} | 
| 48578 | 156 | |
| 157 | ||
| 158 | section {* Invoking the build process \label{sec:tool-build} *}
 | |
| 159 | ||
| 48602 | 160 | text {* The @{tool_def build} tool invokes the build process for
 | 
| 48578 | 161 | Isabelle sessions. It manages dependencies between sessions, | 
| 162 | related sources of theories and auxiliary files, and target heap | |
| 163 | images. Accordingly, it runs instances of the prover process with | |
| 164 | optional document preparation. Its command-line usage | |
| 165 |   is:\footnote{Isabelle/Scala provides the same functionality via
 | |
| 166 |   \texttt{isabelle.Build.build}.}
 | |
| 48602 | 167 | \begin{ttbox}
 | 
| 168 | Usage: isabelle build [OPTIONS] [SESSIONS ...] | |
| 48578 | 169 | |
| 170 | Options are: | |
| 171 | -a select all sessions | |
| 172 | -b build heap images | |
| 48595 | 173 | -c clean build | 
| 48578 | 174 | -d DIR include session directory with ROOT file | 
| 175 | -g NAME select session group NAME | |
| 176 | -j INT maximum number of parallel jobs (default 1) | |
| 177 | -n no build -- test dependencies only | |
| 178 | -o OPTION override session configuration OPTION | |
| 179 | (via NAME=VAL or NAME) | |
| 180 | -s system build mode: produce output in ISABELLE_HOME | |
| 181 | -v verbose | |
| 182 | ||
| 183 | Build and manage Isabelle sessions, depending on implicit | |
| 184 | ISABELLE_BUILD_OPTIONS="..." | |
| 185 | ||
| 186 | ML_PLATFORM="..." | |
| 187 | ML_HOME="..." | |
| 188 | ML_SYSTEM="..." | |
| 189 | ML_OPTIONS="..." | |
| 190 | \end{ttbox}
 | |
| 191 | ||
| 192 | \medskip Isabelle sessions are defined via session ROOT files as | |
| 193 |   described in (\secref{sec:session-root}).  The totality of sessions
 | |
| 194 | is determined by collecting such specifications from all Isabelle | |
| 195 |   component directories (\secref{sec:components}), augmented by more
 | |
| 196 |   directories given via options @{verbatim "-d"}~@{text "DIR"} on the
 | |
| 197 | command line. Each such directory may contain a session | |
| 198 |   \texttt{ROOT} file and an additional catalog file @{verbatim
 | |
| 199 | "etc/sessions"} with further sub-directories (list of lines). Note | |
| 200 |   that a single \texttt{ROOT} file usually defines many sessions;
 | |
| 48591 | 201 | catalogs are only required for extra scalability and modularity | 
| 48578 | 202 | of large libraries. | 
| 203 | ||
| 48604 | 204 | \medskip The subset of sessions to be managed is determined via | 
| 48578 | 205 |   individual @{text "SESSIONS"} given as command-line arguments, or
 | 
| 48604 | 206 |   session groups that are given via one or more options @{verbatim
 | 
| 48578 | 207 |   "-g"}~@{text "NAME"}.  Option @{verbatim "-a"} selects all sessions.
 | 
| 48604 | 208 | The build tool takes session dependencies into account: the set of | 
| 209 | selected sessions is completed by including all ancestors. | |
| 48578 | 210 | |
| 48604 | 211 | \medskip The build process depends on additional options | 
| 212 |   (\secref{sec:system-options}) that are passed to the prover
 | |
| 213 |   eventually.  The settings variable @{setting_ref
 | |
| 214 | ISABELLE_BUILD_OPTIONS} allows to provide additional defaults, e.g.\ | |
| 215 |   \texttt{ISABELLE_BUILD_OPTIONS="document=pdf threads=4"}. Moreover,
 | |
| 216 | the environment of system build options may be augmented on the | |
| 217 |   command line via @{verbatim "-o"}~@{text "name"}@{verbatim
 | |
| 218 |   "="}@{text "value"} or @{verbatim "-o"}~@{text "name"}, which
 | |
| 219 |   abbreviates @{verbatim "-o"}~@{text "name"}@{verbatim"=true"} for
 | |
| 220 |   Boolean options.  Multiple occurrences of @{verbatim "-o"} on the
 | |
| 221 | command-line are applied in the given order. | |
| 48578 | 222 | |
| 223 |   \medskip Option @{verbatim "-b"} ensures that heap images are
 | |
| 224 | produced for all selected sessions. By default, images are only | |
| 225 | saved for inner nodes of the hierarchy of sessions, as required for | |
| 226 | other sessions to continue later on. | |
| 227 | ||
| 48595 | 228 |   \medskip Option @{verbatim "-c"} cleans all descendants of the
 | 
| 229 | selected sessions before performing the specified build operation. | |
| 230 | ||
| 231 |   \medskip Option @{verbatim "-n"} omits the actual build process
 | |
| 232 | after the preparatory stage (including optional cleanup). Note that | |
| 233 | the return code always indicates the status of the set of selected | |
| 234 | sessions. | |
| 235 | ||
| 48578 | 236 |   \medskip Option @{verbatim "-j"} specifies the maximum number of
 | 
| 48604 | 237 | parallel build jobs (prover processes). Each prover process is | 
| 238 | subject to a separate limit of parallel worker threads, cf.\ system | |
| 239 |   option @{system_option_ref threads}.
 | |
| 48578 | 240 | |
| 241 |   \medskip Option @{verbatim "-s"} enables \emph{system mode}, which
 | |
| 242 | means that resulting heap images and log files are stored in | |
| 243 |   @{verbatim "$ISABELLE_HOME/heaps"} instead of the default location
 | |
| 244 |   @{setting ISABELLE_OUTPUT} (which is normally in @{setting
 | |
| 245 | ISABELLE_HOME_USER}, i.e.\ the user's home directory). | |
| 246 | ||
| 247 |   \medskip Option @{verbatim "-v"} enables verbose mode.
 | |
| 248 | *} | |
| 249 | ||
| 250 | subsubsection {* Examples *}
 | |
| 251 | ||
| 252 | text {*
 | |
| 253 | Build a specific logic image: | |
| 254 | \begin{ttbox}
 | |
| 255 | isabelle build -b HOLCF | |
| 256 | \end{ttbox}
 | |
| 257 | ||
| 48604 | 258 | \smallskip Build the main group of logic images: | 
| 48578 | 259 | \begin{ttbox}
 | 
| 260 | isabelle build -b -g main | |
| 261 | \end{ttbox}
 | |
| 262 | ||
| 48595 | 263 | \smallskip Provide a general overview of the status of all Isabelle | 
| 264 | sessions, without building anything: | |
| 48578 | 265 | \begin{ttbox}
 | 
| 266 | isabelle build -a -n -v | |
| 267 | \end{ttbox}
 | |
| 268 | ||
| 48595 | 269 | \smallskip Build all sessions with HTML browser info and PDF | 
| 270 | document preparation: | |
| 48578 | 271 | \begin{ttbox}
 | 
| 272 | isabelle build -a -o browser_info -o document=pdf | |
| 273 | \end{ttbox}
 | |
| 274 | ||
| 48604 | 275 | \smallskip Build all sessions with a maximum of 8 parallel prover | 
| 276 | processes and 4 worker threads each (on a machine with many cores): | |
| 48578 | 277 | \begin{ttbox}
 | 
| 278 | isabelle build -a -j8 -o threads=4 | |
| 279 | \end{ttbox}
 | |
| 48595 | 280 | |
| 281 | \smallskip Build some session images with cleanup of their | |
| 282 | descendants, while retaining their ancestry: | |
| 283 | \begin{ttbox}
 | |
| 284 | isabelle build -b -c HOL-Boogie HOL-SPARK | |
| 285 | \end{ttbox}
 | |
| 286 | ||
| 287 | \smallskip Clean all sessions without building anything: | |
| 288 | \begin{ttbox}
 | |
| 289 | isabelle build -a -n -c | |
| 290 | \end{ttbox}
 | |
| 48578 | 291 | *} | 
| 292 | ||
| 293 | end |