wenzelm@28221: theory Presentation wenzelm@43564: imports Base wenzelm@28221: begin wenzelm@28221: wenzelm@28221: chapter {* Presenting theories \label{ch:present} *} wenzelm@28221: wenzelm@48814: text {* Isabelle provides several ways to present the outcome of wenzelm@48814: formal developments, including WWW-based browsable libraries or wenzelm@48814: actual printable documents. Presentation is centered around the wenzelm@48814: concept of \emph{sessions} (\chref{ch:session}). The global session wenzelm@48814: structure is that of a tree, with Isabelle Pure at its root, further wenzelm@48814: object-logics derived (e.g.\ HOLCF from HOL, and HOL from Pure), and wenzelm@48814: application sessions further on in the hierarchy. wenzelm@28221: wenzelm@48814: The tools @{tool_ref mkroot} and @{tool_ref build} provide the wenzelm@48814: primary means for managing Isabelle sessions, including proper setup wenzelm@48814: for presentation; @{tool build} takes care to have @{executable_ref wenzelm@48814: "isabelle-process"} run any additional stages required for document wenzelm@48814: preparation, notably the @{tool_ref document} and @{tool_ref latex}. wenzelm@48814: The complete tool chain for managing batch-mode Isabelle sessions is wenzelm@48814: illustrated in \figref{fig:session-tools}. wenzelm@28221: wenzelm@28221: \begin{figure}[htbp] wenzelm@28221: \begin{center} wenzelm@28221: \begin{tabular}{lp{0.6\textwidth}} wenzelm@28221: wenzelm@48814: @{tool_ref mkroot} & invoked once by the user to initialize the wenzelm@48814: session @{verbatim ROOT} with optional @{verbatim document} wenzelm@48814: directory; \\ wenzelm@28221: wenzelm@48814: @{tool_ref build} & invoked repeatedly by the user to keep wenzelm@48814: session output up-to-date (HTML, documents etc.); \\ wenzelm@28221: wenzelm@48602: @{executable "isabelle-process"} & run through @{tool_ref wenzelm@48814: build}; \\ wenzelm@28221: wenzelm@48602: @{tool_ref document} & run by the Isabelle process if document wenzelm@48602: preparation is enabled; \\ wenzelm@28221: wenzelm@48602: @{tool_ref latex} & universal {\LaTeX} tool wrapper invoked wenzelm@48602: multiple times by @{tool_ref document}; also useful for manual wenzelm@48602: experiments; \\ wenzelm@28221: wenzelm@28221: \end{tabular} wenzelm@28221: \caption{The tool chain of Isabelle session presentation} \label{fig:session-tools} wenzelm@28221: \end{center} wenzelm@28221: \end{figure} wenzelm@28221: *} wenzelm@28221: wenzelm@28221: wenzelm@28221: section {* Generating theory browser information \label{sec:info} *} wenzelm@28221: wenzelm@28221: text {* wenzelm@28221: \index{theory browsing information|bold} wenzelm@28221: wenzelm@48814: As a side-effect of building sessions, Isabelle is able to generate wenzelm@48814: theory browsing information, including HTML documents that show the wenzelm@48814: theory sources and the relationship with its ancestors and wenzelm@48814: descendants. Besides the HTML file that is generated for every wenzelm@48814: theory, Isabelle stores links to all theories in an index wenzelm@48814: file. These indexes are linked with other indexes to represent the wenzelm@48814: overall tree structure of the sessions. wenzelm@28221: wenzelm@28221: Isabelle also generates graph files that represent the theory wenzelm@48814: dependencies within a session. There is a graph browser Java applet wenzelm@48814: embedded in the generated HTML pages, and also a stand-alone wenzelm@48814: application that allows browsing theory graphs without having to wenzelm@48814: start a WWW client first. The latter version also includes features wenzelm@48814: such as generating Postscript files, which are not available in the wenzelm@48814: applet version. See \secref{sec:browse} for further information. wenzelm@28221: wenzelm@28221: \medskip wenzelm@28221: wenzelm@28221: The easiest way to let Isabelle generate theory browsing information wenzelm@48814: for existing sessions is to invoke @{tool build} with suitable wenzelm@48814: options: wenzelm@28221: wenzelm@28221: \begin{ttbox} wenzelm@48814: isabelle build -o browser_info -v -c FOL wenzelm@28221: \end{ttbox} wenzelm@28221: wenzelm@48814: The presentation output will appear in @{verbatim wenzelm@48814: "$ISABELLE_BROWSER_INFO/FOL"} as reported by the above verbose wenzelm@48814: invocation of the build process. wenzelm@28221: wenzelm@48814: Many Isabelle sessions (such as @{verbatim "HOL-Library"} in @{file wenzelm@48814: "~~/src/HOL/Library"}) also provide actual printable documents. wenzelm@48814: These are prepared automatically as well if enabled like this: wenzelm@28221: \begin{ttbox} wenzelm@48814: isabelle build -o browser_info -o document=pdf -v -c HOL-Library wenzelm@28221: \end{ttbox} wenzelm@48814: wenzelm@48814: Enabling both browser info and document preparation simultaneously wenzelm@48814: causes an appropriate ``document'' link to be included in the HTML wenzelm@48814: index. Documents may be generated independently of browser wenzelm@48814: information as well, see \secref{sec:tool-document} for further wenzelm@48814: details. wenzelm@28221: wenzelm@28221: \bigskip The theory browsing information is stored in a wenzelm@28221: sub-directory directory determined by the @{setting_ref wenzelm@28221: ISABELLE_BROWSER_INFO} setting plus a prefix corresponding to the wenzelm@28221: session identifier (according to the tree structure of sub-sessions wenzelm@48814: by default). In order to present Isabelle applications on the web, wenzelm@48814: the corresponding subdirectory from @{setting ISABELLE_BROWSER_INFO} wenzelm@48814: can be put on a WWW server. wenzelm@28221: *} wenzelm@28221: wenzelm@48814: section {* Preparing session root directories \label{sec:tool-mkroot} *} wenzelm@28221: wenzelm@48814: text {* The @{tool_def mkroot} tool configures a given directory as wenzelm@48814: session root, with some @{verbatim ROOT} file and optional document wenzelm@48814: source directory. Its usage is: wenzelm@28221: \begin{ttbox} wenzelm@48814: Usage: isabelle mkroot [OPTIONS] [DIR] wenzelm@28221: wenzelm@28221: Options are: wenzelm@48814: -d enable document preparation wenzelm@48814: -n NAME alternative session name (default: DIR base name) wenzelm@28221: wenzelm@48814: Prepare session root DIR (default: current directory). wenzelm@28221: \end{ttbox} wenzelm@28221: wenzelm@48814: The results are placed in the given directory @{text dir}, which wenzelm@48814: refers to the current directory by default. The @{tool mkroot} tool wenzelm@48814: is conservative in the sense that it does not overwrite existing wenzelm@48814: files or directories. Earlier attempts to generate a session root wenzelm@48814: need to be deleted manually. wenzelm@28221: wenzelm@48814: \medskip Option @{verbatim "-d"} indicates that the session shall be wenzelm@48814: accompanied by a formal document, with @{text dir}@{verbatim wenzelm@48814: "/document/root.tex"} as its {\LaTeX} entry point (see also wenzelm@48814: \chref{ch:present}). wenzelm@28221: wenzelm@48814: Option @{verbatim "-n"} allows to specify an alternative session wenzelm@48814: name; otherwise the base name of the given directory is used. wenzelm@28221: wenzelm@48814: \medskip The implicit Isabelle settings variable @{setting wenzelm@48814: ISABELLE_LOGIC} specifies the parent session, and @{setting wenzelm@48814: ISABELLE_DOCUMENT_FORMAT} the document format to be filled filled wenzelm@48814: into the generated @{verbatim ROOT} file. *} wenzelm@28221: wenzelm@28221: wenzelm@28221: subsubsection {* Examples *} wenzelm@28221: wenzelm@48814: text {* Produce session @{verbatim Test} (with document preparation) wenzelm@48814: within a separate directory of the same name: wenzelm@28221: \begin{ttbox} wenzelm@48814: isabelle mkroot -d Test && isabelle build -D Test wenzelm@28221: \end{ttbox} wenzelm@28221: wenzelm@48814: \medskip Upgrade the current directory into a session ROOT with wenzelm@48814: document preparation, and build it: wenzelm@48814: \begin{ttbox} wenzelm@48814: isabelle mkroot -d && isabelle build -D . wenzelm@48814: \end{ttbox} wenzelm@28221: *} wenzelm@28221: wenzelm@28221: wenzelm@28221: section {* Preparing Isabelle session documents wenzelm@28221: \label{sec:tool-document} *} wenzelm@28221: wenzelm@48602: text {* The @{tool_def document} tool prepares logic session wenzelm@48602: documents, processing the sources both as provided by the user and wenzelm@48602: generated by Isabelle. Its usage is: wenzelm@28221: \begin{ttbox} wenzelm@48602: Usage: isabelle document [OPTIONS] [DIR] wenzelm@28221: wenzelm@28221: Options are: wenzelm@28221: -c cleanup -- be aggressive in removing old stuff wenzelm@28221: -n NAME specify document name (default 'document') wenzelm@28221: -o FORMAT specify output format: dvi (default), dvi.gz, ps, wenzelm@28221: ps.gz, pdf wenzelm@28221: -t TAGS specify tagged region markup wenzelm@28221: wenzelm@28221: Prepare the theory session document in DIR (default 'document') wenzelm@28221: producing the specified output format. wenzelm@28221: \end{ttbox} wenzelm@48814: This tool is usually run automatically as part of the Isabelle build wenzelm@48814: process, provided document preparation has been enabled via suitable wenzelm@48814: options. It may be manually invoked on the generated browser wenzelm@48814: information document output as well, e.g.\ in case of errors wenzelm@48814: encountered in the batch run. wenzelm@28221: wenzelm@48602: \medskip The @{verbatim "-c"} option tells @{tool document} to wenzelm@48814: dispose the document sources after successful operation! This is wenzelm@28221: the right thing to do for sources generated by an Isabelle process, wenzelm@28221: but take care of your files in manual document preparation! wenzelm@28221: wenzelm@28221: \medskip The @{verbatim "-n"} and @{verbatim "-o"} option specify wenzelm@28221: the final output file name and format, the default is ``@{verbatim wenzelm@28221: document.dvi}''. Note that the result will appear in the parent of wenzelm@28221: the target @{verbatim DIR}. wenzelm@28221: wenzelm@28221: \medskip The @{verbatim "-t"} option tells {\LaTeX} how to interpret wenzelm@28221: tagged Isabelle command regions. Tags are specified as a comma wenzelm@28221: separated list of modifier/name pairs: ``@{verbatim "+"}@{text wenzelm@28221: foo}'' (or just ``@{text foo}'') means to keep, ``@{verbatim wenzelm@28221: "-"}@{text foo}'' to drop, and ``@{verbatim "/"}@{text foo}'' to wenzelm@28221: fold text tagged as @{text foo}. The builtin default is equivalent wenzelm@28221: to the tag specification ``@{verbatim wenzelm@30113: "+theory,+proof,+ML,+visible,-invisible"}''; see also the {\LaTeX} wenzelm@28221: macros @{verbatim "\\isakeeptag"}, @{verbatim "\\isadroptag"}, and wenzelm@40800: @{verbatim "\\isafoldtag"}, in @{file wenzelm@28238: "~~/lib/texinputs/isabelle.sty"}. wenzelm@28221: wenzelm@48814: \medskip Document preparation requires a @{verbatim document} wenzelm@48814: directory within the session sources. This directory is supposed to wenzelm@48814: contain all the files needed to produce the final document --- apart wenzelm@48814: from the actual theories which are generated by Isabelle. wenzelm@28221: wenzelm@48602: \medskip For most practical purposes, @{tool document} is smart wenzelm@48602: enough to create any of the specified output formats, taking wenzelm@28221: @{verbatim root.tex} supplied by the user as a starting point. This wenzelm@28221: even includes multiple runs of {\LaTeX} to accommodate references wenzelm@28221: and bibliographies (the latter assumes @{verbatim root.bib} within wenzelm@28221: the same directory). wenzelm@28221: wenzelm@48657: In more complex situations, a separate @{verbatim build} script for wenzelm@48657: the document sources may be given. It is invoked with command-line wenzelm@48657: arguments for the document format and the document variant name. wenzelm@48657: The script needs to produce corresponding output files, e.g.\ wenzelm@48657: @{verbatim root.pdf} for target format @{verbatim pdf} (and default wenzelm@48657: default variants). The main work can be again delegated to @{tool wenzelm@48814: latex}, but it is also possible to harvest generated {\LaTeX} wenzelm@48814: sources and copy them elsewhere, for example. wenzelm@28221: wenzelm@42009: \medskip When running the session, Isabelle copies the content of wenzelm@42009: the original @{verbatim document} directory into its proper place wenzelm@42009: within @{setting ISABELLE_BROWSER_INFO}, according to the session wenzelm@42009: path and document variant. Then, for any processed theory @{text A} wenzelm@42009: some {\LaTeX} source is generated and put there as @{text wenzelm@42009: A}@{verbatim ".tex"}. Furthermore, a list of all generated theory wenzelm@42009: files is put into @{verbatim session.tex}. Typically, the root wenzelm@42009: {\LaTeX} file provided by the user would include @{verbatim wenzelm@42009: session.tex} to get a document containing all the theories. wenzelm@28221: wenzelm@28221: The {\LaTeX} versions of the theories require some macros defined in wenzelm@40800: @{file "~~/lib/texinputs/isabelle.sty"}. Doing @{verbatim wenzelm@28238: "\\usepackage{isabelle}"} in @{verbatim root.tex} should be fine; wenzelm@48602: the underlying @{tool latex} already includes an appropriate path wenzelm@48602: specification for {\TeX} inputs. wenzelm@28221: wenzelm@28221: If the text contains any references to Isabelle symbols (such as wenzelm@28221: @{verbatim "\\"}@{verbatim ""}) then @{verbatim wenzelm@28238: "isabellesym.sty"} should be included as well. This package wenzelm@28238: contains a standard set of {\LaTeX} macro definitions @{verbatim wenzelm@28221: "\\isasym"}@{text foo} corresponding to @{verbatim "\\"}@{verbatim wenzelm@28838: "<"}@{text foo}@{verbatim ">"}, see \cite{isabelle-implementation} for a wenzelm@28838: complete list of predefined Isabelle symbols. Users may invent wenzelm@28221: further symbols as well, just by providing {\LaTeX} macros in a wenzelm@40800: similar fashion as in @{file "~~/lib/texinputs/isabellesym.sty"} of wenzelm@28238: the distribution. wenzelm@28221: wenzelm@28221: For proper setup of DVI and PDF documents (with hyperlinks and wenzelm@40800: bookmarks), we recommend to include @{file wenzelm@28238: "~~/lib/texinputs/pdfsetup.sty"} as well. wenzelm@28221: wenzelm@48814: \medskip As a final step of Isabelle document preparation, @{tool wenzelm@48814: document}~@{verbatim "-c"} is run on the resulting copy of the wenzelm@48814: @{verbatim document} directory. Thus the actual output document is wenzelm@48814: built and installed in its proper place. The generated sources are wenzelm@48814: deleted after successful run of {\LaTeX} and friends. wenzelm@48814: wenzelm@48814: Some care is needed if the document output location is configured wenzelm@48814: differently, say within a directory whose content is still required wenzelm@48814: afterwards! wenzelm@48814: *} wenzelm@28221: wenzelm@28221: wenzelm@28221: section {* Running {\LaTeX} within the Isabelle environment wenzelm@28221: \label{sec:tool-latex} *} wenzelm@28221: wenzelm@48602: text {* The @{tool_def latex} tool provides the basic interface for wenzelm@28221: Isabelle document preparation. Its usage is: wenzelm@28221: \begin{ttbox} wenzelm@48602: Usage: isabelle latex [OPTIONS] [FILE] wenzelm@28221: wenzelm@28221: Options are: wenzelm@28221: -o FORMAT specify output format: dvi (default), dvi.gz, ps, wenzelm@28221: ps.gz, pdf, bbl, idx, sty, syms wenzelm@28221: wenzelm@28221: Run LaTeX (and related tools) on FILE (default root.tex), wenzelm@28221: producing the specified output format. wenzelm@28221: \end{ttbox} wenzelm@28221: wenzelm@28221: Appropriate {\LaTeX}-related programs are run on the input file, wenzelm@28221: according to the given output format: @{executable latex}, wenzelm@28221: @{executable pdflatex}, @{executable dvips}, @{executable bibtex} wenzelm@28221: (for @{verbatim bbl}), and @{executable makeindex} (for @{verbatim wenzelm@28221: idx}). The actual commands are determined from the settings wenzelm@28221: environment (@{setting ISABELLE_LATEX} etc.). wenzelm@28221: wenzelm@28221: The @{verbatim sty} output format causes the Isabelle style files to wenzelm@28221: be updated from the distribution. This is useful in special wenzelm@28221: situations where the document sources are to be processed another wenzelm@48814: time by separate tools. wenzelm@28221: wenzelm@28221: The @{verbatim syms} output is for internal use; it generates lists wenzelm@28221: of symbols that are available without loading additional {\LaTeX} wenzelm@28221: packages. wenzelm@28221: *} wenzelm@28221: wenzelm@28221: wenzelm@28221: subsubsection {* Examples *} wenzelm@28221: wenzelm@48602: text {* Invoking @{tool latex} by hand may be occasionally useful when wenzelm@48602: debugging failed attempts of the automatic document preparation wenzelm@48602: stage of batch-mode Isabelle. The abortive process leaves the wenzelm@48602: sources at a certain place within @{setting ISABELLE_BROWSER_INFO}, wenzelm@48602: see the runtime error message for details. This enables users to wenzelm@48602: inspect {\LaTeX} runs in further detail, e.g.\ like this: wenzelm@48602: wenzelm@28221: \begin{ttbox} wenzelm@40387: cd ~/.isabelle/IsabelleXXXX/browser_info/HOL/Test/document wenzelm@28504: isabelle latex -o pdf wenzelm@28221: \end{ttbox} wenzelm@28221: *} wenzelm@28221: wenzelm@28221: end