wenzelm@48578: theory Sessions wenzelm@48578: imports Base wenzelm@48578: begin wenzelm@48578: wenzelm@48584: chapter {* Isabelle sessions and build management \label{ch:session} *} wenzelm@48584: wenzelm@48584: text {* An Isabelle \emph{session} consists of a collection of related wenzelm@48604: theories that may be associated with formal documents (see also wenzelm@48604: \chref{ch:present}). There is also a notion of \emph{persistent wenzelm@48604: heap} image to capture the state of a session, similar to wenzelm@48604: object-code in compiled programming languages. Thus the concept of wenzelm@48604: session resembles that of a ``project'' in common IDE environments, wenzelm@48604: but the specific name emphasizes the connection to interactive wenzelm@48604: theorem proving: the session wraps-up the results of wenzelm@48604: user-interaction with the prover in a persistent form. wenzelm@48584: wenzelm@48604: Application sessions are built on a given parent session, which may wenzelm@48604: be built recursively on other parents. Following this path in the wenzelm@48604: hierarchy eventually leads to some major object-logic session like wenzelm@48604: @{text "HOL"}, which itself is based on @{text "Pure"} as the common wenzelm@48604: root of all sessions. wenzelm@48584: wenzelm@48604: Processing sessions may take considerable time. Isabelle build wenzelm@48604: management helps to organize this efficiently. This includes wenzelm@48604: support for parallel build jobs, in addition to the multithreaded wenzelm@48604: theory and proof checking that is already provided by the prover wenzelm@48604: process itself. *} wenzelm@48604: wenzelm@48578: wenzelm@48578: section {* Session ROOT specifications \label{sec:session-root} *} wenzelm@48578: wenzelm@48579: text {* Session specifications reside in files called @{verbatim ROOT} wenzelm@48579: within certain directories, such as the home locations of registered wenzelm@48579: Isabelle components or additional project directories given by the wenzelm@48579: user. wenzelm@48579: wenzelm@48579: The ROOT file format follows the lexical conventions of the wenzelm@48579: \emph{outer syntax} of Isabelle/Isar, see also wenzelm@48579: \cite{isabelle-isar-ref}. This defines common forms like wenzelm@48579: identifiers, names, quoted strings, verbatim text, nested comments wenzelm@48579: etc. The grammar for a single @{syntax session_entry} is given as wenzelm@48579: syntax diagram below; each ROOT file may contain multiple session wenzelm@48579: specifications like this. wenzelm@48579: wenzelm@48582: Isabelle/jEdit (\secref{sec:tool-jedit}) includes a simple editing wenzelm@48582: mode @{verbatim "isabelle-root"} for session ROOT files. wenzelm@48579: wenzelm@48579: @{rail " wenzelm@48579: @{syntax_def session_entry}: @'session' spec '=' (@{syntax name} '+')? body wenzelm@48579: ; wenzelm@48579: body: description? options? ( theories * ) files? wenzelm@48579: ; wenzelm@48579: spec: @{syntax name} '!'? groups? dir? wenzelm@48579: ; wenzelm@48579: groups: '(' (@{syntax name} +) ')' wenzelm@48579: ; wenzelm@48579: dir: @'in' @{syntax name} wenzelm@48579: ; wenzelm@48579: description: @'description' @{syntax text} wenzelm@48579: ; wenzelm@48579: options: @'options' opts wenzelm@48579: ; wenzelm@48605: opts: '[' ( (@{syntax name} '=' value | @{syntax name}) + ',' ) ']' wenzelm@48605: ; wenzelm@48605: value: @{syntax name} | @{syntax real} wenzelm@48579: ; wenzelm@48579: theories: @'theories' opts? ( @{syntax name} + ) wenzelm@48579: ; wenzelm@48579: files: @'files' ( @{syntax name} + ) wenzelm@48579: "} wenzelm@48579: wenzelm@48579: \begin{description} wenzelm@48579: wenzelm@48579: \item \isakeyword{session}~@{text "A = B + body"} defines a new wenzelm@48579: session @{text "A"} based on parent session @{text "B"}, with its wenzelm@48579: content given in @{text body} (theories and auxiliary source files). wenzelm@48579: Note that a parent (like @{text "HOL"}) is mandatory in practical wenzelm@48579: applications: only Isabelle/Pure can bootstrap itself from nothing. wenzelm@48579: wenzelm@48579: All such session specifications together describe a hierarchy (tree) wenzelm@48579: of sessions, with globally unique names. By default, names are wenzelm@48579: derived from parent ones by concatenation, i.e.\ @{text "B\A"} wenzelm@48579: above. Cumulatively, this leads to session paths of the form @{text wenzelm@48579: "X\Y\Z\W"}. Note that in the specification, wenzelm@48579: @{text B} is already such a fully-qualified name, while @{text "A"} wenzelm@48579: is the new base name. wenzelm@48579: wenzelm@48579: \item \isakeyword{session}~@{text "A! = B"} indicates a fresh start wenzelm@48579: in the naming scheme: the session is called just @{text "A"} instead wenzelm@48579: of @{text "B\A"}. Here the name @{text "A"} should be wenzelm@48579: sufficiently long to stand on its own in a potentially large wenzelm@48579: library. wenzelm@48579: wenzelm@48579: \item \isakeyword{session}~@{text "A (groups)"} indicates a wenzelm@48579: collection of groups where the new session is a member. Group names wenzelm@48579: are uninterpreted and merely follow certain conventions. For wenzelm@48579: example, the Isabelle distribution tags some important sessions by wenzelm@48604: the group name called ``@{text "main"}''. Other projects may invent wenzelm@48604: their own conventions, but this requires some care to avoid clashes wenzelm@48604: within this unchecked name space. wenzelm@48579: wenzelm@48579: \item \isakeyword{session}~@{text "A"}~\isakeyword{in}~@{text "dir"} wenzelm@48579: specifies an explicit directory for this session. By default, wenzelm@48579: \isakeyword{session}~@{text "A"} abbreviates wenzelm@48579: \isakeyword{session}~@{text "A"}~\isakeyword{in}~@{text "A"}. This wenzelm@48579: accommodates the common scheme where some base directory contains wenzelm@48579: several sessions in sub-directories of corresponding names. Another wenzelm@48579: common scheme is \isakeyword{session}~@{text wenzelm@48579: "A"}~\isakeyword{in}~@{verbatim "\".\""} to refer to the current wenzelm@48579: directory of the ROOT file. wenzelm@48579: wenzelm@48579: All theories and auxiliary source files are located relatively to wenzelm@48579: the session directory. The prover process is run within the same as wenzelm@48579: its current working directory. wenzelm@48579: wenzelm@48579: \item \isakeyword{description}~@{text "text"} is a free-form wenzelm@48579: annotation for this session. wenzelm@48579: wenzelm@48579: \item \isakeyword{options}~@{text "[x = a, y = b, z]"} defines wenzelm@48604: separate options (\secref{sec:system-options}) that are used when wenzelm@48604: processing this session, but \emph{without} propagation to child wenzelm@48604: sessions. Note that @{text "z"} abbreviates @{text "z = true"} for wenzelm@48604: Boolean options. wenzelm@48579: wenzelm@48579: \item \isakeyword{theories}~@{text "options names"} specifies a wenzelm@48579: block of theories that are processed within an environment that is wenzelm@48604: augmented by the given options, in addition to the global session wenzelm@48604: options given before. Any number of blocks of \isakeyword{theories} wenzelm@48604: may be given. Options are only active for each wenzelm@48604: \isakeyword{theories} block separately. wenzelm@48579: wenzelm@48579: \item \isakeyword{files}~@{text "files"} lists additional source wenzelm@48604: files that are involved in the processing of this session. This wenzelm@48604: should cover anything outside the formal content of the theory wenzelm@48579: sources, say some auxiliary {\TeX} files that are required for wenzelm@48579: document processing. In contrast, files that are specified in wenzelm@48579: formal theory headers as @{keyword "uses"} need not be declared wenzelm@48579: again. wenzelm@48579: wenzelm@48579: \end{description} wenzelm@48580: *} wenzelm@48579: wenzelm@48580: subsubsection {* Examples *} wenzelm@48580: wenzelm@48580: text {* See @{file "~~/src/HOL/ROOT"} for a diversity of practically wenzelm@48580: relevant situations. *} wenzelm@48578: wenzelm@48578: wenzelm@48578: section {* System build options \label{sec:system-options} *} wenzelm@48578: wenzelm@48580: text {* See @{file "~~/etc/options"} for the main defaults provided by wenzelm@48582: the Isabelle distribution. Isabelle/jEdit (\secref{sec:tool-jedit}) wenzelm@48582: includes a simple editing mode @{verbatim "isabelle-options"} for wenzelm@48582: this file-format. wenzelm@48693: wenzelm@48693: The @{tool_def options} tool prints Isabelle system options. Its wenzelm@48693: command-line usage is: wenzelm@48693: \begin{ttbox} wenzelm@48693: Usage: isabelle options [OPTIONS] [MORE_OPTIONS ...] wenzelm@48693: wenzelm@48693: Options are: wenzelm@48693: -b include $ISABELLE_BUILD_OPTIONS wenzelm@48693: -x FILE export to FILE in YXML format wenzelm@48693: wenzelm@48693: Print Isabelle system options, augmented by MORE_OPTIONS given as wenzelm@48693: arguments NAME=VAL or NAME. wenzelm@48693: \end{ttbox} wenzelm@48693: wenzelm@48693: The command line arguments provide additional system options of the wenzelm@48693: form @{text "name"}@{verbatim "="}@{text "value"} or @{text name} wenzelm@48693: for Boolean options. wenzelm@48693: wenzelm@48693: Option @{verbatim "-b"} augments the implicit environment of system wenzelm@48693: options by the ones of @{setting ISABELLE_BUILD_OPTIONS}, cf.\ wenzelm@48693: \secref{sec:tool-build}. wenzelm@48693: wenzelm@48693: Option @{verbatim "-x"} specifies a file to export the result in wenzelm@48693: YXML format, instead of printing it in human-readable form. wenzelm@48580: *} wenzelm@48578: wenzelm@48578: wenzelm@48578: section {* Invoking the build process \label{sec:tool-build} *} wenzelm@48578: wenzelm@48602: text {* The @{tool_def build} tool invokes the build process for wenzelm@48578: Isabelle sessions. It manages dependencies between sessions, wenzelm@48578: related sources of theories and auxiliary files, and target heap wenzelm@48578: images. Accordingly, it runs instances of the prover process with wenzelm@48578: optional document preparation. Its command-line usage wenzelm@48578: is:\footnote{Isabelle/Scala provides the same functionality via wenzelm@48578: \texttt{isabelle.Build.build}.} wenzelm@48602: \begin{ttbox} wenzelm@48602: Usage: isabelle build [OPTIONS] [SESSIONS ...] wenzelm@48578: wenzelm@48578: Options are: wenzelm@48578: -a select all sessions wenzelm@48578: -b build heap images wenzelm@48595: -c clean build wenzelm@48578: -d DIR include session directory with ROOT file wenzelm@48578: -g NAME select session group NAME wenzelm@48578: -j INT maximum number of parallel jobs (default 1) wenzelm@48578: -n no build -- test dependencies only wenzelm@48578: -o OPTION override session configuration OPTION wenzelm@48578: (via NAME=VAL or NAME) wenzelm@48578: -s system build mode: produce output in ISABELLE_HOME wenzelm@48578: -v verbose wenzelm@48578: wenzelm@48578: Build and manage Isabelle sessions, depending on implicit wenzelm@48578: ISABELLE_BUILD_OPTIONS="..." wenzelm@48578: wenzelm@48578: ML_PLATFORM="..." wenzelm@48578: ML_HOME="..." wenzelm@48578: ML_SYSTEM="..." wenzelm@48578: ML_OPTIONS="..." wenzelm@48578: \end{ttbox} wenzelm@48578: wenzelm@48578: \medskip Isabelle sessions are defined via session ROOT files as wenzelm@48578: described in (\secref{sec:session-root}). The totality of sessions wenzelm@48578: is determined by collecting such specifications from all Isabelle wenzelm@48578: component directories (\secref{sec:components}), augmented by more wenzelm@48578: directories given via options @{verbatim "-d"}~@{text "DIR"} on the wenzelm@48578: command line. Each such directory may contain a session wenzelm@48650: \texttt{ROOT} file with several session specifications. wenzelm@48578: wenzelm@48684: Any session root directory may refer recursively to further wenzelm@48684: directories of the same kind, by listing them in a catalog file wenzelm@48684: @{verbatim "ROOTS"} line-by-line. This helps to organize large wenzelm@48684: collections of session specifications, or to make @{verbatim "-d"} wenzelm@48684: command line options persistent (say within @{verbatim wenzelm@48684: "$ISABELLE_HOME_USER/ROOTS"}). wenzelm@48684: wenzelm@48604: \medskip The subset of sessions to be managed is determined via wenzelm@48578: individual @{text "SESSIONS"} given as command-line arguments, or wenzelm@48604: session groups that are given via one or more options @{verbatim wenzelm@48578: "-g"}~@{text "NAME"}. Option @{verbatim "-a"} selects all sessions. wenzelm@48604: The build tool takes session dependencies into account: the set of wenzelm@48604: selected sessions is completed by including all ancestors. wenzelm@48578: wenzelm@48604: \medskip The build process depends on additional options wenzelm@48604: (\secref{sec:system-options}) that are passed to the prover wenzelm@48604: eventually. The settings variable @{setting_ref wenzelm@48604: ISABELLE_BUILD_OPTIONS} allows to provide additional defaults, e.g.\ wenzelm@48604: \texttt{ISABELLE_BUILD_OPTIONS="document=pdf threads=4"}. Moreover, wenzelm@48604: the environment of system build options may be augmented on the wenzelm@48604: command line via @{verbatim "-o"}~@{text "name"}@{verbatim wenzelm@48604: "="}@{text "value"} or @{verbatim "-o"}~@{text "name"}, which wenzelm@48604: abbreviates @{verbatim "-o"}~@{text "name"}@{verbatim"=true"} for wenzelm@48604: Boolean options. Multiple occurrences of @{verbatim "-o"} on the wenzelm@48604: command-line are applied in the given order. wenzelm@48578: wenzelm@48578: \medskip Option @{verbatim "-b"} ensures that heap images are wenzelm@48578: produced for all selected sessions. By default, images are only wenzelm@48578: saved for inner nodes of the hierarchy of sessions, as required for wenzelm@48578: other sessions to continue later on. wenzelm@48578: wenzelm@48595: \medskip Option @{verbatim "-c"} cleans all descendants of the wenzelm@48595: selected sessions before performing the specified build operation. wenzelm@48595: wenzelm@48595: \medskip Option @{verbatim "-n"} omits the actual build process wenzelm@48595: after the preparatory stage (including optional cleanup). Note that wenzelm@48595: the return code always indicates the status of the set of selected wenzelm@48595: sessions. wenzelm@48595: wenzelm@48578: \medskip Option @{verbatim "-j"} specifies the maximum number of wenzelm@48604: parallel build jobs (prover processes). Each prover process is wenzelm@48604: subject to a separate limit of parallel worker threads, cf.\ system wenzelm@48604: option @{system_option_ref threads}. wenzelm@48578: wenzelm@48578: \medskip Option @{verbatim "-s"} enables \emph{system mode}, which wenzelm@48578: means that resulting heap images and log files are stored in wenzelm@48578: @{verbatim "$ISABELLE_HOME/heaps"} instead of the default location wenzelm@48578: @{setting ISABELLE_OUTPUT} (which is normally in @{setting wenzelm@48578: ISABELLE_HOME_USER}, i.e.\ the user's home directory). wenzelm@48578: wenzelm@48578: \medskip Option @{verbatim "-v"} enables verbose mode. wenzelm@48578: *} wenzelm@48578: wenzelm@48578: subsubsection {* Examples *} wenzelm@48578: wenzelm@48578: text {* wenzelm@48578: Build a specific logic image: wenzelm@48578: \begin{ttbox} wenzelm@48578: isabelle build -b HOLCF wenzelm@48578: \end{ttbox} wenzelm@48578: wenzelm@48604: \smallskip Build the main group of logic images: wenzelm@48578: \begin{ttbox} wenzelm@48578: isabelle build -b -g main wenzelm@48578: \end{ttbox} wenzelm@48578: wenzelm@48595: \smallskip Provide a general overview of the status of all Isabelle wenzelm@48595: sessions, without building anything: wenzelm@48578: \begin{ttbox} wenzelm@48578: isabelle build -a -n -v wenzelm@48578: \end{ttbox} wenzelm@48578: wenzelm@48595: \smallskip Build all sessions with HTML browser info and PDF wenzelm@48595: document preparation: wenzelm@48578: \begin{ttbox} wenzelm@48578: isabelle build -a -o browser_info -o document=pdf wenzelm@48578: \end{ttbox} wenzelm@48578: wenzelm@48604: \smallskip Build all sessions with a maximum of 8 parallel prover wenzelm@48604: processes and 4 worker threads each (on a machine with many cores): wenzelm@48578: \begin{ttbox} wenzelm@48578: isabelle build -a -j8 -o threads=4 wenzelm@48578: \end{ttbox} wenzelm@48595: wenzelm@48595: \smallskip Build some session images with cleanup of their wenzelm@48595: descendants, while retaining their ancestry: wenzelm@48595: \begin{ttbox} wenzelm@48595: isabelle build -b -c HOL-Boogie HOL-SPARK wenzelm@48595: \end{ttbox} wenzelm@48595: wenzelm@48595: \smallskip Clean all sessions without building anything: wenzelm@48595: \begin{ttbox} wenzelm@48595: isabelle build -a -n -c wenzelm@48595: \end{ttbox} wenzelm@48578: *} wenzelm@48578: wenzelm@48683: wenzelm@48683: section {* Preparing session root directories \label{sec:tool-mkroot} *} wenzelm@48683: wenzelm@48683: text {* The @{tool_def mkroot} tool prepares Isabelle session source wenzelm@48683: directories, including some @{verbatim ROOT} entry, an example wenzelm@48683: theory file, and some initial configuration for document preparation wenzelm@48683: (see also \chref{ch:present}). The usage of @{tool mkroot} is: wenzelm@48683: wenzelm@48683: \begin{ttbox} wenzelm@48683: Usage: isabelle mkroot NAME wenzelm@48683: wenzelm@48683: Prepare session root directory, adding session NAME with wenzelm@48683: built-in document preparation. wenzelm@48683: \end{ttbox} wenzelm@48683: wenzelm@48683: All session-specific files are placed into a separate sub-directory wenzelm@48683: given as @{verbatim NAME} above. The @{verbatim ROOT} file is in wenzelm@48683: the parent position relative to that --- it could refer to several wenzelm@48683: such sessions. The @{tool mkroot} tool is conservative in the sense wenzelm@48683: that does not overwrite an existing session sub-directory; an wenzelm@48683: already existing @{verbatim ROOT} file is extended. wenzelm@48683: wenzelm@48683: The implicit Isabelle settings variable @{setting ISABELLE_LOGIC} wenzelm@48683: specifies the parent session, and @{setting wenzelm@48683: ISABELLE_DOCUMENT_FORMAT} the document format to be filled filled wenzelm@48683: into the generated @{verbatim ROOT} file. *} wenzelm@48683: wenzelm@48683: wenzelm@48683: subsubsection {* Examples *} wenzelm@48683: wenzelm@48683: text {* The following produces an example session, relatively to the wenzelm@48683: @{verbatim ROOT} in the current directory: wenzelm@48683: \begin{ttbox} wenzelm@48683: isabelle mkroot Test && isabelle build -v -d. Test wenzelm@48683: \end{ttbox} wenzelm@48683: wenzelm@48683: Option @{verbatim "-v"} is not required, but useful to reveal the wenzelm@48683: the location of generated documents. *} wenzelm@48683: wenzelm@48578: end