src/Doc/System/Presentation.thy
author wenzelm
Wed, 04 Nov 2015 20:18:46 +0100
changeset 61575 f18f6e51e901
parent 61504 a7ae3ef886a9
child 61656 cfabbc083977
permissions -rw-r--r--
tuned whitespace;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
     1
(*:wrap=hard:maxLineLen=78:*)
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
     2
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
     3
theory Presentation
43564
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 42009
diff changeset
     4
imports Base
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
     5
begin
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
     6
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
     7
chapter \<open>Presenting theories \label{ch:present}\<close>
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
     8
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
     9
text \<open>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    10
  Isabelle provides several ways to present the outcome of formal
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    11
  developments, including WWW-based browsable libraries or actual printable
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    12
  documents. Presentation is centered around the concept of \<^emph>\<open>sessions\<close>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    13
  (\chref{ch:session}). The global session structure is that of a tree, with
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    14
  Isabelle Pure at its root, further object-logics derived (e.g.\ HOLCF from
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    15
  HOL, and HOL from Pure), and application sessions further on in the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    16
  hierarchy.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    17
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    18
  The tools @{tool_ref mkroot} and @{tool_ref build} provide the primary means
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    19
  for managing Isabelle sessions, including proper setup for presentation;
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    20
  @{tool build} takes care to have @{executable_ref "isabelle_process"} run
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    21
  any additional stages required for document preparation, notably the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    22
  @{tool_ref document} and @{tool_ref latex}. The complete tool chain for
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    23
  managing batch-mode Isabelle sessions is illustrated in
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    24
  \figref{fig:session-tools}.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    25
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    26
  \begin{figure}[htbp]
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    27
  \begin{center}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    28
  \begin{tabular}{lp{0.6\textwidth}}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    29
48814
d488a5f25bf6 some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents: 48722
diff changeset
    30
      @{tool_ref mkroot} & invoked once by the user to initialize the
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61493
diff changeset
    31
      session \<^verbatim>\<open>ROOT\<close> with optional \<^verbatim>\<open>document\<close>
48814
d488a5f25bf6 some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents: 48722
diff changeset
    32
      directory; \\
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    33
48814
d488a5f25bf6 some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents: 48722
diff changeset
    34
      @{tool_ref build} & invoked repeatedly by the user to keep
d488a5f25bf6 some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents: 48722
diff changeset
    35
      session output up-to-date (HTML, documents etc.); \\
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    36
56439
95e2656b3b23 renamed "isabelle-process" to "isabelle_process", with shell function to avoid dynamic path lookups;
wenzelm
parents: 54936
diff changeset
    37
      @{executable "isabelle_process"} & run through @{tool_ref
48814
d488a5f25bf6 some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents: 48722
diff changeset
    38
      build}; \\
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    39
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    40
      @{tool_ref document} & run by the Isabelle process if document
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    41
      preparation is enabled; \\
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    42
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    43
      @{tool_ref latex} & universal {\LaTeX} tool wrapper invoked
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    44
      multiple times by @{tool_ref document}; also useful for manual
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
    45
      experiments; \\
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    46
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    47
  \end{tabular}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    48
  \caption{The tool chain of Isabelle session presentation} \label{fig:session-tools}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    49
  \end{center}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    50
  \end{figure}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
    51
\<close>
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    52
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    53
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
    54
section \<open>Generating theory browser information \label{sec:info}\<close>
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    55
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
    56
text \<open>
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    57
  \index{theory browsing information|bold}
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    58
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    59
  As a side-effect of building sessions, Isabelle is able to generate theory
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    60
  browsing information, including HTML documents that show the theory sources
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    61
  and the relationship with its ancestors and descendants. Besides the HTML
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    62
  file that is generated for every theory, Isabelle stores links to all
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    63
  theories of a session in an index file. As a second hierarchy, groups of
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    64
  sessions are organized as \<^emph>\<open>chapters\<close>, with a separate index. Note that the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    65
  implicit tree structure of the session build hierarchy is \<^emph>\<open>not\<close> relevant
51417
d266f9329368 sessions may be organized via 'chapter' in ROOT;
wenzelm
parents: 51054
diff changeset
    66
  for the presentation.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    67
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    68
  Isabelle also generates graph files that represent the theory dependencies
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    69
  within a session. There is a graph browser Java applet embedded in the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    70
  generated HTML pages, and also a stand-alone application that allows
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    71
  browsing theory graphs without having to start a WWW client first. The
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    72
  latter version also includes features such as generating Postscript files,
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    73
  which are not available in the applet version. See \secref{sec:browse} for
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    74
  further information.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    75
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 58618
diff changeset
    76
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    77
  The easiest way to let Isabelle generate theory browsing information for
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    78
  existing sessions is to invoke @{tool build} with suitable options:
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
    79
  @{verbatim [display] \<open>isabelle build -o browser_info -v -c FOL\<close>}
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    80
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    81
  The presentation output will appear in \<^verbatim>\<open>$ISABELLE_BROWSER_INFO/FOL/FOL\<close> as
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    82
  reported by the above verbose invocation of the build process.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    83
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61493
diff changeset
    84
  Many Isabelle sessions (such as \<^verbatim>\<open>HOL-Library\<close> in @{file
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    85
  "~~/src/HOL/Library"}) also provide actual printable documents. These are
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    86
  prepared automatically as well if enabled like this:
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
    87
  @{verbatim [display] \<open>isabelle build -o browser_info -o document=pdf -v -c HOL-Library\<close>}
48814
d488a5f25bf6 some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents: 48722
diff changeset
    88
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    89
  Enabling both browser info and document preparation simultaneously causes an
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    90
  appropriate ``document'' link to be included in the HTML index. Documents
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    91
  may be generated independently of browser information as well, see
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    92
  \secref{sec:tool-document} for further details.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
    93
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 58618
diff changeset
    94
  \<^bigskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    95
  The theory browsing information is stored in a sub-directory directory
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    96
  determined by the @{setting_ref ISABELLE_BROWSER_INFO} setting plus a prefix
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    97
  corresponding to the session chapter and identifier. In order to present
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    98
  Isabelle applications on the web, the corresponding subdirectory from
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
    99
  @{setting ISABELLE_BROWSER_INFO} can be put on a WWW server.
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   100
\<close>
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   101
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 58618
diff changeset
   102
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
   103
section \<open>Preparing session root directories \label{sec:tool-mkroot}\<close>
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   104
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   105
text \<open>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   106
  The @{tool_def mkroot} tool configures a given directory as session root,
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   107
  with some \<^verbatim>\<open>ROOT\<close> file and optional document source directory. Its usage is:
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   108
  @{verbatim [display]
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   109
\<open>Usage: isabelle mkroot [OPTIONS] [DIR]
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   110
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   111
  Options are:
48814
d488a5f25bf6 some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents: 48722
diff changeset
   112
    -d           enable document preparation
d488a5f25bf6 some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents: 48722
diff changeset
   113
    -n NAME      alternative session name (default: DIR base name)
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   114
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   115
  Prepare session root DIR (default: current directory).\<close>}
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   116
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   117
  The results are placed in the given directory \<open>dir\<close>, which refers to the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   118
  current directory by default. The @{tool mkroot} tool is conservative in the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   119
  sense that it does not overwrite existing files or directories. Earlier
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   120
  attempts to generate a session root need to be deleted manually.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   121
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 58618
diff changeset
   122
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   123
  Option \<^verbatim>\<open>-d\<close> indicates that the session shall be accompanied by a formal
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   124
  document, with \<open>DIR\<close>\<^verbatim>\<open>/document/root.tex\<close> as its {\LaTeX} entry point (see
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   125
  also \chref{ch:present}).
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   126
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   127
  Option \<^verbatim>\<open>-n\<close> allows to specify an alternative session name; otherwise the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   128
  base name of the given directory is used.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   129
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 58618
diff changeset
   130
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   131
  The implicit Isabelle settings variable @{setting ISABELLE_LOGIC} specifies
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   132
  the parent session, and @{setting ISABELLE_DOCUMENT_FORMAT} the document
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   133
  format to be filled filled into the generated \<^verbatim>\<open>ROOT\<close> file.
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61493
diff changeset
   134
\<close>
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   135
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   136
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
   137
subsubsection \<open>Examples\<close>
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   138
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   139
text \<open>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   140
  Produce session \<^verbatim>\<open>Test\<close> (with document preparation) within a separate
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   141
  directory of the same name:
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   142
  @{verbatim [display] \<open>isabelle mkroot -d Test && isabelle build -D Test\<close>}
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   143
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 58618
diff changeset
   144
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   145
  Upgrade the current directory into a session ROOT with document preparation,
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   146
  and build it:
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   147
  @{verbatim [display] \<open>isabelle mkroot -d && isabelle build -D .\<close>}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
   148
\<close>
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   149
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   150
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   151
section \<open>Preparing Isabelle session documents \label{sec:tool-document}\<close>
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   152
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   153
text \<open>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   154
  The @{tool_def document} tool prepares logic session documents, processing
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   155
  the sources as provided by the user and generated by Isabelle. Its usage is:
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   156
  @{verbatim [display]
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   157
\<open>Usage: isabelle document [OPTIONS] [DIR]
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   158
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   159
  Options are:
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   160
    -c           cleanup -- be aggressive in removing old stuff
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   161
    -n NAME      specify document name (default 'document')
52746
eec610972763 discontinued historic document formats;
wenzelm
parents: 52744
diff changeset
   162
    -o FORMAT    specify output format: pdf (default), dvi
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   163
    -t TAGS      specify tagged region markup
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   164
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   165
  Prepare the theory session document in DIR (default 'document')
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   166
  producing the specified output format.\<close>}
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   167
48814
d488a5f25bf6 some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents: 48722
diff changeset
   168
  This tool is usually run automatically as part of the Isabelle build
d488a5f25bf6 some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents: 48722
diff changeset
   169
  process, provided document preparation has been enabled via suitable
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   170
  options. It may be manually invoked on the generated browser information
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   171
  document output as well, e.g.\ in case of errors encountered in the batch
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   172
  run.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   173
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 58618
diff changeset
   174
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   175
  The \<^verbatim>\<open>-c\<close> option tells @{tool document} to dispose the document sources
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   176
  after successful operation! This is the right thing to do for sources
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   177
  generated by an Isabelle process, but take care of your files in manual
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   178
  document preparation!
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   179
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 58618
diff changeset
   180
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   181
  The \<^verbatim>\<open>-n\<close> and \<^verbatim>\<open>-o\<close> option specify the final output file name and format,
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   182
  the default is ``\<^verbatim>\<open>document.dvi\<close>''. Note that the result will appear in the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   183
  parent of the target \<^verbatim>\<open>DIR\<close>.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   184
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 58618
diff changeset
   185
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   186
  The \<^verbatim>\<open>-t\<close> option tells {\LaTeX} how to interpret tagged Isabelle command
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   187
  regions. Tags are specified as a comma separated list of modifier/name
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   188
  pairs: ``\<^verbatim>\<open>+\<close>\<open>foo\<close>'' (or just ``\<open>foo\<close>'') means to keep, ``\<^verbatim>\<open>-\<close>\<open>foo\<close>'' to
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   189
  drop, and ``\<^verbatim>\<open>/\<close>\<open>foo\<close>'' to fold text tagged as \<open>foo\<close>. The builtin default is
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   190
  equivalent to the tag specification
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   191
  ``\<^verbatim>\<open>+theory,+proof,+ML,+visible,-invisible\<close>''; see also the {\LaTeX} macros
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   192
  \<^verbatim>\<open>\isakeeptag\<close>, \<^verbatim>\<open>\isadroptag\<close>, and \<^verbatim>\<open>\isafoldtag\<close>, in @{file
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   193
  "~~/lib/texinputs/isabelle.sty"}.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   194
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 58618
diff changeset
   195
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   196
  Document preparation requires a \<^verbatim>\<open>document\<close> directory within the session
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   197
  sources. This directory is supposed to contain all the files needed to
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   198
  produce the final document --- apart from the actual theories which are
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   199
  generated by Isabelle.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   200
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 58618
diff changeset
   201
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   202
  For most practical purposes, @{tool document} is smart enough to create any
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   203
  of the specified output formats, taking \<^verbatim>\<open>root.tex\<close> supplied by the user as
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   204
  a starting point. This even includes multiple runs of {\LaTeX} to
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   205
  accommodate references and bibliographies (the latter assumes \<^verbatim>\<open>root.bib\<close>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   206
  within the same directory).
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   207
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   208
  In more complex situations, a separate \<^verbatim>\<open>build\<close> script for the document
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   209
  sources may be given. It is invoked with command-line arguments for the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   210
  document format and the document variant name. The script needs to produce
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   211
  corresponding output files, e.g.\ \<^verbatim>\<open>root.pdf\<close> for target format \<^verbatim>\<open>pdf\<close> (and
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   212
  default variants). The main work can be again delegated to @{tool latex},
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   213
  but it is also possible to harvest generated {\LaTeX} sources and copy them
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   214
  elsewhere.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   215
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 58618
diff changeset
   216
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   217
  When running the session, Isabelle copies the content of the original
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   218
  \<^verbatim>\<open>document\<close> directory into its proper place within @{setting
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   219
  ISABELLE_BROWSER_INFO}, according to the session path and document variant.
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   220
  Then, for any processed theory \<open>A\<close> some {\LaTeX} source is generated and put
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   221
  there as \<open>A\<close>\<^verbatim>\<open>.tex\<close>. Furthermore, a list of all generated theory files is
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   222
  put into \<^verbatim>\<open>session.tex\<close>. Typically, the root {\LaTeX} file provided by the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   223
  user would include \<^verbatim>\<open>session.tex\<close> to get a document containing all the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   224
  theories.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   225
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   226
  The {\LaTeX} versions of the theories require some macros defined in @{file
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   227
  "~~/lib/texinputs/isabelle.sty"}. Doing \<^verbatim>\<open>\usepackage{isabelle}\<close> in
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   228
  \<^verbatim>\<open>root.tex\<close> should be fine; the underlying @{tool latex} already includes an
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   229
  appropriate path specification for {\TeX} inputs.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   230
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   231
  If the text contains any references to Isabelle symbols (such as \<^verbatim>\<open>\<forall>\<close>) then
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   232
  \<^verbatim>\<open>isabellesym.sty\<close> should be included as well. This package contains a
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   233
  standard set of {\LaTeX} macro definitions \<^verbatim>\<open>\isasym\<close>\<open>foo\<close> corresponding to
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   234
  \<^verbatim>\<open>\\<close>\<^verbatim>\<open><\<close>\<open>foo\<close>\<^verbatim>\<open>>\<close>, see @{cite "isabelle-implementation"} for a complete list
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   235
  of predefined Isabelle symbols. Users may invent further symbols as well,
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   236
  just by providing {\LaTeX} macros in a similar fashion as in @{file
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   237
  "~~/lib/texinputs/isabellesym.sty"} of the Isabelle distribution.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   238
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   239
  For proper setup of DVI and PDF documents (with hyperlinks and bookmarks),
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   240
  we recommend to include @{file "~~/lib/texinputs/pdfsetup.sty"} as well.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   241
61406
eb2463fc4d7b more symbols;
wenzelm
parents: 58618
diff changeset
   242
  \<^medskip>
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   243
  As a final step of Isabelle document preparation, @{tool document}~\<^verbatim>\<open>-c\<close> is
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   244
  run on the resulting copy of the \<^verbatim>\<open>document\<close> directory. Thus the actual
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   245
  output document is built and installed in its proper place. The generated
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   246
  sources are deleted after successful run of {\LaTeX} and friends.
48814
d488a5f25bf6 some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents: 48722
diff changeset
   247
d488a5f25bf6 some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents: 48722
diff changeset
   248
  Some care is needed if the document output location is configured
d488a5f25bf6 some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents: 48722
diff changeset
   249
  differently, say within a directory whose content is still required
d488a5f25bf6 some updates of "Presenting theories", using mkroot/build instead of former mkdir/make/usedir (which are still present in "Misc");
wenzelm
parents: 48722
diff changeset
   250
  afterwards!
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
   251
\<close>
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   252
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   253
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
   254
section \<open>Running {\LaTeX} within the Isabelle environment
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
   255
  \label{sec:tool-latex}\<close>
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   256
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   257
text \<open>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   258
  The @{tool_def latex} tool provides the basic interface for Isabelle
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   259
  document preparation. Its usage is:
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   260
  @{verbatim [display]
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   261
\<open>Usage: isabelle latex [OPTIONS] [FILE]
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   262
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   263
  Options are:
52746
eec610972763 discontinued historic document formats;
wenzelm
parents: 52744
diff changeset
   264
    -o FORMAT    specify output format: pdf (default), dvi,
eec610972763 discontinued historic document formats;
wenzelm
parents: 52744
diff changeset
   265
                 bbl, idx, sty, syms
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   266
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   267
  Run LaTeX (and related tools) on FILE (default root.tex),
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   268
  producing the specified output format.\<close>}
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   269
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   270
  Appropriate {\LaTeX}-related programs are run on the input file, according
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   271
  to the given output format: @{executable latex}, @{executable pdflatex},
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   272
  @{executable dvips}, @{executable bibtex} (for \<^verbatim>\<open>bbl\<close>), and @{executable
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   273
  makeindex} (for \<^verbatim>\<open>idx\<close>). The actual commands are determined from the
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   274
  settings environment (@{setting ISABELLE_PDFLATEX} etc.).
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   275
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   276
  The \<^verbatim>\<open>sty\<close> output format causes the Isabelle style files to be updated from
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   277
  the distribution. This is useful in special situations where the document
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   278
  sources are to be processed another time by separate tools.
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   279
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   280
  The \<^verbatim>\<open>syms\<close> output is for internal use; it generates lists of symbols that
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   281
  are available without loading additional {\LaTeX} packages.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
   282
\<close>
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   283
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   284
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
   285
subsubsection \<open>Examples\<close>
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   286
61575
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   287
text \<open>
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   288
  Invoking @{tool latex} by hand may be occasionally useful when debugging
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   289
  failed attempts of the automatic document preparation stage of batch-mode
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   290
  Isabelle. The abortive process leaves the sources at a certain place within
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   291
  @{setting ISABELLE_BROWSER_INFO}, see the runtime error message for details.
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   292
  This enables users to inspect {\LaTeX} runs in further detail, e.g.\ like
f18f6e51e901 tuned whitespace;
wenzelm
parents: 61504
diff changeset
   293
  this:
48602
342ca8f3197b more uniform usage of "isabelle tool";
wenzelm
parents: 48590
diff changeset
   294
61407
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   295
  @{verbatim [display]
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   296
\<open>cd "$(isabelle getenv -b ISABELLE_BROWSER_INFO)/Unsorted/Test/document"
7ba7b8103565 @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61406
diff changeset
   297
isabelle latex -o pdf\<close>}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58553
diff changeset
   298
\<close>
28221
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   299
ca9fdab0f971 converted present.tex;
wenzelm
parents:
diff changeset
   300
end