author | wenzelm |
Fri, 04 Jan 2019 23:22:53 +0100 | |
changeset 69593 | 3dda49e08b9d |
parent 68808 | 5467858e9419 |
child 69599 | caa7e406056d |
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 |
|
69593 | 51 |
\<^rail>\<open> |
51417 | 52 |
@{syntax_def session_chapter}: @'chapter' @{syntax name} |
53 |
; |
|
54 |
||
66971 | 55 |
@{syntax_def session_entry}: @'session' @{syntax system_name} groups? dir? '=' \<newline> |
56 |
(@{syntax system_name} '+')? description? options? \<newline> |
|
68292 | 57 |
(sessions?) (theories*) (document_files*) \<newline> (export_files*) |
48579 | 58 |
; |
59 |
groups: '(' (@{syntax name} +) ')' |
|
60 |
; |
|
68808 | 61 |
dir: @'in' @{syntax embedded} |
48579 | 62 |
; |
63 |
description: @'description' @{syntax text} |
|
64 |
; |
|
65 |
options: @'options' opts |
|
66 |
; |
|
48605
e777363440d6
allow negative int values as well, according to real = int | float;
wenzelm
parents:
48604
diff
changeset
|
67 |
opts: '[' ( (@{syntax name} '=' value | @{syntax name}) + ',' ) ']' |
e777363440d6
allow negative int values as well, according to real = int | float;
wenzelm
parents:
48604
diff
changeset
|
68 |
; |
e777363440d6
allow negative int values as well, according to real = int | float;
wenzelm
parents:
48604
diff
changeset
|
69 |
value: @{syntax name} | @{syntax real} |
48579 | 70 |
; |
66916 | 71 |
sessions: @'sessions' (@{syntax system_name}+) |
65374 | 72 |
; |
66970
13857f49d215
clarified ROOT syntax: 'sessions' and 'theories' are optional, but need to be non-empty;
wenzelm
parents:
66916
diff
changeset
|
73 |
theories: @'theories' opts? (theory_entry+) |
48579 | 74 |
; |
66916 | 75 |
theory_entry: @{syntax system_name} ('(' @'global' ')')? |
65504 | 76 |
; |
68808 | 77 |
document_files: @'document_files' ('(' dir ')')? (@{syntax embedded}+) |
68292 | 78 |
; |
68808 | 79 |
export_files: @'export_files' ('(' dir ')')? (@{syntax embedded}+) |
69593 | 80 |
\<close> |
48579 | 81 |
|
61575 | 82 |
\<^descr> \isakeyword{session}~\<open>A = B + body\<close> defines a new session \<open>A\<close> based on |
66759 | 83 |
parent session \<open>B\<close>, with its content given in \<open>body\<close> (imported sessions and |
84 |
theories). Note that a parent (like \<open>HOL\<close>) is mandatory in practical |
|
85 |
applications: only Isabelle/Pure can bootstrap itself from nothing. |
|
48579 | 86 |
|
65504 | 87 |
All such session specifications together describe a hierarchy (graph) of |
61575 | 88 |
sessions, with globally unique names. The new session name \<open>A\<close> should be |
89 |
sufficiently long and descriptive to stand on its own in a potentially large |
|
90 |
library. |
|
48579 | 91 |
|
61575 | 92 |
\<^descr> \isakeyword{session}~\<open>A (groups)\<close> indicates a collection of groups where |
93 |
the new session is a member. Group names are uninterpreted and merely follow |
|
94 |
certain conventions. For example, the Isabelle distribution tags some |
|
95 |
important sessions by the group name called ``\<open>main\<close>''. Other projects may |
|
96 |
invent their own conventions, but this requires some care to avoid clashes |
|
48604 | 97 |
within this unchecked name space. |
48579 | 98 |
|
61575 | 99 |
\<^descr> \isakeyword{session}~\<open>A\<close>~\isakeyword{in}~\<open>dir\<close> specifies an explicit |
100 |
directory for this session; by default this is the current directory of the |
|
101 |
\<^verbatim>\<open>ROOT\<close> file. |
|
48579 | 102 |
|
66759 | 103 |
All theory files are located relatively to the session directory. The prover |
104 |
process is run within the same as its current working directory. |
|
48579 | 105 |
|
61575 | 106 |
\<^descr> \isakeyword{description}~\<open>text\<close> is a free-form annotation for this |
107 |
session. |
|
48579 | 108 |
|
61575 | 109 |
\<^descr> \isakeyword{options}~\<open>[x = a, y = b, z]\<close> defines separate options |
110 |
(\secref{sec:system-options}) that are used when processing this session, |
|
111 |
but \<^emph>\<open>without\<close> propagation to child sessions. Note that \<open>z\<close> abbreviates \<open>z = |
|
112 |
true\<close> for Boolean options. |
|
113 |
||
65504 | 114 |
\<^descr> \isakeyword{sessions}~\<open>names\<close> specifies sessions that are \<^emph>\<open>imported\<close> into |
115 |
the current name space of theories. This allows to refer to a theory \<open>A\<close> |
|
116 |
from session \<open>B\<close> by the qualified name \<open>B.A\<close> --- although it is loaded again |
|
117 |
into the current ML process, which is in contrast to a theory that is |
|
118 |
already present in the \<^emph>\<open>parent\<close> session. |
|
119 |
||
65505 | 120 |
Theories that are imported from other sessions are excluded from the current |
121 |
session document. |
|
122 |
||
61575 | 123 |
\<^descr> \isakeyword{theories}~\<open>options names\<close> specifies a block of theories that |
124 |
are processed within an environment that is augmented by the given options, |
|
125 |
in addition to the global session options given before. Any number of blocks |
|
126 |
of \isakeyword{theories} may be given. Options are only active for each |
|
48604 | 127 |
\isakeyword{theories} block separately. |
48579 | 128 |
|
65374 | 129 |
A theory name that is followed by \<open>(\<close>\isakeyword{global}\<open>)\<close> is treated |
66993 | 130 |
literally in other session specifications or theory imports --- the normal |
131 |
situation is to qualify theory names by the session name; this ensures |
|
132 |
globally unique names in big session graphs. Global theories are usually the |
|
133 |
entry points to major logic sessions: \<open>Pure\<close>, \<open>Main\<close>, \<open>Complex_Main\<close>, |
|
134 |
\<open>HOLCF\<close>, \<open>IFOL\<close>, \<open>FOL\<close>, \<open>ZF\<close>, \<open>ZFC\<close> etc. Regular Isabelle applications |
|
135 |
should not claim any global theory names. |
|
65374 | 136 |
|
61575 | 137 |
\<^descr> \isakeyword{document_files}~\<open>(\<close>\isakeyword{in}~\<open>base_dir) files\<close> lists |
138 |
source files for document preparation, typically \<^verbatim>\<open>.tex\<close> and \<^verbatim>\<open>.sty\<close> for |
|
139 |
{\LaTeX}. Only these explicitly given files are copied from the base |
|
140 |
directory to the document output directory, before formal document |
|
141 |
processing is started (see also \secref{sec:tool-document}). The local path |
|
142 |
structure of the \<open>files\<close> is preserved, which allows to reconstruct the |
|
68292 | 143 |
original directory hierarchy of \<open>base_dir\<close>. The default \<open>base_dir\<close> is |
144 |
\<^verbatim>\<open>document\<close> within the session root directory. |
|
56533
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
wenzelm
parents:
55112
diff
changeset
|
145 |
|
68292 | 146 |
\<^descr> \isakeyword{export_files}~\<open>(\<close>\isakeyword{in}~\<open>target_dir) patterns\<close> writes |
147 |
theory exports to the file-system: the \<open>target_dir\<close> specification is |
|
148 |
relative to the session root directory; its default is \<^verbatim>\<open>export\<close>. Exports |
|
149 |
are selected via \<open>patterns\<close> as in @{tool_ref export} |
|
150 |
(\secref{sec:tool-export}). |
|
58618 | 151 |
\<close> |
48579 | 152 |
|
51055 | 153 |
|
58618 | 154 |
subsubsection \<open>Examples\<close> |
48580 | 155 |
|
61575 | 156 |
text \<open> |
63680 | 157 |
See \<^file>\<open>~~/src/HOL/ROOT\<close> for a diversity of practically relevant situations, |
158 |
although it uses relatively complex quasi-hierarchic naming conventions like |
|
159 |
\<^verbatim>\<open>HOL-SPARK\<close>, \<^verbatim>\<open>HOL-SPARK-Examples\<close>. An alternative is to use unqualified |
|
160 |
names that are relatively long and descriptive, as in the Archive of Formal |
|
67605 | 161 |
Proofs (\<^url>\<open>https://isa-afp.org\<close>), for example. |
61575 | 162 |
\<close> |
48578 | 163 |
|
164 |
||
58618 | 165 |
section \<open>System build options \label{sec:system-options}\<close> |
48578 | 166 |
|
61575 | 167 |
text \<open> |
63680 | 168 |
See \<^file>\<open>~~/etc/options\<close> for the main defaults provided by the Isabelle |
61575 | 169 |
distribution. Isabelle/jEdit @{cite "isabelle-jedit"} includes a simple |
170 |
editing mode \<^verbatim>\<open>isabelle-options\<close> for this file-format. |
|
48693
ceeea46bdeba
"isabelle options" prints Isabelle system options;
wenzelm
parents:
48684
diff
changeset
|
171 |
|
61575 | 172 |
The following options are particularly relevant to build Isabelle sessions, |
173 |
in particular with document preparation (\chref{ch:present}). |
|
51055 | 174 |
|
61575 | 175 |
\<^item> @{system_option_def "browser_info"} controls output of HTML browser |
176 |
info, see also \secref{sec:info}. |
|
51055 | 177 |
|
61575 | 178 |
\<^item> @{system_option_def "document"} specifies the document output format, |
179 |
see @{tool document} option \<^verbatim>\<open>-o\<close> in \secref{sec:tool-document}. In |
|
180 |
practice, the most relevant values are \<^verbatim>\<open>document=false\<close> or |
|
181 |
\<^verbatim>\<open>document=pdf\<close>. |
|
51055 | 182 |
|
61575 | 183 |
\<^item> @{system_option_def "document_output"} specifies an alternative |
184 |
directory for generated output of the document preparation system; the |
|
185 |
default is within the @{setting "ISABELLE_BROWSER_INFO"} hierarchy as |
|
186 |
explained in \secref{sec:info}. See also @{tool mkroot}, which generates a |
|
187 |
default configuration with output readily available to the author of the |
|
188 |
document. |
|
51055 | 189 |
|
61575 | 190 |
\<^item> @{system_option_def "document_variants"} specifies document variants as |
191 |
a colon-separated list of \<open>name=tags\<close> entries, corresponding to @{tool |
|
192 |
document} options \<^verbatim>\<open>-n\<close> and \<^verbatim>\<open>-t\<close>. |
|
51055 | 193 |
|
61575 | 194 |
For example, \<^verbatim>\<open>document_variants=document:outline=/proof,/ML\<close> indicates |
195 |
two documents: the one called \<^verbatim>\<open>document\<close> with default tags, and the other |
|
196 |
called \<^verbatim>\<open>outline\<close> where proofs and ML sections are folded. |
|
51055 | 197 |
|
61575 | 198 |
Document variant names are just a matter of conventions. It is also |
199 |
possible to use different document variant names (without tags) for |
|
200 |
different document root entries, see also \secref{sec:tool-document}. |
|
51055 | 201 |
|
68513 | 202 |
\<^item> @{system_option_def "document_tags"} specifies alternative command tags |
203 |
as a comma-separated list of items: either ``\<open>command\<close>\<^verbatim>\<open>%\<close>\<open>tag\<close>'' for a |
|
204 |
specific command, or ``\<^verbatim>\<open>%\<close>\<open>tag\<close>'' as default for all other commands. This |
|
205 |
is occasionally useful to control the global visibility of commands via |
|
206 |
session options (e.g.\ in \<^verbatim>\<open>ROOT\<close>). |
|
67140 | 207 |
|
61575 | 208 |
\<^item> @{system_option_def "threads"} determines the number of worker threads |
209 |
for parallel checking of theories and proofs. The default \<open>0\<close> means that a |
|
210 |
sensible maximum value is determined by the underlying hardware. For |
|
211 |
machines with many cores or with hyperthreading, this is often requires |
|
212 |
manual adjustment (on the command-line or within personal settings or |
|
213 |
preferences, not within a session \<^verbatim>\<open>ROOT\<close>). |
|
51055 | 214 |
|
63827
b24d0e53dd03
option "checkpoint" helps to fine-tune global heap space management;
wenzelm
parents:
63680
diff
changeset
|
215 |
\<^item> @{system_option_def "checkpoint"} helps to fine-tune the global heap |
b24d0e53dd03
option "checkpoint" helps to fine-tune global heap space management;
wenzelm
parents:
63680
diff
changeset
|
216 |
space management. This is relevant for big sessions that may exhaust the |
b24d0e53dd03
option "checkpoint" helps to fine-tune global heap space management;
wenzelm
parents:
63680
diff
changeset
|
217 |
small 32-bit address space of the ML process (which is used by default). |
b24d0e53dd03
option "checkpoint" helps to fine-tune global heap space management;
wenzelm
parents:
63680
diff
changeset
|
218 |
When the option is enabled for some \isakeyword{theories} block, a full |
b24d0e53dd03
option "checkpoint" helps to fine-tune global heap space management;
wenzelm
parents:
63680
diff
changeset
|
219 |
sharing stage of immutable values in memory happens \<^emph>\<open>before\<close> loading the |
b24d0e53dd03
option "checkpoint" helps to fine-tune global heap space management;
wenzelm
parents:
63680
diff
changeset
|
220 |
specified theories. |
b24d0e53dd03
option "checkpoint" helps to fine-tune global heap space management;
wenzelm
parents:
63680
diff
changeset
|
221 |
|
61575 | 222 |
\<^item> @{system_option_def "condition"} specifies a comma-separated list of |
223 |
process environment variables (or Isabelle settings) that are required for |
|
224 |
the subsequent theories to be processed. Conditions are considered |
|
225 |
``true'' if the corresponding environment value is defined and non-empty. |
|
51055 | 226 |
|
61602 | 227 |
\<^item> @{system_option_def "timeout"} and @{system_option_def "timeout_scale"} |
228 |
specify a real wall-clock timeout for the session as a whole: the two |
|
229 |
values are multiplied and taken as the number of seconds. Typically, |
|
230 |
@{system_option "timeout"} is given for individual sessions, and |
|
231 |
@{system_option "timeout_scale"} as global adjustment to overall hardware |
|
232 |
performance. |
|
233 |
||
234 |
The timer is controlled outside the ML process by the JVM that runs |
|
235 |
Isabelle/Scala. Thus it is relatively reliable in canceling processes that |
|
236 |
get out of control, even if there is a deadlock without CPU time usage. |
|
51055 | 237 |
|
64308 | 238 |
\<^item> @{system_option_def "profiling"} specifies a mode for global ML |
239 |
profiling. Possible values are the empty string (disabled), \<^verbatim>\<open>time\<close> for |
|
69593 | 240 |
\<^ML>\<open>profile_time\<close> and \<^verbatim>\<open>allocations\<close> for \<^ML>\<open>profile_allocations\<close>. |
64308 | 241 |
Results appear near the bottom of the session log file. |
242 |
||
61575 | 243 |
The @{tool_def options} tool prints Isabelle system options. Its |
48693
ceeea46bdeba
"isabelle options" prints Isabelle system options;
wenzelm
parents:
48684
diff
changeset
|
244 |
command-line usage is: |
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
245 |
@{verbatim [display] |
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
246 |
\<open>Usage: isabelle options [OPTIONS] [MORE_OPTIONS ...] |
48693
ceeea46bdeba
"isabelle options" prints Isabelle system options;
wenzelm
parents:
48684
diff
changeset
|
247 |
|
ceeea46bdeba
"isabelle options" prints Isabelle system options;
wenzelm
parents:
48684
diff
changeset
|
248 |
Options are: |
ceeea46bdeba
"isabelle options" prints Isabelle system options;
wenzelm
parents:
48684
diff
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:
50406
diff
changeset
|
251 |
-l list options |
48693
ceeea46bdeba
"isabelle options" prints Isabelle system options;
wenzelm
parents:
48684
diff
changeset
|
252 |
-x FILE export to FILE in YXML format |
ceeea46bdeba
"isabelle options" prints Isabelle system options;
wenzelm
parents:
48684
diff
changeset
|
253 |
|
50531
f841ac0cb757
clarified "isabelle options" command line, to make it more close to "isabelle components";
wenzelm
parents:
50406
diff
changeset
|
254 |
Report Isabelle system options, augmented by MORE_OPTIONS given as |
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
255 |
arguments NAME=VAL or NAME.\<close>} |
48693
ceeea46bdeba
"isabelle options" prints Isabelle system options;
wenzelm
parents:
48684
diff
changeset
|
256 |
|
61575 | 257 |
The command line arguments provide additional system options of the form |
258 |
\<open>name\<close>\<^verbatim>\<open>=\<close>\<open>value\<close> or \<open>name\<close> for Boolean options. |
|
259 |
||
260 |
Option \<^verbatim>\<open>-b\<close> augments the implicit environment of system options by the ones |
|
261 |
of @{setting ISABELLE_BUILD_OPTIONS}, cf.\ \secref{sec:tool-build}. |
|
48693
ceeea46bdeba
"isabelle options" prints Isabelle system options;
wenzelm
parents:
48684
diff
changeset
|
262 |
|
61575 | 263 |
Option \<^verbatim>\<open>-g\<close> prints the value of the given option. Option \<^verbatim>\<open>-l\<close> lists all |
264 |
options with their declaration and current value. |
|
48693
ceeea46bdeba
"isabelle options" prints Isabelle system options;
wenzelm
parents:
48684
diff
changeset
|
265 |
|
61575 | 266 |
Option \<^verbatim>\<open>-x\<close> specifies a file to export the result in YXML format, instead |
267 |
of printing it in human-readable form. |
|
58618 | 268 |
\<close> |
48578 | 269 |
|
270 |
||
58618 | 271 |
section \<open>Invoking the build process \label{sec:tool-build}\<close> |
48578 | 272 |
|
61575 | 273 |
text \<open> |
274 |
The @{tool_def build} tool invokes the build process for Isabelle sessions. |
|
275 |
It manages dependencies between sessions, related sources of theories and |
|
276 |
auxiliary files, and target heap images. Accordingly, it runs instances of |
|
277 |
the prover process with optional document preparation. Its command-line |
|
278 |
usage is:\<^footnote>\<open>Isabelle/Scala provides the same functionality via |
|
61572 | 279 |
\<^verbatim>\<open>isabelle.Build.build\<close>.\<close> |
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
280 |
@{verbatim [display] |
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
281 |
\<open>Usage: isabelle build [OPTIONS] [SESSIONS ...] |
48578 | 282 |
|
283 |
Options are: |
|
66737
2edc0c42c883
option -B for "isabelle build" and "isabelle imports";
wenzelm
parents:
66671
diff
changeset
|
284 |
-B NAME include session NAME and all descendants |
48737
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
wenzelm
parents:
48693
diff
changeset
|
285 |
-D DIR include session directory and select its sessions |
64265 | 286 |
-N cyclic shuffling of NUMA CPU nodes (performance tuning) |
49131 | 287 |
-R operate on requirements of selected sessions |
66745 | 288 |
-S soft build: only observe changes of sources, not heap images |
60106 | 289 |
-X NAME exclude sessions from group NAME and all descendants |
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:
48693
diff
changeset
|
293 |
-d DIR include session directory |
66841 | 294 |
-f fresh build |
48578 | 295 |
-g NAME select session group NAME |
296 |
-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:
59446
diff
changeset
|
297 |
-k KEYWORD check theory sources for conflicts with proposed keywords |
48903 | 298 |
-l list session source files |
48578 | 299 |
-n no build -- test dependencies only |
52056 | 300 |
-o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) |
48578 | 301 |
-s system build mode: produce output in ISABELLE_HOME |
302 |
-v verbose |
|
60106 | 303 |
-x NAME exclude session NAME and all descendants |
48578 | 304 |
|
62596 | 305 |
Build and manage Isabelle sessions, depending on implicit settings: |
306 |
||
48578 | 307 |
ISABELLE_BUILD_OPTIONS="..." |
308 |
||
309 |
ML_PLATFORM="..." |
|
310 |
ML_HOME="..." |
|
311 |
ML_SYSTEM="..." |
|
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
312 |
ML_OPTIONS="..."\<close>} |
48578 | 313 |
|
61406 | 314 |
\<^medskip> |
61575 | 315 |
Isabelle sessions are defined via session ROOT files as described in |
316 |
(\secref{sec:session-root}). The totality of sessions is determined by |
|
317 |
collecting such specifications from all Isabelle component directories |
|
318 |
(\secref{sec:components}), augmented by more directories given via options |
|
319 |
\<^verbatim>\<open>-d\<close>~\<open>DIR\<close> on the command line. Each such directory may contain a session |
|
61503 | 320 |
\<^verbatim>\<open>ROOT\<close> file with several session specifications. |
48578 | 321 |
|
61575 | 322 |
Any session root directory may refer recursively to further directories of |
323 |
the same kind, by listing them in a catalog file \<^verbatim>\<open>ROOTS\<close> line-by-line. This |
|
324 |
helps to organize large collections of session specifications, or to make |
|
66576 | 325 |
\<^verbatim>\<open>-d\<close> command line options persistent (e.g.\ in |
61503 | 326 |
\<^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:
48683
diff
changeset
|
327 |
|
61406 | 328 |
\<^medskip> |
61575 | 329 |
The subset of sessions to be managed is determined via individual \<open>SESSIONS\<close> |
330 |
given as command-line arguments, or session groups that are given via one or |
|
331 |
more options \<^verbatim>\<open>-g\<close>~\<open>NAME\<close>. Option \<^verbatim>\<open>-a\<close> selects all sessions. The build tool |
|
332 |
takes session dependencies into account: the set of selected sessions is |
|
333 |
completed by including all ancestors. |
|
48578 | 334 |
|
61406 | 335 |
\<^medskip> |
68734
c14a2cc9b5ef
isabelle build options -c -x -B refer to imports_graph;
wenzelm
parents:
68523
diff
changeset
|
336 |
One or more options \<^verbatim>\<open>-B\<close>~\<open>NAME\<close> specify base sessions to be included (all |
c14a2cc9b5ef
isabelle build options -c -x -B refer to imports_graph;
wenzelm
parents:
68523
diff
changeset
|
337 |
descendants wrt.\ the session parent or import graph). |
66737
2edc0c42c883
option -B for "isabelle build" and "isabelle imports";
wenzelm
parents:
66671
diff
changeset
|
338 |
|
2edc0c42c883
option -B for "isabelle build" and "isabelle imports";
wenzelm
parents:
66671
diff
changeset
|
339 |
\<^medskip> |
68734
c14a2cc9b5ef
isabelle build options -c -x -B refer to imports_graph;
wenzelm
parents:
68523
diff
changeset
|
340 |
One or more options \<^verbatim>\<open>-x\<close>~\<open>NAME\<close> specify sessions to be excluded (all |
c14a2cc9b5ef
isabelle build options -c -x -B refer to imports_graph;
wenzelm
parents:
68523
diff
changeset
|
341 |
descendants wrt.\ the session parent or import graph). Option \<^verbatim>\<open>-X\<close> is |
c14a2cc9b5ef
isabelle build options -c -x -B refer to imports_graph;
wenzelm
parents:
68523
diff
changeset
|
342 |
analogous to this, but excluded sessions are specified by session group |
c14a2cc9b5ef
isabelle build options -c -x -B refer to imports_graph;
wenzelm
parents:
68523
diff
changeset
|
343 |
membership. |
59892
2a616319c171
added isabelle build option -x, to exclude sessions;
wenzelm
parents:
59891
diff
changeset
|
344 |
|
61406 | 345 |
\<^medskip> |
61575 | 346 |
Option \<^verbatim>\<open>-R\<close> reverses the selection in the sense that it refers to its |
347 |
requirements: all ancestor sessions excluding the original selection. This |
|
348 |
allows to prepare the stage for some build process with different options, |
|
349 |
before running the main build itself (without option \<^verbatim>\<open>-R\<close>). |
|
49131 | 350 |
|
61406 | 351 |
\<^medskip> |
61575 | 352 |
Option \<^verbatim>\<open>-D\<close> is similar to \<^verbatim>\<open>-d\<close>, but selects all sessions that are defined |
353 |
in the given directories. |
|
48737
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
wenzelm
parents:
48693
diff
changeset
|
354 |
|
61406 | 355 |
\<^medskip> |
66745 | 356 |
Option \<^verbatim>\<open>-S\<close> indicates a ``soft build'': the selection is restricted to |
357 |
those sessions that have changed sources (according to actually imported |
|
358 |
theories). The status of heap images is ignored. |
|
359 |
||
360 |
\<^medskip> |
|
61406 | 361 |
The build process depends on additional options |
61575 | 362 |
(\secref{sec:system-options}) that are passed to the prover eventually. The |
363 |
settings variable @{setting_ref ISABELLE_BUILD_OPTIONS} allows to provide |
|
61602 | 364 |
additional defaults, e.g.\ \<^verbatim>\<open>ISABELLE_BUILD_OPTIONS="document=pdf threads=4"\<close>. |
365 |
Moreover, the environment of system build options may be augmented on the |
|
366 |
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 |
|
367 |
\<^verbatim>\<open>-o\<close>~\<open>name\<close>\<^verbatim>\<open>=true\<close> for Boolean options. Multiple occurrences of \<^verbatim>\<open>-o\<close> on |
|
368 |
the command-line are applied in the given order. |
|
48578 | 369 |
|
61406 | 370 |
\<^medskip> |
61575 | 371 |
Option \<^verbatim>\<open>-b\<close> ensures that heap images are produced for all selected |
372 |
sessions. By default, images are only saved for inner nodes of the hierarchy |
|
373 |
of sessions, as required for other sessions to continue later on. |
|
48578 | 374 |
|
61406 | 375 |
\<^medskip> |
68734
c14a2cc9b5ef
isabelle build options -c -x -B refer to imports_graph;
wenzelm
parents:
68523
diff
changeset
|
376 |
Option \<^verbatim>\<open>-c\<close> cleans the selected sessions (all descendants wrt.\ the session |
c14a2cc9b5ef
isabelle build options -c -x -B refer to imports_graph;
wenzelm
parents:
68523
diff
changeset
|
377 |
parent or import graph) before performing the specified build operation. |
48595 | 378 |
|
61406 | 379 |
\<^medskip> |
66841 | 380 |
Option \<^verbatim>\<open>-f\<close> forces a fresh build of all selected sessions and their |
381 |
requirements. |
|
382 |
||
383 |
\<^medskip> |
|
61575 | 384 |
Option \<^verbatim>\<open>-n\<close> omits the actual build process after the preparatory stage |
385 |
(including optional cleanup). Note that the return code always indicates the |
|
386 |
status of the set of selected sessions. |
|
48595 | 387 |
|
61406 | 388 |
\<^medskip> |
61575 | 389 |
Option \<^verbatim>\<open>-j\<close> specifies the maximum number of parallel build jobs (prover |
390 |
processes). Each prover process is subject to a separate limit of parallel |
|
391 |
worker threads, cf.\ system option @{system_option_ref threads}. |
|
48578 | 392 |
|
61406 | 393 |
\<^medskip> |
64265 | 394 |
Option \<^verbatim>\<open>-N\<close> enables cyclic shuffling of NUMA CPU nodes. This may help |
395 |
performance tuning on Linux servers with separate CPU/memory modules. |
|
396 |
||
397 |
\<^medskip> |
|
68523
ccacc84e0251
clarified settings -- avoid hard-wired directories;
wenzelm
parents:
68514
diff
changeset
|
398 |
Option \<^verbatim>\<open>-s\<close> enables \<^emph>\<open>system mode\<close>, which means that session images are |
69593 | 399 |
stored in \<^path>\<open>$ISABELLE_HEAPS_SYSTEM\<close> instead of \<^path>\<open>$ISABELLE_HEAPS\<close>. |
48578 | 400 |
|
61406 | 401 |
\<^medskip> |
61575 | 402 |
Option \<^verbatim>\<open>-v\<close> increases the general level of verbosity. Option \<^verbatim>\<open>-l\<close> lists |
403 |
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:
59446
diff
changeset
|
404 |
|
61406 | 405 |
\<^medskip> |
61575 | 406 |
Option \<^verbatim>\<open>-k\<close> specifies a newly proposed keyword for outer syntax (multiple |
407 |
uses allowed). The theory sources are checked for conflicts wrt.\ this |
|
408 |
hypothetical change of syntax, e.g.\ to reveal occurrences of identifiers |
|
409 |
that need to be quoted. |
|
410 |
\<close> |
|
59891
9ce697050455
added isabelle build option -k, for fast off-line checking of theory sources;
wenzelm
parents:
59446
diff
changeset
|
411 |
|
48578 | 412 |
|
58618 | 413 |
subsubsection \<open>Examples\<close> |
48578 | 414 |
|
58618 | 415 |
text \<open> |
48578 | 416 |
Build a specific logic image: |
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
417 |
@{verbatim [display] \<open>isabelle build -b HOLCF\<close>} |
48578 | 418 |
|
61406 | 419 |
\<^smallskip> |
420 |
Build the main group of logic images: |
|
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
421 |
@{verbatim [display] \<open>isabelle build -b -g main\<close>} |
48578 | 422 |
|
61406 | 423 |
\<^smallskip> |
66748 | 424 |
Build all descendants (and requirements) of \<^verbatim>\<open>FOL\<close> and \<^verbatim>\<open>ZF\<close>: |
425 |
@{verbatim [display] \<open>isabelle build -B FOL -B ZF\<close>} |
|
426 |
||
427 |
\<^smallskip> |
|
428 |
Build all sessions where sources have changed (ignoring heaps): |
|
429 |
@{verbatim [display] \<open>isabelle build -a -S\<close>} |
|
430 |
||
431 |
\<^smallskip> |
|
61575 | 432 |
Provide a general overview of the status of all Isabelle sessions, without |
433 |
building anything: |
|
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
434 |
@{verbatim [display] \<open>isabelle build -a -n -v\<close>} |
48578 | 435 |
|
61406 | 436 |
\<^smallskip> |
61575 | 437 |
Build all sessions with HTML browser info and PDF document preparation: |
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
438 |
@{verbatim [display] \<open>isabelle build -a -o browser_info -o document=pdf\<close>} |
48578 | 439 |
|
61406 | 440 |
\<^smallskip> |
61575 | 441 |
Build all sessions with a maximum of 8 parallel prover processes and 4 |
442 |
worker threads each (on a machine with many cores): |
|
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
443 |
@{verbatim [display] \<open>isabelle build -a -j8 -o threads=4\<close>} |
48595 | 444 |
|
61406 | 445 |
\<^smallskip> |
61575 | 446 |
Build some session images with cleanup of their descendants, while retaining |
447 |
their ancestry: |
|
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
448 |
@{verbatim [display] \<open>isabelle build -b -c HOL-Algebra HOL-Word\<close>} |
48595 | 449 |
|
61406 | 450 |
\<^smallskip> |
451 |
Clean all sessions without building anything: |
|
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
452 |
@{verbatim [display] \<open>isabelle build -a -n -c\<close>} |
48737
f3bbb9ca57d6
added build option -D: include session directory and select its sessions;
wenzelm
parents:
48693
diff
changeset
|
453 |
|
61406 | 454 |
\<^smallskip> |
61575 | 455 |
Build all sessions from some other directory hierarchy, according to the |
456 |
settings variable \<^verbatim>\<open>AFP\<close> that happens to be defined inside the Isabelle |
|
457 |
environment: |
|
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
458 |
@{verbatim [display] \<open>isabelle build -D '$AFP'\<close>} |
49131 | 459 |
|
61406 | 460 |
\<^smallskip> |
61575 | 461 |
Inform about the status of all sessions required for AFP, without building |
462 |
anything yet: |
|
61407
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61406
diff
changeset
|
463 |
@{verbatim [display] \<open>isabelle build -D '$AFP' -R -v -n\<close>} |
58618 | 464 |
\<close> |
48578 | 465 |
|
66671 | 466 |
|
467 |
section \<open>Maintain theory imports wrt.\ session structure\<close> |
|
468 |
||
469 |
text \<open> |
|
470 |
The @{tool_def "imports"} tool helps to maintain theory imports wrt.\ |
|
471 |
session structure. It supports three main operations via options \<^verbatim>\<open>-I\<close>, |
|
472 |
\<^verbatim>\<open>-M\<close>, \<^verbatim>\<open>-U\<close>. Its command-line usage is: @{verbatim [display] |
|
473 |
\<open>Usage: isabelle imports [OPTIONS] [SESSIONS ...] |
|
474 |
||
475 |
Options are: |
|
66737
2edc0c42c883
option -B for "isabelle build" and "isabelle imports";
wenzelm
parents:
66671
diff
changeset
|
476 |
-B NAME include session NAME and all descendants |
66671 | 477 |
-D DIR include session directory and select its sessions |
66851
c75769065548
more informative Imports.Report with actual session imports (minimized);
wenzelm
parents:
66841
diff
changeset
|
478 |
-I operation: report session imports |
66671 | 479 |
-M operation: Mercurial repository check for theory files |
480 |
-R operate on requirements of selected sessions |
|
481 |
-U operation: update theory imports to use session qualifiers |
|
482 |
-X NAME exclude sessions from group NAME and all descendants |
|
483 |
-a select all sessions |
|
484 |
-d DIR include session directory |
|
485 |
-g NAME select session group NAME |
|
486 |
-i incremental update according to session graph structure |
|
487 |
-o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) |
|
488 |
-v verbose |
|
489 |
-x NAME exclude session NAME and all descendants |
|
490 |
||
491 |
Maintain theory imports wrt. session structure. At least one operation |
|
492 |
needs to be specified (see options -I -M -U).\<close>} |
|
493 |
||
494 |
\<^medskip> |
|
495 |
The selection of sessions and session directories works as for @{tool build} |
|
66737
2edc0c42c883
option -B for "isabelle build" and "isabelle imports";
wenzelm
parents:
66671
diff
changeset
|
496 |
via options \<^verbatim>\<open>-B\<close>, \<^verbatim>\<open>-D\<close>, \<^verbatim>\<open>-R\<close>, \<^verbatim>\<open>-X\<close>, \<^verbatim>\<open>-a\<close>, \<^verbatim>\<open>-d\<close>, \<^verbatim>\<open>-g\<close>, \<^verbatim>\<open>-x\<close> (see |
66671 | 497 |
\secref{sec:tool-build}). |
498 |
||
499 |
\<^medskip> |
|
500 |
Option \<^verbatim>\<open>-o\<close> overrides Isabelle system options as for @{tool build} |
|
501 |
(see \secref{sec:tool-build}). |
|
502 |
||
503 |
\<^medskip> |
|
504 |
Option \<^verbatim>\<open>-v\<close> increases the general level of verbosity. |
|
505 |
||
506 |
\<^medskip> |
|
66851
c75769065548
more informative Imports.Report with actual session imports (minimized);
wenzelm
parents:
66841
diff
changeset
|
507 |
Option \<^verbatim>\<open>-I\<close> determines reports session imports: |
c75769065548
more informative Imports.Report with actual session imports (minimized);
wenzelm
parents:
66841
diff
changeset
|
508 |
|
c75769065548
more informative Imports.Report with actual session imports (minimized);
wenzelm
parents:
66841
diff
changeset
|
509 |
\<^descr>[Potential session imports] are derived from old-style use of theory |
c75769065548
more informative Imports.Report with actual session imports (minimized);
wenzelm
parents:
66841
diff
changeset
|
510 |
files from other sessions via the directory structure. After declaring |
c75769065548
more informative Imports.Report with actual session imports (minimized);
wenzelm
parents:
66841
diff
changeset
|
511 |
those as \isakeyword{sessions} in the corresponding \<^verbatim>\<open>ROOT\<close> file entry, a |
c75769065548
more informative Imports.Report with actual session imports (minimized);
wenzelm
parents:
66841
diff
changeset
|
512 |
proper session-qualified theory name can be used (cf.\ option \<^verbatim>\<open>-U\<close>). For |
c75769065548
more informative Imports.Report with actual session imports (minimized);
wenzelm
parents:
66841
diff
changeset
|
513 |
example, adhoc \<^theory_text>\<open>imports\<close> \<^verbatim>\<open>"~~/src/HOL/Library/Multiset"\<close> becomes formal |
c75769065548
more informative Imports.Report with actual session imports (minimized);
wenzelm
parents:
66841
diff
changeset
|
514 |
\<^theory_text>\<open>imports\<close> \<^verbatim>\<open>"HOL-Library.Multiset"\<close> after adding \isakeyword{sessions} |
c75769065548
more informative Imports.Report with actual session imports (minimized);
wenzelm
parents:
66841
diff
changeset
|
515 |
\<^verbatim>\<open>"HOL-Library"\<close> to the \<^verbatim>\<open>ROOT\<close> entry. |
c75769065548
more informative Imports.Report with actual session imports (minimized);
wenzelm
parents:
66841
diff
changeset
|
516 |
|
c75769065548
more informative Imports.Report with actual session imports (minimized);
wenzelm
parents:
66841
diff
changeset
|
517 |
\<^descr>[Actual session imports] are derived from the session qualifiers of all |
c75769065548
more informative Imports.Report with actual session imports (minimized);
wenzelm
parents:
66841
diff
changeset
|
518 |
currently imported theories. This helps to minimize dependencies in the |
c75769065548
more informative Imports.Report with actual session imports (minimized);
wenzelm
parents:
66841
diff
changeset
|
519 |
session import structure to what is actually required. |
66671 | 520 |
|
521 |
\<^medskip> |
|
522 |
Option \<^verbatim>\<open>-M\<close> checks imported theories against the Mercurial repositories of |
|
523 |
the underlying session directories; non-repository directories are ignored. |
|
524 |
This helps to find files that are accidentally ignored, e.g.\ due to |
|
525 |
rearrangements of the session structure. |
|
526 |
||
527 |
\<^medskip> |
|
528 |
Option \<^verbatim>\<open>-U\<close> updates theory imports with old-style directory specifications |
|
529 |
to canonical session-qualified theory names, according to the theory name |
|
530 |
space imported via \isakeyword{sessions} within the \<^verbatim>\<open>ROOT\<close> specification. |
|
531 |
||
532 |
Option \<^verbatim>\<open>-i\<close> modifies the meaning of option \<^verbatim>\<open>-U\<close> to proceed incrementally, |
|
533 |
following to the session graph structure in bottom-up order. This may |
|
534 |
lead to more accurate results in complex session hierarchies. |
|
535 |
\<close> |
|
536 |
||
537 |
subsubsection \<open>Examples\<close> |
|
538 |
||
539 |
text \<open> |
|
540 |
Determine potential session imports for some project directory: |
|
541 |
@{verbatim [display] \<open>isabelle imports -I -D 'some/where/My_Project'\<close>} |
|
542 |
||
543 |
\<^smallskip> |
|
544 |
Mercurial repository check for some project directory: |
|
545 |
@{verbatim [display] \<open>isabelle imports -M -D 'some/where/My_Project'\<close>} |
|
546 |
||
547 |
\<^smallskip> |
|
548 |
Incremental update of theory imports for some project directory: |
|
549 |
@{verbatim [display] \<open>isabelle imports -U -i -D 'some/where/My_Project'\<close>} |
|
550 |
\<close> |
|
551 |
||
68116 | 552 |
|
68152
619de043389f
guard result exports via export_pattern -- avoid bombing client via huge blobs;
wenzelm
parents:
68116
diff
changeset
|
553 |
section \<open>Retrieve theory exports \label{sec:tool-export}\<close> |
68116 | 554 |
|
555 |
text \<open> |
|
556 |
The @{tool_def "export"} tool retrieves theory exports from the session |
|
557 |
database. Its command-line usage is: @{verbatim [display] |
|
558 |
\<open>Usage: isabelle export [OPTIONS] SESSION |
|
559 |
||
560 |
Options are: |
|
68314
2acbf8129d8b
clarified option -O: avoid conflict with build/dump option -D;
wenzelm
parents:
68292
diff
changeset
|
561 |
-O DIR output directory for exported files (default: "export") |
68116 | 562 |
-d DIR include session directory |
563 |
-l list exports |
|
564 |
-n no build of session |
|
565 |
-o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) |
|
566 |
-s system build mode for session image |
|
68514 | 567 |
-x PATTERN extract files matching pattern (e.g.\ "*:**" for all) |
68116 | 568 |
|
569 |
List or export theory exports for SESSION: named blobs produced by |
|
68290 | 570 |
isabelle build. Option -l or -x is required; option -x may be repeated. |
68116 | 571 |
|
572 |
The PATTERN language resembles glob patterns in the shell, with ? and * |
|
573 |
(both excluding ":" and "/"), ** (excluding ":"), and [abc] or [^abc], |
|
574 |
and variants {pattern1,pattern2,pattern3}.\<close>} |
|
575 |
||
576 |
\<^medskip> |
|
577 |
The specified session is updated via @{tool build} |
|
578 |
(\secref{sec:tool-build}), with the same options \<^verbatim>\<open>-d\<close>, \<^verbatim>\<open>-o\<close>, \<^verbatim>\<open>-s\<close>. The |
|
579 |
option \<^verbatim>\<open>-n\<close> suppresses the implicit build process: it means that a |
|
580 |
potentially outdated session database is used! |
|
581 |
||
582 |
\<^medskip> |
|
583 |
Option \<^verbatim>\<open>-l\<close> lists all stored exports, with compound names |
|
584 |
\<open>theory\<close>\<^verbatim>\<open>:\<close>\<open>name\<close>. |
|
585 |
||
586 |
\<^medskip> |
|
587 |
Option \<^verbatim>\<open>-x\<close> extracts stored exports whose compound name matches the given |
|
588 |
pattern. Note that wild cards ``\<^verbatim>\<open>?\<close>'' and ``\<^verbatim>\<open>*\<close>'' do not match the |
|
589 |
separators ``\<^verbatim>\<open>:\<close>'' and ``\<^verbatim>\<open>/\<close>''; the wild card \<^verbatim>\<open>**\<close> matches over directory |
|
590 |
name hierarchies separated by ``\<^verbatim>\<open>/\<close>''. Thus the pattern ``\<^verbatim>\<open>*:**\<close>'' matches |
|
68290 | 591 |
\<^emph>\<open>all\<close> theory exports. Multiple options \<^verbatim>\<open>-x\<close> refer to the union of all |
592 |
specified patterns. |
|
68116 | 593 |
|
68314
2acbf8129d8b
clarified option -O: avoid conflict with build/dump option -D;
wenzelm
parents:
68292
diff
changeset
|
594 |
Option \<^verbatim>\<open>-O\<close> specifies an alternative output directory for option \<^verbatim>\<open>-x\<close>: the |
68116 | 595 |
default is \<^verbatim>\<open>export\<close> within the current directory. Each theory creates its |
596 |
own sub-directory hierarchy, using the session-qualified theory name. |
|
597 |
\<close> |
|
598 |
||
68348 | 599 |
|
600 |
section \<open>Dump PIDE session database \label{sec:tool-dump}\<close> |
|
601 |
||
602 |
text \<open> |
|
603 |
The @{tool_def "dump"} tool dumps information from the cumulative PIDE |
|
604 |
session database (which is processed on the spot). Its command-line usage |
|
605 |
is: @{verbatim [display] |
|
606 |
\<open>Usage: isabelle dump [OPTIONS] [SESSIONS ...] |
|
607 |
||
608 |
Options are: |
|
609 |
-A NAMES dump named aspects (default: ...) |
|
610 |
-B NAME include session NAME and all descendants |
|
611 |
-D DIR include session directory and select its sessions |
|
612 |
-O DIR output directory for dumped files (default: "dump") |
|
613 |
-R operate on requirements of selected sessions |
|
614 |
-X NAME exclude sessions from group NAME and all descendants |
|
615 |
-a select all sessions |
|
616 |
-d DIR include session directory |
|
617 |
-g NAME select session group NAME |
|
618 |
-l NAME logic session name (default ISABELLE_LOGIC="HOL") |
|
619 |
-o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) |
|
620 |
-s system build mode for logic image |
|
621 |
-v verbose |
|
622 |
-x NAME exclude session NAME and all descendants |
|
623 |
||
624 |
Dump cumulative PIDE session database, with the following aspects: |
|
625 |
...\<close>} |
|
626 |
||
627 |
\<^medskip> Options \<^verbatim>\<open>-B\<close>, \<^verbatim>\<open>-D\<close>, \<^verbatim>\<open>-R\<close>, \<^verbatim>\<open>-X\<close>, \<^verbatim>\<open>-a\<close>, \<^verbatim>\<open>-d\<close>, \<^verbatim>\<open>-g\<close>, \<^verbatim>\<open>-x\<close> and the |
|
628 |
remaining command-line arguments specify sessions as in @{tool build} |
|
629 |
(\secref{sec:tool-build}): the cumulative PIDE database of all their loaded |
|
630 |
theories is dumped to the output directory of option \<^verbatim>\<open>-O\<close> (default: \<^verbatim>\<open>dump\<close> |
|
631 |
in the current directory). |
|
632 |
||
633 |
\<^medskip> Option \<^verbatim>\<open>-l\<close> specifies a logic image for the underlying prover process: |
|
634 |
its theories are not processed again, and their PIDE session database is |
|
635 |
excluded from the dump. Option \<^verbatim>\<open>-s\<close> enables \<^emph>\<open>system mode\<close> when building |
|
636 |
the logic image (\secref{sec:tool-build}). |
|
637 |
||
638 |
\<^medskip> Option \<^verbatim>\<open>-o\<close> overrides Isabelle system options as for @{tool build} |
|
639 |
(\secref{sec:tool-build}). |
|
640 |
||
641 |
\<^medskip> Option \<^verbatim>\<open>-v\<close> increases the general level of verbosity. |
|
642 |
||
643 |
\<^medskip> Option \<^verbatim>\<open>-A\<close> specifies named aspects of the dump, as a comma-separated |
|
644 |
list. The default is to dump all known aspects, as given in the command-line |
|
645 |
usage of the tool. The underlying Isabelle/Scala function |
|
646 |
\<^verbatim>\<open>isabelle.Dump.dump()\<close> takes aspects as user-defined operations on the |
|
647 |
final PIDE state and document version. This allows to imitate Prover IDE |
|
648 |
rendering under program control. |
|
649 |
\<close> |
|
650 |
||
651 |
||
652 |
subsubsection \<open>Examples\<close> |
|
653 |
||
654 |
text \<open> |
|
655 |
Dump all Isabelle/ZF sessions (which are rather small): |
|
656 |
@{verbatim [display] \<open>isabelle dump -v -B ZF\<close>} |
|
657 |
||
658 |
\<^smallskip> |
|
659 |
Dump the quite substantial \<^verbatim>\<open>HOL-Analysis\<close> session, using main Isabelle/HOL |
|
660 |
as starting point: |
|
661 |
@{verbatim [display] \<open>isabelle dump -v -l HOL HOL-Analysis\<close>} |
|
662 |
||
663 |
\<^smallskip> |
|
664 |
Dump all sessions connected to HOL-Analysis, including a full bootstrap of |
|
665 |
Isabelle/HOL from Isabelle/Pure: |
|
666 |
@{verbatim [display] \<open>isabelle dump -v -l Pure -B HOL-Analysis\<close>} |
|
667 |
||
668 |
This results in uniform PIDE markup for everything, except for the |
|
669 |
Isabelle/Pure bootstrap process itself. Producing that on the spot requires |
|
670 |
several GB of heap space, both for the Isabelle/Scala and Isabelle/ML |
|
671 |
process (in 64bit mode). Here are some relevant settings (\secref{sec:boot}) |
|
672 |
for such ambitious applications: |
|
673 |
@{verbatim [display] |
|
674 |
\<open>ISABELLE_TOOL_JAVA_OPTIONS="-Xms4g -Xmx32g -Xss16m" |
|
675 |
ML_OPTIONS="--minheap 4G --maxheap 32G" |
|
676 |
\<close>} |
|
677 |
||
678 |
\<close> |
|
679 |
||
48578 | 680 |
end |