src/Doc/System/Sessions.thy
author wenzelm
Sat, 27 Aug 2022 17:46:58 +0200
changeset 76005 a9bbf075f431
parent 75998 c36e5c6f3069
child 76089 13ae8dff47b6
permissions -rw-r--r--
include groups from 'chapter_definition' in session info, based on the state of chapter_defs after processing all ROOT files (thus the declaration order does not matter);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61656
cfabbc083977 more uniform jEdit properties;
wenzelm
parents: 61602
diff changeset
     1
(*:maxLineLen=78:*)
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
     2
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
     3
theory Sessions
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
     4
imports Base
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
     5
begin
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
     6
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
     7
chapter \<open>Isabelle sessions and build management \label{ch:session}\<close>
48584
8026c852cc10 some introduction on sessions;
wenzelm
parents: 48582
diff changeset
     8
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
     9
text \<open>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    10
  An Isabelle \<^emph>\<open>session\<close> consists of a collection of related theories that may
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    11
  be associated with formal documents (\chref{ch:present}). There is also a
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    12
  notion of \<^emph>\<open>persistent heap\<close> image to capture the state of a session,
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    13
  similar to object-code in compiled programming languages. Thus the concept
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    14
  of session resembles that of a ``project'' in common IDE environments, but
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    15
  the specific name emphasizes the connection to interactive theorem proving:
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    16
  the session wraps-up the results of user-interaction with the prover in a
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    17
  persistent form.
48584
8026c852cc10 some introduction on sessions;
wenzelm
parents: 48582
diff changeset
    18
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    19
  Application sessions are built on a given parent session, which may be built
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    20
  recursively on other parents. Following this path in the hierarchy
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    21
  eventually leads to some major object-logic session like \<open>HOL\<close>, which itself
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    22
  is based on \<open>Pure\<close> as the common root of all sessions.
48584
8026c852cc10 some introduction on sessions;
wenzelm
parents: 48582
diff changeset
    23
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    24
  Processing sessions may take considerable time. Isabelle build management
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    25
  helps to organize this efficiently. This includes support for parallel build
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    26
  jobs, in addition to the multithreaded theory and proof checking that is
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    27
  already provided by the prover process itself.
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    28
\<close>
48604
f651323139f0 misc tuning;
wenzelm
parents: 48602
diff changeset
    29
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
    30
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
    31
section \<open>Session ROOT specifications \label{sec:session-root}\<close>
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
    32
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    33
text \<open>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    34
  Session specifications reside in files called \<^verbatim>\<open>ROOT\<close> within certain
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    35
  directories, such as the home locations of registered Isabelle components or
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    36
  additional project directories given by the user.
48579
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
    37
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    38
  The ROOT file format follows the lexical conventions of the \<^emph>\<open>outer syntax\<close>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    39
  of Isabelle/Isar, see also @{cite "isabelle-isar-ref"}. This defines common
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    40
  forms like identifiers, names, quoted strings, verbatim text, nested
75986
27d98da31985 support 'chapter_definition' with description for presentation purposes;
wenzelm
parents: 75559
diff changeset
    41
  comments etc. The grammar for @{syntax chapter_def}, @{syntax chapter_entry}
27d98da31985 support 'chapter_definition' with description for presentation purposes;
wenzelm
parents: 75559
diff changeset
    42
  and @{syntax session_entry} is given as syntax diagram below. Each ROOT file
27d98da31985 support 'chapter_definition' with description for presentation purposes;
wenzelm
parents: 75559
diff changeset
    43
  may contain multiple specifications like this. Chapters help to organize
27d98da31985 support 'chapter_definition' with description for presentation purposes;
wenzelm
parents: 75559
diff changeset
    44
  browser info (\secref{sec:info}), but have no formal meaning. The default
76005
a9bbf075f431 include groups from 'chapter_definition' in session info, based on the state of chapter_defs after processing all ROOT files (thus the declaration order does not matter);
wenzelm
parents: 75998
diff changeset
    45
  chapter is ``\<open>Unsorted\<close>''. Chapter definitions, which are optional, allow to
a9bbf075f431 include groups from 'chapter_definition' in session info, based on the state of chapter_defs after processing all ROOT files (thus the declaration order does not matter);
wenzelm
parents: 75998
diff changeset
    46
  associate additional information.
48579
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
    47
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    48
  Isabelle/jEdit @{cite "isabelle-jedit"} includes a simple editing mode
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    49
  \<^verbatim>\<open>isabelle-root\<close> for session ROOT files, which is enabled by default for any
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    50
  file of that name.
48579
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
    51
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68808
diff changeset
    52
  \<^rail>\<open>
76005
a9bbf075f431 include groups from 'chapter_definition' in session info, based on the state of chapter_defs after processing all ROOT files (thus the declaration order does not matter);
wenzelm
parents: 75998
diff changeset
    53
    @{syntax_def chapter_def}: @'chapter_definition' @{syntax name} \<newline>
a9bbf075f431 include groups from 'chapter_definition' in session info, based on the state of chapter_defs after processing all ROOT files (thus the declaration order does not matter);
wenzelm
parents: 75998
diff changeset
    54
      groups? description?
75986
27d98da31985 support 'chapter_definition' with description for presentation purposes;
wenzelm
parents: 75559
diff changeset
    55
    ;
27d98da31985 support 'chapter_definition' with description for presentation purposes;
wenzelm
parents: 75559
diff changeset
    56
27d98da31985 support 'chapter_definition' with description for presentation purposes;
wenzelm
parents: 75559
diff changeset
    57
    @{syntax_def chapter_entry}: @'chapter' @{syntax name}
51417
d266f9329368 sessions may be organized via 'chapter' in ROOT;
wenzelm
parents: 51055
diff changeset
    58
    ;
d266f9329368 sessions may be organized via 'chapter' in ROOT;
wenzelm
parents: 51055
diff changeset
    59
66971
43b2aac6053c tuned diagram;
wenzelm
parents: 66970
diff changeset
    60
    @{syntax_def session_entry}: @'session' @{syntax system_name} groups? dir? '=' \<newline>
70678
36c8c32346cb clarified syntax: 'directories' and 'theories' belong together;
wenzelm
parents: 70677
diff changeset
    61
      (@{syntax system_name} '+')? description? options? \<newline>
36c8c32346cb clarified syntax: 'directories' and 'theories' belong together;
wenzelm
parents: 70677
diff changeset
    62
      sessions? directories? (theories*) \<newline>
72600
2fa4f25d9d07 official support for document theories from other sessions;
wenzelm
parents: 72574
diff changeset
    63
      (document_theories?) (document_files*) \<newline>
2fa4f25d9d07 official support for document theories from other sessions;
wenzelm
parents: 72574
diff changeset
    64
      (export_files*)
48579
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
    65
    ;
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
    66
    groups: '(' (@{syntax name} +) ')'
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
    67
    ;
68808
5467858e9419 more uniform cartouche syntax;
wenzelm
parents: 68734
diff changeset
    68
    dir: @'in' @{syntax embedded}
48579
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
    69
    ;
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
    70
    description: @'description' @{syntax text}
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
    71
    ;
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
    72
    options: @'options' opts
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
    73
    ;
48605
e777363440d6 allow negative int values as well, according to real = int | float;
wenzelm
parents: 48604
diff changeset
    74
    opts: '[' ( (@{syntax name} '=' value | @{syntax name}) + ',' ) ']'
e777363440d6 allow negative int values as well, according to real = int | float;
wenzelm
parents: 48604
diff changeset
    75
    ;
e777363440d6 allow negative int values as well, according to real = int | float;
wenzelm
parents: 48604
diff changeset
    76
    value: @{syntax name} | @{syntax real}
48579
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
    77
    ;
66916
aca50a1572c5 more documentation;
wenzelm
parents: 66851
diff changeset
    78
    sessions: @'sessions' (@{syntax system_name}+)
65374
a5b38d8d3c1e tuned syntax;
wenzelm
parents: 64308
diff changeset
    79
    ;
70681
a6c0f2d106c8 disallow overlapping session directories;
wenzelm
parents: 70678
diff changeset
    80
    directories: @'directories' (dir+)
70678
36c8c32346cb clarified syntax: 'directories' and 'theories' belong together;
wenzelm
parents: 70677
diff changeset
    81
    ;
66970
13857f49d215 clarified ROOT syntax: 'sessions' and 'theories' are optional, but need to be non-empty;
wenzelm
parents: 66916
diff changeset
    82
    theories: @'theories' opts? (theory_entry+)
48579
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
    83
    ;
66916
aca50a1572c5 more documentation;
wenzelm
parents: 66851
diff changeset
    84
    theory_entry: @{syntax system_name} ('(' @'global' ')')?
65504
b80477da30eb some documentation;
wenzelm
parents: 65374
diff changeset
    85
    ;
72600
2fa4f25d9d07 official support for document theories from other sessions;
wenzelm
parents: 72574
diff changeset
    86
    document_theories: @'document_theories' (@{syntax name}+)
2fa4f25d9d07 official support for document theories from other sessions;
wenzelm
parents: 72574
diff changeset
    87
    ;
68808
5467858e9419 more uniform cartouche syntax;
wenzelm
parents: 68734
diff changeset
    88
    document_files: @'document_files' ('(' dir ')')? (@{syntax embedded}+)
68292
7ca0c23179e6 support 'export_files' in session ROOT;
wenzelm
parents: 68290
diff changeset
    89
    ;
69671
2486792eaf61 support pruning of export names;
wenzelm
parents: 69610
diff changeset
    90
    export_files: @'export_files' ('(' dir ')')? ('[' nat ']')? \<newline>
2486792eaf61 support pruning of export names;
wenzelm
parents: 69610
diff changeset
    91
      (@{syntax embedded}+)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68808
diff changeset
    92
  \<close>
48579
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
    93
76005
a9bbf075f431 include groups from 'chapter_definition' in session info, based on the state of chapter_defs after processing all ROOT files (thus the declaration order does not matter);
wenzelm
parents: 75998
diff changeset
    94
  \<^descr> \isakeyword{chapter{\isacharunderscorekeyword}definition}~\<open>A (groups)\<close>
a9bbf075f431 include groups from 'chapter_definition' in session info, based on the state of chapter_defs after processing all ROOT files (thus the declaration order does not matter);
wenzelm
parents: 75998
diff changeset
    95
  associates a collection of groups with chapter \<open>A\<close>. All sessions that belong
a9bbf075f431 include groups from 'chapter_definition' in session info, based on the state of chapter_defs after processing all ROOT files (thus the declaration order does not matter);
wenzelm
parents: 75998
diff changeset
    96
  to this chapter will automatically become members of these groups.
a9bbf075f431 include groups from 'chapter_definition' in session info, based on the state of chapter_defs after processing all ROOT files (thus the declaration order does not matter);
wenzelm
parents: 75998
diff changeset
    97
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
    98
  \<^descr> \isakeyword{session}~\<open>A = B + body\<close> defines a new session \<open>A\<close> based on
66759
918f15c9367a discontinued obsolete 'files' in session ROOT;
wenzelm
parents: 66748
diff changeset
    99
  parent session \<open>B\<close>, with its content given in \<open>body\<close> (imported sessions and
918f15c9367a discontinued obsolete 'files' in session ROOT;
wenzelm
parents: 66748
diff changeset
   100
  theories). Note that a parent (like \<open>HOL\<close>) is mandatory in practical
918f15c9367a discontinued obsolete 'files' in session ROOT;
wenzelm
parents: 66748
diff changeset
   101
  applications: only Isabelle/Pure can bootstrap itself from nothing.
48579
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
   102
65504
b80477da30eb some documentation;
wenzelm
parents: 65374
diff changeset
   103
  All such session specifications together describe a hierarchy (graph) of
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   104
  sessions, with globally unique names. The new session name \<open>A\<close> should be
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   105
  sufficiently long and descriptive to stand on its own in a potentially large
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   106
  library.
48579
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
   107
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   108
  \<^descr> \isakeyword{session}~\<open>A (groups)\<close> indicates a collection of groups where
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   109
  the new session is a member. Group names are uninterpreted and merely follow
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   110
  certain conventions. For example, the Isabelle distribution tags some
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   111
  important sessions by the group name called ``\<open>main\<close>''. Other projects may
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   112
  invent their own conventions, but this requires some care to avoid clashes
48604
f651323139f0 misc tuning;
wenzelm
parents: 48602
diff changeset
   113
  within this unchecked name space.
48579
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
   114
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   115
  \<^descr> \isakeyword{session}~\<open>A\<close>~\isakeyword{in}~\<open>dir\<close> specifies an explicit
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   116
  directory for this session; by default this is the current directory of the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   117
  \<^verbatim>\<open>ROOT\<close> file.
48579
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
   118
66759
918f15c9367a discontinued obsolete 'files' in session ROOT;
wenzelm
parents: 66748
diff changeset
   119
  All theory files are located relatively to the session directory. The prover
918f15c9367a discontinued obsolete 'files' in session ROOT;
wenzelm
parents: 66748
diff changeset
   120
  process is run within the same as its current working directory.
48579
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
   121
76005
a9bbf075f431 include groups from 'chapter_definition' in session info, based on the state of chapter_defs after processing all ROOT files (thus the declaration order does not matter);
wenzelm
parents: 75998
diff changeset
   122
  \<^descr> \isakeyword{description}~\<open>text\<close> is a free-form description for this
a9bbf075f431 include groups from 'chapter_definition' in session info, based on the state of chapter_defs after processing all ROOT files (thus the declaration order does not matter);
wenzelm
parents: 75998
diff changeset
   123
  session (or chapter), e.g. for presentation purposes.
48579
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
   124
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   125
  \<^descr> \isakeyword{options}~\<open>[x = a, y = b, z]\<close> defines separate options
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   126
  (\secref{sec:system-options}) that are used when processing this session,
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   127
  but \<^emph>\<open>without\<close> propagation to child sessions. Note that \<open>z\<close> abbreviates \<open>z =
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   128
  true\<close> for Boolean options.
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   129
65504
b80477da30eb some documentation;
wenzelm
parents: 65374
diff changeset
   130
  \<^descr> \isakeyword{sessions}~\<open>names\<close> specifies sessions that are \<^emph>\<open>imported\<close> into
b80477da30eb some documentation;
wenzelm
parents: 65374
diff changeset
   131
  the current name space of theories. This allows to refer to a theory \<open>A\<close>
b80477da30eb some documentation;
wenzelm
parents: 65374
diff changeset
   132
  from session \<open>B\<close> by the qualified name \<open>B.A\<close> --- although it is loaded again
b80477da30eb some documentation;
wenzelm
parents: 65374
diff changeset
   133
  into the current ML process, which is in contrast to a theory that is
b80477da30eb some documentation;
wenzelm
parents: 65374
diff changeset
   134
  already present in the \<^emph>\<open>parent\<close> session.
b80477da30eb some documentation;
wenzelm
parents: 65374
diff changeset
   135
65505
741fad555d82 exclude theories from other sessions;
wenzelm
parents: 65504
diff changeset
   136
  Theories that are imported from other sessions are excluded from the current
741fad555d82 exclude theories from other sessions;
wenzelm
parents: 65504
diff changeset
   137
  session document.
741fad555d82 exclude theories from other sessions;
wenzelm
parents: 65504
diff changeset
   138
70678
36c8c32346cb clarified syntax: 'directories' and 'theories' belong together;
wenzelm
parents: 70677
diff changeset
   139
  \<^descr> \isakeyword{directories}~\<open>dirs\<close> specifies additional directories for
36c8c32346cb clarified syntax: 'directories' and 'theories' belong together;
wenzelm
parents: 70677
diff changeset
   140
  import of theory files via \isakeyword{theories} within \<^verbatim>\<open>ROOT\<close> or
36c8c32346cb clarified syntax: 'directories' and 'theories' belong together;
wenzelm
parents: 70677
diff changeset
   141
  \<^theory_text>\<open>imports\<close> within a theory; \<open>dirs\<close> are relative to the main session
36c8c32346cb clarified syntax: 'directories' and 'theories' belong together;
wenzelm
parents: 70677
diff changeset
   142
  directory (cf.\ \isakeyword{session} \dots \isakeyword{in}~\<open>dir\<close>). These
70681
a6c0f2d106c8 disallow overlapping session directories;
wenzelm
parents: 70678
diff changeset
   143
  directories need to be exclusively assigned to a unique session, without
a6c0f2d106c8 disallow overlapping session directories;
wenzelm
parents: 70678
diff changeset
   144
  implicit sharing of file-system locations.
70678
36c8c32346cb clarified syntax: 'directories' and 'theories' belong together;
wenzelm
parents: 70677
diff changeset
   145
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   146
  \<^descr> \isakeyword{theories}~\<open>options names\<close> specifies a block of theories that
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   147
  are processed within an environment that is augmented by the given options,
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   148
  in addition to the global session options given before. Any number of blocks
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   149
  of \isakeyword{theories} may be given. Options are only active for each
48604
f651323139f0 misc tuning;
wenzelm
parents: 48602
diff changeset
   150
  \isakeyword{theories} block separately.
48579
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
   151
65374
a5b38d8d3c1e tuned syntax;
wenzelm
parents: 64308
diff changeset
   152
  A theory name that is followed by \<open>(\<close>\isakeyword{global}\<open>)\<close> is treated
66993
2c2a346cfe70 clarified situation of global theory names;
wenzelm
parents: 66971
diff changeset
   153
  literally in other session specifications or theory imports --- the normal
2c2a346cfe70 clarified situation of global theory names;
wenzelm
parents: 66971
diff changeset
   154
  situation is to qualify theory names by the session name; this ensures
2c2a346cfe70 clarified situation of global theory names;
wenzelm
parents: 66971
diff changeset
   155
  globally unique names in big session graphs. Global theories are usually the
2c2a346cfe70 clarified situation of global theory names;
wenzelm
parents: 66971
diff changeset
   156
  entry points to major logic sessions: \<open>Pure\<close>, \<open>Main\<close>, \<open>Complex_Main\<close>,
2c2a346cfe70 clarified situation of global theory names;
wenzelm
parents: 66971
diff changeset
   157
  \<open>HOLCF\<close>, \<open>IFOL\<close>, \<open>FOL\<close>, \<open>ZF\<close>, \<open>ZFC\<close> etc. Regular Isabelle applications
2c2a346cfe70 clarified situation of global theory names;
wenzelm
parents: 66971
diff changeset
   158
  should not claim any global theory names.
65374
a5b38d8d3c1e tuned syntax;
wenzelm
parents: 64308
diff changeset
   159
72600
2fa4f25d9d07 official support for document theories from other sessions;
wenzelm
parents: 72574
diff changeset
   160
  \<^descr> \isakeyword{document_theories}~\<open>names\<close> specifies theories from other
2fa4f25d9d07 official support for document theories from other sessions;
wenzelm
parents: 72574
diff changeset
   161
  sessions that should be included in the generated document source directory.
2fa4f25d9d07 official support for document theories from other sessions;
wenzelm
parents: 72574
diff changeset
   162
  These theories need to be explicit imports in the current session, or
72769
haftmann
parents: 72648
diff changeset
   163
  implicit imports from the underlying hierarchy of parent sessions. The
72600
2fa4f25d9d07 official support for document theories from other sessions;
wenzelm
parents: 72574
diff changeset
   164
  generated \<^verbatim>\<open>session.tex\<close> file is not affected: the session's {\LaTeX} setup
2fa4f25d9d07 official support for document theories from other sessions;
wenzelm
parents: 72574
diff changeset
   165
  needs to \<^verbatim>\<open>\input{\<close>\<open>\<dots>\<close>\<^verbatim>\<open>}\<close> generated \<^verbatim>\<open>.tex\<close> files separately.
2fa4f25d9d07 official support for document theories from other sessions;
wenzelm
parents: 72574
diff changeset
   166
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   167
  \<^descr> \isakeyword{document_files}~\<open>(\<close>\isakeyword{in}~\<open>base_dir) files\<close> lists
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   168
  source files for document preparation, typically \<^verbatim>\<open>.tex\<close> and \<^verbatim>\<open>.sty\<close> for
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   169
  {\LaTeX}. Only these explicitly given files are copied from the base
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   170
  directory to the document output directory, before formal document
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   171
  processing is started (see also \secref{sec:tool-document}). The local path
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   172
  structure of the \<open>files\<close> is preserved, which allows to reconstruct the
68292
7ca0c23179e6 support 'export_files' in session ROOT;
wenzelm
parents: 68290
diff changeset
   173
  original directory hierarchy of \<open>base_dir\<close>. The default \<open>base_dir\<close> is
7ca0c23179e6 support 'export_files' in session ROOT;
wenzelm
parents: 68290
diff changeset
   174
  \<^verbatim>\<open>document\<close> within the session root directory.
56533
cd8b6d849b6a explicit 'document_files' in session ROOT specifications;
wenzelm
parents: 55112
diff changeset
   175
69671
2486792eaf61 support pruning of export names;
wenzelm
parents: 69610
diff changeset
   176
  \<^descr> \isakeyword{export_files}~\<open>(\<close>\isakeyword{in}~\<open>target_dir) [number]
69811
18f61ce86425 clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents: 69755
diff changeset
   177
  patterns\<close> specifies theory exports that may get written to the file-system,
18f61ce86425 clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents: 69755
diff changeset
   178
  e.g. via @{tool_ref build} with option \<^verbatim>\<open>-e\<close> (\secref{sec:tool-build}). The
18f61ce86425 clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents: 69755
diff changeset
   179
  \<open>target_dir\<close> specification is relative to the session root directory; its
18f61ce86425 clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents: 69755
diff changeset
   180
  default is \<^verbatim>\<open>export\<close>. Exports are selected via \<open>patterns\<close> as in @{tool_ref
18f61ce86425 clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents: 69755
diff changeset
   181
  export} (\secref{sec:tool-export}). The number given in brackets (default:
75161
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   182
  0) specifies the prefix of elements that should be removed from each name:
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   183
  it allows to reduce the resulting directory hierarchy at the danger of
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   184
  overwriting files due to loss of uniqueness.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
   185
\<close>
48579
0b95a13ed90a more on "Session ROOT specifications";
wenzelm
parents: 48578
diff changeset
   186
51055
5c4be88f8747 some explanations of important build options;
wenzelm
parents: 50546
diff changeset
   187
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
   188
subsubsection \<open>Examples\<close>
48580
9df76dd45900 some description of main build options;
wenzelm
parents: 48579
diff changeset
   189
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   190
text \<open>
63680
6e1e8b5abbfa more symbols;
wenzelm
parents: 63669
diff changeset
   191
  See \<^file>\<open>~~/src/HOL/ROOT\<close> for a diversity of practically relevant situations,
6e1e8b5abbfa more symbols;
wenzelm
parents: 63669
diff changeset
   192
  although it uses relatively complex quasi-hierarchic naming conventions like
6e1e8b5abbfa more symbols;
wenzelm
parents: 63669
diff changeset
   193
  \<^verbatim>\<open>HOL-SPARK\<close>, \<^verbatim>\<open>HOL-SPARK-Examples\<close>. An alternative is to use unqualified
6e1e8b5abbfa more symbols;
wenzelm
parents: 63669
diff changeset
   194
  names that are relatively long and descriptive, as in the Archive of Formal
67605
3dd0dfe04fcb corrected some URLs
Lars Hupel <lars.hupel@mytum.de>
parents: 67140
diff changeset
   195
  Proofs (\<^url>\<open>https://isa-afp.org\<close>), for example.
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   196
\<close>
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   197
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   198
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
   199
section \<open>System build options \label{sec:system-options}\<close>
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   200
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   201
text \<open>
63680
6e1e8b5abbfa more symbols;
wenzelm
parents: 63669
diff changeset
   202
  See \<^file>\<open>~~/etc/options\<close> for the main defaults provided by the Isabelle
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   203
  distribution. Isabelle/jEdit @{cite "isabelle-jedit"} includes a simple
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   204
  editing mode \<^verbatim>\<open>isabelle-options\<close> for this file-format.
48693
ceeea46bdeba "isabelle options" prints Isabelle system options;
wenzelm
parents: 48684
diff changeset
   205
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   206
  The following options are particularly relevant to build Isabelle sessions,
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   207
  in particular with document preparation (\chref{ch:present}).
51055
5c4be88f8747 some explanations of important build options;
wenzelm
parents: 50546
diff changeset
   208
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   209
    \<^item> @{system_option_def "browser_info"} controls output of HTML browser
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   210
    info, see also \secref{sec:info}.
51055
5c4be88f8747 some explanations of important build options;
wenzelm
parents: 50546
diff changeset
   211
72574
d892f6d66402 build documents in Isabelle/Scala, based on generated tex files as session exports;
wenzelm
parents: 72515
diff changeset
   212
    \<^item> @{system_option_def "document"} controls document output for a
73826
72900f34dbb3 allow system option short form NAME for NAME=true for type string, not just bool;
wenzelm
parents: 73777
diff changeset
   213
    particular session or theory; \<^verbatim>\<open>document=pdf\<close> or \<^verbatim>\<open>document=true\<close> means
72900f34dbb3 allow system option short form NAME for NAME=true for type string, not just bool;
wenzelm
parents: 73777
diff changeset
   214
    enabled, \<^verbatim>\<open>document=""\<close> or \<^verbatim>\<open>document=false\<close> means disabled (especially
72900f34dbb3 allow system option short form NAME for NAME=true for type string, not just bool;
wenzelm
parents: 73777
diff changeset
   215
    for particular theories).
51055
5c4be88f8747 some explanations of important build options;
wenzelm
parents: 50546
diff changeset
   216
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   217
    \<^item> @{system_option_def "document_output"} specifies an alternative
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   218
    directory for generated output of the document preparation system; the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   219
    default is within the @{setting "ISABELLE_BROWSER_INFO"} hierarchy as
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   220
    explained in \secref{sec:info}. See also @{tool mkroot}, which generates a
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   221
    default configuration with output readily available to the author of the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   222
    document.
51055
5c4be88f8747 some explanations of important build options;
wenzelm
parents: 50546
diff changeset
   223
74873
0ab2ed1489eb more documentation about document build options;
wenzelm
parents: 74827
diff changeset
   224
    \<^item> @{system_option_def "document_echo"} informs about document file names
0ab2ed1489eb more documentation about document build options;
wenzelm
parents: 74827
diff changeset
   225
    during session presentation.
0ab2ed1489eb more documentation about document build options;
wenzelm
parents: 74827
diff changeset
   226
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   227
    \<^item> @{system_option_def "document_variants"} specifies document variants as
75161
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   228
    a colon-separated list of \<open>name=tags\<close> entries. The default name is
72574
d892f6d66402 build documents in Isabelle/Scala, based on generated tex files as session exports;
wenzelm
parents: 72515
diff changeset
   229
    \<^verbatim>\<open>document\<close>, without additional tags.
51055
5c4be88f8747 some explanations of important build options;
wenzelm
parents: 50546
diff changeset
   230
72574
d892f6d66402 build documents in Isabelle/Scala, based on generated tex files as session exports;
wenzelm
parents: 72515
diff changeset
   231
    Tags are specified as a comma separated list of modifier/name pairs and
d892f6d66402 build documents in Isabelle/Scala, based on generated tex files as session exports;
wenzelm
parents: 72515
diff changeset
   232
    tell {\LaTeX} how to interpret certain Isabelle command regions:
d892f6d66402 build documents in Isabelle/Scala, based on generated tex files as session exports;
wenzelm
parents: 72515
diff changeset
   233
    ``\<^verbatim>\<open>+\<close>\<open>foo\<close>'' (or just ``\<open>foo\<close>'') means to keep, ``\<^verbatim>\<open>-\<close>\<open>foo\<close>'' to drop,
d892f6d66402 build documents in Isabelle/Scala, based on generated tex files as session exports;
wenzelm
parents: 72515
diff changeset
   234
    and ``\<^verbatim>\<open>/\<close>\<open>foo\<close>'' to fold text tagged as \<open>foo\<close>. The builtin default is
d892f6d66402 build documents in Isabelle/Scala, based on generated tex files as session exports;
wenzelm
parents: 72515
diff changeset
   235
    equivalent to the tag specification
d892f6d66402 build documents in Isabelle/Scala, based on generated tex files as session exports;
wenzelm
parents: 72515
diff changeset
   236
    ``\<^verbatim>\<open>+document,+theory,+proof,+ML,+visible,-invisible,+important,+unimportant\<close>'';
d892f6d66402 build documents in Isabelle/Scala, based on generated tex files as session exports;
wenzelm
parents: 72515
diff changeset
   237
    see also the {\LaTeX} macros \<^verbatim>\<open>\isakeeptag\<close>, \<^verbatim>\<open>\isadroptag\<close>, and
d892f6d66402 build documents in Isabelle/Scala, based on generated tex files as session exports;
wenzelm
parents: 72515
diff changeset
   238
    \<^verbatim>\<open>\isafoldtag\<close>, in \<^file>\<open>~~/lib/texinputs/isabelle.sty\<close>.
d892f6d66402 build documents in Isabelle/Scala, based on generated tex files as session exports;
wenzelm
parents: 72515
diff changeset
   239
d892f6d66402 build documents in Isabelle/Scala, based on generated tex files as session exports;
wenzelm
parents: 72515
diff changeset
   240
    In contrast, \<^verbatim>\<open>document_variants=document:outline=/proof,/ML\<close> indicates
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   241
    two documents: the one called \<^verbatim>\<open>document\<close> with default tags, and the other
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   242
    called \<^verbatim>\<open>outline\<close> where proofs and ML sections are folded.
51055
5c4be88f8747 some explanations of important build options;
wenzelm
parents: 50546
diff changeset
   243
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   244
    Document variant names are just a matter of conventions. It is also
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   245
    possible to use different document variant names (without tags) for
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   246
    different document root entries, see also \secref{sec:tool-document}.
51055
5c4be88f8747 some explanations of important build options;
wenzelm
parents: 50546
diff changeset
   247
68513
88b0e63d58a5 updated documentation;
wenzelm
parents: 68348
diff changeset
   248
    \<^item> @{system_option_def "document_tags"} specifies alternative command tags
88b0e63d58a5 updated documentation;
wenzelm
parents: 68348
diff changeset
   249
    as a comma-separated list of items: either ``\<open>command\<close>\<^verbatim>\<open>%\<close>\<open>tag\<close>'' for a
88b0e63d58a5 updated documentation;
wenzelm
parents: 68348
diff changeset
   250
    specific command, or ``\<^verbatim>\<open>%\<close>\<open>tag\<close>'' as default for all other commands. This
88b0e63d58a5 updated documentation;
wenzelm
parents: 68348
diff changeset
   251
    is occasionally useful to control the global visibility of commands via
88b0e63d58a5 updated documentation;
wenzelm
parents: 68348
diff changeset
   252
    session options (e.g.\ in \<^verbatim>\<open>ROOT\<close>).
67140
386a31d6d17a more documentation;
wenzelm
parents: 66993
diff changeset
   253
74873
0ab2ed1489eb more documentation about document build options;
wenzelm
parents: 74827
diff changeset
   254
    \<^item> @{system_option_def "document_comment_latex"} enables regular {\LaTeX}
0ab2ed1489eb more documentation about document build options;
wenzelm
parents: 74827
diff changeset
   255
    \<^verbatim>\<open>comment.sty\<close>, instead of the historic version for plain {\TeX}
0ab2ed1489eb more documentation about document build options;
wenzelm
parents: 74827
diff changeset
   256
    (default). The latter is much faster, but in conflict with {\LaTeX}
0ab2ed1489eb more documentation about document build options;
wenzelm
parents: 74827
diff changeset
   257
    classes like Dagstuhl
0ab2ed1489eb more documentation about document build options;
wenzelm
parents: 74827
diff changeset
   258
    LIPIcs\<^footnote>\<open>\<^url>\<open>https://github.com/dagstuhl-publishing/styles\<close>\<close>.
0ab2ed1489eb more documentation about document build options;
wenzelm
parents: 74827
diff changeset
   259
73743
813a08dff3fd explicit option document_bibliography;
wenzelm
parents: 73735
diff changeset
   260
    \<^item> @{system_option_def "document_bibliography"} explicitly enables the use
813a08dff3fd explicit option document_bibliography;
wenzelm
parents: 73735
diff changeset
   261
    of \<^verbatim>\<open>bibtex\<close>; the default is to check the presence of \<^verbatim>\<open>root.bib\<close>, but it
813a08dff3fd explicit option document_bibliography;
wenzelm
parents: 73735
diff changeset
   262
    could have a different name.
813a08dff3fd explicit option document_bibliography;
wenzelm
parents: 73735
diff changeset
   263
74873
0ab2ed1489eb more documentation about document build options;
wenzelm
parents: 74827
diff changeset
   264
    \<^item> @{system_option_def "document_heading_prefix"} specifies a prefix for
0ab2ed1489eb more documentation about document build options;
wenzelm
parents: 74827
diff changeset
   265
    the {\LaTeX} macro names generated from Isar commands like \<^theory_text>\<open>chapter\<close>,
0ab2ed1489eb more documentation about document build options;
wenzelm
parents: 74827
diff changeset
   266
    \<^theory_text>\<open>section\<close> etc. The default is \<^verbatim>\<open>isamarkup\<close>, e.g. \<^theory_text>\<open>section\<close> becomes
0ab2ed1489eb more documentation about document build options;
wenzelm
parents: 74827
diff changeset
   267
    \<^verbatim>\<open>\isamarkupsection\<close>.
0ab2ed1489eb more documentation about document build options;
wenzelm
parents: 74827
diff changeset
   268
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   269
    \<^item> @{system_option_def "threads"} determines the number of worker threads
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   270
    for parallel checking of theories and proofs. The default \<open>0\<close> means that a
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   271
    sensible maximum value is determined by the underlying hardware. For
75161
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   272
    machines with many cores or with hyperthreading, this sometimes requires
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   273
    manual adjustment (on the command-line or within personal settings or
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   274
    preferences, not within a session \<^verbatim>\<open>ROOT\<close>).
51055
5c4be88f8747 some explanations of important build options;
wenzelm
parents: 50546
diff changeset
   275
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   276
    \<^item> @{system_option_def "condition"} specifies a comma-separated list of
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   277
    process environment variables (or Isabelle settings) that are required for
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   278
    the subsequent theories to be processed. Conditions are considered
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   279
    ``true'' if the corresponding environment value is defined and non-empty.
51055
5c4be88f8747 some explanations of important build options;
wenzelm
parents: 50546
diff changeset
   280
61602
a2f0f659a3c2 added option timeout_scale;
wenzelm
parents: 61575
diff changeset
   281
    \<^item> @{system_option_def "timeout"} and @{system_option_def "timeout_scale"}
a2f0f659a3c2 added option timeout_scale;
wenzelm
parents: 61575
diff changeset
   282
    specify a real wall-clock timeout for the session as a whole: the two
a2f0f659a3c2 added option timeout_scale;
wenzelm
parents: 61575
diff changeset
   283
    values are multiplied and taken as the number of seconds. Typically,
a2f0f659a3c2 added option timeout_scale;
wenzelm
parents: 61575
diff changeset
   284
    @{system_option "timeout"} is given for individual sessions, and
a2f0f659a3c2 added option timeout_scale;
wenzelm
parents: 61575
diff changeset
   285
    @{system_option "timeout_scale"} as global adjustment to overall hardware
a2f0f659a3c2 added option timeout_scale;
wenzelm
parents: 61575
diff changeset
   286
    performance.
a2f0f659a3c2 added option timeout_scale;
wenzelm
parents: 61575
diff changeset
   287
a2f0f659a3c2 added option timeout_scale;
wenzelm
parents: 61575
diff changeset
   288
    The timer is controlled outside the ML process by the JVM that runs
a2f0f659a3c2 added option timeout_scale;
wenzelm
parents: 61575
diff changeset
   289
    Isabelle/Scala. Thus it is relatively reliable in canceling processes that
a2f0f659a3c2 added option timeout_scale;
wenzelm
parents: 61575
diff changeset
   290
    get out of control, even if there is a deadlock without CPU time usage.
51055
5c4be88f8747 some explanations of important build options;
wenzelm
parents: 50546
diff changeset
   291
64308
b00508facb4f added system option "profiling";
wenzelm
parents: 64265
diff changeset
   292
    \<^item> @{system_option_def "profiling"} specifies a mode for global ML
b00508facb4f added system option "profiling";
wenzelm
parents: 64265
diff changeset
   293
    profiling. Possible values are the empty string (disabled), \<^verbatim>\<open>time\<close> for
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68808
diff changeset
   294
    \<^ML>\<open>profile_time\<close> and \<^verbatim>\<open>allocations\<close> for \<^ML>\<open>profile_allocations\<close>.
64308
b00508facb4f added system option "profiling";
wenzelm
parents: 64265
diff changeset
   295
    Results appear near the bottom of the session log file.
b00508facb4f added system option "profiling";
wenzelm
parents: 64265
diff changeset
   296
73777
52e43a93d51f clarified system_log: make this work independently of the particular "isabelle build" command-line (e.g. "isabelle mirabelle");
wenzelm
parents: 73774
diff changeset
   297
    \<^item> @{system_option_def system_log} specifies an optional log file for
52e43a93d51f clarified system_log: make this work independently of the particular "isabelle build" command-line (e.g. "isabelle mirabelle");
wenzelm
parents: 73774
diff changeset
   298
    low-level messages produced by \<^ML>\<open>Output.system_message\<close> in
74827
c1b5d6e6ff74 clarified system option standard values: avoid oddities like "isabelle build -o document_output" producing directories named "true";
wenzelm
parents: 74733
diff changeset
   299
    Isabelle/ML; the standard value ``\<^verbatim>\<open>-\<close>'' refers to console progress of the
c1b5d6e6ff74 clarified system option standard values: avoid oddities like "isabelle build -o document_output" producing directories named "true";
wenzelm
parents: 74733
diff changeset
   300
    build job.
73777
52e43a93d51f clarified system_log: make this work independently of the particular "isabelle build" command-line (e.g. "isabelle mirabelle");
wenzelm
parents: 73774
diff changeset
   301
69854
cc0b3e177b49 system option "system_heaps" supersedes various command-line options for "system build mode";
wenzelm
parents: 69811
diff changeset
   302
    \<^item> @{system_option_def "system_heaps"} determines the directories for
cc0b3e177b49 system option "system_heaps" supersedes various command-line options for "system build mode";
wenzelm
parents: 69811
diff changeset
   303
    session heap images: \<^path>\<open>$ISABELLE_HEAPS\<close> is the user directory and
cc0b3e177b49 system option "system_heaps" supersedes various command-line options for "system build mode";
wenzelm
parents: 69811
diff changeset
   304
    \<^path>\<open>$ISABELLE_HEAPS_SYSTEM\<close> the system directory (usually within the
cc0b3e177b49 system option "system_heaps" supersedes various command-line options for "system build mode";
wenzelm
parents: 69811
diff changeset
   305
    Isabelle application). For \<^verbatim>\<open>system_heaps=false\<close>, heaps are stored in the
cc0b3e177b49 system option "system_heaps" supersedes various command-line options for "system build mode";
wenzelm
parents: 69811
diff changeset
   306
    user directory and may be loaded from both directories. For
cc0b3e177b49 system option "system_heaps" supersedes various command-line options for "system build mode";
wenzelm
parents: 69811
diff changeset
   307
    \<^verbatim>\<open>system_heaps=true\<close>, store and load happens only in the system directory.
cc0b3e177b49 system option "system_heaps" supersedes various command-line options for "system build mode";
wenzelm
parents: 69811
diff changeset
   308
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   309
  The @{tool_def options} tool prints Isabelle system options. Its
48693
ceeea46bdeba "isabelle options" prints Isabelle system options;
wenzelm
parents: 48684
diff changeset
   310
  command-line usage is:
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   311
  @{verbatim [display]
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   312
\<open>Usage: isabelle options [OPTIONS] [MORE_OPTIONS ...]
48693
ceeea46bdeba "isabelle options" prints Isabelle system options;
wenzelm
parents: 48684
diff changeset
   313
ceeea46bdeba "isabelle options" prints Isabelle system options;
wenzelm
parents: 48684
diff changeset
   314
  Options are:
ceeea46bdeba "isabelle options" prints Isabelle system options;
wenzelm
parents: 48684
diff changeset
   315
    -b           include $ISABELLE_BUILD_OPTIONS
52735
842b5e7dcac8 support isabelle options -g;
wenzelm
parents: 52056
diff changeset
   316
    -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
   317
    -l           list options
48693
ceeea46bdeba "isabelle options" prints Isabelle system options;
wenzelm
parents: 48684
diff changeset
   318
    -x FILE      export to FILE in YXML format
ceeea46bdeba "isabelle options" prints Isabelle system options;
wenzelm
parents: 48684
diff changeset
   319
50531
f841ac0cb757 clarified "isabelle options" command line, to make it more close to "isabelle components";
wenzelm
parents: 50406
diff changeset
   320
  Report Isabelle system options, augmented by MORE_OPTIONS given as
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   321
  arguments NAME=VAL or NAME.\<close>}
48693
ceeea46bdeba "isabelle options" prints Isabelle system options;
wenzelm
parents: 48684
diff changeset
   322
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   323
  The command line arguments provide additional system options of the form
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   324
  \<open>name\<close>\<^verbatim>\<open>=\<close>\<open>value\<close> or \<open>name\<close> for Boolean options.
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   325
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   326
  Option \<^verbatim>\<open>-b\<close> augments the implicit environment of system options by the ones
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   327
  of @{setting ISABELLE_BUILD_OPTIONS}, cf.\ \secref{sec:tool-build}.
48693
ceeea46bdeba "isabelle options" prints Isabelle system options;
wenzelm
parents: 48684
diff changeset
   328
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   329
  Option \<^verbatim>\<open>-g\<close> prints the value of the given option. Option \<^verbatim>\<open>-l\<close> lists all
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   330
  options with their declaration and current value.
48693
ceeea46bdeba "isabelle options" prints Isabelle system options;
wenzelm
parents: 48684
diff changeset
   331
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   332
  Option \<^verbatim>\<open>-x\<close> specifies a file to export the result in YXML format, instead
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   333
  of printing it in human-readable form.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
   334
\<close>
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   335
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   336
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
   337
section \<open>Invoking the build process \label{sec:tool-build}\<close>
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   338
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   339
text \<open>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   340
  The @{tool_def build} tool invokes the build process for Isabelle sessions.
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   341
  It manages dependencies between sessions, related sources of theories and
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   342
  auxiliary files, and target heap images. Accordingly, it runs instances of
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   343
  the prover process with optional document preparation. Its command-line
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   344
  usage is:\<^footnote>\<open>Isabelle/Scala provides the same functionality via
71896
ce06d6456cc8 clarified signature --- fit within limit of 22 arguments;
wenzelm
parents: 71893
diff changeset
   345
  \<^scala_method>\<open>isabelle.Build.build\<close>.\<close>
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   346
  @{verbatim [display]
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   347
\<open>Usage: isabelle build [OPTIONS] [SESSIONS ...]
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   348
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   349
  Options are:
66737
2edc0c42c883 option -B for "isabelle build" and "isabelle imports";
wenzelm
parents: 66671
diff changeset
   350
    -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
   351
    -D DIR       include session directory and select its sessions
64265
8eb6365f5916 isabelle build -N;
wenzelm
parents: 63827
diff changeset
   352
    -N           cyclic shuffling of NUMA CPU nodes (performance tuning)
72648
1cbac4ae934d more explicit presentation directory;
wenzelm
parents: 72600
diff changeset
   353
    -P DIR       enable HTML/PDF presentation in directory (":" for default)
71807
cdfa8f027bb9 tuned messages;
wenzelm
parents: 70863
diff changeset
   354
    -R           refer to requirements of selected sessions
66745
e7ac579b883c option -S for "isabelle build";
wenzelm
parents: 66737
diff changeset
   355
    -S           soft build: only observe changes of sources, not heap images
60106
e0d1d9203275 allow to exclude session groups;
wenzelm
parents: 59892
diff changeset
   356
    -X NAME      exclude sessions from group NAME and all descendants
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   357
    -a           select all sessions
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   358
    -b           build heap images
48595
231e6fa96dbb added build option -c;
wenzelm
parents: 48594
diff changeset
   359
    -c           clean build
48737
f3bbb9ca57d6 added build option -D: include session directory and select its sessions;
wenzelm
parents: 48693
diff changeset
   360
    -d DIR       include session directory
69811
18f61ce86425 clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents: 69755
diff changeset
   361
    -e           export files from session specification into file-system
66841
5c32a072ca8b added isablle build option -f;
wenzelm
parents: 66759
diff changeset
   362
    -f           fresh build
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   363
    -g NAME      select session group NAME
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   364
    -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
   365
    -k KEYWORD   check theory sources for conflicts with proposed keywords
48903
1621b3f26095 added build option -l (list files);
wenzelm
parents: 48814
diff changeset
   366
    -l           list session source files
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   367
    -n           no build -- test dependencies only
52056
fc458f304f93 added isabelle-process option -o;
wenzelm
parents: 51417
diff changeset
   368
    -o OPTION    override Isabelle system OPTION (via NAME=VAL or NAME)
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   369
    -v           verbose
60106
e0d1d9203275 allow to exclude session groups;
wenzelm
parents: 59892
diff changeset
   370
    -x NAME      exclude session NAME and all descendants
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   371
62596
cf79f8866bc3 tuned messages;
wenzelm
parents: 62289
diff changeset
   372
  Build and manage Isabelle sessions, depending on implicit settings:
cf79f8866bc3 tuned messages;
wenzelm
parents: 62289
diff changeset
   373
71982
cea6087e8a70 tuned message;
wenzelm
parents: 71896
diff changeset
   374
  ISABELLE_TOOL_JAVA_OPTIONS="..."
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   375
  ISABELLE_BUILD_OPTIONS="..."
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   376
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   377
  ML_PLATFORM="..."
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   378
  ML_HOME="..."
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   379
  ML_SYSTEM="..."
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   380
  ML_OPTIONS="..."\<close>}
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   381
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   382
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   383
  Isabelle sessions are defined via session ROOT files as described in
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   384
  (\secref{sec:session-root}). The totality of sessions is determined by
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   385
  collecting such specifications from all Isabelle component directories
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   386
  (\secref{sec:components}), augmented by more directories given via options
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   387
  \<^verbatim>\<open>-d\<close>~\<open>DIR\<close> on the command line. Each such directory may contain a session
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61493
diff changeset
   388
  \<^verbatim>\<open>ROOT\<close> file with several session specifications.
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   389
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   390
  Any session root directory may refer recursively to further directories of
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   391
  the same kind, by listing them in a catalog file \<^verbatim>\<open>ROOTS\<close> line-by-line. This
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   392
  helps to organize large collections of session specifications, or to make
66576
wenzelm
parents: 65505
diff changeset
   393
  \<^verbatim>\<open>-d\<close> command line options persistent (e.g.\ in
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61493
diff changeset
   394
  \<^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
   395
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   396
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   397
  The subset of sessions to be managed is determined via individual \<open>SESSIONS\<close>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   398
  given as command-line arguments, or session groups that are given via one or
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   399
  more options \<^verbatim>\<open>-g\<close>~\<open>NAME\<close>. Option \<^verbatim>\<open>-a\<close> selects all sessions. The build tool
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   400
  takes session dependencies into account: the set of selected sessions is
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   401
  completed by including all ancestors.
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   402
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   403
  \<^medskip>
68734
c14a2cc9b5ef isabelle build options -c -x -B refer to imports_graph;
wenzelm
parents: 68523
diff changeset
   404
  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
   405
  descendants wrt.\ the session parent or import graph).
66737
2edc0c42c883 option -B for "isabelle build" and "isabelle imports";
wenzelm
parents: 66671
diff changeset
   406
2edc0c42c883 option -B for "isabelle build" and "isabelle imports";
wenzelm
parents: 66671
diff changeset
   407
  \<^medskip>
68734
c14a2cc9b5ef isabelle build options -c -x -B refer to imports_graph;
wenzelm
parents: 68523
diff changeset
   408
  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
   409
  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
   410
  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
   411
  membership.
59892
2a616319c171 added isabelle build option -x, to exclude sessions;
wenzelm
parents: 59891
diff changeset
   412
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   413
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   414
  Option \<^verbatim>\<open>-R\<close> reverses the selection in the sense that it refers to its
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   415
  requirements: all ancestor sessions excluding the original selection. This
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   416
  allows to prepare the stage for some build process with different options,
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   417
  before running the main build itself (without option \<^verbatim>\<open>-R\<close>).
49131
aa1e2ba3c697 added build option -R;
wenzelm
parents: 48985
diff changeset
   418
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   419
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   420
  Option \<^verbatim>\<open>-D\<close> is similar to \<^verbatim>\<open>-d\<close>, but selects all sessions that are defined
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   421
  in the given directories.
48737
f3bbb9ca57d6 added build option -D: include session directory and select its sessions;
wenzelm
parents: 48693
diff changeset
   422
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   423
  \<^medskip>
66745
e7ac579b883c option -S for "isabelle build";
wenzelm
parents: 66737
diff changeset
   424
  Option \<^verbatim>\<open>-S\<close> indicates a ``soft build'': the selection is restricted to
e7ac579b883c option -S for "isabelle build";
wenzelm
parents: 66737
diff changeset
   425
  those sessions that have changed sources (according to actually imported
e7ac579b883c option -S for "isabelle build";
wenzelm
parents: 66737
diff changeset
   426
  theories). The status of heap images is ignored.
e7ac579b883c option -S for "isabelle build";
wenzelm
parents: 66737
diff changeset
   427
e7ac579b883c option -S for "isabelle build";
wenzelm
parents: 66737
diff changeset
   428
  \<^medskip>
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   429
  The build process depends on additional options
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   430
  (\secref{sec:system-options}) that are passed to the prover eventually. The
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   431
  settings variable @{setting_ref ISABELLE_BUILD_OPTIONS} allows to provide
73826
72900f34dbb3 allow system option short form NAME for NAME=true for type string, not just bool;
wenzelm
parents: 73777
diff changeset
   432
  additional defaults, e.g.\ \<^verbatim>\<open>ISABELLE_BUILD_OPTIONS="document=pdf
72900f34dbb3 allow system option short form NAME for NAME=true for type string, not just bool;
wenzelm
parents: 73777
diff changeset
   433
  threads=4"\<close>. Moreover, the environment of system build options may be
72900f34dbb3 allow system option short form NAME for NAME=true for type string, not just bool;
wenzelm
parents: 73777
diff changeset
   434
  augmented on the command line via \<^verbatim>\<open>-o\<close>~\<open>name\<close>\<^verbatim>\<open>=\<close>\<open>value\<close> or \<^verbatim>\<open>-o\<close>~\<open>name\<close>,
72900f34dbb3 allow system option short form NAME for NAME=true for type string, not just bool;
wenzelm
parents: 73777
diff changeset
   435
  which abbreviates \<^verbatim>\<open>-o\<close>~\<open>name\<close>\<^verbatim>\<open>=true\<close> for Boolean or string options.
72900f34dbb3 allow system option short form NAME for NAME=true for type string, not just bool;
wenzelm
parents: 73777
diff changeset
   436
  Multiple occurrences of \<^verbatim>\<open>-o\<close> on the command-line are applied in the given
72900f34dbb3 allow system option short form NAME for NAME=true for type string, not just bool;
wenzelm
parents: 73777
diff changeset
   437
  order.
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   438
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   439
  \<^medskip>
72648
1cbac4ae934d more explicit presentation directory;
wenzelm
parents: 72600
diff changeset
   440
  Option \<^verbatim>\<open>-P\<close> enables PDF/HTML presentation in the given directory, where
1cbac4ae934d more explicit presentation directory;
wenzelm
parents: 72600
diff changeset
   441
  ``\<^verbatim>\<open>-P:\<close>'' refers to the default @{setting_ref ISABELLE_BROWSER_INFO} (or
73012
238ddf525da4 clarified HTML presentation, e.g. avoid bulky jobs like HOL or HOL-Analysis in applications;
wenzelm
parents: 72878
diff changeset
   442
  @{setting_ref ISABELLE_BROWSER_INFO_SYSTEM}). This applies only to
75161
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   443
  explicitly selected sessions; note that option \<^verbatim>\<open>-R\<close> allows to select all
73012
238ddf525da4 clarified HTML presentation, e.g. avoid bulky jobs like HOL or HOL-Analysis in applications;
wenzelm
parents: 72878
diff changeset
   444
  requirements separately.
72648
1cbac4ae934d more explicit presentation directory;
wenzelm
parents: 72600
diff changeset
   445
1cbac4ae934d more explicit presentation directory;
wenzelm
parents: 72600
diff changeset
   446
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   447
  Option \<^verbatim>\<open>-b\<close> ensures that heap images are produced for all selected
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   448
  sessions. By default, images are only saved for inner nodes of the hierarchy
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   449
  of sessions, as required for other sessions to continue later on.
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   450
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   451
  \<^medskip>
68734
c14a2cc9b5ef isabelle build options -c -x -B refer to imports_graph;
wenzelm
parents: 68523
diff changeset
   452
  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
   453
  parent or import graph) before performing the specified build operation.
48595
231e6fa96dbb added build option -c;
wenzelm
parents: 48594
diff changeset
   454
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   455
  \<^medskip>
69811
18f61ce86425 clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents: 69755
diff changeset
   456
  Option \<^verbatim>\<open>-e\<close> executes the \isakeyword{export_files} directives from the ROOT
18f61ce86425 clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents: 69755
diff changeset
   457
  specification of all explicitly selected sessions: the status of the session
18f61ce86425 clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents: 69755
diff changeset
   458
  build database needs to be OK, but the session could have been built
18f61ce86425 clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents: 69755
diff changeset
   459
  earlier. Using \isakeyword{export_files}, a session may serve as abstract
18f61ce86425 clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents: 69755
diff changeset
   460
  interface for add-on build artefacts, but these are only materialized on
18f61ce86425 clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents: 69755
diff changeset
   461
  explicit request: without option \<^verbatim>\<open>-e\<close> there is no effect on the physical
18f61ce86425 clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents: 69755
diff changeset
   462
  file-system yet.
18f61ce86425 clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents: 69755
diff changeset
   463
18f61ce86425 clarified 'export_files' in session ROOT: require explicit "isabelle build -e";
wenzelm
parents: 69755
diff changeset
   464
  \<^medskip>
66841
5c32a072ca8b added isablle build option -f;
wenzelm
parents: 66759
diff changeset
   465
  Option \<^verbatim>\<open>-f\<close> forces a fresh build of all selected sessions and their
5c32a072ca8b added isablle build option -f;
wenzelm
parents: 66759
diff changeset
   466
  requirements.
5c32a072ca8b added isablle build option -f;
wenzelm
parents: 66759
diff changeset
   467
5c32a072ca8b added isablle build option -f;
wenzelm
parents: 66759
diff changeset
   468
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   469
  Option \<^verbatim>\<open>-n\<close> omits the actual build process after the preparatory stage
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   470
  (including optional cleanup). Note that the return code always indicates the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   471
  status of the set of selected sessions.
48595
231e6fa96dbb added build option -c;
wenzelm
parents: 48594
diff changeset
   472
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   473
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   474
  Option \<^verbatim>\<open>-j\<close> specifies the maximum number of parallel build jobs (prover
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   475
  processes). Each prover process is subject to a separate limit of parallel
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   476
  worker threads, cf.\ system option @{system_option_ref threads}.
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   477
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   478
  \<^medskip>
64265
8eb6365f5916 isabelle build -N;
wenzelm
parents: 63827
diff changeset
   479
  Option \<^verbatim>\<open>-N\<close> enables cyclic shuffling of NUMA CPU nodes. This may help
8eb6365f5916 isabelle build -N;
wenzelm
parents: 63827
diff changeset
   480
  performance tuning on Linux servers with separate CPU/memory modules.
8eb6365f5916 isabelle build -N;
wenzelm
parents: 63827
diff changeset
   481
8eb6365f5916 isabelle build -N;
wenzelm
parents: 63827
diff changeset
   482
  \<^medskip>
75161
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   483
  Option \<^verbatim>\<open>-v\<close> increases the general level of verbosity.
59891
9ce697050455 added isabelle build option -k, for fast off-line checking of theory sources;
wenzelm
parents: 59446
diff changeset
   484
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   485
  \<^medskip>
75161
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   486
  Option \<^verbatim>\<open>-l\<close> lists the source files that contribute to a session.
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   487
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   488
  \<^medskip>
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   489
  Option \<^verbatim>\<open>-k\<close> specifies a newly proposed keyword for outer syntax. It is
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   490
  possible to use option \<^verbatim>\<open>-k\<close> repeatedly to check multiple keywords. The
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   491
  theory sources are checked for conflicts wrt.\ this hypothetical change of
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   492
  syntax, e.g.\ to reveal occurrences of identifiers that need to be quoted.
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   493
\<close>
59891
9ce697050455 added isabelle build option -k, for fast off-line checking of theory sources;
wenzelm
parents: 59446
diff changeset
   494
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   495
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
   496
subsubsection \<open>Examples\<close>
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   497
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
   498
text \<open>
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   499
  Build a specific logic image:
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   500
  @{verbatim [display] \<open>  isabelle build -b HOLCF\<close>}
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   501
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   502
  \<^smallskip>
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   503
  Build the main group of logic images:
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   504
  @{verbatim [display] \<open>  isabelle build -b -g main\<close>}
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   505
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   506
  \<^smallskip>
66748
3efac90a11a7 more documentation;
wenzelm
parents: 66745
diff changeset
   507
  Build all descendants (and requirements) of \<^verbatim>\<open>FOL\<close> and \<^verbatim>\<open>ZF\<close>:
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   508
  @{verbatim [display] \<open>  isabelle build -B FOL -B ZF\<close>}
66748
3efac90a11a7 more documentation;
wenzelm
parents: 66745
diff changeset
   509
3efac90a11a7 more documentation;
wenzelm
parents: 66745
diff changeset
   510
  \<^smallskip>
3efac90a11a7 more documentation;
wenzelm
parents: 66745
diff changeset
   511
  Build all sessions where sources have changed (ignoring heaps):
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   512
  @{verbatim [display] \<open>  isabelle build -a -S\<close>}
66748
3efac90a11a7 more documentation;
wenzelm
parents: 66745
diff changeset
   513
3efac90a11a7 more documentation;
wenzelm
parents: 66745
diff changeset
   514
  \<^smallskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   515
  Provide a general overview of the status of all Isabelle sessions, without
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   516
  building anything:
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   517
  @{verbatim [display] \<open>  isabelle build -a -n -v\<close>}
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   518
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   519
  \<^smallskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   520
  Build all sessions with HTML browser info and PDF document preparation:
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   521
  @{verbatim [display] \<open>  isabelle build -a -o browser_info -o document\<close>}
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   522
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   523
  \<^smallskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   524
  Build all sessions with a maximum of 8 parallel prover processes and 4
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   525
  worker threads each (on a machine with many cores):
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   526
  @{verbatim [display] \<open>  isabelle build -a -j8 -o threads=4\<close>}
48595
231e6fa96dbb added build option -c;
wenzelm
parents: 48594
diff changeset
   527
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   528
  \<^smallskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   529
  Build some session images with cleanup of their descendants, while retaining
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   530
  their ancestry:
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   531
  @{verbatim [display] \<open>  isabelle build -b -c HOL-Library HOL-Algebra\<close>}
48595
231e6fa96dbb added build option -c;
wenzelm
parents: 48594
diff changeset
   532
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   533
  \<^smallskip>
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   534
  Clean all sessions without building anything:
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   535
  @{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
   536
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   537
  \<^smallskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   538
  Build all sessions from some other directory hierarchy, according to the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   539
  settings variable \<^verbatim>\<open>AFP\<close> that happens to be defined inside the Isabelle
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   540
  environment:
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   541
  @{verbatim [display] \<open>  isabelle build -D '$AFP'\<close>}
49131
aa1e2ba3c697 added build option -R;
wenzelm
parents: 48985
diff changeset
   542
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 60106
diff changeset
   543
  \<^smallskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   544
  Inform about the status of all sessions required for AFP, without building
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61572
diff changeset
   545
  anything yet:
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   546
  @{verbatim [display] \<open>  isabelle build -D '$AFP' -R -v -n\<close>}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
   547
\<close>
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   548
66671
41b64e53b6a1 more documentation;
wenzelm
parents: 66576
diff changeset
   549
72876
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   550
section \<open>Print messages from build database \label{sec:tool-log}\<close>
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   551
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   552
text \<open>
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   553
  The @{tool_def "log"} tool prints prover messages from the build
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   554
  database of the given session. Its command-line usage is:
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   555
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   556
  @{verbatim [display]
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   557
\<open>Usage: isabelle log [OPTIONS] SESSION
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   558
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   559
  Options are:
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   560
    -T NAME      restrict to given theories (multiple options possible)
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   561
    -U           output Unicode symbols
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   562
    -m MARGIN    margin for pretty printing (default: 76.0)
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   563
    -o OPTION    override Isabelle system OPTION (via NAME=VAL or NAME)
73837
f72335f6a9ed clarified documentation: tracing messages are not shown here;
wenzelm
parents: 73826
diff changeset
   564
    -v           print all messages, including information etc.
72876
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   565
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   566
  Print messages from the build database of the given session, without any
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   567
  checks against current sources: results from a failed build can be
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   568
  printed as well.\<close>}
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   569
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   570
  The specified session database is taken as is, independently of the current
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   571
  session structure and theories sources. The order of messages follows the
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   572
  source positions of source files; thus the erratic evaluation of parallel
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   573
  processing rarely matters. There is \<^emph>\<open>no\<close> implicit build process involved,
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   574
  so it is possible to retrieve error messages from a failed session as well.
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   575
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   576
  \<^medskip> Option \<^verbatim>\<open>-o\<close> allows to change system options, as in @{tool build}
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   577
  (\secref{sec:tool-build}). This may affect the storage space for the build
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   578
  database, notably via @{system_option system_heaps}, or @{system_option
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   579
  build_database_server} and its relatives.
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   580
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   581
  \<^medskip> Option \<^verbatim>\<open>-T\<close> restricts output to given theories: multiple entries are
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   582
  possible by repeating this option on the command-line. The default is to
75161
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   583
  refer to \<^emph>\<open>all\<close> theories used in the original session build process.
72876
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   584
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   585
  \<^medskip> Options \<^verbatim>\<open>-m\<close> and \<^verbatim>\<open>-U\<close> modify pretty printing and output of Isabelle
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   586
  symbols. The default is for an old-fashioned ASCII terminal at 80 characters
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   587
  per line (76 + 4 characters to prefix warnings or errors).
72878
80465b791f95 clarified messages;
wenzelm
parents: 72876
diff changeset
   588
73837
f72335f6a9ed clarified documentation: tracing messages are not shown here;
wenzelm
parents: 73826
diff changeset
   589
  \<^medskip> Option \<^verbatim>\<open>-v\<close> prints all messages from the session database that are
f72335f6a9ed clarified documentation: tracing messages are not shown here;
wenzelm
parents: 73826
diff changeset
   590
  normally inlined into the source text, including information messages etc.
72876
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   591
\<close>
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   592
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   593
subsubsection \<open>Examples\<close>
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   594
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   595
text \<open>
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   596
  Print messages from theory \<^verbatim>\<open>HOL.Nat\<close> of session \<^verbatim>\<open>HOL\<close>, using Unicode
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   597
  rendering of Isabelle symbols and a margin of 100 characters:
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   598
  @{verbatim [display] \<open>  isabelle log -T HOL.Nat -U -m 100 HOL\<close>}
72876
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   599
\<close>
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   600
626fcaebd049 NEWS and documentation for "isabelle log";
wenzelm
parents: 72769
diff changeset
   601
68152
619de043389f guard result exports via export_pattern -- avoid bombing client via huge blobs;
wenzelm
parents: 68116
diff changeset
   602
section \<open>Retrieve theory exports \label{sec:tool-export}\<close>
68116
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   603
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   604
text \<open>
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   605
  The @{tool_def "export"} tool retrieves theory exports from the session
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   606
  database. Its command-line usage is: @{verbatim [display]
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   607
\<open>Usage: isabelle export [OPTIONS] SESSION
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   608
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   609
  Options are:
68314
2acbf8129d8b clarified option -O: avoid conflict with build/dump option -D;
wenzelm
parents: 68292
diff changeset
   610
    -O DIR       output directory for exported files (default: "export")
68116
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   611
    -d DIR       include session directory
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   612
    -l           list exports
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   613
    -n           no build of session
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   614
    -o OPTION    override Isabelle system OPTION (via NAME=VAL or NAME)
69671
2486792eaf61 support pruning of export names;
wenzelm
parents: 69610
diff changeset
   615
    -p NUM       prune path of exported files by NUM elements
75161
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   616
    -x PATTERN   extract files matching pattern (e.g. "*:**" for all)
68116
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   617
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   618
  List or export theory exports for SESSION: named blobs produced by
68290
f1f5ccc85b25 support multiple patterns;
wenzelm
parents: 68219
diff changeset
   619
  isabelle build. Option -l or -x is required; option -x may be repeated.
68116
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   620
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   621
  The PATTERN language resembles glob patterns in the shell, with ? and *
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   622
  (both excluding ":" and "/"), ** (excluding ":"), and [abc] or [^abc],
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   623
  and variants {pattern1,pattern2,pattern3}.\<close>}
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   624
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   625
  \<^medskip>
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   626
  The specified session is updated via @{tool build}
69854
cc0b3e177b49 system option "system_heaps" supersedes various command-line options for "system build mode";
wenzelm
parents: 69811
diff changeset
   627
  (\secref{sec:tool-build}), with the same options \<^verbatim>\<open>-d\<close>, \<^verbatim>\<open>-o\<close>. The option
cc0b3e177b49 system option "system_heaps" supersedes various command-line options for "system build mode";
wenzelm
parents: 69811
diff changeset
   628
  \<^verbatim>\<open>-n\<close> suppresses the implicit build process: it means that a potentially
cc0b3e177b49 system option "system_heaps" supersedes various command-line options for "system build mode";
wenzelm
parents: 69811
diff changeset
   629
  outdated session database is used!
68116
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   630
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   631
  \<^medskip>
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   632
  Option \<^verbatim>\<open>-l\<close> lists all stored exports, with compound names
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   633
  \<open>theory\<close>\<^verbatim>\<open>:\<close>\<open>name\<close>.
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   634
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   635
  \<^medskip>
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   636
  Option \<^verbatim>\<open>-x\<close> extracts stored exports whose compound name matches the given
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   637
  pattern. Note that wild cards ``\<^verbatim>\<open>?\<close>'' and ``\<^verbatim>\<open>*\<close>'' do not match the
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   638
  separators ``\<^verbatim>\<open>:\<close>'' and ``\<^verbatim>\<open>/\<close>''; the wild card \<^verbatim>\<open>**\<close> matches over directory
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   639
  name hierarchies separated by ``\<^verbatim>\<open>/\<close>''. Thus the pattern ``\<^verbatim>\<open>*:**\<close>'' matches
68290
f1f5ccc85b25 support multiple patterns;
wenzelm
parents: 68219
diff changeset
   640
  \<^emph>\<open>all\<close> theory exports. Multiple options \<^verbatim>\<open>-x\<close> refer to the union of all
f1f5ccc85b25 support multiple patterns;
wenzelm
parents: 68219
diff changeset
   641
  specified patterns.
68116
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   642
68314
2acbf8129d8b clarified option -O: avoid conflict with build/dump option -D;
wenzelm
parents: 68292
diff changeset
   643
  Option \<^verbatim>\<open>-O\<close> specifies an alternative output directory for option \<^verbatim>\<open>-x\<close>: the
68116
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   644
  default is \<^verbatim>\<open>export\<close> within the current directory. Each theory creates its
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   645
  own sub-directory hierarchy, using the session-qualified theory name.
69671
2486792eaf61 support pruning of export names;
wenzelm
parents: 69610
diff changeset
   646
2486792eaf61 support pruning of export names;
wenzelm
parents: 69610
diff changeset
   647
  Option \<^verbatim>\<open>-p\<close> specifies the number of elements that should be pruned from
2486792eaf61 support pruning of export names;
wenzelm
parents: 69610
diff changeset
   648
  each name: it allows to reduce the resulting directory hierarchy at the
2486792eaf61 support pruning of export names;
wenzelm
parents: 69610
diff changeset
   649
  danger of overwriting files due to loss of uniqueness.
68116
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   650
\<close>
ac82ee617a75 command-line tool "isabelle export";
wenzelm
parents: 67605
diff changeset
   651
68348
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   652
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   653
section \<open>Dump PIDE session database \label{sec:tool-dump}\<close>
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   654
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   655
text \<open>
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   656
  The @{tool_def "dump"} tool dumps information from the cumulative PIDE
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   657
  session database (which is processed on the spot). Its command-line usage
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   658
  is: @{verbatim [display]
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   659
\<open>Usage: isabelle dump [OPTIONS] [SESSIONS ...]
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   660
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   661
  Options are:
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   662
    -A NAMES     dump named aspects (default: ...)
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   663
    -B NAME      include session NAME and all descendants
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   664
    -D DIR       include session directory and select its sessions
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   665
    -O DIR       output directory for dumped files (default: "dump")
71807
cdfa8f027bb9 tuned messages;
wenzelm
parents: 70863
diff changeset
   666
    -R           refer to requirements of selected sessions
68348
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   667
    -X NAME      exclude sessions from group NAME and all descendants
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   668
    -a           select all sessions
70862
a4ccd277e9c4 clarified treatment of base logic image;
wenzelm
parents: 70861
diff changeset
   669
    -b NAME      base logic image (default "Pure")
68348
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   670
    -d DIR       include session directory
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   671
    -g NAME      select session group NAME
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   672
    -o OPTION    override Isabelle system OPTION (via NAME=VAL or NAME)
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   673
    -v           verbose
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   674
    -x NAME      exclude session NAME and all descendants
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   675
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   676
  Dump cumulative PIDE session database, with the following aspects:
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   677
    ...\<close>}
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   678
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   679
  \<^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
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   680
  remaining command-line arguments specify sessions as in @{tool build}
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   681
  (\secref{sec:tool-build}): the cumulative PIDE database of all their loaded
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   682
  theories is dumped to the output directory of option \<^verbatim>\<open>-O\<close> (default: \<^verbatim>\<open>dump\<close>
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   683
  in the current directory).
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   684
70859
6e6254bbce1f split into standard partitions, for improved scalability;
wenzelm
parents: 70858
diff changeset
   685
  \<^medskip> Option \<^verbatim>\<open>-b\<close> specifies an optional base logic image, for improved
70862
a4ccd277e9c4 clarified treatment of base logic image;
wenzelm
parents: 70861
diff changeset
   686
  scalability of the PIDE session. Its theories are only processed if it is
a4ccd277e9c4 clarified treatment of base logic image;
wenzelm
parents: 70861
diff changeset
   687
  included in the overall session selection.
70859
6e6254bbce1f split into standard partitions, for improved scalability;
wenzelm
parents: 70858
diff changeset
   688
68348
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   689
  \<^medskip> Option \<^verbatim>\<open>-o\<close> overrides Isabelle system options as for @{tool build}
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   690
  (\secref{sec:tool-build}).
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   691
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   692
  \<^medskip> Option \<^verbatim>\<open>-v\<close> increases the general level of verbosity.
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   693
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   694
  \<^medskip> Option \<^verbatim>\<open>-A\<close> specifies named aspects of the dump, as a comma-separated
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   695
  list. The default is to dump all known aspects, as given in the command-line
71893
a27747c85700 more antiquotations;
wenzelm
parents: 71808
diff changeset
   696
  usage of the tool. The underlying Isabelle/Scala operation
a27747c85700 more antiquotations;
wenzelm
parents: 71808
diff changeset
   697
  \<^scala_method>\<open>isabelle.Dump.dump\<close> takes aspects as user-defined
a27747c85700 more antiquotations;
wenzelm
parents: 71808
diff changeset
   698
  operations on the final PIDE state and document version. This allows to
a27747c85700 more antiquotations;
wenzelm
parents: 71808
diff changeset
   699
  imitate Prover IDE rendering under program control.
68348
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   700
\<close>
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   701
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   702
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   703
subsubsection \<open>Examples\<close>
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   704
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   705
text \<open>
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   706
  Dump all Isabelle/ZF sessions (which are rather small):
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   707
  @{verbatim [display] \<open>  isabelle dump -v -B ZF\<close>}
68348
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   708
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   709
  \<^smallskip>
70862
a4ccd277e9c4 clarified treatment of base logic image;
wenzelm
parents: 70861
diff changeset
   710
  Dump the quite substantial \<^verbatim>\<open>HOL-Analysis\<close> session, with full bootstrap
a4ccd277e9c4 clarified treatment of base logic image;
wenzelm
parents: 70861
diff changeset
   711
  from Isabelle/Pure:
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   712
  @{verbatim [display] \<open>  isabelle dump -v HOL-Analysis\<close>}
68348
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   713
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   714
  \<^smallskip>
70862
a4ccd277e9c4 clarified treatment of base logic image;
wenzelm
parents: 70861
diff changeset
   715
  Dump all sessions connected to HOL-Analysis, using main Isabelle/HOL as
a4ccd277e9c4 clarified treatment of base logic image;
wenzelm
parents: 70861
diff changeset
   716
  basis:
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   717
  @{verbatim [display] \<open>  isabelle dump -v -b HOL -B HOL-Analysis\<close>}
68348
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   718
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   719
  This results in uniform PIDE markup for everything, except for the
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   720
  Isabelle/Pure bootstrap process itself. Producing that on the spot requires
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   721
  several GB of heap space, both for the Isabelle/Scala and Isabelle/ML
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   722
  process (in 64bit mode). Here are some relevant settings (\secref{sec:boot})
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   723
  for such ambitious applications:
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   724
  @{verbatim [display]
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   725
\<open>  ISABELLE_TOOL_JAVA_OPTIONS="-Xms4g -Xmx32g -Xss16m"
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   726
  ML_OPTIONS="--minheap 4G --maxheap 32G"
68348
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   727
\<close>}
69599
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   728
\<close>
68348
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   729
69599
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   730
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   731
section \<open>Update theory sources based on PIDE markup \label{sec:tool-update}\<close>
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   732
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   733
text \<open>
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   734
  The @{tool_def "update"} tool updates theory sources based on markup that is
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   735
  produced from a running PIDE session (similar to @{tool dump}
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   736
  \secref{sec:tool-dump}). Its command-line usage is: @{verbatim [display]
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   737
\<open>Usage: isabelle update [OPTIONS] [SESSIONS ...]
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   738
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   739
  Options are:
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   740
    -B NAME      include session NAME and all descendants
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   741
    -D DIR       include session directory and select its sessions
71807
cdfa8f027bb9 tuned messages;
wenzelm
parents: 70863
diff changeset
   742
    -R           refer to requirements of selected sessions
69599
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   743
    -X NAME      exclude sessions from group NAME and all descendants
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   744
    -a           select all sessions
70863
d1299774543d clarified "isabelle update" options -- more like "isabelle dump";
wenzelm
parents: 70862
diff changeset
   745
    -b NAME      base logic image (default "Pure")
69599
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   746
    -d DIR       include session directory
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   747
    -g NAME      select session group NAME
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   748
    -o OPTION    override Isabelle system OPTION (via NAME=VAL or NAME)
75161
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   749
    -u OPT       override "update" option: shortcut for "-o update_OPT"
69599
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   750
    -v           verbose
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   751
    -x NAME      exclude session NAME and all descendants
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   752
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   753
  Update theory sources based on PIDE markup.\<close>}
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   754
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   755
  \<^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
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   756
  remaining command-line arguments specify sessions as in @{tool build}
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   757
  (\secref{sec:tool-build}) or @{tool dump} (\secref{sec:tool-dump}).
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   758
70863
d1299774543d clarified "isabelle update" options -- more like "isabelle dump";
wenzelm
parents: 70862
diff changeset
   759
  \<^medskip> Option \<^verbatim>\<open>-b\<close> specifies an optional base logic image, for improved
d1299774543d clarified "isabelle update" options -- more like "isabelle dump";
wenzelm
parents: 70862
diff changeset
   760
  scalability of the PIDE session. Its theories are only processed if it is
d1299774543d clarified "isabelle update" options -- more like "isabelle dump";
wenzelm
parents: 70862
diff changeset
   761
  included in the overall session selection.
69599
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   762
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   763
  \<^medskip> Option \<^verbatim>\<open>-v\<close> increases the general level of verbosity.
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   764
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   765
  \<^medskip> Option \<^verbatim>\<open>-o\<close> overrides Isabelle system options as for @{tool build}
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   766
  (\secref{sec:tool-build}). Option \<^verbatim>\<open>-u\<close> refers to specific \<^verbatim>\<open>update\<close>
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   767
  options, by relying on naming convention: ``\<^verbatim>\<open>-u\<close>~\<open>OPT\<close>'' is a shortcut for
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   768
  ``\<^verbatim>\<open>-o\<close>~\<^verbatim>\<open>update_\<close>\<open>OPT\<close>''.
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   769
75161
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   770
  \<^medskip> The following \<^verbatim>\<open>update\<close> options are supported:
69599
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   771
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   772
    \<^item> @{system_option update_inner_syntax_cartouches} to update inner syntax
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   773
    (types, terms, etc.)~to use cartouches, instead of double-quoted strings
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   774
    or atomic identifiers. For example, ``\<^theory_text>\<open>lemma \<doublequote>x =
69610
10644973cdde proper example for inner syntax, not name;
wenzelm
parents: 69603
diff changeset
   775
    x\<doublequote>\<close>'' is replaced by ``\<^theory_text>\<open>lemma \<open>x = x\<close>\<close>'', and ``\<^theory_text>\<open>assume
10644973cdde proper example for inner syntax, not name;
wenzelm
parents: 69603
diff changeset
   776
    A\<close>'' is replaced by ``\<^theory_text>\<open>assume \<open>A\<close>\<close>''.
69599
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   777
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   778
    \<^item> @{system_option update_mixfix_cartouches} to update mixfix templates to
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   779
    use cartouches instead of double-quoted strings. For example, ``\<^theory_text>\<open>(infixl
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   780
    \<doublequote>+\<doublequote> 65)\<close>'' is replaced by ``\<^theory_text>\<open>(infixl \<open>+\<close>
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   781
    65)\<close>''.
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   782
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   783
    \<^item> @{system_option update_control_cartouches} to update antiquotations to
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   784
    use the compact form with control symbol and cartouche argument. For
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   785
    example, ``\<open>@{term \<doublequote>x + y\<doublequote>}\<close>'' is replaced by
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   786
    ``\<open>\<^term>\<open>x + y\<close>\<close>'' (the control symbol is literally \<^verbatim>\<open>\<^term>\<close>.)
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   787
69603
67ae2e164c0f support for isabelle update -u path_cartouches;
wenzelm
parents: 69602
diff changeset
   788
    \<^item> @{system_option update_path_cartouches} to update file-system paths to
67ae2e164c0f support for isabelle update -u path_cartouches;
wenzelm
parents: 69602
diff changeset
   789
    use cartouches: this depends on language markup provided by semantic
67ae2e164c0f support for isabelle update -u path_cartouches;
wenzelm
parents: 69602
diff changeset
   790
    processing of parsed input.
67ae2e164c0f support for isabelle update -u path_cartouches;
wenzelm
parents: 69602
diff changeset
   791
69599
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   792
  It is also possible to produce custom updates in Isabelle/ML, by reporting
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   793
  \<^ML>\<open>Markup.update\<close> with the precise source position and a replacement
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   794
  text. This operation should be made conditional on specific system options,
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   795
  similar to the ones above. Searching the above option names in ML sources of
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   796
  \<^dir>\<open>$ISABELLE_HOME/src/Pure\<close> provides some examples.
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   797
69602
48e973251070 clarified documentation;
wenzelm
parents: 69601
diff changeset
   798
  Updates can be in conflict by producing nested or overlapping edits: this
48e973251070 clarified documentation;
wenzelm
parents: 69601
diff changeset
   799
  may require to run @{tool update} multiple times.
69599
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   800
\<close>
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   801
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   802
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   803
subsubsection \<open>Examples\<close>
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   804
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   805
text \<open>
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   806
  Update some cartouche notation in all theory sources required for session
70863
d1299774543d clarified "isabelle update" options -- more like "isabelle dump";
wenzelm
parents: 70862
diff changeset
   807
  \<^verbatim>\<open>HOL-Analysis\<close> (and ancestors):
69599
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   808
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   809
  @{verbatim [display] \<open>  isabelle update -u mixfix_cartouches HOL-Analysis\<close>}
69599
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   810
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   811
  \<^smallskip> Update the same for all application sessions based on \<^verbatim>\<open>HOL-Analysis\<close> ---
75161
95612f330c93 misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents: 74873
diff changeset
   812
  using its image as starting point (for reduced resource requirements):
69599
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   813
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   814
  @{verbatim [display] \<open>  isabelle update -u mixfix_cartouches -b HOL-Analysis -B HOL-Analysis\<close>}
69599
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   815
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   816
  \<^smallskip> Update sessions that build on \<^verbatim>\<open>HOL-Proofs\<close>, which need to be run
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   817
  separately with special options as follows:
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   818
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   819
  @{verbatim [display] \<open>  isabelle update -u mixfix_cartouches -l HOL-Proofs -B HOL-Proofs
69601
c51a9bd4cf09 redundant (see isabelle.Dump.make_options);
wenzelm
parents: 69599
diff changeset
   820
  -o record_proofs=2\<close>}
69599
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   821
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   822
  \<^smallskip> See also the end of \secref{sec:tool-dump} for hints on increasing
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   823
  Isabelle/ML heap sizes for very big PIDE processes that include many
caa7e406056d documentation on "isabelle update";
wenzelm
parents: 69593
diff changeset
   824
  sessions, notably from the Archive of Formal Proofs.
68348
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   825
\<close>
2ac3a5c07dfa documentation for "isabelle dump";
wenzelm
parents: 68314
diff changeset
   826
71808
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   827
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   828
section \<open>Explore sessions structure\<close>
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   829
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   830
text \<open>
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   831
  The @{tool_def "sessions"} tool explores the sessions structure. Its
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   832
  command-line usage is:
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   833
  @{verbatim [display]
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   834
\<open>Usage: isabelle sessions [OPTIONS] [SESSIONS ...]
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   835
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   836
  Options are:
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   837
    -B NAME      include session NAME and all descendants
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   838
    -D DIR       include session directory and select its sessions
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   839
    -R           refer to requirements of selected sessions
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   840
    -X NAME      exclude sessions from group NAME and all descendants
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   841
    -a           select all sessions
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   842
    -d DIR       include session directory
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   843
    -g NAME      select session group NAME
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   844
    -x NAME      exclude session NAME and all descendants
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   845
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   846
  Explore the structure of Isabelle sessions and print result names in
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   847
  topological order (on stdout).\<close>}
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   848
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   849
  Arguments and options for session selection resemble @{tool build}
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   850
  (\secref{sec:tool-build}).
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   851
\<close>
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   852
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   853
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   854
subsubsection \<open>Examples\<close>
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   855
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   856
text \<open>
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   857
  All sessions of the Isabelle distribution:
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   858
  @{verbatim [display] \<open>  isabelle sessions -a\<close>}
71808
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   859
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   860
  \<^medskip>
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   861
  Sessions that are based on \<^verbatim>\<open>ZF\<close> (and required by it):
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   862
  @{verbatim [display] \<open>  isabelle sessions -B ZF\<close>}
71808
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   863
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   864
  \<^medskip>
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   865
  All sessions of Isabelle/AFP (based in directory \<^path>\<open>AFP\<close>):
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   866
  @{verbatim [display] \<open>  isabelle sessions -D AFP/thys\<close>}
71808
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   867
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   868
  \<^medskip>
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   869
  Sessions required by Isabelle/AFP (based in directory \<^path>\<open>AFP\<close>):
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   870
  @{verbatim [display] \<open>  isabelle sessions -R -D AFP/thys\<close>}
71808
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   871
\<close>
e2ad50885887 added "isabelle sessions" tool;
wenzelm
parents: 71807
diff changeset
   872
75556
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   873
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   874
section \<open>Synchronize source repositories and session images for Isabelle and AFP\<close>
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   875
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   876
text \<open>
75557
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   877
  The @{tool_def sync} tool synchronizes a local Isabelle and AFP source
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   878
  repository, possibly with prebuilt \<^verbatim>\<open>.jar\<close> files and session images. Its
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   879
  command-line usage is: @{verbatim [display]
75556
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   880
\<open>Usage: isabelle sync [OPTIONS] TARGET
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   881
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   882
  Options are:
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   883
    -A ROOT      include AFP with given root directory (":" for $AFP_BASE)
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   884
    -H           purge heaps directory on target
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   885
    -I NAME      include session heap image and build database
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   886
                 (based on accidental local state)
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   887
    -J           preserve *.jar files
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   888
    -S           robust (but less portable) treatment of spaces in
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   889
                 file and directory names on the target
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   890
    -T           thorough treatment of file content and directory times
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   891
    -a REV       explicit AFP revision (default: state of working directory)
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   892
    -n           no changes: dry-run
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   893
    -r REV       explicit revision (default: state of working directory)
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   894
    -p PORT      explicit SSH port (default: 22)
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   895
    -v           verbose
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   896
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   897
  Synchronize Isabelle + AFP repositories, based on "isabelle hg_sync".\<close>}
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   898
75557
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   899
  The approach is to apply @{tool hg_sync} (see \secref{sec:tool-hg-sync}) on
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   900
  the underlying Isabelle repository, and an optional AFP repository.
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   901
  Consequently, the Isabelle installation needs to be a Mercurial repository
75556
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   902
  clone: a regular download of the Isabelle distribution is not sufficient!
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   903
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   904
  On the target side, AFP is placed into @{setting ISABELLE_HOME} as immediate
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   905
  sub-directory with the literal name \<^verbatim>\<open>AFP\<close>; thus it can be easily included
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   906
  elsewhere, e.g. @{tool build}~\<^verbatim>\<open>-d\<close>~\<^verbatim>\<open>'~~/AFP'\<close> on the remote side.
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   907
75559
5340239ff468 clarified options of "isabelle hg_sync" vs. "isabelle sync";
wenzelm
parents: 75558
diff changeset
   908
  \<^medskip> Options \<^verbatim>\<open>-S\<close>, \<^verbatim>\<open>-T\<close>, \<^verbatim>\<open>-n\<close>, \<^verbatim>\<open>-p\<close>, \<^verbatim>\<open>-v\<close> are the same as the underlying
75556
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   909
  @{tool hg_sync}.
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   910
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   911
  \<^medskip> Options \<^verbatim>\<open>-r\<close> and \<^verbatim>\<open>-a\<close> are the same as option \<^verbatim>\<open>-r\<close> for @{tool hg_sync},
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   912
  but for the Isabelle and AFP repositories, respectively. The AFP version is
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   913
  only used if a corresponding repository is given via option \<^verbatim>\<open>-A\<close>, either
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   914
  with explicit root directory, or as \<^verbatim>\<open>-A:\<close> to refer to \<^verbatim>\<open>$AFP_BASE\<close> (this
75557
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   915
  assumes AFP as component of the local Isabelle installation). If no AFP
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   916
  repository is given, an existing \<^verbatim>\<open>AFP\<close> directory on the target remains
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   917
  unchanged.
75556
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   918
75557
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   919
  \<^medskip> Option \<^verbatim>\<open>-J\<close> uploads existing \<^verbatim>\<open>.jar\<close> files from the working directory,
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   920
  which are usually Isabelle/Scala/Java modules under control of @{tool
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   921
  scala_build} via \<^verbatim>\<open>etc/build.props\<close> (see also \secref{sec:scala-build}).
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   922
  Thus the dependency management is accurate: bad uploads will be rebuilt
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   923
  eventually (or ignored). This might fail for very old Isabelle versions,
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   924
  when going into the past via option \<^verbatim>\<open>-r\<close>: here it is better to omit option
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   925
  \<^verbatim>\<open>-J\<close> and thus purge \<^verbatim>\<open>.jar\<close> files on the target (because they do not belong
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   926
  to the repository).
75556
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   927
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   928
  \<^medskip> Option \<^verbatim>\<open>-I\<close> uploads a collection of session images. The set of \<^verbatim>\<open>-I\<close>
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   929
  options specifies the end-points in the session build graph, including all
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   930
  required ancestors. The result collection is uploaded using the underlying
75557
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   931
  \<^verbatim>\<open>rsync\<close> policies, so unchanged images are not sent again. Session images
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   932
  are assembled within the target \<^verbatim>\<open>heaps\<close> directory: this scheme fits
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   933
  together with @{tool build}~\<^verbatim>\<open>-o system_heaps\<close>. Images are taken as-is from
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   934
  the local Isabelle installation, regardless of option \<^verbatim>\<open>-r\<close>. Upload of bad
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   935
  images could waste time and space, but running e.g. @{tool build} on the
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   936
  target will check dependencies accurately and rebuild outdated images on
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   937
  demand.
75556
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   938
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   939
  \<^medskip> Option \<^verbatim>\<open>-H\<close> tells the underlying \<^verbatim>\<open>rsync\<close> process to purge the \<^verbatim>\<open>heaps\<close>
75557
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   940
  directory on the target, before uploading new images via option \<^verbatim>\<open>-I\<close>. The
75556
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   941
  default is to work monotonically: old material that is not overwritten
75557
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   942
  remains unchanged. Over time, this may lead to unused garbage, due to
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   943
  changes in session names or the Poly/ML version. Option \<^verbatim>\<open>-H\<close> helps to avoid
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   944
  wasting file-system space.
75556
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   945
\<close>
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   946
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   947
subsubsection \<open>Examples\<close>
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   948
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   949
text \<open>
75557
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   950
  For quick testing of Isabelle + AFP on a remote machine, upload changed
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   951
  sources, jars, and local sessions images for \<^verbatim>\<open>HOL\<close>:
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   952
75556
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   953
  @{verbatim [display] \<open>  isabelle sync -A: -I HOL -J testmachine:test/isabelle_afp\<close>}
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   954
  Assuming that the local \<^verbatim>\<open>HOL\<close> hierarchy has been up-to-date, and the local
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   955
  and remote ML platforms coincide, a remote @{tool build} will proceed
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   956
  without building \<^verbatim>\<open>HOL\<close> again.
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   957
75557
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   958
  \<^medskip> Here is a variation for extra-clean testing of Isabelle + AFP: no option
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   959
  \<^verbatim>\<open>-J\<close>, but option \<^verbatim>\<open>-T\<close> to disable the default ``quick check'' of \<^verbatim>\<open>rsync\<close>
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   960
  (which only inspects file sizes and date stamps); existing heaps are
df14a62129e9 misc tuning;
wenzelm
parents: 75556
diff changeset
   961
  deleted:
75558
cf69c9112d09 tuned layout;
wenzelm
parents: 75557
diff changeset
   962
  @{verbatim [display] \<open>  isabelle sync -A: -T -H testmachine:test/isabelle_afp\<close>}
75556
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   963
\<close>
1f6fc2416a48 clarified document structure;
wenzelm
parents: 75161
diff changeset
   964
48578
21361b6189a6 some description of isabelle build;
wenzelm
parents:
diff changeset
   965
end