doc-src/System/Thy/Presentation.thy
author wenzelm
Tue, 07 Aug 2012 23:43:05 +0200
changeset 48722 a5e3ba7cbb2a
parent 48657 63ef2f0cf8bb
child 48814 d488a5f25bf6
permissions -rw-r--r--
discontinued obsolete IsaMakefile and ROOT.ML files from the Isabelle distribution;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
     1
theory Presentation
43564
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 42009
diff changeset
     2
imports Base
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
     3
begin
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
     4
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
     5
chapter {* Presenting theories \label{ch:present} *}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
     6
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
     7
text {*
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
     8
  Isabelle provides several ways to present the outcome of formal
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
     9
  developments, including WWW-based browsable libraries or actual
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    10
  printable documents.  Presentation is centered around the concept of
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    11
  \emph{logic sessions}.  The global session structure is that of a
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    12
  tree, with Isabelle Pure at its root, further object-logics derived
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    13
  (e.g.\ HOLCF from HOL, and HOL from Pure), and application sessions
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    14
  in leaf positions (usually without a separate image).
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    15
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    16
  The tools @{tool_ref mkdir} and @{tool_ref make} provide the primary
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    17
  means for managing Isabelle sessions, including proper setup for
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    18
  presentation.  Here @{tool_ref usedir} takes care to let
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    19
  @{executable_ref "isabelle-process"} process run any additional
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    20
  stages required for document preparation, notably the tools
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    21
  @{tool_ref document} and @{tool_ref latex}.  The complete tool chain
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    22
  for managing batch-mode Isabelle sessions is illustrated in
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    23
  \figref{fig:session-tools}.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    24
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    25
  \begin{figure}[htbp]
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    26
  \begin{center}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    27
  \begin{tabular}{lp{0.6\textwidth}}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    28
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    29
      @{tool_ref mkdir} & invoked once by the user to create the
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    30
      initial source setup (common @{verbatim IsaMakefile} plus a
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    31
      single session directory); \\
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    32
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    33
      @{tool make} & invoked repeatedly by the user to keep session
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    34
      output up-to-date (HTML, documents etc.); \\
28238
398bf960d3d4 misc tuning and modernization;
wenzelm
parents: 28225
diff changeset
    35
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    36
      @{tool usedir} & part of the standard @{verbatim IsaMakefile}
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    37
      entry of a session; \\
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    38
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    39
      @{executable "isabelle-process"} & run through @{tool_ref
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    40
      usedir}; \\
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    41
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    42
      @{tool_ref document} & run by the Isabelle process if document
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    43
      preparation is enabled; \\
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    44
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    45
      @{tool_ref latex} & universal {\LaTeX} tool wrapper invoked
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    46
      multiple times by @{tool_ref document}; also useful for manual
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    47
      experiments; \\
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    48
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    49
  \end{tabular}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    50
  \caption{The tool chain of Isabelle session presentation} \label{fig:session-tools}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    51
  \end{center}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    52
  \end{figure}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    53
*}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    54
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    55
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    56
section {* Generating theory browser information \label{sec:info} *}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    57
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    58
text {*
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    59
  \index{theory browsing information|bold}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    60
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    61
  As a side-effect of running a logic sessions, Isabelle is able to
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    62
  generate theory browsing information, including HTML documents that
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    63
  show a theory's definition, the theorems proved in its ML file and
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    64
  the relationship with its ancestors and descendants.  Besides the
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    65
  HTML file that is generated for every theory, Isabelle stores links
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    66
  to all theories in an index file. These indexes are linked with
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    67
  other indexes to represent the overall tree structure of logic
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    68
  sessions.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    69
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    70
  Isabelle also generates graph files that represent the theory
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    71
  hierarchy of a logic.  There is a graph browser Java applet embedded
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    72
  in the generated HTML pages, and also a stand-alone application that
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    73
  allows browsing theory graphs without having to start a WWW client
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    74
  first.  The latter version also includes features such as generating
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    75
  Postscript files, which are not available in the applet version.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    76
  See \secref{sec:browse} for further information.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    77
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    78
  \medskip
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    79
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    80
  The easiest way to let Isabelle generate theory browsing information
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    81
  for existing sessions is to append ``@{verbatim "-i true"}'' to the
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    82
  @{setting_ref ISABELLE_USEDIR_OPTIONS} before invoking @{tool make}.
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    83
  For example, add something like this to your Isabelle settings file
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    84
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    85
\begin{ttbox}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    86
ISABELLE_USEDIR_OPTIONS="-i true"
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    87
\end{ttbox}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    88
40800
330eb65c9469 Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents: 40387
diff changeset
    89
  and then change into the @{file "~~/src/FOL"} directory and run
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    90
  @{tool make}, or even @{tool make}~@{verbatim all}.  The
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    91
  presentation output will appear in @{verbatim
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    92
  "ISABELLE_BROWSER_INFO/FOL"}, which usually refers to something like
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    93
  @{verbatim "~/.isabelle/IsabelleXXXX/browser_info/FOL"}.  Note that
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    94
  option @{verbatim "-v true"} will make the internal runs of @{tool
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    95
  usedir} more explicit about such details.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    96
40800
330eb65c9469 Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents: 40387
diff changeset
    97
  Many standard Isabelle sessions (such as @{file "~~/src/HOL/ex"})
28238
398bf960d3d4 misc tuning and modernization;
wenzelm
parents: 28225
diff changeset
    98
  also provide actual printable documents.  These are prepared
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    99
  automatically as well if enabled like this, using the @{verbatim
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   100
  "-d"} option
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   101
\begin{ttbox}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   102
ISABELLE_USEDIR_OPTIONS="-i true -d dvi"
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   103
\end{ttbox}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   104
  Enabling options @{verbatim "-i"} and @{verbatim "-d"}
28225
wenzelm
parents: 28221
diff changeset
   105
  simultaneously as shown above causes an appropriate ``document''
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   106
  link to be included in the HTML index.  Documents (or raw document
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   107
  sources) may be generated independently of browser information as
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   108
  well, see \secref{sec:tool-document} for further details.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   109
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   110
  \bigskip The theory browsing information is stored in a
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   111
  sub-directory directory determined by the @{setting_ref
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   112
  ISABELLE_BROWSER_INFO} setting plus a prefix corresponding to the
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   113
  session identifier (according to the tree structure of sub-sessions
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   114
  by default).  A complete WWW view of all standard object-logics and
28225
wenzelm
parents: 28221
diff changeset
   115
  examples of the Isabelle distribution is available at the usual
wenzelm
parents: 28221
diff changeset
   116
  Isabelle sites:
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   117
  \begin{center}\small
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   118
  \begin{tabular}{l}
28225
wenzelm
parents: 28221
diff changeset
   119
    \url{http://isabelle.in.tum.de/dist/library/} \\
wenzelm
parents: 28221
diff changeset
   120
    \url{http://www.cl.cam.ac.uk/research/hvg/Isabelle/dist/library/} \\
wenzelm
parents: 28221
diff changeset
   121
    \url{http://mirror.cse.unsw.edu.au/pub/isabelle/dist/library/} \\
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   122
  \end{tabular}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   123
  \end{center}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   124
  
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   125
  \medskip In order to present your own theories on the web, simply
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   126
  copy the corresponding subdirectory from @{setting
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   127
  ISABELLE_BROWSER_INFO} to your WWW server, having generated browser
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   128
  info like this:
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   129
\begin{ttbox}
28504
7ad7d7d6df47 simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents: 28238
diff changeset
   130
isabelle usedir -i true HOL Foo
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   131
\end{ttbox}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   132
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   133
  This assumes that directory @{verbatim Foo} contains some @{verbatim
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   134
  ROOT.ML} file to load all your theories, and HOL is your parent
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   135
  logic image (@{tool_ref mkdir} assists in setting up Isabelle
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   136
  session directories.  Theory browser information for HOL should have
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   137
  been generated already beforehand.  Alternatively, one may specify
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   138
  an external link to an existing body of HTML data by giving @{tool
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   139
  usedir} a @{verbatim "-P"} option like this:
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   140
\begin{ttbox}
28504
7ad7d7d6df47 simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents: 28238
diff changeset
   141
isabelle usedir -i true -P http://isabelle.in.tum.de/library/ HOL Foo
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   142
\end{ttbox}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   143
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   144
  \medskip For production use, @{tool usedir} is usually invoked in an
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   145
  appropriate @{verbatim IsaMakefile}, via @{tool make}.  There is a
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   146
  separate @{tool mkdir} tool to provide easy setup of all this, with
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   147
  only minimal manual editing required.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   148
\begin{ttbox}
28504
7ad7d7d6df47 simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents: 28238
diff changeset
   149
isabelle mkdir HOL Foo && isabelle make
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   150
\end{ttbox}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   151
  See \secref{sec:tool-mkdir} for more information on preparing
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   152
  Isabelle session directories, including the setup for documents.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   153
*}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   154
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   155
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   156
section {* Creating Isabelle session directories
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   157
  \label{sec:tool-mkdir} *}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   158
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   159
text {* The @{tool_def mkdir} tool prepares Isabelle session source
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   160
  directories, including a sensible default setup of @{verbatim
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   161
  IsaMakefile}, @{verbatim ROOT.ML}, and a @{verbatim document}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   162
  directory with a minimal @{verbatim root.tex} that is sufficient to
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   163
  print all theories of the session (in the order of appearance); see
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   164
  \secref{sec:tool-document} for further information on Isabelle
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   165
  document preparation.  The usage of @{tool mkdir} is:
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   166
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   167
\begin{ttbox}
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   168
Usage: isabelle mkdir [OPTIONS] [LOGIC] NAME
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   169
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   170
  Options are:
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   171
    -I FILE      alternative IsaMakefile output
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   172
    -P           include parent logic target
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   173
    -b           setup build mode (session outputs heap image)
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   174
    -q           quiet mode
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   175
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   176
  Prepare session directory, including IsaMakefile and document source,
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   177
  with parent LOGIC (default ISABELLE_LOGIC=\$ISABELLE_LOGIC)
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   178
\end{ttbox}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   179
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   180
  The @{tool mkdir} tool is conservative in the sense that any
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   181
  existing @{verbatim IsaMakefile} etc.\ is left unchanged.  Thus it
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   182
  is safe to invoke it multiple times, although later runs may not
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   183
  have the desired effect.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   184
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   185
  Note that @{tool mkdir} is unable to change @{verbatim IsaMakefile}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   186
  incrementally --- manual changes are required for multiple
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   187
  sub-sessions.  On order to get an initial working session, the only
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   188
  editing needed is to add appropriate @{ML use_thy} calls to the
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   189
  generated @{verbatim ROOT.ML} file.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   190
*}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   191
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   192
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   193
subsubsection {* Options *}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   194
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   195
text {*
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   196
  The @{verbatim "-I"} option specifies an alternative to @{verbatim
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   197
  IsaMakefile} for dependencies.  Note that ``@{verbatim "-"}'' refers
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   198
  to \emph{stdout}, i.e.\ ``@{verbatim "-I-"}'' provides an easy way
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   199
  to peek at @{tool mkdir}'s idea of @{tool make} setup required for
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   200
  some particular of Isabelle session.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   201
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   202
  \medskip The @{verbatim "-P"} option includes a target for the
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   203
  parent @{verbatim LOGIC} session in the generated @{verbatim
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   204
  IsaMakefile}.  The corresponding sources are assumed to be located
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   205
  within the Isabelle distribution.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   206
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   207
  \medskip The @{verbatim "-b"} option sets up the current directory
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   208
  as the base for a new session that provides an actual logic image,
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   209
  as opposed to one that only runs several theories based on an
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   210
  existing image.  Note that in the latter case, everything except
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   211
  @{verbatim IsaMakefile} would be placed into a separate directory
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   212
  @{verbatim NAME}, rather than the current one.  See
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   213
  \secref{sec:tool-usedir} for further information on \emph{build
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   214
  mode} vs.\ \emph{example mode} of @{tool usedir}.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   215
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   216
  \medskip The @{verbatim "-q"} option enables quiet mode, suppressing
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   217
  further notes on how to proceed.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   218
*}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   219
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   220
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   221
subsubsection {* Examples *}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   222
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   223
text {*
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   224
  The standard setup of a single ``example session'' based on the
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   225
  default logic, with proper document generation is generated like
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   226
  this:
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   227
\begin{ttbox}
28504
7ad7d7d6df47 simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents: 28238
diff changeset
   228
isabelle mkdir Foo && isabelle make
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   229
\end{ttbox}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   230
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   231
  \noindent The theory sources should be put into the @{verbatim Foo}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   232
  directory, and its @{verbatim ROOT.ML} should be edited to load all
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   233
  required theories.  Invoking @{tool make} again would run the whole
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   234
  session, generating browser information and the document
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   235
  automatically.  The @{verbatim IsaMakefile} is typically tuned
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   236
  manually later, e.g.\ adding source dependencies, or changing the
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   237
  options passed to @{tool usedir}.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   238
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   239
  \medskip Large projects may demand further sessions, potentially
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   240
  with separate logic images being created.  This usually requires
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   241
  manual editing of the generated @{verbatim IsaMakefile}, which is
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   242
  meant to cover all of the sub-session directories at the same time
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   243
  (this is the deeper reasong why @{verbatim IsaMakefile} is not made
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   244
  part of the initial session directory created by @{tool mkdir}).
48722
a5e3ba7cbb2a discontinued obsolete IsaMakefile and ROOT.ML files from the Isabelle distribution;
wenzelm
parents: 48657
diff changeset
   245
*}
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   246
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   247
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   248
section {* Running Isabelle sessions \label{sec:tool-usedir} *}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   249
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   250
text {* The @{tool_def usedir} tool builds object-logic images, or
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   251
  runs example sessions based on existing logics. Its usage is:
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   252
\begin{ttbox}
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   253
Usage: isabelle usedir [OPTIONS] LOGIC NAME
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   254
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   255
  Options are:
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   256
    -C BOOL      copy existing document directory to -D PATH (default true)
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   257
    -D PATH      dump generated document sources into PATH
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   258
    -M MAX       multithreading: maximum number of worker threads (default 1)
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   259
    -P PATH      set path for remote theory browsing information
41819
2d84831dc1a0 scale parallel_proofs_threshold with max_threads_value to improve saturation of cores;
wenzelm
parents: 41703
diff changeset
   260
    -Q INT       set threshold for sub-proof parallelization (default 50)
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   261
    -T LEVEL     multithreading: trace level (default 0)
42004
e06351ffb475 tuned terminology for document variants;
wenzelm
parents: 41819
diff changeset
   262
    -V VARIANT   declare alternative document VARIANT
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   263
    -b           build mode (output heap image, using current dir)
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   264
    -d FORMAT    build document as FORMAT (default false)
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   265
    -f NAME      use ML file NAME (default ROOT.ML)
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   266
    -g BOOL      generate session graph image for document (default false)
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   267
    -i BOOL      generate theory browser information (default false)
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   268
    -m MODE      add print mode for output
32061
11f8ee55662d parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
wenzelm
parents: 31688
diff changeset
   269
    -p LEVEL     set level of detail for proof objects (default 0)
11f8ee55662d parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
wenzelm
parents: 31688
diff changeset
   270
    -q LEVEL     set level of parallel proof checking (default 1)
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   271
    -r           reset session path
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   272
    -s NAME      override session NAME
31688
f27cc190083b usedir: internal timing option;
wenzelm
parents: 31317
diff changeset
   273
    -t BOOL      internal session timing (default false)
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   274
    -v BOOL      be verbose (default false)
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   275
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   276
  Build object-logic or run examples. Also creates browsing
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   277
  information (HTML etc.) according to settings.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   278
47978
f8f503a1782a less specific sample usage;
wenzelm
parents: 43564
diff changeset
   279
  ISABELLE_USEDIR_OPTIONS=...
29435
a5f84ac14609 added parallel_proofs flag (default true, cf. usedir option -Q), which can be disabled in low-memory situations;
wenzelm
parents: 28914
diff changeset
   280
47978
f8f503a1782a less specific sample usage;
wenzelm
parents: 43564
diff changeset
   281
  ML_PLATFORM=...
f8f503a1782a less specific sample usage;
wenzelm
parents: 43564
diff changeset
   282
  ML_HOME=...
f8f503a1782a less specific sample usage;
wenzelm
parents: 43564
diff changeset
   283
  ML_SYSTEM=...
f8f503a1782a less specific sample usage;
wenzelm
parents: 43564
diff changeset
   284
  ML_OPTIONS=...
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   285
\end{ttbox}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   286
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   287
  Note that the value of the @{setting_ref ISABELLE_USEDIR_OPTIONS}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   288
  setting is implicitly prefixed to \emph{any} @{tool usedir}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   289
  call. Since the @{verbatim IsaMakefile}s of all object-logics
28238
398bf960d3d4 misc tuning and modernization;
wenzelm
parents: 28225
diff changeset
   290
  distributed with Isabelle just invoke @{tool usedir} for the real
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   291
  work, one may control compilation options globally via above
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   292
  variable. In particular, generation of \rmindex{HTML} browsing
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   293
  information and document preparation is controlled here.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   294
*}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   295
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   296
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   297
subsubsection {* Options *}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   298
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   299
text {*
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   300
  Basically, there are two different modes of operation: \emph{build
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   301
  mode} (enabled through the @{verbatim "-b"} option) and
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   302
  \emph{example mode} (default).
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   303
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   304
  Calling @{tool usedir} with @{verbatim "-b"} runs @{executable
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   305
  "isabelle-process"} with input image @{verbatim LOGIC} and output to
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   306
  @{verbatim NAME}, as provided on the command line. This will be a
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   307
  batch session, running @{verbatim ROOT.ML} from the current
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   308
  directory and then quitting.  It is assumed that @{verbatim ROOT.ML}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   309
  contains all ML commands required to build the logic.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   310
28238
398bf960d3d4 misc tuning and modernization;
wenzelm
parents: 28225
diff changeset
   311
  In example mode, @{tool usedir} runs a read-only session of
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   312
  @{verbatim LOGIC} and automatically runs @{verbatim ROOT.ML} from
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   313
  within directory @{verbatim NAME}.  It assumes that this file
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   314
  contains appropriate ML commands to run the desired examples.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   315
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   316
  \medskip The @{verbatim "-i"} option controls theory browser data
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   317
  generation. It may be explicitly turned on or off --- as usual, the
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   318
  last occurrence of @{verbatim "-i"} on the command line wins.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   319
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   320
  The @{verbatim "-P"} option specifies a path (or actual URL) to be
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   321
  prefixed to any \emph{non-local} reference of existing theories.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   322
  Thus user sessions may easily link to existing Isabelle libraries
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   323
  already present on the WWW.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   324
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   325
  The @{verbatim "-m"} options specifies additional print modes to be
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   326
  activated temporarily while the session is processed.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   327
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   328
  \medskip The @{verbatim "-d"} option controls document preparation.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   329
  Valid arguments are @{verbatim false} (do not prepare any document;
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   330
  this is default), or any of @{verbatim dvi}, @{verbatim dvi.gz},
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   331
  @{verbatim ps}, @{verbatim ps.gz}, @{verbatim pdf}.  The logic
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   332
  session has to provide a properly setup @{verbatim document}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   333
  directory.  See \secref{sec:tool-document} and
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   334
  \secref{sec:tool-latex} for more details.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   335
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   336
  \medskip The @{verbatim "-V"} option declares alternative document
42004
e06351ffb475 tuned terminology for document variants;
wenzelm
parents: 41819
diff changeset
   337
  variants, consisting of name/tags pairs (cf.\ options @{verbatim
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   338
  "-n"} and @{verbatim "-t"} of @{tool_ref document}).  The standard
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   339
  document is equivalent to ``@{verbatim
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   340
  "document=theory,proof,ML"}'', which means that all theory begin/end
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   341
  commands, proof body texts, and ML code will be presented
48616
be8002ee43d8 document variant NAME may use different LaTeX entry point document/root_NAME.tex if that file exists;
wenzelm
parents: 48602
diff changeset
   342
  faithfully.
be8002ee43d8 document variant NAME may use different LaTeX entry point document/root_NAME.tex if that file exists;
wenzelm
parents: 48602
diff changeset
   343
be8002ee43d8 document variant NAME may use different LaTeX entry point document/root_NAME.tex if that file exists;
wenzelm
parents: 48602
diff changeset
   344
  An alternative variant ``@{verbatim "outline=/proof/ML"}'' would
be8002ee43d8 document variant NAME may use different LaTeX entry point document/root_NAME.tex if that file exists;
wenzelm
parents: 48602
diff changeset
   345
  fold proof and ML parts, replacing the original text by a short
be8002ee43d8 document variant NAME may use different LaTeX entry point document/root_NAME.tex if that file exists;
wenzelm
parents: 48602
diff changeset
   346
  place-holder.  The form ``@{text name}@{verbatim "=-"},'' means to
be8002ee43d8 document variant NAME may use different LaTeX entry point document/root_NAME.tex if that file exists;
wenzelm
parents: 48602
diff changeset
   347
  remove document @{text name} from the list of variants to be
be8002ee43d8 document variant NAME may use different LaTeX entry point document/root_NAME.tex if that file exists;
wenzelm
parents: 48602
diff changeset
   348
  processed.  Any number of @{verbatim "-V"} options may be given;
be8002ee43d8 document variant NAME may use different LaTeX entry point document/root_NAME.tex if that file exists;
wenzelm
parents: 48602
diff changeset
   349
  later declarations have precedence over earlier ones.
be8002ee43d8 document variant NAME may use different LaTeX entry point document/root_NAME.tex if that file exists;
wenzelm
parents: 48602
diff changeset
   350
be8002ee43d8 document variant NAME may use different LaTeX entry point document/root_NAME.tex if that file exists;
wenzelm
parents: 48602
diff changeset
   351
  Some document variant @{text name} may use an alternative {\LaTeX}
be8002ee43d8 document variant NAME may use different LaTeX entry point document/root_NAME.tex if that file exists;
wenzelm
parents: 48602
diff changeset
   352
  entry point called @{verbatim "document/root_"}@{text
be8002ee43d8 document variant NAME may use different LaTeX entry point document/root_NAME.tex if that file exists;
wenzelm
parents: 48602
diff changeset
   353
  "name"}@{verbatim ".tex"} if that file exists; otherwise the common
be8002ee43d8 document variant NAME may use different LaTeX entry point document/root_NAME.tex if that file exists;
wenzelm
parents: 48602
diff changeset
   354
  @{verbatim "document/root.tex"} is used.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   355
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   356
  \medskip The @{verbatim "-g"} option produces images of the theory
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   357
  dependency graph (cf.\ \secref{sec:browse}) for inclusion in the
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   358
  generated document, both as @{verbatim session_graph.eps} and
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   359
  @{verbatim session_graph.pdf} at the same time.  To include this in
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   360
  the final {\LaTeX} document one could say @{verbatim
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   361
  "\\includegraphics{session_graph}"} in @{verbatim
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   362
  "document/root.tex"} (omitting the file-name extension enables
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   363
  {\LaTeX} to select to correct version, either for the DVI or PDF
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   364
  output path).
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   365
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   366
  \medskip The @{verbatim "-D"} option causes the generated document
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   367
  sources to be dumped at location @{verbatim PATH}; this path is
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   368
  relative to the session's main directory.  If the @{verbatim "-C"}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   369
  option is true, this will include a copy of an existing @{verbatim
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   370
  document} directory as provided by the user.  For example, @{tool
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   371
  usedir}~@{verbatim "-D generated HOL Foo"} produces a complete set
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   372
  of document sources at @{verbatim "Foo/generated"}.  Subsequent
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   373
  invocation of @{tool document}~@{verbatim "Foo/generated"} (see also
28238
398bf960d3d4 misc tuning and modernization;
wenzelm
parents: 28225
diff changeset
   374
  \secref{sec:tool-document}) will process the final result
398bf960d3d4 misc tuning and modernization;
wenzelm
parents: 28225
diff changeset
   375
  independently of an Isabelle job.  This decoupled mode of operation
398bf960d3d4 misc tuning and modernization;
wenzelm
parents: 28225
diff changeset
   376
  facilitates debugging of serious {\LaTeX} errors, for example.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   377
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   378
  \medskip The @{verbatim "-p"} option determines the level of detail
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   379
  for internal proof objects, see also the \emph{Isabelle Reference
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   380
  Manual}~\cite{isabelle-ref}.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   381
32061
11f8ee55662d parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
wenzelm
parents: 31688
diff changeset
   382
  \medskip The @{verbatim "-q"} option specifies the level of parallel
11f8ee55662d parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
wenzelm
parents: 31688
diff changeset
   383
  proof checking: @{verbatim 0} no proofs, @{verbatim 1} toplevel
11f8ee55662d parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
wenzelm
parents: 31688
diff changeset
   384
  proofs (default), @{verbatim 2} toplevel and nested Isar proofs.
41703
d27950860514 parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents: 40800
diff changeset
   385
  The option @{verbatim "-Q"} specifies a threshold for @{verbatim
d27950860514 parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents: 40800
diff changeset
   386
  "-q2"}: nested proofs are only parallelized when the current number
41819
2d84831dc1a0 scale parallel_proofs_threshold with max_threads_value to improve saturation of cores;
wenzelm
parents: 41703
diff changeset
   387
  of forked proofs falls below the given value (default 50),
2d84831dc1a0 scale parallel_proofs_threshold with max_threads_value to improve saturation of cores;
wenzelm
parents: 41703
diff changeset
   388
  multiplied by the number of worker threads (see option @{verbatim
2d84831dc1a0 scale parallel_proofs_threshold with max_threads_value to improve saturation of cores;
wenzelm
parents: 41703
diff changeset
   389
  "-M"}).
32061
11f8ee55662d parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
wenzelm
parents: 31688
diff changeset
   390
31688
f27cc190083b usedir: internal timing option;
wenzelm
parents: 31317
diff changeset
   391
  \medskip The @{verbatim "-t"} option produces a more detailed
f27cc190083b usedir: internal timing option;
wenzelm
parents: 31317
diff changeset
   392
  internal timing report of the session.
f27cc190083b usedir: internal timing option;
wenzelm
parents: 31317
diff changeset
   393
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   394
  \medskip The @{verbatim "-v"} option causes additional information
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   395
  to be printed while running the session, notably the location of
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   396
  prepared documents.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   397
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   398
  \medskip The @{verbatim "-M"} option specifies the maximum number of
41819
2d84831dc1a0 scale parallel_proofs_threshold with max_threads_value to improve saturation of cores;
wenzelm
parents: 41703
diff changeset
   399
  parallel worker threads used for processing independent tasks when
2d84831dc1a0 scale parallel_proofs_threshold with max_threads_value to improve saturation of cores;
wenzelm
parents: 41703
diff changeset
   400
  checking theory sources (multithreading only works on suitable ML
2d84831dc1a0 scale parallel_proofs_threshold with max_threads_value to improve saturation of cores;
wenzelm
parents: 41703
diff changeset
   401
  platforms).  The special value of @{verbatim 0} or @{verbatim max}
2d84831dc1a0 scale parallel_proofs_threshold with max_threads_value to improve saturation of cores;
wenzelm
parents: 41703
diff changeset
   402
  refers to the number of actual CPU cores of the underlying machine,
2d84831dc1a0 scale parallel_proofs_threshold with max_threads_value to improve saturation of cores;
wenzelm
parents: 41703
diff changeset
   403
  which is a good starting point for optimal performance tuning.  The
2d84831dc1a0 scale parallel_proofs_threshold with max_threads_value to improve saturation of cores;
wenzelm
parents: 41703
diff changeset
   404
  @{verbatim "-T"} option determines the level of detail in tracing
2d84831dc1a0 scale parallel_proofs_threshold with max_threads_value to improve saturation of cores;
wenzelm
parents: 41703
diff changeset
   405
  output concerning the internal locking and scheduling in
2d84831dc1a0 scale parallel_proofs_threshold with max_threads_value to improve saturation of cores;
wenzelm
parents: 41703
diff changeset
   406
  multithreaded operation.  This may be helpful in isolating
2d84831dc1a0 scale parallel_proofs_threshold with max_threads_value to improve saturation of cores;
wenzelm
parents: 41703
diff changeset
   407
  performance bottle-necks, e.g.\ due to excessive wait states when
2d84831dc1a0 scale parallel_proofs_threshold with max_threads_value to improve saturation of cores;
wenzelm
parents: 41703
diff changeset
   408
  locking critical code sections.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   409
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   410
  \medskip Any @{tool usedir} session is named by some \emph{session
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   411
  identifier}. These accumulate, documenting the way sessions depend
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   412
  on others. For example, consider @{verbatim "Pure/FOL/ex"}, which
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   413
  refers to the examples of FOL, which in turn is built upon Pure.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   414
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   415
  The current session's identifier is by default just the base name of
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   416
  the @{verbatim LOGIC} argument (in build mode), or of the @{verbatim
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   417
  NAME} argument (in example mode). This may be overridden explicitly
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   418
  via the @{verbatim "-s"} option.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   419
*}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   420
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   421
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   422
section {* Preparing Isabelle session documents
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   423
  \label{sec:tool-document} *}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   424
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   425
text {* The @{tool_def document} tool prepares logic session
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   426
  documents, processing the sources both as provided by the user and
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   427
  generated by Isabelle.  Its usage is:
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   428
\begin{ttbox}
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   429
Usage: isabelle document [OPTIONS] [DIR]
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   430
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   431
  Options are:
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   432
    -c           cleanup -- be aggressive in removing old stuff
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   433
    -n NAME      specify document name (default 'document')
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   434
    -o FORMAT    specify output format: dvi (default), dvi.gz, ps,
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   435
                 ps.gz, pdf
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   436
    -t TAGS      specify tagged region markup
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   437
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   438
  Prepare the theory session document in DIR (default 'document')
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   439
  producing the specified output format.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   440
\end{ttbox}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   441
  This tool is usually run automatically as part of the corresponding
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   442
  Isabelle batch process, provided document preparation has been
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   443
  enabled (cf.\ the @{verbatim "-d"} option of @{tool_ref usedir}).
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   444
  It may be manually invoked on the generated browser information
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   445
  document output as well, e.g.\ in case of errors encountered in the
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   446
  batch run.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   447
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   448
  \medskip The @{verbatim "-c"} option tells @{tool document} to
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   449
  dispose the document sources after successful operation.  This is
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   450
  the right thing to do for sources generated by an Isabelle process,
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   451
  but take care of your files in manual document preparation!
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   452
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   453
  \medskip The @{verbatim "-n"} and @{verbatim "-o"} option specify
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   454
  the final output file name and format, the default is ``@{verbatim
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   455
  document.dvi}''.  Note that the result will appear in the parent of
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   456
  the target @{verbatim DIR}.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   457
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   458
  \medskip The @{verbatim "-t"} option tells {\LaTeX} how to interpret
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   459
  tagged Isabelle command regions.  Tags are specified as a comma
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   460
  separated list of modifier/name pairs: ``@{verbatim "+"}@{text
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   461
  foo}'' (or just ``@{text foo}'') means to keep, ``@{verbatim
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   462
  "-"}@{text foo}'' to drop, and ``@{verbatim "/"}@{text foo}'' to
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   463
  fold text tagged as @{text foo}.  The builtin default is equivalent
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   464
  to the tag specification ``@{verbatim
30113
5ea17e90b08a isabelle document: adapted (postulated) defaults for tags to actual isabelle.sty;
wenzelm
parents: 29435
diff changeset
   465
  "+theory,+proof,+ML,+visible,-invisible"}''; see also the {\LaTeX}
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   466
  macros @{verbatim "\\isakeeptag"}, @{verbatim "\\isadroptag"}, and
40800
330eb65c9469 Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents: 40387
diff changeset
   467
  @{verbatim "\\isafoldtag"}, in @{file
28238
398bf960d3d4 misc tuning and modernization;
wenzelm
parents: 28225
diff changeset
   468
  "~~/lib/texinputs/isabelle.sty"}.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   469
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   470
  \medskip Document preparation requires a properly setup ``@{verbatim
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   471
  document}'' directory within the logic session sources.  This
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   472
  directory is supposed to contain all the files needed to produce the
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   473
  final document --- apart from the actual theories which are
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   474
  generated by Isabelle.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   475
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   476
  \medskip For most practical purposes, @{tool document} is smart
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   477
  enough to create any of the specified output formats, taking
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   478
  @{verbatim root.tex} supplied by the user as a starting point.  This
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   479
  even includes multiple runs of {\LaTeX} to accommodate references
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   480
  and bibliographies (the latter assumes @{verbatim root.bib} within
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   481
  the same directory).
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   482
48657
63ef2f0cf8bb simplified custom document/build script, instead of old-style document/IsaMakefile;
wenzelm
parents: 48616
diff changeset
   483
  In more complex situations, a separate @{verbatim build} script for
63ef2f0cf8bb simplified custom document/build script, instead of old-style document/IsaMakefile;
wenzelm
parents: 48616
diff changeset
   484
  the document sources may be given.  It is invoked with command-line
63ef2f0cf8bb simplified custom document/build script, instead of old-style document/IsaMakefile;
wenzelm
parents: 48616
diff changeset
   485
  arguments for the document format and the document variant name.
63ef2f0cf8bb simplified custom document/build script, instead of old-style document/IsaMakefile;
wenzelm
parents: 48616
diff changeset
   486
  The script needs to produce corresponding output files, e.g.\
63ef2f0cf8bb simplified custom document/build script, instead of old-style document/IsaMakefile;
wenzelm
parents: 48616
diff changeset
   487
  @{verbatim root.pdf} for target format @{verbatim pdf} (and default
63ef2f0cf8bb simplified custom document/build script, instead of old-style document/IsaMakefile;
wenzelm
parents: 48616
diff changeset
   488
  default variants).  The main work can be again delegated to @{tool
63ef2f0cf8bb simplified custom document/build script, instead of old-style document/IsaMakefile;
wenzelm
parents: 48616
diff changeset
   489
  latex}.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   490
42009
b008525c4399 parallel preparation of document variants, within separate directories;
wenzelm
parents: 42004
diff changeset
   491
  \medskip When running the session, Isabelle copies the content of
b008525c4399 parallel preparation of document variants, within separate directories;
wenzelm
parents: 42004
diff changeset
   492
  the original @{verbatim document} directory into its proper place
b008525c4399 parallel preparation of document variants, within separate directories;
wenzelm
parents: 42004
diff changeset
   493
  within @{setting ISABELLE_BROWSER_INFO}, according to the session
b008525c4399 parallel preparation of document variants, within separate directories;
wenzelm
parents: 42004
diff changeset
   494
  path and document variant.  Then, for any processed theory @{text A}
b008525c4399 parallel preparation of document variants, within separate directories;
wenzelm
parents: 42004
diff changeset
   495
  some {\LaTeX} source is generated and put there as @{text
b008525c4399 parallel preparation of document variants, within separate directories;
wenzelm
parents: 42004
diff changeset
   496
  A}@{verbatim ".tex"}.  Furthermore, a list of all generated theory
b008525c4399 parallel preparation of document variants, within separate directories;
wenzelm
parents: 42004
diff changeset
   497
  files is put into @{verbatim session.tex}.  Typically, the root
b008525c4399 parallel preparation of document variants, within separate directories;
wenzelm
parents: 42004
diff changeset
   498
  {\LaTeX} file provided by the user would include @{verbatim
b008525c4399 parallel preparation of document variants, within separate directories;
wenzelm
parents: 42004
diff changeset
   499
  session.tex} to get a document containing all the theories.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   500
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   501
  The {\LaTeX} versions of the theories require some macros defined in
40800
330eb65c9469 Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents: 40387
diff changeset
   502
  @{file "~~/lib/texinputs/isabelle.sty"}.  Doing @{verbatim
28238
398bf960d3d4 misc tuning and modernization;
wenzelm
parents: 28225
diff changeset
   503
  "\\usepackage{isabelle}"} in @{verbatim root.tex} should be fine;
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   504
  the underlying @{tool latex} already includes an appropriate path
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   505
  specification for {\TeX} inputs.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   506
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   507
  If the text contains any references to Isabelle symbols (such as
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   508
  @{verbatim "\\"}@{verbatim "<forall>"}) then @{verbatim
28238
398bf960d3d4 misc tuning and modernization;
wenzelm
parents: 28225
diff changeset
   509
  "isabellesym.sty"} should be included as well.  This package
398bf960d3d4 misc tuning and modernization;
wenzelm
parents: 28225
diff changeset
   510
  contains a standard set of {\LaTeX} macro definitions @{verbatim
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   511
  "\\isasym"}@{text foo} corresponding to @{verbatim "\\"}@{verbatim
28838
d5db6dfcb34a moved table of standard Isabelle symbols to isar-ref manual;
wenzelm
parents: 28504
diff changeset
   512
  "<"}@{text foo}@{verbatim ">"}, see \cite{isabelle-implementation} for a
d5db6dfcb34a moved table of standard Isabelle symbols to isar-ref manual;
wenzelm
parents: 28504
diff changeset
   513
  complete list of predefined Isabelle symbols.  Users may invent
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   514
  further symbols as well, just by providing {\LaTeX} macros in a
40800
330eb65c9469 Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents: 40387
diff changeset
   515
  similar fashion as in @{file "~~/lib/texinputs/isabellesym.sty"} of
28238
398bf960d3d4 misc tuning and modernization;
wenzelm
parents: 28225
diff changeset
   516
  the distribution.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   517
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   518
  For proper setup of DVI and PDF documents (with hyperlinks and
40800
330eb65c9469 Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents: 40387
diff changeset
   519
  bookmarks), we recommend to include @{file
28238
398bf960d3d4 misc tuning and modernization;
wenzelm
parents: 28225
diff changeset
   520
  "~~/lib/texinputs/pdfsetup.sty"} as well.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   521
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   522
  \medskip As a final step of document preparation within Isabelle,
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   523
  @{tool document}~@{verbatim "-c"} is run on the resulting @{verbatim
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   524
  document} directory.  Thus the actual output document is built and
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   525
  installed in its proper place (as linked by the session's @{verbatim
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   526
  index.html} if option @{verbatim "-i"} of @{tool_ref usedir} has
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   527
  been enabled, cf.\ \secref{sec:info}).  The generated sources are
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   528
  deleted after successful run of {\LaTeX} and friends.  Note that a
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   529
  separate copy of the sources may be retained by passing an option
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   530
  @{verbatim "-D"} to @{tool usedir} when running the session.  *}
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   531
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   532
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   533
section {* Running {\LaTeX} within the Isabelle environment
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   534
  \label{sec:tool-latex} *}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   535
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   536
text {* The @{tool_def latex} tool provides the basic interface for
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   537
  Isabelle document preparation.  Its usage is:
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   538
\begin{ttbox}
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   539
Usage: isabelle latex [OPTIONS] [FILE]
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   540
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   541
  Options are:
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   542
    -o FORMAT    specify output format: dvi (default), dvi.gz, ps,
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   543
                 ps.gz, pdf, bbl, idx, sty, syms
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   544
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   545
  Run LaTeX (and related tools) on FILE (default root.tex),
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   546
  producing the specified output format.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   547
\end{ttbox}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   548
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   549
  Appropriate {\LaTeX}-related programs are run on the input file,
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   550
  according to the given output format: @{executable latex},
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   551
  @{executable pdflatex}, @{executable dvips}, @{executable bibtex}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   552
  (for @{verbatim bbl}), and @{executable makeindex} (for @{verbatim
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   553
  idx}).  The actual commands are determined from the settings
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   554
  environment (@{setting ISABELLE_LATEX} etc.).
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   555
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   556
  The @{verbatim sty} output format causes the Isabelle style files to
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   557
  be updated from the distribution.  This is useful in special
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   558
  situations where the document sources are to be processed another
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   559
  time by separate tools (cf.\ option @{verbatim "-D"} of @{tool
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   560
  usedir}).
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   561
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   562
  The @{verbatim syms} output is for internal use; it generates lists
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   563
  of symbols that are available without loading additional {\LaTeX}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   564
  packages.
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   565
*}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   566
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   567
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   568
subsubsection {* Examples *}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   569
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   570
text {* Invoking @{tool latex} by hand may be occasionally useful when
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   571
  debugging failed attempts of the automatic document preparation
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   572
  stage of batch-mode Isabelle.  The abortive process leaves the
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   573
  sources at a certain place within @{setting ISABELLE_BROWSER_INFO},
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   574
  see the runtime error message for details.  This enables users to
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   575
  inspect {\LaTeX} runs in further detail, e.g.\ like this:
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   576
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   577
\begin{ttbox}
40387
e4c9e0dad473 moved ISABELLE_IDENTIFIER from ISABELLE_OUTPUT further up to ISABELLE_HOME_USER;
wenzelm
parents: 35587
diff changeset
   578
  cd ~/.isabelle/IsabelleXXXX/browser_info/HOL/Test/document
28504
7ad7d7d6df47 simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents: 28238
diff changeset
   579
  isabelle latex -o pdf
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   580
\end{ttbox}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   581
*}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   582
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   583
end