# HG changeset patch # User wenzelm # Date 1521665428 -3600 # Node ID d13b2dd20f5ef7c658d5f75b0fb7870968ff0369 # Parent a72f01c63262a59f4e834b79d27d1b9b0e6129c0 more documentation; diff -r a72f01c63262 -r d13b2dd20f5e src/Doc/System/Server.thy --- a/src/Doc/System/Server.thy Wed Mar 21 21:31:40 2018 +0100 +++ b/src/Doc/System/Server.thy Wed Mar 21 21:50:28 2018 +0100 @@ -224,7 +224,7 @@ \ -subsection \Input and output messages\ +subsection \Input and output messages \label{sec:input-output-messages}\ text \ Server input and output messages have a uniform format as follows: @@ -346,9 +346,677 @@ asynchronous messages properly. The synchronous command \<^verbatim>\cancel\~\id\ tells the specified task to terminate - prematurely: usually causing a \<^verbatim>\FAILED\ result with error message - \<^verbatim>\Interrupt\, but this is not guaranteed: the cancel event may come too - late or the task may just ignore it. + prematurely: usually causing a \<^verbatim>\FAILED\ result, but this is not guaranteed: + the cancel event may come too late or the task may just ignore it. +\ + + +section \Server commands and results\ + +text \ + Here follows an overview of particular Isabelle server commands with their + results, which are usually represented as JSON values. The general format of + input and output messages is described in + \secref{sec:input-output-messages}. The relevant Isabelle/Scala source files + are: + + \<^medskip> + \begin{tabular}{l} + \<^file>\$ISABELLE_HOME/src/Pure/Tools/server_commands.scala\ \\ + \<^file>\$ISABELLE_HOME/src/Pure/Tools/server.scala\ + \end{tabular} +\ + + +subsection \Types for JSON values\ + +text \ + In order to specify concrete JSON types for command arguments and result + messages, the following type definition language shall be used: + + \<^rail>\ + @{syntax type_def}: @'type' @{syntax name} '=' @{syntax type} + ; + @{syntax type}: @{syntax name} | @{syntax value} | 'any' | 'null' | + 'bool' | 'int' | 'long' | 'double' | 'string' | '[' @{syntax type} ']' | + '{' (@{syntax field_type} ',' *) '}' | + @{syntax type} '\' @{syntax type} | + @{syntax type} '|' @{syntax type} + ; + @{syntax field_type}: @{syntax name} ('?'?) ':' @{syntax type} + \ + + This is a simplified version of interfaces in + TypeScript.\<^footnote>\\<^url>\https://www.typescriptlang.org/docs/handbook/interfaces.html\\ + The meaning of these types is specified wrt. the implementation in + Isabelle/Scala as follows. + + \<^item> A \name\ refers to a type defined elsewhere. The environment of type + definitions is given informally: put into proper foundational order, it + needs to specify a strongly normalizing system; type definitions may not be + recursive. + + \<^item> A \value\ in JSON notation represents the singleton type of the given + item. For example, the string \<^verbatim>\"error"\ can be used is the type for a slot + that is guaranteed to contain that constant. + + \<^item> Type \any\ is the super type of all other types: it is an untyped slot in + the specification and corresponds to \<^verbatim>\Any\ or \<^verbatim>\JSON.T\ in Isabelle/Scala. + + \<^item> Type \null\ is the type of the improper value \null\; it corresponds to + type \<^verbatim>\Null\ in Scala and is normally not used in Isabelle/Scala.\<^footnote>\See also + ``Null References: The Billion Dollar Mistake'' by Tony Hoare + \<^url>\https://www.infoq.com/presentations/Null-References-The-Billion-Dollar-Mistake-Tony-Hoare\.\ + + \<^item> Type \bool\ is the type of the truth values \true\ and \false\; it + corresponds to \<^verbatim>\Boolean\ in Scala. + + \<^item> Types \int\, \long\, \double\ are specific versions of the generic + \number\ type, corresponding to \<^verbatim>\Int\, \<^verbatim>\Long\, \<^verbatim>\Double\ in Scala, but + \<^verbatim>\Long\ is limited to 53 bit precision.\<^footnote>\Implementations of JSON typically + standardize \number\ to \<^verbatim>\Double\, which can absorb \<^verbatim>\Int\ faithfully, but + not all of \<^verbatim>\Long\.\ + + \<^item> Type \string\ represents Unicode text; it corresponds to type \<^verbatim>\String\ in + Scala. + + \<^item> Type \[t]\ is the array (or list) type over \t\; it corresponds to + \<^verbatim>\List[t]\ in Scala. The list type is co-variant as usual (i.e.\ monotonic + wrt. the subtype order). + + \<^item> Object types describe the possible content of JSON records, with field + names and types. A question mark after a field name means that it is + optional: in Scala this could refer to an explicit type \<^verbatim>\Option[t]\. For + example, \{a: int, b?: string}\ corresponds to a Scala case class with + arguments \<^verbatim>\a: Int\, \<^verbatim>\b: Option[String]\. + + Alternatively, optional fields can have a default value. If nothing else is + specified, the standard default value is the ``empty'' value of a type, + i.e.\ \<^verbatim>\0\ for the number types, \<^verbatim>\false\ for \bool\, or the empty string, + array, object etc. + + Object types are \<^emph>\permissive\ in the sense that only the specified field + names need to conform to the given types, but unspecified fields may be + present as well. + + \<^item> The type expression \t\<^sub>1 \ t\<^sub>2\ only works for two object types with + disjoint field names: it is the concatenation of the respective @{syntax + field_type} specifications taken together. For example: \{task: string} \ + {ok: bool}\ is the equivalent to \{task: string, ok: bool}\. + + \<^item> The type expression \t\<^sub>1 | t\<^sub>2\ is the disjoint union of two types, either + one of the two cases may occur. + + + These types correspond to JSON values in an obvious manner, which is not + further described here. For example, the JSON array \<^verbatim>\[1, 2, 3]\ conforms to + types \[int]\, \[long]\, \[double]\, \[any]\, \any\. + + Note that JSON objects require field names to be quoted, but the type + language omits quotes for clarity. Thus \<^verbatim>\{"a": 42, "b": "xyz"}\ conforms to + the type \{a: int, b: string}\, for example. + + \<^medskip> + The absence of an argument or result is represented by type \<^verbatim>\Unit\ in + Isabelle/Scala: it is written as empty text in the message \argument\ + (\secref{sec:input-output-messages}). This is not part of the JSON language. + + Server replies have name tags like \<^verbatim>\OK\, \<^verbatim>\ERROR\: these are used literally + together with type specifications to indicate the particular name with the + type of its argument, e.g.\ \<^verbatim>\OK\~\[string]\ for a regular result that is a + list (JSON array) of strings. + + \<^bigskip> + Here are some common type definitions, for use in the subsequent + specifications of command arguments and results. + + \<^item> \<^bold>\type\~\position = {line?: int, offset?: int, end_offset?: int, file?: + string, id?: long}\ describes a source position within Isabelle source text. + Only the \line\ and \file\ fields make immediate sense to external programs. + Detailed \offset\ and \end_offset\ positions are counted according to + Isabelle symbols, see @{ML_type Symbol.symbol} in Isabelle/ML @{cite + "isabelle-implementation"}. The position \id\ belongs to the internal + representation of command transactions in the Isabelle/PIDE protocol. + + \<^item> \<^bold>\type\~\message = {kind?: string, message: string, pos?: position}\ where + the \kind\ provides some hint about the role and importance of the message. + The main message kinds are \<^verbatim>\writeln\ (for regular output), \<^verbatim>\warning\, and + \<^verbatim>\error\. The table \<^verbatim>\Markup.messages\ in Isabelle/Scala defines further + message kinds for more specific applications. + + \<^item> \<^bold>\type\~\error_message = {kind:\~\<^verbatim>\"error"\\, message: string}\ refers to + error messages in particular. These occur routinely with \<^verbatim>\ERROR\ or + \<^verbatim>\FAILED\ replies, but also as initial command syntax errors (which are + omitted in the command specifications below). + + \<^item> \<^bold>\type\~\theory_progress = {kind:\~\<^verbatim>\"writeln"\\, message: string, theory: + string, session: string}\ reports formal progress in loading theories (e.g.\ + when building a session image). Apart from a regular output message, it also + reveals the formal theory name (e.g.\ \<^verbatim>\HOL.Nat\) and session name (e.g.\ + \<^verbatim>\HOL\). Note that some rare theory names lack a proper session prefix, e.g. + theory \<^verbatim>\Main\ in session \<^verbatim>\HOL\. + + \<^item> \<^bold>\type\~\timing = {elapsed: double, cpu: double, gc: double}\ refers to + common Isabelle timing information in seconds, usually with a precision of + three digits after the point (whole milliseconds). + + \<^item> \<^bold>\type\~\uuid = string\ refers to a Universally Unique Identifier (UUID) + as plain text.\<^footnote>\See \<^url>\https://www.ietf.org/rfc/rfc4122.txt\ and + \<^url>\https://docs.oracle.com/javase/8/docs/api/java/util/UUID.html\.\ Such + identifiers are created as private random numbers of the server and only + revealed to the client that creates a certain resource (e.g.\ task or + session). A client may disclose this information for use in a different + client connection: e.g.\ this allows to share sessions between multiple + connections. + + Client commands need to provide syntactically wellformed UUIDs: this is + trivial to achieve by using only identifiers that have been produced by the + server beforehand. + + \<^item> \<^bold>\type\~\task = {task: uuid}\ identifies a newly created asynchronous task + and thus allows the client to control it by the \<^verbatim>\cancel\ command. The same + task identification is included in messages and the final result produced by + this task. +\ + + +subsection \Command \<^verbatim>\help\\ + +text \ + \begin{tabular}{ll} + regular result: & \<^verbatim>\OK\ \[string]\ \\ + \end{tabular} + \<^medskip> + + The \<^verbatim>\help\ command has no argument and returns the list of command names + known to the server. This is occasionally useful for interactive + experimentation (see also @{tool client} in \secref{sec:tool-client}). +\ + + +subsection \Command \<^verbatim>\echo\\ + +text \ + \begin{tabular}{ll} + argument: & \any\ \\ + regular result: & \<^verbatim>\OK\ \any\ \\ + \end{tabular} + \<^medskip> + + The \<^verbatim>\echo\ command is the identity function: it returns its argument as + regular result. This is occasionally useful for testing and interactive + experimentation (see also @{tool client} in \secref{sec:tool-client}). + + The type of \<^verbatim>\echo\ is actually more general than given above: \<^verbatim>\Unit\, + \<^verbatim>\XML.Elem\, \<^verbatim>\JSON.T\ are supported uniformly. Note that \<^verbatim>\XML.Elem\ might + be difficult to type on the console in its YXML syntax + (\secref{sec:yxml-vs-xml}). +\ + + +subsection \Command \<^verbatim>\shutdown\\ + +text \ + \begin{tabular}{ll} + regular result: & \<^verbatim>\OK\ \\ + \end{tabular} + \<^medskip> + + The \<^verbatim>\shutdown\ command has no argument and result value. It forces a + shutdown of the connected server process, stopping all open sessions and + closing the server socket. This may disrupt pending commands on other + connections! + + \<^medskip> + Likewise, the command-line invocation \<^verbatim>\isabelle server -x\ opens a server + connection and issues a \<^verbatim>\shutdown\ command (see also + \secref{sec:tool-server}). +\ + + +subsection \Command \<^verbatim>\cancel\\ + +text \ + \begin{tabular}{ll} + argument: & \uuid\ \\ + regular result: & \<^verbatim>\OK\ \\ + \end{tabular} + \<^medskip> + + The command invocation \<^verbatim>\cancel "\\id\\<^verbatim>\"\ attempts to cancel the specified + task. + + Cancellation is merely a hint that the client prefers an ongoing process to + be stopped. The command always succeeds formally, but it may get ignored by + a task that is still running, or it might refer to a non-existing or + no-longer existing task without producing an error. + + Successful cancellation typically leads to an asynchronous failure of type + \<^verbatim>\FAILED {\\task: string, message: \\<^verbatim>\"Interrupt"}\ --- or a slightly + different message, depending how the task handles the event. +\ + + +subsection \Command \<^verbatim>\session_build\ \label{sec:command-session-build}\ + +text \ + \begin{tabular}{lll} + argument: & \session_build_args\ \\ + immediate result: & \<^verbatim>\OK\ \task\ \\ + notifications: & \<^verbatim>\NOTE\ \task \ (theory_progress | message)\ \\ + regular result: & \<^verbatim>\FINISHED\ \task \ session_build_results\ \\ + error result: & \<^verbatim>\FAILED\ \task \ error_message \ session_build_results\ \\[2ex] + \end{tabular} + + \begin{tabular}{lll} + \<^bold>\type\ \session_build_args =\ \\ + \quad\{session: string,\ \\ + \quad~~\preferences?: string,\ & \<^bold>\default:\ server preferences \\ + \quad~~\options?: [string],\ \\ + \quad~~\dirs?: [string],\ \\ + \quad~~\ancestor_session?: string,\ \\ + \quad~~\all_known?: bool,\ \\ + \quad~~\focus_session?: bool,\ \\ + \quad~~\required_session?: bool,\ \\ + \quad~~\system_mode?: bool,\ \\ + \quad~~\verbose?: bool}\ \\[2ex] + + \<^bold>\type\ \session_build_results =\ \\ + \quad\{ok: bool,\ \\ + \quad~~\return_code: int,\ \\ + \quad~~\sessions: [session_build_result]}\ \\[2ex] + + \<^bold>\type\ \session_build_result =\ \\ + \quad\{session: string,\ \\ + \quad~~\ok: bool,\ \\ + \quad~~\return_code: int,\ \\ + \quad~~\timeout: bool,\ \\ + \quad~~\timing: timing}\ \\ + \end{tabular} +\ + +text \ + The \<^verbatim>\session_build\ command prepares given a session image for interactive + use of theories. This is a limited version of command-line tool @{tool + build} (\secref{sec:tool-build}), with specific options to request a formal + context of an interactive session: it also resembles options of @{tool + jedit} @{cite "isabelle-jedit"}. + + The build process is asynchronous, with notifications that inform about the + progress of loaded theories. Some further human-readable messages may be + output as well. + + Coordination of independent build processes is at the discretion of the + client: as for like @{tool build} there are no built-in measures against + conflicting builds with overlapping hierarchies of session images. +\ + + +subsubsection \Arguments\ + +text \ + The field \session\ is mandatory. It specifies the target session name: + either directly (default) or indirectly (if \required_session\ is + \<^verbatim>\true\). + + \<^medskip> + The environment of Isabelle system options is determined from \preferences\ + that are augmented by \options\, which is a list individual updates of the + form the \name\\<^verbatim>\=\\value\ or \name\ (the latter abbreviates + \name\\<^verbatim>\=true\); see also command-line option \<^verbatim>\-o\ for @{tool build}. The + preferences are loaded from the file + \<^path>\$ISABELLE_HOME_USER/etc/preferences\ by default, but the client may + provide alternative contents for it (as text, not a file-name). This could + be relevant in situations where client and server run in different + operating-system contexts. + + \<^medskip> + The \dirs\ field specifies additional directories for session ROOT files + (\secref{sec:session-root}). This augments the name space of available + sessions; see also option \<^verbatim>\-d\ in @{tool build} or @{tool jedit}. + + \<^medskip> + The \ancestor_session\ field specifies an alternative image as starting + point for the target image. The default is to use the parent session + according to the ROOT entry; see also option \<^verbatim>\-A\ in @{tool jedit}. This + can lead to a more light-weight build process, by skipping intermediate + session images of the hierarchy that are not used later on. + + \<^medskip> + The \all_known\ field set to \<^verbatim>\true\ includes all existing sessions from + reachable ROOT entries in the name space of theories with a proper + session-qualified name (instead of referring to the file-system location). + This could be relevant for the \<^verbatim>\use_theories\ command + (\secref{sec:command-use-theories}) to refer to arbitrary theories from + other sessions. Note that considerable time is required to explore all + accessible session directories and theory dependencies. + + \<^medskip> + The \focus_session\ field set to \<^verbatim>\true\ to focuses on the target session: + the accessible name space of theories is restricted to sessions that are + connected to it. + + \<^medskip> + The \required_session\ field set to \<^verbatim>\true\ indicates that the target image + should not be the \session\, but its parent (or ancestor according to + \ancestor_session\). Thus it prepares a session context where theories from + the \session\ itself can be loaded. + + \<^medskip> + The \system_mode\ field may be set to \<^verbatim>\true\ to store session images and + log files in @{path "$ISABELLE_HOME/heaps"} instead of the default location + @{setting ISABELLE_OUTPUT} (which is normally in @{setting + ISABELLE_HOME_USER}, i.e.\ the user's home directory). See also option \<^verbatim>\-s\ + in @{tool build} and @{tool jedit}. + + \<^medskip> + The \verbose\ field may be set to \<^verbatim>\true\ for extra verbosity. The effect is + similar to option \<^verbatim>\-v\ in @{tool build}. +\ + + +subsubsection \Intermediate output\ + +text \ + The asynchronous notifications of command \<^verbatim>\session_build\ mainly serve as + progress indicator: the output resembles that of the session build window of + Isabelle/jEdit after startup @{cite "isabelle-jedit"}. + + For the client it is usually sufficient to print the messages in plain text, + but note that \theory_progress\ also reveals the internal \theory\ and + \session\ names. +\ + + +subsubsection \Results\ + +text \ + The overall \session_build_results\ contain both a summary and and entry + \session_build_result\ for each session in the build hierarchy, leading up + to the specified target session. The result is always provided, + independently of overall success (\<^verbatim>\FINISHED\ task) or failure (\<^verbatim>\FAILED\ + task). + + The \ok\ field tells abstractly, whether all required session builds came + out as \ok\, i.e.\ \return_code\. A non-zero \return_code\ indicates an + error according to usual POSIX conventions for process exit. + + For individual \session_build_result\ entries, there are additional fields: + \timeout\ to tell if the build process was aborted after running too long, + and \timing\ for the overall process timing in the usual Isabelle format + with elapsed, CPU, gc time. +\ + + +subsubsection \Examples\ + +text \ + Build of a session image from the Isabelle distribution: + @{verbatim [display] \session_build {"session": "HOL-Word"}\} + + Build a session image from the Archive of Formal Proofs: + @{verbatim [display] \session_build {"session": "Coinductive", "dirs": ["$AFP_BASE/thys"]}\} + + Build of a session image of \<^verbatim>\HOL-Number_Theory\ directly on top of \<^verbatim>\HOL\: + @{verbatim [display] \session_build {"session": "HOL-Number_Theory", "ancestor_session": "HOL"}\} +\ + + +subsection \Command \<^verbatim>\session_start\ \label{sec:command-session-start}\ + +text \ + \begin{tabular}{lll} + argument: & \session_build_args \ {print_mode?: [string]}\ \\ + immediate result: & \<^verbatim>\OK\ \task\ \\ + notifications: & \<^verbatim>\NOTE\ \task \ (theory_progress | message)\ \\ + regular result: & \<^verbatim>\FINISHED\ \task \ session_start_result\ \\ + error result: & \<^verbatim>\FAILED\ \task \ error_message\ \\[2ex] + \end{tabular} + + \begin{tabular}{l} + \<^bold>\type\ \session_start_result = {session: string, session_id: uuid}\ + \end{tabular} + + \<^medskip> + The \<^verbatim>\session_start\ command starts a new Isabelle/PIDE session with + underlying Isabelle/ML process, based on a session image that it produces on + demand in the same manner as \<^verbatim>\session_build\. Thus it accepts all + \session_build_args\ and produces similar notifications, but the detailed + \session_build_results\ are omitted. + + The session build and startup process is asynchronous: when the task is + finished, the session remains active for commands such as \<^verbatim>\use_theories\ + (\secref{sec:command-use-theories}), until a \<^verbatim>\session_stop\ or \<^verbatim>\shutdown\ + command is sent to the server. + + Sessions are independent of client connections: it is possible to start a + session and later apply \<^verbatim>\use_theories\ on different connections, as long as + the internal session identifier is known. +\ + + +subsubsection \Arguments\ + +text \ + Most of the arguments are the same as for \<^verbatim>\session_build\ + (\secref{sec:command-session-build}). + + \<^medskip> + The \print_mode\ field adds identifiers of print modes to be made active for + this session. For example, \<^verbatim>\"print_mode": ["ASCII"]\ prefers ASCII + replacement syntax over mathematical Isabelle symbols. See also option \<^verbatim>\-m\ + in @{tool process} (\secref{sec:tool-process}). +\ + + +subsubsection \Results\ + +text \ + The \session\ field provides the active session name for clarity. + + \<^medskip> + The \session_id\ field provides the internal identification of the session + object within the sever process. It can remain active as long as the server + is running, independently of the current client connection. +\ + + +subsection \Examples\ + +text \ + Start a default Isabelle/HOL session: + @{verbatim [display] \session_start {"session": "HOL"}\} +\ + + +subsection \Command \<^verbatim>\session_stop\\ + +text \ + \begin{tabular}{ll} + argument: & \{session_id?: uuid}\ \\ + immediate result: & \<^verbatim>\OK\ \task\ \\ + regular result: & \<^verbatim>\FINISHED\ \task \ session_stop_result\ \\ + error result: & \<^verbatim>\FAILED\ \task \ error_message \ session_stop_result\ \\[2ex] + \end{tabular} + + \begin{tabular}{l} + \<^bold>\type\ \session_stop_result = {ok: bool, return_code: int}\ + \end{tabular} + + \<^medskip> + The \<^verbatim>\session_stop\ command forces a shutdown of the identified + Isabelle/PIDE session. This asynchronous tasks usually finishes quickly. + Failure only happens unusual situations, according to the return code of the + underlying Isabelle/ML process. +\ + + +subsubsection \Arguments\ + +text \ + The \session_id\ field is the UUID originally created by the server for the + intended session. +\ + + +subsubsection \Results\ + +text \ + The \ok\ field tells abstractly, whether the Isabelle/ML process terminated + properly. The \return_code\ expresses this information according to usual + POSIX conventions for process exit. +\ + + +subsection \Command \<^verbatim>\use_theories\ \label{sec:command-use-theories}\ + +text \ + \begin{tabular}{ll} + argument: & \use_theories_arguments\ \\ + regular result: & \<^verbatim>\OK\ \use_theories_results\ \\ + \end{tabular} + + \begin{tabular}{ll} + \<^bold>\type\ \theory_name = string | {name: string, pos: position}\ \\ + \end{tabular} + + \begin{tabular}{ll} + \<^bold>\type\ \use_theories_arguments =\ \\ + \quad\{session_id: uuid,\ \\ + \quad~~\theories: [theory_name],\ \\ + \quad~~\qualifier?: string,\ \\ + \quad~~\master_dir?: string,\ \\ + \quad~~\pretty_margin?: double\ & \<^bold>\default:\ \<^verbatim>\76\ \\ + \quad~~\unicode_symbols?: bool}\ \\[2ex] + + \<^bold>\type\ \node_status =\ \\ + \quad\{ok: bool,\ \\ + \quad~~\total: int,\ \\ + \quad~~\unprocessed: int,\ \\ + \quad~~\running: int,\ \\ + \quad~~\warned: int,\ \\ + \quad~~\failed: int,\ \\ + \quad~~\finished: int,\ \\ + \quad~~\consolidated: bool}\ \\[2ex] + + \<^bold>\type\ \node_result =\ \\ + \quad\{node_name: string,\ \\ + \quad~~\theory: string,\ \\ + \quad~~\status: node_status,\ \\ + \quad~~\messages: [message]}\ \\[2ex] + + \<^bold>\type\ \use_theories_results =\ \\ + \quad\{ok: bool,\ \\ + \quad~~\errors: [message],\ \\ + \quad~~\nodes: [node_result]}\ \\[2ex] + \end{tabular} + + \<^medskip> + + The \<^verbatim>\use_theories\ command updates the identified session by adding the + current version of given theory files to it: theory dependencies are + resolved implicitly. The command succeeds eventually, when all theories have + been \<^emph>\consolidated\ in the sense that the outermost command structure has + finished (or failed) and the final \<^theory_text>\end\ command of each theory has been + checked. +\ + + +subsubsection \Arguments\ + +text \ + The \session_id\ is the identifier provided by the server, when the session + was created (possibly on a different client connection). + + \<^medskip> + The \theories\ field specifies theory names as in \<^theory_text>\theory imports\ or in + session ROOT \<^verbatim>\theories\: an explicit source position for these theory name + specifications may be given as well, which is occasionally useful for + precise error locations. + + \<^medskip> + The \qualifier\ field provides an alternative session qualifier for theories + that are not formally recognized as a member of a particular session. The + default is \<^verbatim>\Draft\ as in Isabelle/jEdit. There is rarely a need to change + that, as theory nodes are already uniquely identified by their physical + file-system location. + + \<^medskip> + The \master_dir\ field explicit specifies the formal master directory of the + imported theory. Normally this is determined implicitly from the parent + directory of the physical theory file location. + + \<^medskip> + The \pretty_margin\ field specifies the line width for pretty-printing. The + default is for classic console output. Formatting happens at the end of + \<^verbatim>\use_theories\, when all prover messages produced are exported to the + client. + + \<^medskip> + The \unicode_symbols\ field set to \<^verbatim>\true\ means that message output is + rendered for direct output on a Unicode capable channel, ideally with the + Isabelle fonts as in Isabelle/jEdit (or Isabelle HTML output). The default + is to keep the symbolic representation of Isabelle text, e.g.\ \<^verbatim>\\\ instead + of its rendering as \\\. This means the client needs to perform its own + rendering to present it to the end-user. +\ + +subsubsection \Results\ + +text \ + The \ok\ field indicates overall success of processing the specified + theories with all their dependencies. + + When \ok\ is \<^verbatim>\false\, the \errors\ field lists all errors cumulatively, + including position information for the theory nodes where the error happened + (this includes imported theories). + + \<^medskip> + The \nodes\ field provides more detailed information about each imported + theory node. Here the individual fields are as follows: + + \<^item> \node_name\: the physical name for the theory node, based on its + file-system location;\<^footnote>\Clients may recognize the originally provided + file-names after careful normalization in the file-system of the server.\ + + \<^item> \theory\: the logical theory name, e.g.\ \<^verbatim>\HOL-Library.AList\ or + \<^verbatim>\Draft.Test\; + + \<^item> \status\: the overall node status, e.g.\ see the visualization in the + \Theories\ panel of Isabelle/jEdit @{cite "isabelle-jedit"}; + + \<^item> \messages\: the main bulk of prover messages produced in this theory note + (\<^verbatim>\writeln\, \<^verbatim>\warning\, \<^verbatim>\error\, etc.). +\ + + +subsubsection \Examples\ + +text \ + Process some example theory from the Isabelle distribution, within the + context of an already started session for Isabelle/HOL (see also + \secref{sec:command-session-start}): + @{verbatim [display] \use_theories {"session_id": ..., "theories": ["~~/src/HOL/Isar_Examples/Drinker"]}\} + + \<^medskip> + Process some example theories that import other theories via + session-qualified theory names: + + @{verbatim [display] \session_start {"session": "HOL", "all_known": true} +use_theories {"session_id": ..., "theories": ["~~/src/HOL/Unix/Unix"]} +session_stop {"session_id": ...}\} + + The option \all_known\ has been used here to the full name space of + session-qualified theory names accessible in this session. This is also the + default in Isabelle/jEdit, although it requires significant startup time. + + \<^medskip> + Process some example theories in the context of their (single) parent + session: + + @{verbatim [display] \session_start {"session": "HOL-Library"} +use_theories {"session_id": ..., "theories": ["~~/src/HOL/Unix/Unix"]} +session_stop {"session_id": ...}\} \ end