diff --git a/src/Doc/System/Presentation.thy b/src/Doc/System/Presentation.thy --- a/src/Doc/System/Presentation.thy +++ b/src/Doc/System/Presentation.thy @@ -1,221 +1,221 @@ (*:maxLineLen=78:*) theory Presentation imports Base begin chapter \Presenting theories \label{ch:present}\ text \ Isabelle provides several ways to present the outcome of formal developments, including WWW-based browsable libraries or actual printable documents. Presentation is centered around the concept of \<^emph>\sessions\ (\chref{ch:session}). The global session structure is that of a tree, with Isabelle Pure at its root, further object-logics derived (e.g.\ HOLCF from HOL, and HOL from Pure), and application sessions further on in the hierarchy. The command-line tools @{tool_ref mkroot} and @{tool_ref build} provide the primary means for managing Isabelle sessions, including options for presentation: ``\<^verbatim>\document=pdf\'' generates PDF output from the theory session, and ``\<^verbatim>\document_output=dir\'' emits a copy of the document sources with the PDF into the given directory (relative to the session directory). Alternatively, @{tool_ref document} may be used to turn the generated {\LaTeX} sources of a session (exports from its build database) into PDF, using suitable invocations of @{tool_ref latex}. \ section \Generating HTML browser information \label{sec:info}\ text \ As a side-effect of building sessions, Isabelle is able to generate theory browsing information, including HTML documents that show the theory sources and the relationship with its ancestors and descendants. Besides the HTML file that is generated for every theory, Isabelle stores links to all theories of a session in an index file. As a second hierarchy, groups of sessions are organized as \<^emph>\chapters\, with a separate index. Note that the implicit tree structure of the session build hierarchy is \<^emph>\not\ relevant for the presentation. \<^medskip> To generate theory browsing information for an existing session, just invoke @{tool build} with suitable options: @{verbatim [display] \isabelle build -o browser_info -v -c FOL\} The presentation output will appear in \<^verbatim>\$ISABELLE_BROWSER_INFO/FOL/FOL\ as reported by the above verbose invocation of the build process. Many Isabelle sessions (such as \<^session>\HOL-Library\ in \<^dir>\~~/src/HOL/Library\) also provide printable documents in PDF. These are prepared automatically as well if enabled like this: @{verbatim [display] \isabelle build -o browser_info -o document=pdf -v -c HOL-Library\} Enabling both browser info and document preparation simultaneously causes an appropriate ``document'' link to be included in the HTML index. Documents may be generated independently of browser information as well, see \secref{sec:tool-document} for further details. \<^bigskip> The theory browsing information is stored in a sub-directory directory determined by the @{setting_ref ISABELLE_BROWSER_INFO} setting plus a prefix corresponding to the session chapter and identifier. In order to present Isabelle applications on the web, the corresponding subdirectory from @{setting ISABELLE_BROWSER_INFO} can be put on a WWW server. \ section \Preparing session root directories \label{sec:tool-mkroot}\ text \ The @{tool_def mkroot} tool configures a given directory as session root, with some \<^verbatim>\ROOT\ file and optional document source directory. Its usage is: @{verbatim [display] \Usage: isabelle mkroot [OPTIONS] [DIRECTORY] Options are: -A LATEX provide author in LaTeX notation (default: user name) -I init Mercurial repository and add generated files -T LATEX provide title in LaTeX notation (default: session name) -n NAME alternative session name (default: directory base name) Prepare session root directory (default: current directory). \} The results are placed in the given directory \dir\, which refers to the current directory by default. The @{tool mkroot} tool is conservative in the sense that it does not overwrite existing files or directories. Earlier attempts to generate a session root need to be deleted manually. The generated session template will be accompanied by a formal document, with \DIRECTORY\\<^verbatim>\/document/root.tex\ as its {\LaTeX} entry point (see also \chref{ch:present}). Options \<^verbatim>\-T\ and \<^verbatim>\-A\ specify the document title and author explicitly, using {\LaTeX} source notation. Option \<^verbatim>\-I\ initializes a Mercurial repository in the target directory, and adds all generated files (without commit). Option \<^verbatim>\-n\ specifies an alternative session name; otherwise the base name of the given directory is used. \<^medskip> The implicit Isabelle settings variable @{setting ISABELLE_LOGIC} specifies the parent session. \ subsubsection \Examples\ text \ Produce session \<^verbatim>\Test\ within a separate directory of the same name: @{verbatim [display] \isabelle mkroot Test && isabelle build -D Test\} \<^medskip> Upgrade the current directory into a session ROOT with document preparation, and build it: @{verbatim [display] \isabelle mkroot && isabelle build -D .\} \ section \Preparing Isabelle session documents \label{sec:tool-document}\ text \ The @{tool_def document} tool prepares logic session documents. Its usage is: @{verbatim [display] \Usage: isabelle document [OPTIONS] SESSION Options are: -O set option "document_output", relative to current directory -V verbose latex -d DIR include session directory - -n no build of session -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -v verbose build Prepare the theory document of a session.\} Generated {\LaTeX} sources are taken from the session build database: - @{tool_ref build} is invoked beforehand to ensure that it is up-to-date, but - option \<^verbatim>\-n\ suppresses that. Further files are generated on the spot, - notably essential Isabelle style files, and \<^verbatim>\session.tex\ to input all - theory sources from the session (excluding imports from other sessions). + @{tool_ref build} is invoked beforehand to ensure that it is up-to-date. + Further files are generated on the spot, notably essential Isabelle style + files, and \<^verbatim>\session.tex\ to input all theory sources from the session + (excluding imports from other sessions). - \<^medskip> Options \<^verbatim>\-d\, \<^verbatim>\-o\, \<^verbatim>\-v\ have the same meaning as for @{tool build}. + \<^medskip> Options \<^verbatim>\-d\, \<^verbatim>\-o\, \<^verbatim>\-v\ have the same meaning as for @{tool + build}. \<^medskip> Option \<^verbatim>\-V\ prints full output of {\LaTeX} tools. \<^medskip> Option \<^verbatim>\-O\~\dir\ specifies the @{system_option document_output} option relative to the current directory, while \<^verbatim>\-o document_output=\\dir\ is relative to the session directory. For example, for output directory ``\<^verbatim>\output\'' and the default document variant ``\<^verbatim>\document\'', the generated document sources are placed into the subdirectory \<^verbatim>\output/document/\ and the resulting PDF into \<^verbatim>\output/document.pdf\. \<^medskip> Isabelle is usually smart enough to create the PDF from the given \<^verbatim>\root.tex\ and optional \<^verbatim>\root.bib\ (bibliography) and \<^verbatim>\root.idx\ (index) using standard {\LaTeX} tools. Alternatively, \isakeyword{document\_files} in the session \<^verbatim>\ROOT\ may include an executable \<^verbatim>\build\ script to take care of that. It is invoked with command-line arguments for the document format (\<^verbatim>\pdf\) and the document variant name. The script needs to produce corresponding output files, e.g.\ \<^verbatim>\root.pdf\ for default document variants (the main work can be delegated to @{tool latex}). \ subsubsection \Examples\ text \ Produce the document from session \<^verbatim>\FOL\ with full verbosity, and a copy in the current directory (subdirectory \<^verbatim>\document\ and file \<^verbatim>\document.pdf)\: @{verbatim [display] \isabelle document -v -V -O. FOL\} \ section \Running {\LaTeX} within the Isabelle environment \label{sec:tool-latex}\ text \ The @{tool_def latex} tool provides the basic interface for Isabelle document preparation. Its usage is: @{verbatim [display] \Usage: isabelle latex [OPTIONS] [FILE] Options are: -o FORMAT specify output format: pdf (default), bbl, idx, sty Run LaTeX (and related tools) on FILE (default root.tex), producing the specified output format.\} Appropriate {\LaTeX}-related programs are run on the input file, according to the given output format: @{executable pdflatex}, @{executable bibtex} (for \<^verbatim>\bbl\), and @{executable makeindex} (for \<^verbatim>\idx\). The actual commands are determined from the settings environment (@{setting ISABELLE_PDFLATEX} etc.). The \<^verbatim>\sty\ output format causes the Isabelle style files to be updated from the distribution. This is useful in special situations where the document sources are to be processed another time by separate tools. \ subsubsection \Examples\ text \ Invoking @{tool latex} by hand may be occasionally useful when debugging failed attempts of the automatic document preparation stage of batch-mode Isabelle. The abortive process leaves the sources at a certain place within @{setting ISABELLE_BROWSER_INFO}, see the runtime error message for details. This enables users to inspect {\LaTeX} runs in further detail, e.g.\ like this: @{verbatim [display] \cd "$(isabelle getenv -b ISABELLE_BROWSER_INFO)/Unsorted/Test/document" isabelle latex -o pdf\} \ end diff --git a/src/Doc/System/Sessions.thy b/src/Doc/System/Sessions.thy --- a/src/Doc/System/Sessions.thy +++ b/src/Doc/System/Sessions.thy @@ -1,777 +1,783 @@ (*:maxLineLen=78:*) theory Sessions imports Base begin chapter \Isabelle sessions and build management \label{ch:session}\ text \ An Isabelle \<^emph>\session\ consists of a collection of related theories that may be associated with formal documents (\chref{ch:present}). There is also a notion of \<^emph>\persistent heap\ image to capture the state of a session, similar to object-code in compiled programming languages. Thus the concept of session resembles that of a ``project'' in common IDE environments, but the specific name emphasizes the connection to interactive theorem proving: the session wraps-up the results of user-interaction with the prover in a persistent form. Application sessions are built on a given parent session, which may be built recursively on other parents. Following this path in the hierarchy eventually leads to some major object-logic session like \HOL\, which itself is based on \Pure\ as the common root of all sessions. Processing sessions may take considerable time. Isabelle build management helps to organize this efficiently. This includes support for parallel build jobs, in addition to the multithreaded theory and proof checking that is already provided by the prover process itself. \ section \Session ROOT specifications \label{sec:session-root}\ text \ Session specifications reside in files called \<^verbatim>\ROOT\ within certain directories, such as the home locations of registered Isabelle components or additional project directories given by the user. The ROOT file format follows the lexical conventions of the \<^emph>\outer syntax\ of Isabelle/Isar, see also @{cite "isabelle-isar-ref"}. This defines common forms like identifiers, names, quoted strings, verbatim text, nested comments etc. The grammar for @{syntax session_chapter} and @{syntax session_entry} is given as syntax diagram below; each ROOT file may contain multiple specifications like this. Chapters help to organize browser info (\secref{sec:info}), but have no formal meaning. The default chapter is ``\Unsorted\''. Isabelle/jEdit @{cite "isabelle-jedit"} includes a simple editing mode \<^verbatim>\isabelle-root\ for session ROOT files, which is enabled by default for any file of that name. \<^rail>\ @{syntax_def session_chapter}: @'chapter' @{syntax name} ; @{syntax_def session_entry}: @'session' @{syntax system_name} groups? dir? '=' \ (@{syntax system_name} '+')? description? options? \ sessions? directories? (theories*) \ (document_theories?) (document_files*) \ (export_files*) ; groups: '(' (@{syntax name} +) ')' ; dir: @'in' @{syntax embedded} ; description: @'description' @{syntax text} ; options: @'options' opts ; opts: '[' ( (@{syntax name} '=' value | @{syntax name}) + ',' ) ']' ; value: @{syntax name} | @{syntax real} ; sessions: @'sessions' (@{syntax system_name}+) ; directories: @'directories' (dir+) ; theories: @'theories' opts? (theory_entry+) ; theory_entry: @{syntax system_name} ('(' @'global' ')')? ; document_theories: @'document_theories' (@{syntax name}+) ; document_files: @'document_files' ('(' dir ')')? (@{syntax embedded}+) ; export_files: @'export_files' ('(' dir ')')? ('[' nat ']')? \ (@{syntax embedded}+) \ \<^descr> \isakeyword{session}~\A = B + body\ defines a new session \A\ based on parent session \B\, with its content given in \body\ (imported sessions and theories). Note that a parent (like \HOL\) is mandatory in practical applications: only Isabelle/Pure can bootstrap itself from nothing. All such session specifications together describe a hierarchy (graph) of sessions, with globally unique names. The new session name \A\ should be sufficiently long and descriptive to stand on its own in a potentially large library. \<^descr> \isakeyword{session}~\A (groups)\ indicates a collection of groups where the new session is a member. Group names are uninterpreted and merely follow certain conventions. For example, the Isabelle distribution tags some important sessions by the group name called ``\main\''. Other projects may invent their own conventions, but this requires some care to avoid clashes within this unchecked name space. \<^descr> \isakeyword{session}~\A\~\isakeyword{in}~\dir\ specifies an explicit directory for this session; by default this is the current directory of the \<^verbatim>\ROOT\ file. All theory files are located relatively to the session directory. The prover process is run within the same as its current working directory. \<^descr> \isakeyword{description}~\text\ is a free-form annotation for this session. \<^descr> \isakeyword{options}~\[x = a, y = b, z]\ defines separate options (\secref{sec:system-options}) that are used when processing this session, but \<^emph>\without\ propagation to child sessions. Note that \z\ abbreviates \z = true\ for Boolean options. \<^descr> \isakeyword{sessions}~\names\ specifies sessions that are \<^emph>\imported\ into the current name space of theories. This allows to refer to a theory \A\ from session \B\ by the qualified name \B.A\ --- although it is loaded again into the current ML process, which is in contrast to a theory that is already present in the \<^emph>\parent\ session. Theories that are imported from other sessions are excluded from the current session document. \<^descr> \isakeyword{directories}~\dirs\ specifies additional directories for import of theory files via \isakeyword{theories} within \<^verbatim>\ROOT\ or \<^theory_text>\imports\ within a theory; \dirs\ are relative to the main session directory (cf.\ \isakeyword{session} \dots \isakeyword{in}~\dir\). These directories need to be exclusively assigned to a unique session, without implicit sharing of file-system locations. \<^descr> \isakeyword{theories}~\options names\ specifies a block of theories that are processed within an environment that is augmented by the given options, in addition to the global session options given before. Any number of blocks of \isakeyword{theories} may be given. Options are only active for each \isakeyword{theories} block separately. A theory name that is followed by \(\\isakeyword{global}\)\ is treated literally in other session specifications or theory imports --- the normal situation is to qualify theory names by the session name; this ensures globally unique names in big session graphs. Global theories are usually the entry points to major logic sessions: \Pure\, \Main\, \Complex_Main\, \HOLCF\, \IFOL\, \FOL\, \ZF\, \ZFC\ etc. Regular Isabelle applications should not claim any global theory names. \<^descr> \isakeyword{document_theories}~\names\ specifies theories from other sessions that should be included in the generated document source directory. These theories need to be explicit imports in the current session, or impliciti imports from the underlying hierarchy of parent sessions. The generated \<^verbatim>\session.tex\ file is not affected: the session's {\LaTeX} setup needs to \<^verbatim>\\input{\\\\\<^verbatim>\}\ generated \<^verbatim>\.tex\ files separately. \<^descr> \isakeyword{document_files}~\(\\isakeyword{in}~\base_dir) files\ lists source files for document preparation, typically \<^verbatim>\.tex\ and \<^verbatim>\.sty\ for {\LaTeX}. Only these explicitly given files are copied from the base directory to the document output directory, before formal document processing is started (see also \secref{sec:tool-document}). The local path structure of the \files\ is preserved, which allows to reconstruct the original directory hierarchy of \base_dir\. The default \base_dir\ is \<^verbatim>\document\ within the session root directory. \<^descr> \isakeyword{export_files}~\(\\isakeyword{in}~\target_dir) [number] patterns\ specifies theory exports that may get written to the file-system, e.g. via @{tool_ref build} with option \<^verbatim>\-e\ (\secref{sec:tool-build}). The \target_dir\ specification is relative to the session root directory; its default is \<^verbatim>\export\. Exports are selected via \patterns\ as in @{tool_ref export} (\secref{sec:tool-export}). The number given in brackets (default: 0) specifies elements that should be pruned from each name: it allows to reduce the resulting directory hierarchy at the danger of overwriting files due to loss of uniqueness. \ subsubsection \Examples\ text \ See \<^file>\~~/src/HOL/ROOT\ for a diversity of practically relevant situations, although it uses relatively complex quasi-hierarchic naming conventions like \<^verbatim>\HOL-SPARK\, \<^verbatim>\HOL-SPARK-Examples\. An alternative is to use unqualified names that are relatively long and descriptive, as in the Archive of Formal Proofs (\<^url>\https://isa-afp.org\), for example. \ section \System build options \label{sec:system-options}\ text \ See \<^file>\~~/etc/options\ for the main defaults provided by the Isabelle distribution. Isabelle/jEdit @{cite "isabelle-jedit"} includes a simple editing mode \<^verbatim>\isabelle-options\ for this file-format. The following options are particularly relevant to build Isabelle sessions, in particular with document preparation (\chref{ch:present}). \<^item> @{system_option_def "browser_info"} controls output of HTML browser info, see also \secref{sec:info}. \<^item> @{system_option_def "document"} controls document output for a particular session or theory; \<^verbatim>\document=pdf\ means enabled, \<^verbatim>\document=false\ means disabled (especially for particular theories). \<^item> @{system_option_def "document_output"} specifies an alternative directory for generated output of the document preparation system; the default is within the @{setting "ISABELLE_BROWSER_INFO"} hierarchy as explained in \secref{sec:info}. See also @{tool mkroot}, which generates a default configuration with output readily available to the author of the document. \<^item> @{system_option_def "document_variants"} specifies document variants as a colon-separated list of \name=tags\ entries. The default name \<^verbatim>\document\, without additional tags. Tags are specified as a comma separated list of modifier/name pairs and tell {\LaTeX} how to interpret certain Isabelle command regions: ``\<^verbatim>\+\\foo\'' (or just ``\foo\'') means to keep, ``\<^verbatim>\-\\foo\'' to drop, and ``\<^verbatim>\/\\foo\'' to fold text tagged as \foo\. The builtin default is equivalent to the tag specification ``\<^verbatim>\+document,+theory,+proof,+ML,+visible,-invisible,+important,+unimportant\''; see also the {\LaTeX} macros \<^verbatim>\\isakeeptag\, \<^verbatim>\\isadroptag\, and \<^verbatim>\\isafoldtag\, in \<^file>\~~/lib/texinputs/isabelle.sty\. In contrast, \<^verbatim>\document_variants=document:outline=/proof,/ML\ indicates two documents: the one called \<^verbatim>\document\ with default tags, and the other called \<^verbatim>\outline\ where proofs and ML sections are folded. Document variant names are just a matter of conventions. It is also possible to use different document variant names (without tags) for different document root entries, see also \secref{sec:tool-document}. \<^item> @{system_option_def "document_tags"} specifies alternative command tags as a comma-separated list of items: either ``\command\\<^verbatim>\%\\tag\'' for a specific command, or ``\<^verbatim>\%\\tag\'' as default for all other commands. This is occasionally useful to control the global visibility of commands via session options (e.g.\ in \<^verbatim>\ROOT\). \<^item> @{system_option_def "threads"} determines the number of worker threads for parallel checking of theories and proofs. The default \0\ means that a sensible maximum value is determined by the underlying hardware. For machines with many cores or with hyperthreading, this is often requires manual adjustment (on the command-line or within personal settings or preferences, not within a session \<^verbatim>\ROOT\). \<^item> @{system_option_def "condition"} specifies a comma-separated list of process environment variables (or Isabelle settings) that are required for the subsequent theories to be processed. Conditions are considered ``true'' if the corresponding environment value is defined and non-empty. \<^item> @{system_option_def "timeout"} and @{system_option_def "timeout_scale"} specify a real wall-clock timeout for the session as a whole: the two values are multiplied and taken as the number of seconds. Typically, @{system_option "timeout"} is given for individual sessions, and @{system_option "timeout_scale"} as global adjustment to overall hardware performance. The timer is controlled outside the ML process by the JVM that runs Isabelle/Scala. Thus it is relatively reliable in canceling processes that get out of control, even if there is a deadlock without CPU time usage. \<^item> @{system_option_def "profiling"} specifies a mode for global ML profiling. Possible values are the empty string (disabled), \<^verbatim>\time\ for \<^ML>\profile_time\ and \<^verbatim>\allocations\ for \<^ML>\profile_allocations\. Results appear near the bottom of the session log file. \<^item> @{system_option_def "system_heaps"} determines the directories for session heap images: \<^path>\$ISABELLE_HEAPS\ is the user directory and \<^path>\$ISABELLE_HEAPS_SYSTEM\ the system directory (usually within the Isabelle application). For \<^verbatim>\system_heaps=false\, heaps are stored in the user directory and may be loaded from both directories. For \<^verbatim>\system_heaps=true\, store and load happens only in the system directory. The @{tool_def options} tool prints Isabelle system options. Its command-line usage is: @{verbatim [display] \Usage: isabelle options [OPTIONS] [MORE_OPTIONS ...] Options are: -b include $ISABELLE_BUILD_OPTIONS -g OPTION get value of OPTION -l list options -x FILE export to FILE in YXML format Report Isabelle system options, augmented by MORE_OPTIONS given as arguments NAME=VAL or NAME.\} The command line arguments provide additional system options of the form \name\\<^verbatim>\=\\value\ or \name\ for Boolean options. Option \<^verbatim>\-b\ augments the implicit environment of system options by the ones of @{setting ISABELLE_BUILD_OPTIONS}, cf.\ \secref{sec:tool-build}. Option \<^verbatim>\-g\ prints the value of the given option. Option \<^verbatim>\-l\ lists all options with their declaration and current value. Option \<^verbatim>\-x\ specifies a file to export the result in YXML format, instead of printing it in human-readable form. \ section \Invoking the build process \label{sec:tool-build}\ text \ The @{tool_def build} tool invokes the build process for Isabelle sessions. It manages dependencies between sessions, related sources of theories and auxiliary files, and target heap images. Accordingly, it runs instances of the prover process with optional document preparation. Its command-line usage is:\<^footnote>\Isabelle/Scala provides the same functionality via \<^scala_method>\isabelle.Build.build\.\ @{verbatim [display] \Usage: isabelle build [OPTIONS] [SESSIONS ...] Options are: -B NAME include session NAME and all descendants -D DIR include session directory and select its sessions -N cyclic shuffling of NUMA CPU nodes (performance tuning) + -P DIR enable HTML/PDF presentation in directory (":" for default) -R refer to requirements of selected sessions -S soft build: only observe changes of sources, not heap images -X NAME exclude sessions from group NAME and all descendants -a select all sessions -b build heap images -c clean build -d DIR include session directory -e export files from session specification into file-system -f fresh build -g NAME select session group NAME -j INT maximum number of parallel jobs (default 1) -k KEYWORD check theory sources for conflicts with proposed keywords -l list session source files -n no build -- test dependencies only -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -v verbose -x NAME exclude session NAME and all descendants Build and manage Isabelle sessions, depending on implicit settings: ISABELLE_TOOL_JAVA_OPTIONS="..." ISABELLE_BUILD_OPTIONS="..." ML_PLATFORM="..." ML_HOME="..." ML_SYSTEM="..." ML_OPTIONS="..."\} \<^medskip> Isabelle sessions are defined via session ROOT files as described in (\secref{sec:session-root}). The totality of sessions is determined by collecting such specifications from all Isabelle component directories (\secref{sec:components}), augmented by more directories given via options \<^verbatim>\-d\~\DIR\ on the command line. Each such directory may contain a session \<^verbatim>\ROOT\ file with several session specifications. Any session root directory may refer recursively to further directories of the same kind, by listing them in a catalog file \<^verbatim>\ROOTS\ line-by-line. This helps to organize large collections of session specifications, or to make \<^verbatim>\-d\ command line options persistent (e.g.\ in \<^verbatim>\$ISABELLE_HOME_USER/ROOTS\). \<^medskip> The subset of sessions to be managed is determined via individual \SESSIONS\ given as command-line arguments, or session groups that are given via one or more options \<^verbatim>\-g\~\NAME\. Option \<^verbatim>\-a\ selects all sessions. The build tool takes session dependencies into account: the set of selected sessions is completed by including all ancestors. \<^medskip> One or more options \<^verbatim>\-B\~\NAME\ specify base sessions to be included (all descendants wrt.\ the session parent or import graph). \<^medskip> One or more options \<^verbatim>\-x\~\NAME\ specify sessions to be excluded (all descendants wrt.\ the session parent or import graph). Option \<^verbatim>\-X\ is analogous to this, but excluded sessions are specified by session group membership. \<^medskip> Option \<^verbatim>\-R\ reverses the selection in the sense that it refers to its requirements: all ancestor sessions excluding the original selection. This allows to prepare the stage for some build process with different options, before running the main build itself (without option \<^verbatim>\-R\). \<^medskip> Option \<^verbatim>\-D\ is similar to \<^verbatim>\-d\, but selects all sessions that are defined in the given directories. \<^medskip> Option \<^verbatim>\-S\ indicates a ``soft build'': the selection is restricted to those sessions that have changed sources (according to actually imported theories). The status of heap images is ignored. \<^medskip> The build process depends on additional options (\secref{sec:system-options}) that are passed to the prover eventually. The settings variable @{setting_ref ISABELLE_BUILD_OPTIONS} allows to provide additional defaults, e.g.\ \<^verbatim>\ISABELLE_BUILD_OPTIONS="document=pdf threads=4"\. Moreover, the environment of system build options may be augmented on the command line via \<^verbatim>\-o\~\name\\<^verbatim>\=\\value\ or \<^verbatim>\-o\~\name\, which abbreviates \<^verbatim>\-o\~\name\\<^verbatim>\=true\ for Boolean options. Multiple occurrences of \<^verbatim>\-o\ on the command-line are applied in the given order. \<^medskip> + Option \<^verbatim>\-P\ enables PDF/HTML presentation in the given directory, where + ``\<^verbatim>\-P:\'' refers to the default @{setting_ref ISABELLE_BROWSER_INFO} (or + @{setting_ref ISABELLE_BROWSER_INFO_SYSTEM}). + + \<^medskip> Option \<^verbatim>\-b\ ensures that heap images are produced for all selected sessions. By default, images are only saved for inner nodes of the hierarchy of sessions, as required for other sessions to continue later on. \<^medskip> Option \<^verbatim>\-c\ cleans the selected sessions (all descendants wrt.\ the session parent or import graph) before performing the specified build operation. \<^medskip> Option \<^verbatim>\-e\ executes the \isakeyword{export_files} directives from the ROOT specification of all explicitly selected sessions: the status of the session build database needs to be OK, but the session could have been built earlier. Using \isakeyword{export_files}, a session may serve as abstract interface for add-on build artefacts, but these are only materialized on explicit request: without option \<^verbatim>\-e\ there is no effect on the physical file-system yet. \<^medskip> Option \<^verbatim>\-f\ forces a fresh build of all selected sessions and their requirements. \<^medskip> Option \<^verbatim>\-n\ omits the actual build process after the preparatory stage (including optional cleanup). Note that the return code always indicates the status of the set of selected sessions. \<^medskip> Option \<^verbatim>\-j\ specifies the maximum number of parallel build jobs (prover processes). Each prover process is subject to a separate limit of parallel worker threads, cf.\ system option @{system_option_ref threads}. \<^medskip> Option \<^verbatim>\-N\ enables cyclic shuffling of NUMA CPU nodes. This may help performance tuning on Linux servers with separate CPU/memory modules. \<^medskip> Option \<^verbatim>\-v\ increases the general level of verbosity. Option \<^verbatim>\-l\ lists the source files that contribute to a session. \<^medskip> Option \<^verbatim>\-k\ specifies a newly proposed keyword for outer syntax (multiple uses allowed). The theory sources are checked for conflicts wrt.\ this hypothetical change of syntax, e.g.\ to reveal occurrences of identifiers that need to be quoted. \ subsubsection \Examples\ text \ Build a specific logic image: @{verbatim [display] \isabelle build -b HOLCF\} \<^smallskip> Build the main group of logic images: @{verbatim [display] \isabelle build -b -g main\} \<^smallskip> Build all descendants (and requirements) of \<^verbatim>\FOL\ and \<^verbatim>\ZF\: @{verbatim [display] \isabelle build -B FOL -B ZF\} \<^smallskip> Build all sessions where sources have changed (ignoring heaps): @{verbatim [display] \isabelle build -a -S\} \<^smallskip> Provide a general overview of the status of all Isabelle sessions, without building anything: @{verbatim [display] \isabelle build -a -n -v\} \<^smallskip> Build all sessions with HTML browser info and PDF document preparation: @{verbatim [display] \isabelle build -a -o browser_info -o document=pdf\} \<^smallskip> Build all sessions with a maximum of 8 parallel prover processes and 4 worker threads each (on a machine with many cores): @{verbatim [display] \isabelle build -a -j8 -o threads=4\} \<^smallskip> Build some session images with cleanup of their descendants, while retaining their ancestry: @{verbatim [display] \isabelle build -b -c HOL-Library HOL-Algebra\} \<^smallskip> Clean all sessions without building anything: @{verbatim [display] \isabelle build -a -n -c\} \<^smallskip> Build all sessions from some other directory hierarchy, according to the settings variable \<^verbatim>\AFP\ that happens to be defined inside the Isabelle environment: @{verbatim [display] \isabelle build -D '$AFP'\} \<^smallskip> Inform about the status of all sessions required for AFP, without building anything yet: @{verbatim [display] \isabelle build -D '$AFP' -R -v -n\} \ section \Retrieve theory exports \label{sec:tool-export}\ text \ The @{tool_def "export"} tool retrieves theory exports from the session database. Its command-line usage is: @{verbatim [display] \Usage: isabelle export [OPTIONS] SESSION Options are: -O DIR output directory for exported files (default: "export") -d DIR include session directory -l list exports -n no build of session -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -p NUM prune path of exported files by NUM elements -x PATTERN extract files matching pattern (e.g.\ "*:**" for all) List or export theory exports for SESSION: named blobs produced by isabelle build. Option -l or -x is required; option -x may be repeated. The PATTERN language resembles glob patterns in the shell, with ? and * (both excluding ":" and "/"), ** (excluding ":"), and [abc] or [^abc], and variants {pattern1,pattern2,pattern3}.\} \<^medskip> The specified session is updated via @{tool build} (\secref{sec:tool-build}), with the same options \<^verbatim>\-d\, \<^verbatim>\-o\. The option \<^verbatim>\-n\ suppresses the implicit build process: it means that a potentially outdated session database is used! \<^medskip> Option \<^verbatim>\-l\ lists all stored exports, with compound names \theory\\<^verbatim>\:\\name\. \<^medskip> Option \<^verbatim>\-x\ extracts stored exports whose compound name matches the given pattern. Note that wild cards ``\<^verbatim>\?\'' and ``\<^verbatim>\*\'' do not match the separators ``\<^verbatim>\:\'' and ``\<^verbatim>\/\''; the wild card \<^verbatim>\**\ matches over directory name hierarchies separated by ``\<^verbatim>\/\''. Thus the pattern ``\<^verbatim>\*:**\'' matches \<^emph>\all\ theory exports. Multiple options \<^verbatim>\-x\ refer to the union of all specified patterns. Option \<^verbatim>\-O\ specifies an alternative output directory for option \<^verbatim>\-x\: the default is \<^verbatim>\export\ within the current directory. Each theory creates its own sub-directory hierarchy, using the session-qualified theory name. Option \<^verbatim>\-p\ specifies the number of elements that should be pruned from each name: it allows to reduce the resulting directory hierarchy at the danger of overwriting files due to loss of uniqueness. \ section \Dump PIDE session database \label{sec:tool-dump}\ text \ The @{tool_def "dump"} tool dumps information from the cumulative PIDE session database (which is processed on the spot). Its command-line usage is: @{verbatim [display] \Usage: isabelle dump [OPTIONS] [SESSIONS ...] Options are: -A NAMES dump named aspects (default: ...) -B NAME include session NAME and all descendants -D DIR include session directory and select its sessions -O DIR output directory for dumped files (default: "dump") -R refer to requirements of selected sessions -X NAME exclude sessions from group NAME and all descendants -a select all sessions -b NAME base logic image (default "Pure") -d DIR include session directory -g NAME select session group NAME -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -v verbose -x NAME exclude session NAME and all descendants Dump cumulative PIDE session database, with the following aspects: ...\} \<^medskip> Options \<^verbatim>\-B\, \<^verbatim>\-D\, \<^verbatim>\-R\, \<^verbatim>\-X\, \<^verbatim>\-a\, \<^verbatim>\-d\, \<^verbatim>\-g\, \<^verbatim>\-x\ and the remaining command-line arguments specify sessions as in @{tool build} (\secref{sec:tool-build}): the cumulative PIDE database of all their loaded theories is dumped to the output directory of option \<^verbatim>\-O\ (default: \<^verbatim>\dump\ in the current directory). \<^medskip> Option \<^verbatim>\-b\ specifies an optional base logic image, for improved scalability of the PIDE session. Its theories are only processed if it is included in the overall session selection. \<^medskip> Option \<^verbatim>\-o\ overrides Isabelle system options as for @{tool build} (\secref{sec:tool-build}). \<^medskip> Option \<^verbatim>\-v\ increases the general level of verbosity. \<^medskip> Option \<^verbatim>\-A\ specifies named aspects of the dump, as a comma-separated list. The default is to dump all known aspects, as given in the command-line usage of the tool. The underlying Isabelle/Scala operation \<^scala_method>\isabelle.Dump.dump\ takes aspects as user-defined operations on the final PIDE state and document version. This allows to imitate Prover IDE rendering under program control. \ subsubsection \Examples\ text \ Dump all Isabelle/ZF sessions (which are rather small): @{verbatim [display] \isabelle dump -v -B ZF\} \<^smallskip> Dump the quite substantial \<^verbatim>\HOL-Analysis\ session, with full bootstrap from Isabelle/Pure: @{verbatim [display] \isabelle dump -v HOL-Analysis\} \<^smallskip> Dump all sessions connected to HOL-Analysis, using main Isabelle/HOL as basis: @{verbatim [display] \isabelle dump -v -b HOL -B HOL-Analysis\} This results in uniform PIDE markup for everything, except for the Isabelle/Pure bootstrap process itself. Producing that on the spot requires several GB of heap space, both for the Isabelle/Scala and Isabelle/ML process (in 64bit mode). Here are some relevant settings (\secref{sec:boot}) for such ambitious applications: @{verbatim [display] \ISABELLE_TOOL_JAVA_OPTIONS="-Xms4g -Xmx32g -Xss16m" ML_OPTIONS="--minheap 4G --maxheap 32G" \} \ section \Update theory sources based on PIDE markup \label{sec:tool-update}\ text \ The @{tool_def "update"} tool updates theory sources based on markup that is produced from a running PIDE session (similar to @{tool dump} \secref{sec:tool-dump}). Its command-line usage is: @{verbatim [display] \Usage: isabelle update [OPTIONS] [SESSIONS ...] Options are: -B NAME include session NAME and all descendants -D DIR include session directory and select its sessions -R refer to requirements of selected sessions -X NAME exclude sessions from group NAME and all descendants -a select all sessions -b NAME base logic image (default "Pure") -d DIR include session directory -g NAME select session group NAME -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -u OPT overide update option: shortcut for "-o update_OPT" -v verbose -x NAME exclude session NAME and all descendants Update theory sources based on PIDE markup.\} \<^medskip> Options \<^verbatim>\-B\, \<^verbatim>\-D\, \<^verbatim>\-R\, \<^verbatim>\-X\, \<^verbatim>\-a\, \<^verbatim>\-d\, \<^verbatim>\-g\, \<^verbatim>\-x\ and the remaining command-line arguments specify sessions as in @{tool build} (\secref{sec:tool-build}) or @{tool dump} (\secref{sec:tool-dump}). \<^medskip> Option \<^verbatim>\-b\ specifies an optional base logic image, for improved scalability of the PIDE session. Its theories are only processed if it is included in the overall session selection. \<^medskip> Option \<^verbatim>\-v\ increases the general level of verbosity. \<^medskip> Option \<^verbatim>\-o\ overrides Isabelle system options as for @{tool build} (\secref{sec:tool-build}). Option \<^verbatim>\-u\ refers to specific \<^verbatim>\update\ options, by relying on naming convention: ``\<^verbatim>\-u\~\OPT\'' is a shortcut for ``\<^verbatim>\-o\~\<^verbatim>\update_\\OPT\''. \<^medskip> The following update options are supported: \<^item> @{system_option update_inner_syntax_cartouches} to update inner syntax (types, terms, etc.)~to use cartouches, instead of double-quoted strings or atomic identifiers. For example, ``\<^theory_text>\lemma \x = x\\'' is replaced by ``\<^theory_text>\lemma \x = x\\'', and ``\<^theory_text>\assume A\'' is replaced by ``\<^theory_text>\assume \A\\''. \<^item> @{system_option update_mixfix_cartouches} to update mixfix templates to use cartouches instead of double-quoted strings. For example, ``\<^theory_text>\(infixl \+\ 65)\'' is replaced by ``\<^theory_text>\(infixl \+\ 65)\''. \<^item> @{system_option update_control_cartouches} to update antiquotations to use the compact form with control symbol and cartouche argument. For example, ``\@{term \x + y\}\'' is replaced by ``\\<^term>\x + y\\'' (the control symbol is literally \<^verbatim>\\<^term>\.) \<^item> @{system_option update_path_cartouches} to update file-system paths to use cartouches: this depends on language markup provided by semantic processing of parsed input. It is also possible to produce custom updates in Isabelle/ML, by reporting \<^ML>\Markup.update\ with the precise source position and a replacement text. This operation should be made conditional on specific system options, similar to the ones above. Searching the above option names in ML sources of \<^dir>\$ISABELLE_HOME/src/Pure\ provides some examples. Updates can be in conflict by producing nested or overlapping edits: this may require to run @{tool update} multiple times. \ subsubsection \Examples\ text \ Update some cartouche notation in all theory sources required for session \<^verbatim>\HOL-Analysis\ (and ancestors): @{verbatim [display] \isabelle update -u mixfix_cartouches HOL-Analysis\} \<^smallskip> Update the same for all application sessions based on \<^verbatim>\HOL-Analysis\ --- using its image is taken starting point (for reduced resource requirements): @{verbatim [display] \isabelle update -u mixfix_cartouches -b HOL-Analysis -B HOL-Analysis\} \<^smallskip> Update sessions that build on \<^verbatim>\HOL-Proofs\, which need to be run separately with special options as follows: @{verbatim [display] \isabelle update -u mixfix_cartouches -l HOL-Proofs -B HOL-Proofs -o record_proofs=2\} \<^smallskip> See also the end of \secref{sec:tool-dump} for hints on increasing Isabelle/ML heap sizes for very big PIDE processes that include many sessions, notably from the Archive of Formal Proofs. \ section \Explore sessions structure\ text \ The @{tool_def "sessions"} tool explores the sessions structure. Its command-line usage is: @{verbatim [display] \Usage: isabelle sessions [OPTIONS] [SESSIONS ...] Options are: -B NAME include session NAME and all descendants -D DIR include session directory and select its sessions -R refer to requirements of selected sessions -X NAME exclude sessions from group NAME and all descendants -a select all sessions -d DIR include session directory -g NAME select session group NAME -x NAME exclude session NAME and all descendants Explore the structure of Isabelle sessions and print result names in topological order (on stdout).\} Arguments and options for session selection resemble @{tool build} (\secref{sec:tool-build}). \ subsubsection \Examples\ text \ All sessions of the Isabelle distribution: @{verbatim [display] \isabelle sessions -a\} \<^medskip> Sessions that are based on \<^verbatim>\ZF\ (and required by it): @{verbatim [display] \isabelle sessions -B ZF\} \<^medskip> All sessions of Isabelle/AFP (based in directory \<^path>\AFP\): @{verbatim [display] \isabelle sessions -D AFP/thys\} \<^medskip> Sessions required by Isabelle/AFP (based in directory \<^path>\AFP\): @{verbatim [display] \isabelle sessions -R -D AFP/thys\} \ end diff --git a/src/Pure/Admin/build_doc.scala b/src/Pure/Admin/build_doc.scala --- a/src/Pure/Admin/build_doc.scala +++ b/src/Pure/Admin/build_doc.scala @@ -1,101 +1,101 @@ /* Title: Pure/Admin/build_doc.scala Author: Makarius Build Isabelle documentation. */ package isabelle object Build_Doc { /* build_doc */ def build_doc( options: Options, progress: Progress = new Progress, all_docs: Boolean = false, max_jobs: Int = 1, docs: List[String] = Nil) { val store = Sessions.store(options) val sessions_structure = Sessions.load_structure(options) val selected = for { session <- sessions_structure.build_topological_order info = sessions_structure(session) if info.groups.contains("doc") doc = info.options.string("document_variants") if all_docs || docs.contains(doc) } yield (doc, session) val documents = selected.map(_._1) val selection = Sessions.Selection(sessions = selected.map(_._2)) docs.filter(doc => !documents.contains(doc)) match { case Nil => case bad => error("No documentation session for " + commas_quote(bad)) } progress.echo("Build started for sessions " + commas_quote(selection.sessions)) Build.build(options, selection = selection, progress = progress, max_jobs = max_jobs).ok || error("Build failed") progress.echo("Build started for documentation " + commas_quote(documents)) val doc_options = options + "document=pdf" + "document_output=~~/doc" + "document_output_sources=false" val deps = Sessions.load_structure(doc_options).selection_deps(selection) val errs = Par_List.map((doc_session: (String, String)) => try { - Present.build_documents(doc_session._2, deps, store, progress = progress) + Presentation.build_documents(doc_session._2, deps, store, progress = progress) None } catch { case Exn.Interrupt.ERROR(msg) => val sep = if (msg.contains('\n')) "\n" else " " Some("Documentation " + doc_session._1 + " failed:" + sep + msg) }, selected).flatten if (errs.nonEmpty) error(cat_lines(errs)) } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("build_doc", "build Isabelle documentation", args => { var all_docs = false var max_jobs = 1 var options = Options.init() val getopts = Getopts(""" Usage: isabelle build_doc [OPTIONS] [DOCS ...] Options are: -a select all documentation sessions -j INT maximum number of parallel jobs (default 1) -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) Build Isabelle documentation from documentation sessions with suitable document_variants entry. """, "a" -> (_ => all_docs = true), "j:" -> (arg => max_jobs = Value.Int.parse(arg)), "o:" -> (arg => options = options + arg)) val docs = getopts(args) if (!all_docs && docs.isEmpty) getopts.usage() val progress = new Console_Progress() progress.interrupt_handler { build_doc(options, progress, all_docs, max_jobs, docs) } }) } diff --git a/src/Pure/General/sha1.scala b/src/Pure/General/sha1.scala --- a/src/Pure/General/sha1.scala +++ b/src/Pure/General/sha1.scala @@ -1,71 +1,73 @@ /* Title: Pure/General/sha1.scala Author: Makarius Digest strings according to SHA-1 (see RFC 3174). */ package isabelle import java.io.{File => JFile, FileInputStream} import java.security.MessageDigest object SHA1 { final class Digest private[SHA1](val rep: String) { override def hashCode: Int = rep.hashCode override def equals(that: Any): Boolean = that match { case other: Digest => rep == other.rep case _ => false } override def toString: String = rep } private def make_result(digest: MessageDigest): Digest = { val result = new StringBuilder for (b <- digest.digest()) { val i = b.asInstanceOf[Int] & 0xFF if (i < 16) result += '0' result ++= Integer.toHexString(i) } new Digest(result.toString) } def fake(rep: String): Digest = new Digest(rep) def digest(file: JFile): Digest = { val digest = MessageDigest.getInstance("SHA") using(new FileInputStream(file))(stream => { val buf = new Array[Byte](65536) var m = 0 do { m = stream.read(buf, 0, buf.length) if (m != -1) digest.update(buf, 0, m) } while (m != -1) }) make_result(digest) } def digest(path: Path): Digest = digest(path.file) def digest(bytes: Array[Byte]): Digest = { val digest = MessageDigest.getInstance("SHA") digest.update(bytes) make_result(digest) } def digest(bytes: Bytes): Digest = bytes.sha1_digest def digest(string: String): Digest = digest(Bytes(string)) + def digest_set(digests: List[Digest]): Digest = + digest(cat_lines(digests.map(_.toString).sorted)) val digest_length: Int = digest("").rep.length } diff --git a/src/Pure/PIDE/resources.scala b/src/Pure/PIDE/resources.scala --- a/src/Pure/PIDE/resources.scala +++ b/src/Pure/PIDE/resources.scala @@ -1,448 +1,448 @@ /* Title: Pure/PIDE/resources.scala Author: Makarius Resources for theories and auxiliary files. */ package isabelle import scala.util.parsing.input.Reader import java.io.{File => JFile} class Resources( val sessions_structure: Sessions.Structure, val session_base: Sessions.Base, val log: Logger = No_Logger, command_timings: List[Properties.T] = Nil) { resources => /* init session */ def init_session_yxml: String = { import XML.Encode._ YXML.string_of_body( pair(list(pair(string, int)), pair(list(pair(string, properties)), pair(list(pair(string, string)), pair(list(pair(string, string)), pair(list(pair(string, list(string))), pair(list(properties), pair(list(string), pair(list(pair(string, string)), list(string)))))))))( (Symbol.codes, - (resources.sessions_structure.session_positions, - (resources.sessions_structure.dest_session_directories, - (resources.sessions_structure.session_chapters, - (resources.sessions_structure.bibtex_entries, + (sessions_structure.session_positions, + (sessions_structure.dest_session_directories, + (sessions_structure.session_chapters, + (sessions_structure.bibtex_entries, (command_timings, - (resources.session_base.doc_names, - (resources.session_base.global_theories.toList, - resources.session_base.loaded_theories.keys)))))))))) + (session_base.doc_names, + (session_base.global_theories.toList, + session_base.loaded_theories.keys)))))))))) } /* file formats */ def make_theory_name(name: Document.Node.Name): Option[Document.Node.Name] = File_Format.registry.get(name).flatMap(_.make_theory_name(resources, name)) def make_theory_content(thy_name: Document.Node.Name): Option[String] = File_Format.registry.get_theory(thy_name).flatMap(_.make_theory_content(resources, thy_name)) - def make_preview(snapshot: Document.Snapshot): Option[Present.Preview] = + def make_preview(snapshot: Document.Snapshot): Option[Presentation.Preview] = File_Format.registry.get(snapshot.node_name).flatMap(_.make_preview(snapshot)) def is_hidden(name: Document.Node.Name): Boolean = !name.is_theory || name.theory == Sessions.root_name || File_Format.registry.is_theory(name) def thy_path(path: Path): Path = path.ext("thy") /* file-system operations */ def append(dir: String, source_path: Path): String = (Path.explode(dir) + source_path).expand.implode def append(node_name: Document.Node.Name, source_path: Path): String = append(node_name.master_dir, source_path) def make_theory_node(dir: String, file: Path, theory: String): Document.Node.Name = { val node = append(dir, file) val master_dir = append(dir, file.dir) Document.Node.Name(node, master_dir, theory) } def loaded_theory_node(theory: String): Document.Node.Name = Document.Node.Name(theory, "", theory) /* source files of Isabelle/ML bootstrap */ def source_file(raw_name: String): Option[String] = { if (Path.is_wellformed(raw_name)) { if (Path.is_valid(raw_name)) { def check(p: Path): Option[Path] = if (p.is_file) Some(p) else None val path = Path.explode(raw_name) val path1 = if (path.is_absolute || path.is_current) check(path) else { check(Path.explode("~~/src/Pure") + path) orElse (if (Isabelle_System.getenv("ML_SOURCES") == "") None else check(Path.explode("$ML_SOURCES") + path)) } Some(File.platform_path(path1 getOrElse path)) } else None } else Some(raw_name) } /* theory files */ def loaded_files(syntax: Outer_Syntax, name: Document.Node.Name): () => List[Path] = { val (is_utf8, raw_text) = with_thy_reader(name, reader => (Scan.reader_is_utf8(reader), reader.source.toString)) () => { if (syntax.load_commands_in(raw_text)) { val text = Symbol.decode(Scan.reader_decode_utf8(is_utf8, raw_text)) val spans = syntax.parse_spans(text) val dir = Path.explode(name.master_dir) spans.iterator.flatMap(Command.span_files(syntax, _)._1). map(a => (dir + Path.explode(a)).expand).toList } else Nil } } def pure_files(syntax: Outer_Syntax): List[Path] = { val pure_dir = Path.explode("~~/src/Pure") val roots = for { (name, _) <- Thy_Header.ml_roots } yield (pure_dir + Path.explode(name)).expand val files = for { (path, (_, theory)) <- roots zip Thy_Header.ml_roots file <- loaded_files(syntax, Document.Node.Name(path.implode, path.dir.implode, theory))() } yield file roots ::: files } def theory_name(qualifier: String, theory: String): String = if (Long_Name.is_qualified(theory) || session_base.global_theories.isDefinedAt(theory)) theory else Long_Name.qualify(qualifier, theory) def find_theory_node(theory: String): Option[Document.Node.Name] = { val thy_file = thy_path(Path.basic(Long_Name.base_name(theory))) val session = session_base.theory_qualifier(theory) val dirs = sessions_structure.get(session) match { case Some(info) => info.dirs case None => Nil } dirs.collectFirst({ case dir if (dir + thy_file).is_file => make_theory_node("", dir + thy_file, theory) }) } def import_name(qualifier: String, dir: String, s: String): Document.Node.Name = { val theory = theory_name(qualifier, Thy_Header.import_name(s)) def theory_node = make_theory_node(dir, thy_path(Path.explode(s)), theory) if (!Thy_Header.is_base_name(s)) theory_node else if (session_base.loaded_theory(theory)) loaded_theory_node(theory) else { find_theory_node(theory) match { case Some(node_name) => node_name case None => if (Long_Name.is_qualified(s)) loaded_theory_node(theory) else theory_node } } } def import_name(name: Document.Node.Name, s: String): Document.Node.Name = import_name(session_base.theory_qualifier(name), name.master_dir, s) def import_name(info: Sessions.Info, s: String): Document.Node.Name = import_name(info.name, info.dir.implode, s) def find_theory(file: JFile): Option[Document.Node.Name] = { for { qualifier <- session_base.session_directories.get(File.canonical(file).getParentFile) theory_base <- proper_string(Thy_Header.theory_name(file.getName)) theory = theory_name(qualifier, theory_base) theory_node <- find_theory_node(theory) if File.eq(theory_node.path.file, file) } yield theory_node } def complete_import_name(context_name: Document.Node.Name, s: String): List[String] = { val context_session = session_base.theory_qualifier(context_name) val context_dir = try { Some(context_name.master_dir_path) } catch { case ERROR(_) => None } (for { (session, (info, _)) <- sessions_structure.imports_graph.iterator dir <- (if (session == context_session) context_dir.toList else info.dirs).iterator theory <- Thy_Header.try_read_dir(dir).iterator if Completion.completed(s)(theory) } yield { if (session == context_session || session_base.global_theories.isDefinedAt(theory)) theory else Long_Name.qualify(session, theory) }).toList.sorted } def with_thy_reader[A](name: Document.Node.Name, f: Reader[Char] => A): A = { val path = name.path if (path.is_file) using(Scan.byte_reader(path.file))(f) else if (name.node == name.theory) error("Cannot load theory " + quote(name.theory)) else error ("Cannot load theory file " + path) } def check_thy_reader(node_name: Document.Node.Name, reader: Reader[Char], start: Token.Pos = Token.Pos.command, strict: Boolean = true): Document.Node.Header = { if (node_name.is_theory && reader.source.length > 0) { try { val header = Thy_Header.read(reader, start, strict) val base_name = node_name.theory_base_name if (Long_Name.is_qualified(header.name)) { error("Bad theory name " + quote(header.name) + " with qualification" + Position.here(header.pos)) } if (base_name != header.name) { error("Bad theory name " + quote(header.name) + " for file " + thy_path(Path.basic(base_name)) + Position.here(header.pos) + Completion.report_theories(header.pos, List(base_name))) } val imports_pos = header.imports_pos.map({ case (s, pos) => val name = import_name(node_name, s) if (Sessions.exclude_theory(name.theory_base_name)) error("Bad theory name " + quote(name.theory_base_name) + Position.here(pos)) (name, pos) }) Document.Node.Header(imports_pos, header.keywords, header.abbrevs) } catch { case exn: Throwable => Document.Node.bad_header(Exn.message(exn)) } } else Document.Node.no_header } def check_thy(name: Document.Node.Name, start: Token.Pos = Token.Pos.command, strict: Boolean = true): Document.Node.Header = with_thy_reader(name, check_thy_reader(name, _, start, strict)) /* special header */ def special_header(name: Document.Node.Name): Option[Document.Node.Header] = { val imports = if (name.theory == Sessions.root_name) List(import_name(name, Sessions.theory_name)) else if (Thy_Header.is_ml_root(name.theory)) List(import_name(name, Thy_Header.ML_BOOTSTRAP)) else if (Thy_Header.is_bootstrap(name.theory)) List(import_name(name, Thy_Header.PURE)) else Nil if (imports.isEmpty) None else Some(Document.Node.Header(imports.map((_, Position.none)))) } /* blobs */ def undefined_blobs(nodes: Document.Nodes): List[Document.Node.Name] = (for { (node_name, node) <- nodes.iterator if !session_base.loaded_theory(node_name) cmd <- node.load_commands.iterator name <- cmd.blobs_undefined.iterator } yield name).toList /* document changes */ def parse_change( reparse_limit: Int, previous: Document.Version, doc_blobs: Document.Blobs, edits: List[Document.Edit_Text], consolidate: List[Document.Node.Name]): Session.Change = Thy_Syntax.parse_change(resources, reparse_limit, previous, doc_blobs, edits, consolidate) def commit(change: Session.Change) { } /* theory and file dependencies */ def dependencies( thys: List[(Document.Node.Name, Position.T)], progress: Progress = new Progress): Dependencies[Unit] = Dependencies.empty[Unit].require_thys((), thys, progress = progress) def session_dependencies(info: Sessions.Info, progress: Progress = new Progress) : Dependencies[Options] = { (Dependencies.empty[Options] /: info.theories)({ case (dependencies, (options, thys)) => dependencies.require_thys(options, for { (thy, pos) <- thys } yield (import_name(info, thy), pos), progress = progress) }) } object Dependencies { def empty[A]: Dependencies[A] = new Dependencies[A](Nil, Map.empty) private def show_path(names: List[Document.Node.Name]): String = names.map(name => quote(name.theory)).mkString(" via ") private def cycle_msg(names: List[Document.Node.Name]): String = "Cyclic dependency of " + show_path(names) private def required_by(initiators: List[Document.Node.Name]): String = if (initiators.isEmpty) "" else "\n(required by " + show_path(initiators.reverse) + ")" } final class Dependencies[A] private( rev_entries: List[Document.Node.Entry], seen: Map[Document.Node.Name, A]) { private def cons(entry: Document.Node.Entry): Dependencies[A] = new Dependencies[A](entry :: rev_entries, seen) def require_thy(adjunct: A, thy: (Document.Node.Name, Position.T), initiators: List[Document.Node.Name] = Nil, progress: Progress = new Progress): Dependencies[A] = { val (name, pos) = thy def message: String = "The error(s) above occurred for theory " + quote(name.theory) + Dependencies.required_by(initiators) + Position.here(pos) if (seen.isDefinedAt(name)) this else { val dependencies1 = new Dependencies[A](rev_entries, seen + (name -> adjunct)) if (session_base.loaded_theory(name)) dependencies1 else { try { if (initiators.contains(name)) error(Dependencies.cycle_msg(initiators)) progress.expose_interrupt() val header = try { check_thy(name, Token.Pos.file(name.node)).cat_errors(message) } catch { case ERROR(msg) => cat_error(msg, message) } val entry = Document.Node.Entry(name, header) dependencies1.require_thys(adjunct, header.imports_pos, initiators = name :: initiators, progress = progress).cons(entry) } catch { case e: Throwable => dependencies1.cons(Document.Node.Entry(name, Document.Node.bad_header(Exn.message(e)))) } } } } def require_thys(adjunct: A, thys: List[(Document.Node.Name, Position.T)], progress: Progress = new Progress, initiators: List[Document.Node.Name] = Nil): Dependencies[A] = (this /: thys)(_.require_thy(adjunct, _, progress = progress, initiators = initiators)) def entries: List[Document.Node.Entry] = rev_entries.reverse def theories: List[Document.Node.Name] = entries.map(_.name) def theories_adjunct: List[(Document.Node.Name, A)] = theories.map(name => (name, seen(name))) def errors: List[String] = entries.flatMap(_.header.errors) def check_errors: Dependencies[A] = errors match { case Nil => this case errs => error(cat_lines(errs)) } lazy val theory_graph: Document.Node.Name.Graph[Unit] = { val regular = theories.toSet val irregular = (for { entry <- entries.iterator imp <- entry.header.imports if !regular(imp) } yield imp).toSet Document.Node.Name.make_graph( irregular.toList.map(name => ((name, ()), Nil)) ::: entries.map(entry => ((entry.name, ()), entry.header.imports))) } lazy val loaded_theories: Graph[String, Outer_Syntax] = (session_base.loaded_theories /: entries)({ case (graph, entry) => val name = entry.name.theory val imports = entry.header.imports.map(_.theory) val graph1 = (graph /: (name :: imports))(_.default_node(_, Outer_Syntax.empty)) val graph2 = (graph1 /: imports)(_.add_edge(_, name)) val syntax0 = if (name == Thy_Header.PURE) List(Thy_Header.bootstrap_syntax) else Nil val syntax1 = (name :: graph2.imm_preds(name).toList).map(graph2.get_node) val syntax = Outer_Syntax.merge(syntax0 ::: syntax1) + entry.header graph2.map_node(name, _ => syntax) }) def loaded_files(pure: Boolean): List[(String, List[Path])] = { val loaded_files = theories.map(_.theory) zip Par_List.map((e: () => List[Path]) => e(), theories.map(name => resources.loaded_files(loaded_theories.get_node(name.theory), name))) if (pure) { val pure_files = resources.pure_files(overall_syntax) loaded_files.map({ case (name, files) => (name, if (name == Thy_Header.PURE) pure_files ::: files else files) }) } else loaded_files } def imported_files: List[Path] = { val base_theories = loaded_theories.all_preds(theories.map(_.theory)). filter(session_base.loaded_theories.defined) base_theories.map(theory => session_base.known_theories(theory).name.path) ::: base_theories.flatMap(session_base.known_loaded_files(_)) } lazy val overall_syntax: Outer_Syntax = Outer_Syntax.merge(loaded_theories.maximals.map(loaded_theories.get_node)) override def toString: String = entries.toString } } diff --git a/src/Pure/ROOT.ML b/src/Pure/ROOT.ML --- a/src/Pure/ROOT.ML +++ b/src/Pure/ROOT.ML @@ -1,356 +1,355 @@ (* Title: Pure/ROOT.ML Author: Makarius Main entry point for the Isabelle/Pure bootstrap process. Note: When this file is open in the Prover IDE, the ML files of Isabelle/Pure can be explored interactively. This is a separate copy of Pure within Pure: it does not affect the running logic session. *) chapter "Isabelle/Pure bootstrap"; ML_file "ML/ml_name_space.ML"; section "Bootstrap phase 0: Poly/ML setup"; ML_file "ML/ml_init.ML"; ML_file "ML/ml_system.ML"; ML_file "System/distribution.ML"; ML_file "General/basics.ML"; ML_file "General/symbol_explode.ML"; ML_file "Concurrent/multithreading.ML"; ML_file "Concurrent/unsynchronized.ML"; ML_file "Concurrent/synchronized.ML"; ML_file "Concurrent/counter.ML"; ML_file "ML/ml_heap.ML"; ML_file "ML/ml_profiling.ML"; ML_file "ML/ml_print_depth0.ML"; ML_file "ML/ml_pretty.ML"; ML_file "ML/ml_compiler0.ML"; section "Bootstrap phase 1: towards ML within position context"; subsection "Library of general tools"; ML_file "library.ML"; ML_file "General/print_mode.ML"; ML_file "General/alist.ML"; ML_file "General/table.ML"; ML_file "General/random.ML"; ML_file "General/value.ML"; ML_file "General/properties.ML"; ML_file "General/output.ML"; ML_file "PIDE/markup.ML"; ML_file "General/utf8.ML"; ML_file "General/scan.ML"; ML_file "General/source.ML"; ML_file "General/symbol.ML"; ML_file "General/position.ML"; ML_file "General/symbol_pos.ML"; ML_file "General/input.ML"; ML_file "General/comment.ML"; ML_file "General/antiquote.ML"; ML_file "ML/ml_lex.ML"; ML_file "ML/ml_compiler1.ML"; section "Bootstrap phase 2: towards ML within Isar context"; subsection "Library of general tools"; ML_file "General/integer.ML"; ML_file "General/stack.ML"; ML_file "General/queue.ML"; ML_file "General/heap.ML"; ML_file "General/same.ML"; ML_file "General/ord_list.ML"; ML_file "General/balanced_tree.ML"; ML_file "General/linear_set.ML"; ML_file "General/buffer.ML"; ML_file "General/pretty.ML"; ML_file "General/rat.ML"; ML_file "PIDE/xml.ML"; ML_file "General/path.ML"; ML_file "General/url.ML"; ML_file "System/bash_syntax.ML"; ML_file "General/file.ML"; ML_file "General/long_name.ML"; ML_file "General/binding.ML"; ML_file "General/socket_io.ML"; ML_file "General/seq.ML"; ML_file "General/timing.ML"; ML_file "General/sha1.ML"; ML_file "PIDE/byte_message.ML"; ML_file "PIDE/yxml.ML"; ML_file "PIDE/protocol_message.ML"; ML_file "PIDE/document_id.ML"; ML_file "General/change_table.ML"; ML_file "General/graph.ML"; ML_file "System/options.ML"; subsection "Fundamental structures"; ML_file "name.ML"; ML_file "term.ML"; ML_file "context.ML"; ML_file "config.ML"; ML_file "context_position.ML"; ML_file "soft_type_system.ML"; subsection "Concurrency within the ML runtime"; ML_file "ML/exn_properties.ML"; ML_file_no_debug "ML/exn_debugger.ML"; ML_file "Concurrent/thread_data_virtual.ML"; ML_file "Concurrent/isabelle_thread.ML"; ML_file "Concurrent/single_assignment.ML"; ML_file "Concurrent/par_exn.ML"; ML_file "Concurrent/task_queue.ML"; ML_file "Concurrent/future.ML"; ML_file "Concurrent/event_timer.ML"; ML_file "Concurrent/timeout.ML"; ML_file "Concurrent/lazy.ML"; ML_file "Concurrent/par_list.ML"; ML_file "Concurrent/mailbox.ML"; ML_file "Concurrent/cache.ML"; ML_file "PIDE/active.ML"; ML_file "Thy/export.ML"; subsection "Inner syntax"; ML_file "Syntax/type_annotation.ML"; ML_file "Syntax/term_position.ML"; ML_file "Syntax/lexicon.ML"; ML_file "Syntax/ast.ML"; ML_file "Syntax/syntax_ext.ML"; ML_file "Syntax/parser.ML"; ML_file "Syntax/syntax_trans.ML"; ML_file "Syntax/mixfix.ML"; ML_file "Syntax/printer.ML"; ML_file "Syntax/syntax.ML"; subsection "Core of tactical proof system"; ML_file "term_ord.ML"; ML_file "term_subst.ML"; ML_file "General/completion.ML"; ML_file "General/name_space.ML"; ML_file "sorts.ML"; ML_file "type.ML"; ML_file "logic.ML"; ML_file "Syntax/simple_syntax.ML"; ML_file "net.ML"; ML_file "item_net.ML"; ML_file "envir.ML"; ML_file "consts.ML"; ML_file "term_xml.ML"; ML_file "primitive_defs.ML"; ML_file "sign.ML"; ML_file "defs.ML"; ML_file "term_sharing.ML"; ML_file "pattern.ML"; ML_file "unify.ML"; ML_file "theory.ML"; ML_file "proofterm.ML"; ML_file "thm.ML"; ML_file "more_pattern.ML"; ML_file "more_unify.ML"; ML_file "more_thm.ML"; ML_file "facts.ML"; ML_file "thm_name.ML"; ML_file "global_theory.ML"; ML_file "pure_thy.ML"; ML_file "drule.ML"; ML_file "morphism.ML"; ML_file "variable.ML"; ML_file "conv.ML"; ML_file "goal_display.ML"; ML_file "tactical.ML"; ML_file "search.ML"; ML_file "tactic.ML"; ML_file "raw_simplifier.ML"; ML_file "conjunction.ML"; ML_file "assumption.ML"; subsection "Isar -- Intelligible Semi-Automated Reasoning"; (*ML support and global execution*) ML_file "ML/ml_syntax.ML"; ML_file "ML/ml_env.ML"; ML_file "ML/ml_options.ML"; ML_file "ML/ml_print_depth.ML"; ML_file_no_debug "Isar/runtime.ML"; ML_file "PIDE/execution.ML"; ML_file "ML/ml_compiler.ML"; ML_file "skip_proof.ML"; ML_file "goal.ML"; (*outer syntax*) ML_file "Isar/keyword.ML"; ML_file "Isar/token.ML"; ML_file "Isar/parse.ML"; ML_file "Thy/document_source.ML"; ML_file "Thy/thy_header.ML"; ML_file "Thy/document_marker.ML"; (*proof context*) ML_file "Isar/object_logic.ML"; ML_file "Isar/rule_cases.ML"; ML_file "Isar/auto_bind.ML"; ML_file "type_infer.ML"; ML_file "Syntax/local_syntax.ML"; ML_file "Isar/proof_context.ML"; ML_file "type_infer_context.ML"; ML_file "Syntax/syntax_phases.ML"; ML_file "Isar/args.ML"; (*theory specifications*) ML_file "Isar/local_defs.ML"; ML_file "Isar/local_theory.ML"; ML_file "Isar/entity.ML"; ML_file "PIDE/command_span.ML"; ML_file "Thy/thy_element.ML"; ML_file "Thy/markdown.ML"; ML_file "Thy/html.ML"; ML_file "Thy/latex.ML"; (*ML with context and antiquotations*) ML_file "ML/ml_context.ML"; ML_file "ML/ml_antiquotation.ML"; ML_file "ML/ml_compiler2.ML"; ML_file "ML/ml_pid.ML"; section "Bootstrap phase 3: towards theory Pure and final ML toplevel setup"; (*basic proof engine*) ML_file "par_tactical.ML"; ML_file "context_tactic.ML"; ML_file "Isar/proof_display.ML"; ML_file "Isar/attrib.ML"; ML_file "Isar/context_rules.ML"; ML_file "Isar/method.ML"; ML_file "Isar/proof.ML"; ML_file "Isar/element.ML"; ML_file "Isar/obtain.ML"; ML_file "Isar/subgoal.ML"; ML_file "Isar/calculation.ML"; (*local theories and targets*) ML_file "Isar/locale.ML"; ML_file "Isar/generic_target.ML"; ML_file "Isar/bundle.ML"; ML_file "Isar/overloading.ML"; ML_file "axclass.ML"; ML_file "Isar/class.ML"; ML_file "Isar/named_target.ML"; ML_file "Isar/expression.ML"; ML_file "Isar/interpretation.ML"; ML_file "Isar/class_declaration.ML"; ML_file "Isar/target_context.ML"; ML_file "Isar/experiment.ML"; ML_file "simplifier.ML"; ML_file "Tools/plugin.ML"; (*executable theory content*) ML_file "Isar/code.ML"; (*specifications*) ML_file "Isar/spec_rules.ML"; ML_file "Isar/specification.ML"; ML_file "Isar/parse_spec.ML"; ML_file "Isar/typedecl.ML"; (*toplevel transactions*) ML_file "Isar/proof_node.ML"; ML_file "Isar/toplevel.ML"; (*proof term operations*) ML_file "Proof/proof_rewrite_rules.ML"; ML_file "Proof/proof_syntax.ML"; ML_file "Proof/proof_checker.ML"; ML_file "Proof/extraction.ML"; (*Isabelle system*) ML_file "System/bash.ML"; ML_file "System/isabelle_system.ML"; (*theory documents*) ML_file "Thy/term_style.ML"; ML_file "Isar/outer_syntax.ML"; ML_file "ML/ml_antiquotations.ML"; ML_file "Thy/document_antiquotation.ML"; ML_file "Thy/thy_output.ML"; ML_file "Thy/document_antiquotations.ML"; ML_file "General/graph_display.ML"; ML_file "pure_syn.ML"; ML_file "PIDE/command.ML"; ML_file "PIDE/query_operation.ML"; ML_file "PIDE/resources.ML"; -ML_file "Thy/present.ML"; ML_file "Thy/thy_info.ML"; ML_file "thm_deps.ML"; ML_file "Thy/export_theory.ML"; ML_file "Thy/sessions.ML"; ML_file "PIDE/session.ML"; ML_file "PIDE/document.ML"; (*theory and proof operations*) ML_file "Isar/isar_cmd.ML"; subsection "Isabelle/Isar system"; ML_file "System/command_line.ML"; ML_file "System/message_channel.ML"; ML_file "System/isabelle_process.ML"; ML_file "System/scala.ML"; ML_file "System/scala_compiler.ML"; ML_file "Thy/bibtex.ML"; ML_file "PIDE/protocol.ML"; ML_file "General/output_primitives_virtual.ML"; subsection "Miscellaneous tools and packages for Pure Isabelle"; ML_file "ML/ml_pp.ML"; ML_file "ML/ml_thms.ML"; ML_file "ML/ml_file.ML"; ML_file "Tools/build.ML"; ML_file "Tools/named_thms.ML"; ML_file "Tools/print_operation.ML"; ML_file "Tools/rail.ML"; ML_file "Tools/rule_insts.ML"; ML_file "Tools/thy_deps.ML"; ML_file "Tools/class_deps.ML"; ML_file "Tools/find_theorems.ML"; ML_file "Tools/find_consts.ML"; ML_file "Tools/simplifier_trace.ML"; ML_file_no_debug "Tools/debugger.ML"; ML_file "Tools/named_theorems.ML"; ML_file "Tools/jedit.ML"; ML_file "Tools/ghc.ML"; ML_file "Tools/generated_files.ML" diff --git a/src/Pure/System/isabelle_tool.scala b/src/Pure/System/isabelle_tool.scala --- a/src/Pure/System/isabelle_tool.scala +++ b/src/Pure/System/isabelle_tool.scala @@ -1,188 +1,188 @@ /* Title: Pure/System/isabelle_tool.scala Author: Makarius Author: Lars Hupel Isabelle system tools: external executables or internal Scala functions. */ package isabelle import java.net.URLClassLoader import scala.reflect.runtime.universe import scala.tools.reflect.{ToolBox,ToolBoxError} object Isabelle_Tool { /* Scala source tools */ abstract class Body extends Function[List[String], Unit] private def compile(path: Path): Body = { def err(msg: String): Nothing = cat_error(msg, "The error(s) above occurred in Isabelle/Scala tool " + path) val source = File.read(path) val class_loader = new URLClassLoader(Array(), getClass.getClassLoader) val tool_box = universe.runtimeMirror(class_loader).mkToolBox() try { val symbol = tool_box.parse(source) match { case tree: universe.ModuleDef => tool_box.define(tree) case _ => err("Source does not describe a module (Scala object)") } tool_box.compile(universe.Ident(symbol))() match { case body: Body => body case _ => err("Ill-typed source: Isabelle_Tool.Body expected") } } catch { case e: ToolBoxError => if (tool_box.frontEnd.hasErrors) { val infos = tool_box.frontEnd.infos.toList val msgs = infos.map(info => "Error in line " + info.pos.line + ":\n" + info.msg) err(msgs.mkString("\n")) } else err(e.toString) } } /* external tools */ private def dirs(): List[Path] = Path.split(Isabelle_System.getenv_strict("ISABELLE_TOOLS")) private def is_external(dir: Path, file_name: String): Boolean = { val file = (dir + Path.explode(file_name)).file try { file.isFile && file.canRead && (file_name.endsWith(".scala") || file.canExecute) && !file_name.endsWith("~") && !file_name.endsWith(".orig") } catch { case _: SecurityException => false } } private def list_external(): List[(String, String)] = { val Description = """.*\bDESCRIPTION: *(.*)""".r for { dir <- dirs() if dir.is_dir file_name <- File.read_dir(dir) if is_external(dir, file_name) } yield { val source = File.read(dir + Path.explode(file_name)) val name = Library.perhaps_unsuffix(".scala", file_name) val description = split_lines(source).collectFirst({ case Description(s) => s }) getOrElse "" (name, description) } } private def find_external(name: String): Option[List[String] => Unit] = dirs().collectFirst({ case dir if is_external(dir, name + ".scala") => compile(dir + Path.explode(name + ".scala")) case dir if is_external(dir, name) => (args: List[String]) => { val tool = dir + Path.explode(name) val result = Isabelle_System.bash(File.bash_path(tool) + " " + Bash.strings(args)) sys.exit(result.print_stdout.rc) } }) /* internal tools */ private lazy val internal_tools: List[Isabelle_Tool] = Isabelle_System.make_services(classOf[Isabelle_Scala_Tools]).flatMap(_.tools) private def list_internal(): List[(String, String)] = for (tool <- internal_tools.toList) yield (tool.name, tool.description) private def find_internal(name: String): Option[List[String] => Unit] = internal_tools.collectFirst({ case tool if tool.name == name => args => Command_Line.tool { tool.body(args) } }) /* command line entry point */ def main(args: Array[String]) { Command_Line.tool { args.toList match { case Nil | List("-?") => val tool_descriptions = (list_external() ::: list_internal()).sortBy(_._1). map({ case (a, "") => a case (a, b) => a + " - " + b }) Getopts(""" Usage: isabelle TOOL [ARGS ...] Start Isabelle TOOL with ARGS; pass "-?" for tool-specific help. Available tools:""" + tool_descriptions.mkString("\n ", "\n ", "\n")).usage case tool_name :: tool_args => find_external(tool_name) orElse find_internal(tool_name) match { case Some(tool) => tool(tool_args) case None => error("Unknown Isabelle tool: " + quote(tool_name)) } } } } } sealed case class Isabelle_Tool(name: String, description: String, body: List[String] => Unit) class Isabelle_Scala_Tools(val tools: Isabelle_Tool*) extends Isabelle_System.Service class Tools extends Isabelle_Scala_Tools( Build.isabelle_tool, Build_Docker.isabelle_tool, Doc.isabelle_tool, Dump.isabelle_tool, Export.isabelle_tool, ML_Process.isabelle_tool, Mercurial.isabelle_tool, Mkroot.isabelle_tool, Options.isabelle_tool, Phabricator.isabelle_tool1, Phabricator.isabelle_tool2, Phabricator.isabelle_tool3, Phabricator.isabelle_tool4, - Present.isabelle_tool, + Presentation.isabelle_tool, Profiling_Report.isabelle_tool, Server.isabelle_tool, Sessions.isabelle_tool, Scala_Project.isabelle_tool, Update.isabelle_tool, Update_Cartouches.isabelle_tool, Update_Comments.isabelle_tool, Update_Header.isabelle_tool, Update_Then.isabelle_tool, Update_Theorems.isabelle_tool, isabelle.vscode.Grammar.isabelle_tool, isabelle.vscode.Server.isabelle_tool) class Admin_Tools extends Isabelle_Scala_Tools( Build_CSDP.isabelle_tool, Build_Cygwin.isabelle_tool, Build_Doc.isabelle_tool, Build_E.isabelle_tool, Build_Fonts.isabelle_tool, Build_JDK.isabelle_tool, Build_PolyML.isabelle_tool1, Build_PolyML.isabelle_tool2, Build_SPASS.isabelle_tool, Build_SQLite.isabelle_tool, Build_Status.isabelle_tool, Build_VeriT.isabelle_tool, Build_Zipperposition.isabelle_tool, Check_Sources.isabelle_tool, Components.isabelle_tool, isabelle.vscode.Build_VSCode.isabelle_tool) diff --git a/src/Pure/Thy/bibtex.scala b/src/Pure/Thy/bibtex.scala --- a/src/Pure/Thy/bibtex.scala +++ b/src/Pure/Thy/bibtex.scala @@ -1,699 +1,699 @@ /* Title: Pure/Thy/bibtex.scala Author: Makarius BibTeX support. */ package isabelle import java.io.{File => JFile} import scala.collection.mutable import scala.util.parsing.combinator.RegexParsers import scala.util.parsing.input.Reader object Bibtex { /** file format **/ def is_bibtex(name: String): Boolean = name.endsWith(".bib") class File_Format extends isabelle.File_Format { val format_name: String = "bibtex" val file_ext: String = "bib" override def theory_suffix: String = "bibtex_file" override def theory_content(name: String): String = """theory "bib" imports Pure begin bibtex_file """ + quote(name) + """ end""" - override def make_preview(snapshot: Document.Snapshot): Option[Present.Preview] = + override def make_preview(snapshot: Document.Snapshot): Option[Presentation.Preview] = { val name = snapshot.node_name if (detect(name.node)) { val title = "Bibliography " + quote(snapshot.node_name.path.file_name) val content = Isabelle_System.with_tmp_file("bib", "bib") { bib => File.write(bib, snapshot.node.source) Bibtex.html_output(List(bib), style = "unsort", title = title) } - Some(Present.Preview(title, content)) + Some(Presentation.Preview(title, content)) } else None } } /** bibtex errors **/ def bibtex_errors(dir: Path, root_name: String): List[String] = { val log_path = dir + Path.explode(root_name).ext("blg") if (log_path.is_file) { val Error1 = """^(I couldn't open database file .+)$""".r val Error2 = """^(.+)---line (\d+) of file (.+)""".r Line.logical_lines(File.read(log_path)).flatMap(line => line match { case Error1(msg) => Some("Bibtex error: " + msg) case Error2(msg, Value.Int(l), file) => val path = File.standard_path(file) if (Path.is_wellformed(path)) { val pos = Position.Line_File(l, (dir + Path.explode(path)).canonical.implode) Some("Bibtex error" + Position.here(pos) + ":\n " + msg) } else None case _ => None }) } else Nil } /** check database **/ def check_database(database: String): (List[(String, Position.T)], List[(String, Position.T)]) = { val chunks = parse(Line.normalize(database)) var chunk_pos = Map.empty[String, Position.T] val tokens = new mutable.ListBuffer[(Token, Position.T)] var line = 1 var offset = 1 def make_pos(length: Int): Position.T = Position.Offset(offset) ::: Position.End_Offset(offset + length) ::: Position.Line(line) def advance_pos(tok: Token) { for (s <- Symbol.iterator(tok.source)) { if (Symbol.is_newline(s)) line += 1 offset += 1 } } def get_line_pos(l: Int): Position.T = if (0 < l && l <= tokens.length) tokens(l - 1)._2 else Position.none for (chunk <- chunks) { val name = chunk.name if (name != "" && !chunk_pos.isDefinedAt(name)) { chunk_pos += (name -> make_pos(chunk.heading_length)) } for (tok <- chunk.tokens) { tokens += (tok.copy(source = tok.source.replace("\n", " ")) -> make_pos(tok.source.length)) advance_pos(tok) } } Isabelle_System.with_tmp_dir("bibtex")(tmp_dir => { File.write(tmp_dir + Path.explode("root.bib"), tokens.iterator.map(p => p._1.source).mkString("", "\n", "\n")) File.write(tmp_dir + Path.explode("root.aux"), "\\bibstyle{plain}\n\\bibdata{root}\n\\citation{*}") Isabelle_System.bash("\"$ISABELLE_BIBTEX\" root", cwd = tmp_dir.file) val Error = """^(.*)---line (\d+) of file root.bib$""".r val Warning = """^Warning--(.+)$""".r val Warning_Line = """--line (\d+) of file root.bib$""".r val Warning_in_Chunk = """^Warning--(.+) in (.+)$""".r val log_file = tmp_dir + Path.explode("root.blg") val lines = if (log_file.is_file) Line.logical_lines(File.read(log_file)) else Nil val (errors, warnings) = if (lines.isEmpty) (Nil, Nil) else { lines.zip(lines.tail ::: List("")).flatMap( { case (Error(msg, Value.Int(l)), _) => Some((true, (msg, get_line_pos(l)))) case (Warning_in_Chunk(msg, name), _) if chunk_pos.isDefinedAt(name) => Some((false, (Word.capitalize(msg + " in entry " + quote(name)), chunk_pos(name)))) case (Warning(msg), Warning_Line(Value.Int(l))) => Some((false, (Word.capitalize(msg), get_line_pos(l)))) case (Warning(msg), _) => Some((false, (Word.capitalize(msg), Position.none))) case _ => None } ).partition(_._1) } (errors.map(_._2), warnings.map(_._2)) }) } object Check_Database extends Scala.Fun("bibtex_check_database") { def apply(database: String): String = { import XML.Encode._ YXML.string_of_body(pair(list(pair(string, properties)), list(pair(string, properties)))( check_database(database))) } } /** document model **/ /* entries */ def entries(text: String): List[Text.Info[String]] = { val result = new mutable.ListBuffer[Text.Info[String]] var offset = 0 for (chunk <- Bibtex.parse(text)) { val end_offset = offset + chunk.source.length if (chunk.name != "" && !chunk.is_command) result += Text.Info(Text.Range(offset, end_offset), chunk.name) offset = end_offset } result.toList } def entries_iterator[A, B <: Document.Model](models: Map[A, B]) : Iterator[Text.Info[(String, B)]] = { for { (_, model) <- models.iterator info <- model.bibtex_entries.iterator } yield info.map((_, model)) } /* completion */ def completion[A, B <: Document.Model]( history: Completion.History, rendering: Rendering, caret: Text.Offset, models: Map[A, B]): Option[Completion.Result] = { for { Text.Info(r, name) <- rendering.citation(rendering.before_caret_range(caret)) name1 <- Completion.clean_name(name) original <- rendering.model.get_text(r) original1 <- Completion.clean_name(Library.perhaps_unquote(original)) entries = (for { Text.Info(_, (entry, _)) <- entries_iterator(models) if entry.toLowerCase.containsSlice(name1.toLowerCase) && entry != original1 } yield entry).toList if entries.nonEmpty items = entries.sorted.map({ case entry => val full_name = Long_Name.qualify(Markup.CITATION, entry) val description = List(entry, "(BibTeX entry)") val replacement = quote(entry) Completion.Item(r, original, full_name, description, replacement, 0, false) }).sorted(history.ordering).take(rendering.options.int("completion_limit")) } yield Completion.Result(r, original, false, items) } /** content **/ private val months = List( "jan", "feb", "mar", "apr", "may", "jun", "jul", "aug", "sep", "oct", "nov", "dec") def is_month(s: String): Boolean = months.contains(s.toLowerCase) private val commands = List("preamble", "string") def is_command(s: String): Boolean = commands.contains(s.toLowerCase) sealed case class Entry( kind: String, required: List[String], optional_crossref: List[String], optional_other: List[String]) { val optional_standard: List[String] = List("url", "doi", "ee") def is_required(s: String): Boolean = required.contains(s.toLowerCase) def is_optional(s: String): Boolean = optional_crossref.contains(s.toLowerCase) || optional_other.contains(s.toLowerCase) || optional_standard.contains(s.toLowerCase) def fields: List[String] = required ::: optional_crossref ::: optional_other ::: optional_standard def template: String = "@" + kind + "{,\n" + fields.map(x => " " + x + " = {},\n").mkString + "}\n" } val known_entries: List[Entry] = List( Entry("Article", List("author", "title"), List("journal", "year"), List("volume", "number", "pages", "month", "note")), Entry("InProceedings", List("author", "title"), List("booktitle", "year"), List("editor", "volume", "number", "series", "pages", "month", "address", "organization", "publisher", "note")), Entry("InCollection", List("author", "title", "booktitle"), List("publisher", "year"), List("editor", "volume", "number", "series", "type", "chapter", "pages", "edition", "month", "address", "note")), Entry("InBook", List("author", "editor", "title", "chapter"), List("publisher", "year"), List("volume", "number", "series", "type", "address", "edition", "month", "pages", "note")), Entry("Proceedings", List("title", "year"), List(), List("booktitle", "editor", "volume", "number", "series", "address", "month", "organization", "publisher", "note")), Entry("Book", List("author", "editor", "title"), List("publisher", "year"), List("volume", "number", "series", "address", "edition", "month", "note")), Entry("Booklet", List("title"), List(), List("author", "howpublished", "address", "month", "year", "note")), Entry("PhdThesis", List("author", "title", "school", "year"), List(), List("type", "address", "month", "note")), Entry("MastersThesis", List("author", "title", "school", "year"), List(), List("type", "address", "month", "note")), Entry("TechReport", List("author", "title", "institution", "year"), List(), List("type", "number", "address", "month", "note")), Entry("Manual", List("title"), List(), List("author", "organization", "address", "edition", "month", "year", "note")), Entry("Unpublished", List("author", "title", "note"), List(), List("month", "year")), Entry("Misc", List(), List(), List("author", "title", "howpublished", "month", "year", "note"))) def get_entry(kind: String): Option[Entry] = known_entries.find(entry => entry.kind.toLowerCase == kind.toLowerCase) def is_entry(kind: String): Boolean = get_entry(kind).isDefined /** tokens and chunks **/ object Token { object Kind extends Enumeration { val COMMAND = Value("command") val ENTRY = Value("entry") val KEYWORD = Value("keyword") val NAT = Value("natural number") val STRING = Value("string") val NAME = Value("name") val IDENT = Value("identifier") val SPACE = Value("white space") val COMMENT = Value("ignored text") val ERROR = Value("bad input") } } sealed case class Token(kind: Token.Kind.Value, source: String) { def is_kind: Boolean = kind == Token.Kind.COMMAND || kind == Token.Kind.ENTRY || kind == Token.Kind.IDENT def is_name: Boolean = kind == Token.Kind.NAME || kind == Token.Kind.IDENT def is_ignored: Boolean = kind == Token.Kind.SPACE || kind == Token.Kind.COMMENT def is_malformed: Boolean = kind == Token.Kind.ERROR def is_open: Boolean = kind == Token.Kind.KEYWORD && (source == "{" || source == "(") } case class Chunk(kind: String, tokens: List[Token]) { val source = tokens.map(_.source).mkString private val content: Option[List[Token]] = tokens match { case Token(Token.Kind.KEYWORD, "@") :: body if body.nonEmpty => (body.init.filterNot(_.is_ignored), body.last) match { case (tok :: Token(Token.Kind.KEYWORD, "{") :: toks, Token(Token.Kind.KEYWORD, "}")) if tok.is_kind => Some(toks) case (tok :: Token(Token.Kind.KEYWORD, "(") :: toks, Token(Token.Kind.KEYWORD, ")")) if tok.is_kind => Some(toks) case _ => None } case _ => None } def name: String = content match { case Some(tok :: _) if tok.is_name => tok.source case _ => "" } def heading_length: Int = if (name == "") 1 else (0 /: tokens.takeWhile(tok => !tok.is_open)){ case (n, tok) => n + tok.source.length } def is_ignored: Boolean = kind == "" && tokens.forall(_.is_ignored) def is_malformed: Boolean = kind == "" || tokens.exists(_.is_malformed) def is_command: Boolean = Bibtex.is_command(kind) && name != "" && content.isDefined def is_entry: Boolean = Bibtex.is_entry(kind) && name != "" && content.isDefined } /** parsing **/ // context of partial line-oriented scans abstract class Line_Context case object Ignored extends Line_Context case object At extends Line_Context case class Item_Start(kind: String) extends Line_Context case class Item_Open(kind: String, end: String) extends Line_Context case class Item(kind: String, end: String, delim: Delimited) extends Line_Context case class Delimited(quoted: Boolean, depth: Int) val Closed = Delimited(false, 0) private def token(kind: Token.Kind.Value)(source: String): Token = Token(kind, source) private def keyword(source: String): Token = Token(Token.Kind.KEYWORD, source) // See also https://ctan.org/tex-archive/biblio/bibtex/base/bibtex.web // module @. object Parsers extends RegexParsers { /* white space and comments */ override val whiteSpace = "".r private val space = """[ \t\n\r]+""".r ^^ token(Token.Kind.SPACE) private val spaces = rep(space) /* ignored text */ private val ignored: Parser[Chunk] = rep1("""(?i)([^@]+|@[ \t]*comment)""".r) ^^ { case ss => Chunk("", List(Token(Token.Kind.COMMENT, ss.mkString))) } private def ignored_line: Parser[(Chunk, Line_Context)] = ignored ^^ { case a => (a, Ignored) } /* delimited string: outermost "..." or {...} and body with balanced {...} */ // see also bibtex.web: scan_a_field_token_and_eat_white, scan_balanced_braces private def delimited_depth(delim: Delimited): Parser[(String, Delimited)] = new Parser[(String, Delimited)] { require(if (delim.quoted) delim.depth > 0 else delim.depth >= 0) def apply(in: Input) = { val start = in.offset val end = in.source.length var i = start var q = delim.quoted var d = delim.depth var finished = false while (!finished && i < end) { val c = in.source.charAt(i) if (c == '"' && d == 0) { i += 1; d = 1; q = true } else if (c == '"' && d == 1 && q) { i += 1; d = 0; q = false; finished = true } else if (c == '{') { i += 1; d += 1 } else if (c == '}') { if (d == 1 && !q || d > 1) { i += 1; d -= 1; if (d == 0) finished = true } else {i = start; finished = true } } else if (d > 0) i += 1 else finished = true } if (i == start) Failure("bad input", in) else { val s = in.source.subSequence(start, i).toString Success((s, Delimited(q, d)), in.drop(i - start)) } } }.named("delimited_depth") private def delimited: Parser[Token] = delimited_depth(Closed) ^? { case (s, delim) if delim == Closed => Token(Token.Kind.STRING, s) } private def recover_delimited: Parser[Token] = """["{][^@]*""".r ^^ token(Token.Kind.ERROR) def delimited_line(ctxt: Item): Parser[(Chunk, Line_Context)] = delimited_depth(ctxt.delim) ^^ { case (s, delim1) => (Chunk(ctxt.kind, List(Token(Token.Kind.STRING, s))), ctxt.copy(delim = delim1)) } | recover_delimited ^^ { case a => (Chunk(ctxt.kind, List(a)), Ignored) } /* other tokens */ private val at = "@" ^^ keyword private val nat = "[0-9]+".r ^^ token(Token.Kind.NAT) private val name = """[\x21-\x7f&&[^"#%'(),={}]]+""".r ^^ token(Token.Kind.NAME) private val identifier = """[\x21-\x7f&&[^"#%'(),={}0-9]][\x21-\x7f&&[^"#%'(),={}]]*""".r private val ident = identifier ^^ token(Token.Kind.IDENT) val other_token = "[=#,]".r ^^ keyword | (nat | (ident | space)) /* body */ private val body = delimited | (recover_delimited | other_token) private def body_line(ctxt: Item) = if (ctxt.delim.depth > 0) delimited_line(ctxt) else delimited_line(ctxt) | other_token ^^ { case a => (Chunk(ctxt.kind, List(a)), ctxt) } | ctxt.end ^^ { case a => (Chunk(ctxt.kind, List(keyword(a))), Ignored) } /* items: command or entry */ private val item_kind = identifier ^^ { case a => val kind = if (is_command(a)) Token.Kind.COMMAND else if (is_entry(a)) Token.Kind.ENTRY else Token.Kind.IDENT Token(kind, a) } private val item_begin = "{" ^^ { case a => ("}", keyword(a)) } | "(" ^^ { case a => (")", keyword(a)) } private def item_name(kind: String) = kind.toLowerCase match { case "preamble" => failure("") case "string" => identifier ^^ token(Token.Kind.NAME) case _ => name } private val item_start = at ~ spaces ~ item_kind ~ spaces ^^ { case a ~ b ~ c ~ d => (c.source, List(a) ::: b ::: List(c) ::: d) } private val item: Parser[Chunk] = (item_start ~ item_begin ~ spaces) into { case (kind, a) ~ ((end, b)) ~ c => opt(item_name(kind)) ~ rep(body) ~ opt(end ^^ keyword) ^^ { case d ~ e ~ f => Chunk(kind, a ::: List(b) ::: c ::: d.toList ::: e ::: f.toList) } } private val recover_item: Parser[Chunk] = at ~ "[^@]*".r ^^ { case a ~ b => Chunk("", List(a, Token(Token.Kind.ERROR, b))) } /* chunks */ val chunk: Parser[Chunk] = ignored | (item | recover_item) def chunk_line(ctxt: Line_Context): Parser[(Chunk, Line_Context)] = { ctxt match { case Ignored => ignored_line | at ^^ { case a => (Chunk("", List(a)), At) } case At => space ^^ { case a => (Chunk("", List(a)), ctxt) } | item_kind ^^ { case a => (Chunk(a.source, List(a)), Item_Start(a.source)) } | recover_item ^^ { case a => (a, Ignored) } | ignored_line case Item_Start(kind) => space ^^ { case a => (Chunk(kind, List(a)), ctxt) } | item_begin ^^ { case (end, a) => (Chunk(kind, List(a)), Item_Open(kind, end)) } | recover_item ^^ { case a => (a, Ignored) } | ignored_line case Item_Open(kind, end) => space ^^ { case a => (Chunk(kind, List(a)), ctxt) } | item_name(kind) ^^ { case a => (Chunk(kind, List(a)), Item(kind, end, Closed)) } | body_line(Item(kind, end, Closed)) | ignored_line case item_ctxt: Item => body_line(item_ctxt) | ignored_line case _ => failure("") } } } /* parse */ def parse(input: CharSequence): List[Chunk] = Parsers.parseAll(Parsers.rep(Parsers.chunk), Scan.char_reader(input)) match { case Parsers.Success(result, _) => result case _ => error("Unexpected failure to parse input:\n" + input.toString) } def parse_line(input: CharSequence, context: Line_Context): (List[Chunk], Line_Context) = { var in: Reader[Char] = Scan.char_reader(input) val chunks = new mutable.ListBuffer[Chunk] var ctxt = context while (!in.atEnd) { Parsers.parse(Parsers.chunk_line(ctxt), in) match { case Parsers.Success((x, c), rest) => chunks += x; ctxt = c; in = rest case Parsers.NoSuccess(_, rest) => error("Unepected failure to parse input:\n" + rest.source.toString) } } (chunks.toList, ctxt) } /** HTML output **/ private val output_styles = List( "" -> "html-n", "plain" -> "html-n", "alpha" -> "html-a", "named" -> "html-n", "paragraph" -> "html-n", "unsort" -> "html-u", "unsortlist" -> "html-u") def html_output(bib: List[Path], title: String = "Bibliography", body: Boolean = false, citations: List[String] = List("*"), style: String = "", chronological: Boolean = false): String = { Isabelle_System.with_tmp_dir("bibtex")(tmp_dir => { /* database files */ val bib_files = bib.map(_.drop_ext) val bib_names = { val names0 = bib_files.map(_.file_name) if (Library.duplicates(names0).isEmpty) names0 else names0.zipWithIndex.map({ case (name, i) => (i + 1).toString + "-" + name }) } for ((a, b) <- bib_files zip bib_names) { File.copy(a.ext("bib"), tmp_dir + Path.basic(b).ext("bib")) } /* style file */ val bst = output_styles.toMap.get(style) match { case Some(base) => base + (if (chronological) "c" else "") + ".bst" case None => error("Bad style for bibtex HTML output: " + quote(style) + "\n(expected: " + commas_quote(output_styles.map(_._1)) + ")") } File.copy(Path.explode("$BIB2XHTML_HOME/bst") + Path.explode(bst), tmp_dir) /* result */ val in_file = Path.explode("bib.aux") val out_file = Path.explode("bib.html") File.write(tmp_dir + in_file, bib_names.mkString("\\bibdata{", ",", "}\n") + citations.map(cite => "\\citation{" + cite + "}\n").mkString) Isabelle_System.bash( "\"$BIB2XHTML_HOME/main/bib2xhtml.pl\" -B \"$ISABELLE_BIBTEX\"" + " -u -s " + Bash.string(proper_string(style) getOrElse "empty") + (if (chronological) " -c" else "") + (if (title != "") " -h " + Bash.string(title) + " " else "") + " " + File.bash_path(in_file) + " " + File.bash_path(out_file), cwd = tmp_dir.file).check val html = File.read(tmp_dir + out_file) if (body) { cat_lines( split_lines(html). dropWhile(line => !line.startsWith("