wenzelm@67904: (*:maxLineLen=78:*) wenzelm@67904: wenzelm@67904: theory Server wenzelm@67904: imports Base wenzelm@67904: begin wenzelm@67904: wenzelm@67904: chapter \The Isabelle server\ wenzelm@67904: wenzelm@67904: text \ wenzelm@67904: An Isabelle session requires at least two processes, which are both rather wenzelm@67904: heavy: Isabelle/Scala for the system infrastructure and Isabelle/ML for the wenzelm@67904: logic session (e.g.\ HOL). In principle, these processes can be invoked wenzelm@67904: directly on the command-line, e.g.\ via @{tool java}, @{tool scala}, @{tool wenzelm@67904: process}, @{tool console}, but this approach is inadequate for reactive wenzelm@67904: applications that require quick responses from the prover. wenzelm@67904: wenzelm@67904: In contrast, the Isabelle server exposes Isabelle/Scala as a wenzelm@67904: ``terminate-stay-resident'' application that manages multiple logic wenzelm@67904: \<^emph>\sessions\ and concurrent tasks to use \<^emph>\theories\. This provides an wenzelm@67904: analogous to @{ML Thy_Info.use_theories} in Isabelle/ML, but with full wenzelm@67904: concurrency and Isabelle/PIDE markup. wenzelm@67904: wenzelm@67904: The client/server arrangement via TCP sockets also opens possibilities for wenzelm@67904: remote Isabelle services that are accessed by local applications, e.g.\ via wenzelm@67904: an SSH tunnel. wenzelm@67904: \ wenzelm@67904: wenzelm@67904: wenzelm@67904: section \Command-line tools\ wenzelm@67904: wenzelm@67904: subsection \Server \label{sec:tool-server}\ wenzelm@67904: wenzelm@67904: text \ wenzelm@67904: The @{tool_def server} tool manages resident server processes: wenzelm@67904: @{verbatim [display] wenzelm@67904: \Usage: isabelle server [OPTIONS] wenzelm@67904: wenzelm@67904: Options are: wenzelm@67904: -L FILE logging on FILE wenzelm@67904: -c console interaction with specified server wenzelm@67904: -l list servers (alternative operation) wenzelm@67904: -n NAME explicit server name (default: isabelle) wenzelm@67904: -p PORT explicit server port wenzelm@67904: -s assume existing server, no implicit startup wenzelm@67904: -x exit specified server (alternative operation) wenzelm@67904: wenzelm@67904: Manage resident Isabelle servers.\} wenzelm@67904: wenzelm@67904: The main operation of \<^verbatim>\isabelle server\ is to ensure that a named server is wenzelm@67904: running, either by finding an already running process (according to the wenzelm@67904: central database file @{path "$ISABELLE_HOME_USER/servers.db"}) or by wenzelm@67904: becoming itself a new server that accepts connections on a particular TCP wenzelm@67904: socket. The server name and its address are printed as initial output line. wenzelm@67904: If another server instance is already running, the current wenzelm@67904: \<^verbatim>\isabelle server\ process will terminate; otherwise, it keeps running as a wenzelm@67904: new server process until an explicit \<^verbatim>\shutdown\ command is received. wenzelm@67904: Further details of the server socket protocol are explained in wenzelm@67904: \secref{sec:server-protocol}. wenzelm@67904: wenzelm@67904: Other server management operations are invoked via options \<^verbatim>\-l\ and \<^verbatim>\-x\ wenzelm@67904: (see below). wenzelm@67904: wenzelm@67904: \<^medskip> wenzelm@67904: Option \<^verbatim>\-n\ specifies an alternative server name: at most one process for wenzelm@67904: each name may run, but each server instance supports multiple connections wenzelm@67904: and logic sessions. wenzelm@67904: wenzelm@67904: \<^medskip> wenzelm@67904: Option \<^verbatim>\-p\ specifies an explicit TCP port for the server socket (which is wenzelm@67904: always on \<^verbatim>\localhost\): the default is to let the operating system assign a wenzelm@67904: free port number. wenzelm@67904: wenzelm@67904: \<^medskip> wenzelm@67904: Option \<^verbatim>\-s\ strictly assumes that the specified server process is already wenzelm@67904: running, skipping the optional server startup phase. wenzelm@67904: wenzelm@67904: \<^medskip> wenzelm@67904: Option \<^verbatim>\-c\ connects the console in/out channels after the initial check wenzelm@67918: for a suitable server process. Also note that the @{tool client} tool wenzelm@67918: (\secref{sec:tool-client}) provides a command-line editor to interact with wenzelm@67918: the server. wenzelm@67904: wenzelm@67904: \<^medskip> wenzelm@67904: Option \<^verbatim>\-L\ specifies a log file for exceptional output of internal server wenzelm@67904: and session operations. wenzelm@67918: wenzelm@67918: \<^medskip> wenzelm@67918: Operation \<^verbatim>\-l\ lists all active server processes with their connection wenzelm@67918: details. wenzelm@67918: wenzelm@67918: \<^medskip> wenzelm@67918: Operation \<^verbatim>\-x\ exits the specified server process by sending it a wenzelm@67918: \<^verbatim>\shutdown\ command. wenzelm@67904: \ wenzelm@67904: wenzelm@67904: wenzelm@67904: subsection \Client \label{sec:tool-client}\ wenzelm@67904: wenzelm@67904: text \ wenzelm@67918: The @{tool_def client} tool provides console interaction for Isabelle wenzelm@67918: servers: wenzelm@67904: @{verbatim [display] wenzelm@67904: \Usage: isabelle client [OPTIONS] wenzelm@67904: wenzelm@67904: Options are: wenzelm@67904: -n NAME explicit server name wenzelm@67904: -p PORT explicit server port wenzelm@67904: wenzelm@67904: Console interaction for Isabelle server (with line-editor).\} wenzelm@67904: wenzelm@67918: This is a wrapper to \<^verbatim>\isabelle server -s -c\ for interactive wenzelm@67918: experimentation, which uses @{setting ISABELLE_LINE_EDITOR} if available. wenzelm@67904: The server name is sufficient for identification, as the client can wenzelm@67904: determine the connection details from the local database of active servers. wenzelm@67904: wenzelm@67904: \<^medskip> wenzelm@67904: Option \<^verbatim>\-n\ specifies an explicit server name as in @{tool server}. wenzelm@67904: wenzelm@67904: \<^medskip> wenzelm@67904: Option \<^verbatim>\-p\ specifies an explicit server port as in @{tool server}. wenzelm@67904: \ wenzelm@67904: wenzelm@67904: wenzelm@67904: subsection \Examples\ wenzelm@67904: wenzelm@67904: text \ wenzelm@67904: Ensure that a particular server instance is running in the background: wenzelm@67904: @{verbatim [display] \isabelle server -n test &\} wenzelm@67904: wenzelm@67918: The first line of output presents the connection details:\<^footnote>\This information wenzelm@67918: may be used in other TCP clients, without access to Isabelle/Scala and the wenzelm@67918: underlying database of running servers.\ wenzelm@67904: @{verbatim [display] \server "test" = 127.0.0.1:4711 (password "XYZ")\} wenzelm@67904: wenzelm@67904: \<^medskip> wenzelm@67904: List available server processes: wenzelm@67904: @{verbatim [display] \isabelle server -l\} wenzelm@67904: wenzelm@67904: \<^medskip> wenzelm@67904: Connect the command-line client to the above test server: wenzelm@67904: @{verbatim [display] \isabelle client -n test\} wenzelm@67904: wenzelm@67904: Interaction now works on a line-by-line basis, with commands like \<^verbatim>\help\ or wenzelm@67904: \<^verbatim>\echo\. For example, some JSON values may be echoed like this: wenzelm@67904: wenzelm@67904: @{verbatim [display] wenzelm@67904: \echo 42 wenzelm@67904: echo [1, 2, 3] wenzelm@67904: echo {"a": "text", "b": true, "c": 42}\} wenzelm@67904: wenzelm@67904: Closing the connection (via CTRL-D) leaves the server running: it is wenzelm@67904: possible to reconnect again, and have multiple connections at the same time. wenzelm@67904: wenzelm@67904: \<^medskip> wenzelm@67918: Exit the named server on the command-line: wenzelm@67904: @{verbatim [display] \isabelle server -n test -x\} wenzelm@67904: \ wenzelm@67904: wenzelm@67904: wenzelm@67904: section \Protocol messages \label{sec:server-protocol}\ wenzelm@67904: wenzelm@67904: text \ wenzelm@67904: The Isabelle server listens on a regular TCP socket, using a line-oriented wenzelm@67918: protocol of structured messages. Input \<^emph>\commands\ and output \<^emph>\results\ wenzelm@67904: (via \<^verbatim>\OK\ or \<^verbatim>\ERROR\) are strictly alternating on the toplevel, but wenzelm@67904: commands may also return a \<^emph>\task\ identifier to indicate an ongoing wenzelm@67904: asynchronous process that is joined later (via \<^verbatim>\FINISHED\ or \<^verbatim>\FAILED\). wenzelm@67904: Asynchronous \<^verbatim>\NOTE\ messages may occur at any time: they are independent of wenzelm@67918: the main command-result protocol. wenzelm@67904: wenzelm@67904: For example, the synchronous \<^verbatim>\echo\ command immediately returns its wenzelm@67904: argument as \<^verbatim>\OK\ result. In contrast, the asynchronous \<^verbatim>\session_build\ wenzelm@67904: command returns \<^verbatim>\OK {"task":\\id\\<^verbatim>\}\ and continues in the background. It wenzelm@67904: will eventually produce \<^verbatim>\FINISHED {"task":\\id\\<^verbatim>\,\\\\\<^verbatim>\}\ or wenzelm@67904: \<^verbatim>\FAILED {"task":\\id\\<^verbatim>\,\\\\\<^verbatim>\}\ with the final result. Intermediately, it wenzelm@67904: may emit asynchronous messages of the form \<^verbatim>\NOTE {"task":\\id\\<^verbatim>\,\\\\\<^verbatim>\}\ wenzelm@67904: to inform about its progress. Due to the explicit task identifier, the wenzelm@67904: client can show these messages in the proper context, e.g.\ a GUI window for wenzelm@67918: this particular session build job. wenzelm@67904: wenzelm@67904: \medskip Subsequently, the protocol message formats are described in further wenzelm@67904: detail. wenzelm@67904: \ wenzelm@67904: wenzelm@67904: wenzelm@67904: subsection \Byte messages\ wenzelm@67904: wenzelm@67904: text \ wenzelm@67904: The client-server connection is a raw byte-channel for bidirectional wenzelm@67904: communication, but the Isabelle server always works with messages of a wenzelm@67904: particular length. Messages are written as a single chunk that is flushed wenzelm@67904: immediately. wenzelm@67904: wenzelm@67918: Message boundaries are determined as follows: wenzelm@67904: wenzelm@67904: \<^item> A \<^emph>\short message\ consists of a single line: it is a sequence of wenzelm@67904: arbitrary bytes excluding CR (13) and LF (10), and terminated by CR-LF or wenzelm@67904: just LF. wenzelm@67904: wenzelm@67918: \<^item> A \<^emph>\long message\ starts with a single that consists only of decimal wenzelm@67918: digits: these are interpreted as length of the subsequent block of wenzelm@67918: arbitrary bytes. A final line-terminator (as above) may be included here, wenzelm@67904: but is not required. wenzelm@67904: wenzelm@67904: Messages in JSON format (see below) always fit on a single line, due to wenzelm@67904: escaping of newline characters within string literals. This is convenient wenzelm@67904: for interactive experimentation, but it can impact performance for very long wenzelm@67904: messages. If the message byte-length is given on the preceding line, the wenzelm@67918: server can read the message more efficiently as a single block. wenzelm@67904: \ wenzelm@67904: wenzelm@67904: wenzelm@67904: subsection \Text messages\ wenzelm@67904: wenzelm@67904: text \ wenzelm@67904: Messages are read and written as byte streams (with byte lengths), but the wenzelm@67904: content is always interpreted as plain text in terms of the UTF-8 wenzelm@67904: encoding.\<^footnote>\See also the ``UTF-8 Everywhere Manifesto'' wenzelm@67904: \<^url>\http://utf8everywhere.org\.\ wenzelm@67904: wenzelm@67904: Note that line-endings and other formatting characters are invariant wrt. wenzelm@67904: UTF-8 representation of text: thus implementations are free to determine the wenzelm@67904: overall message structure before or after applying the text encoding. wenzelm@67904: \ wenzelm@67904: wenzelm@67904: wenzelm@67917: subsection \Input and output messages \label{sec:input-output-messages}\ wenzelm@67904: wenzelm@67904: text \ wenzelm@67904: Server input and output messages have a uniform format as follows: wenzelm@67904: wenzelm@67904: \<^item> \name argument\ such that: wenzelm@67904: wenzelm@67904: \<^item> \name\ is the longest prefix consisting of ASCII letters, digits, wenzelm@67918: ``\<^verbatim>\_\'', ``\<^verbatim>\.\'', wenzelm@67904: wenzelm@67904: \<^item> the separator between \name\ and \argument\ is the longest possible wenzelm@67904: sequence of ASCII blanks (it could be empty, e.g.\ when the argument wenzelm@67904: starts with a quote or bracket), wenzelm@67904: wenzelm@67918: \<^item> \argument\ is the rest of the message without line terminator. wenzelm@67904: wenzelm@67904: \<^medskip> wenzelm@67904: Input messages are sent from the client to the server. Here the \name\ wenzelm@67904: specifies a \<^emph>\server command\: the list of known commands may be wenzelm@67904: retrieved via the \<^verbatim>\help\ command. wenzelm@67904: wenzelm@67904: \<^medskip> wenzelm@67904: Output messages are sent from the server to the client. Here the \name\ wenzelm@67904: specifies the \<^emph>\server reply\, which always has a specific meaning as wenzelm@67904: follows: wenzelm@67904: wenzelm@67904: \<^item> synchronous results: \<^verbatim>\OK\ or \<^verbatim>\ERROR\ wenzelm@67904: \<^item> asynchronous results: \<^verbatim>\FINISHED\ or \<^verbatim>\FAILED\ wenzelm@67904: \<^item> intermediate notifications: \<^verbatim>\NOTE\ wenzelm@67904: wenzelm@67904: \<^medskip> wenzelm@67904: The \argument\ format is uniform for both input and output messages: wenzelm@67904: wenzelm@67904: \<^item> empty argument (Scala type \<^verbatim>\Unit\) wenzelm@67904: \<^item> XML element in YXML notation (Scala type \<^verbatim>\XML.Elem\) wenzelm@67904: \<^item> JSON value (Scala type \<^verbatim>\JSON.T\) wenzelm@67904: wenzelm@67904: JSON values may consist of objects (records), arrays (lists), strings, wenzelm@67904: numbers, bools, null.\<^footnote>\See also the official specification wenzelm@67904: \<^url>\https://www.json.org\ and unofficial explorations ``Parsing JSON is a wenzelm@67904: Minefield'' \<^url>\http://seriot.ch/parsing_json.php\.\ Since JSON requires wenzelm@67904: explicit quotes and backslash-escapes to represent arbitrary text, the YXML wenzelm@67904: notation for XML trees (\secref{sec:yxml-vs-xml}) works better wenzelm@67904: for large messages with a lot of PIDE markup. wenzelm@67904: wenzelm@67904: Nonetheless, the most common commands use JSON by default: big chunks of wenzelm@67904: text (theory sources etc.) are taken from the underlying file-system and wenzelm@67904: results are pre-formatted for plain-text output, without PIDE markup wenzelm@67904: information. This is a concession to simplicity: the server imitates the wenzelm@67904: appearance of command-line tools on top of the Isabelle/PIDE infrastructure. wenzelm@67904: \ wenzelm@67904: wenzelm@67904: wenzelm@67904: subsection \Initial password exchange\ wenzelm@67904: wenzelm@67904: text \ wenzelm@67904: Whenever a new client opens the server socket, the initial message needs to wenzelm@67904: be its unique password. The server replies either with \<^verbatim>\OK\ (and some wenzelm@67904: information about the Isabelle version) or by silent disconnection of what wenzelm@67918: is considered an illegal connection attempt. Note that @{tool client} wenzelm@67918: already presents the correct password internally. wenzelm@67904: wenzelm@67904: Server passwords are created as Universally Unique Identifier (UUID) in wenzelm@67904: Isabelle/Scala and stored in a per-user database, with restricted wenzelm@67904: file-system access only for the current user. The Isabelle/Scala server wenzelm@67904: implementation is careful to expose the password only on private output wenzelm@67904: channels, and not on a process command-line (which is accessible to other wenzelm@67904: users, e.g.\ via the \<^verbatim>\ps\ command). wenzelm@67904: \ wenzelm@67904: wenzelm@67904: wenzelm@67904: subsection \Synchronous commands\ wenzelm@67904: wenzelm@67904: text \ wenzelm@67904: A \<^emph>\synchronous command\ corresponds to regular function application in wenzelm@67904: Isabelle/Scala, with single argument and result (regular or error). Both the wenzelm@67904: argument and the result may consist of type \<^verbatim>\Unit\, \<^verbatim>\XML.Elem\, \<^verbatim>\JSON.T\. wenzelm@67904: An error result typically consists of a JSON object with error message and wenzelm@67918: potentially further result fields (this resembles exceptions in Scala). wenzelm@67904: wenzelm@67904: These are the protocol exchanges for both cases of command execution: wenzelm@67904: \begin{center} wenzelm@67904: \begin{tabular}{rl} wenzelm@67904: \<^bold>\input:\ & \command argument\ \\ wenzelm@67904: (a) regular \<^bold>\output:\ & \<^verbatim>\OK\ \result\ \\ wenzelm@67904: (b) error \<^bold>\output:\ & \<^verbatim>\ERROR\ \result\ \\ wenzelm@67904: \end{tabular} wenzelm@67904: \end{center} wenzelm@67904: \ wenzelm@67904: wenzelm@67904: wenzelm@67904: subsection \Asynchronous commands\ wenzelm@67904: wenzelm@67904: text \ wenzelm@67904: An \<^emph>\asynchronous command\ corresponds to an ongoing process that finishes wenzelm@67918: or fails eventually, while emitting arbitrary notifications in between. wenzelm@67918: Formally, it starts as synchronous command with immediate result \<^verbatim>\OK\ wenzelm@67918: giving the \<^verbatim>\task\ identifier, or an immediate \<^verbatim>\ERROR\ that indicates bad wenzelm@67918: command syntax. For a running task, the termination is indicated later by wenzelm@67918: \<^verbatim>\FINISHED\ or \<^verbatim>\FAILED\, together with its ultimate result value. wenzelm@67904: wenzelm@67904: These are the protocol exchanges for various cases of command task wenzelm@67904: execution: wenzelm@67904: wenzelm@67904: \begin{center} wenzelm@67904: \begin{tabular}{rl} wenzelm@67904: \<^bold>\input:\ & \command argument\ \\ wenzelm@67904: immediate \<^bold>\output:\ & \<^verbatim>\OK {"task":\\id\\<^verbatim>\}\ \\ wenzelm@67904: intermediate \<^bold>\output:\ & \<^verbatim>\NOTE {"task":\\id\\<^verbatim>\,\\\\\<^verbatim>\}\ \\ wenzelm@67904: (a) regular \<^bold>\output:\ & \<^verbatim>\FINISHED {"task":\\id\\<^verbatim>\,\\\\\<^verbatim>\}\ \\ wenzelm@67904: (b) error \<^bold>\output:\ & \<^verbatim>\FAILED {"task":\\id\\<^verbatim>\,\\\\\<^verbatim>\}\ \\[3ex] wenzelm@67904: \<^bold>\input:\ & \command argument\ \\ wenzelm@67904: immediate \<^bold>\output:\ & \<^verbatim>\ERROR\~\\\ \\ wenzelm@67904: \end{tabular} wenzelm@67904: \end{center} wenzelm@67904: wenzelm@67904: All asynchronous messages are decorated with the task identifier that was wenzelm@67918: revealed in the immediate (synchronous) result. Thus the client can wenzelm@67918: invoke further asynchronous commands and still dispatch the resulting stream of wenzelm@67904: asynchronous messages properly. wenzelm@67904: wenzelm@67904: The synchronous command \<^verbatim>\cancel\~\id\ tells the specified task to terminate wenzelm@67917: prematurely: usually causing a \<^verbatim>\FAILED\ result, but this is not guaranteed: wenzelm@67918: the cancel event may come too late or the running process may just ignore wenzelm@67918: it. wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67918: section \Types for JSON values \label{sec:json-types}\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67917: In order to specify concrete JSON types for command arguments and result wenzelm@67917: messages, the following type definition language shall be used: wenzelm@67917: wenzelm@67917: \<^rail>\ wenzelm@67917: @{syntax type_def}: @'type' @{syntax name} '=' @{syntax type} wenzelm@67917: ; wenzelm@67917: @{syntax type}: @{syntax name} | @{syntax value} | 'any' | 'null' | wenzelm@67917: 'bool' | 'int' | 'long' | 'double' | 'string' | '[' @{syntax type} ']' | wenzelm@67917: '{' (@{syntax field_type} ',' *) '}' | wenzelm@67917: @{syntax type} '\' @{syntax type} | wenzelm@67918: @{syntax type} '|' @{syntax type} | wenzelm@67918: '(' @{syntax type} ')' wenzelm@67917: ; wenzelm@67917: @{syntax field_type}: @{syntax name} ('?'?) ':' @{syntax type} wenzelm@67917: \ wenzelm@67917: wenzelm@67918: This is a simplified variation of TypeScript wenzelm@67918: interfaces.\<^footnote>\\<^url>\https://www.typescriptlang.org/docs/handbook/interfaces.html\\ wenzelm@67918: The meaning of these types is specified wrt. the Isabelle/Scala wenzelm@67918: implementation as follows. wenzelm@67917: wenzelm@67917: \<^item> A \name\ refers to a type defined elsewhere. The environment of type wenzelm@67917: definitions is given informally: put into proper foundational order, it wenzelm@67918: needs to specify a strongly normalizing system of syntactic abbreviations; wenzelm@67918: type definitions may not be recursive. wenzelm@67917: wenzelm@67917: \<^item> A \value\ in JSON notation represents the singleton type of the given wenzelm@67918: item. For example, the string \<^verbatim>\"error"\ can be used as type for a slot that wenzelm@67918: is guaranteed to contain that constant. wenzelm@67917: wenzelm@67917: \<^item> Type \any\ is the super type of all other types: it is an untyped slot in wenzelm@67917: the specification and corresponds to \<^verbatim>\Any\ or \<^verbatim>\JSON.T\ in Isabelle/Scala. wenzelm@67917: wenzelm@67917: \<^item> Type \null\ is the type of the improper value \null\; it corresponds to wenzelm@67917: type \<^verbatim>\Null\ in Scala and is normally not used in Isabelle/Scala.\<^footnote>\See also wenzelm@67917: ``Null References: The Billion Dollar Mistake'' by Tony Hoare wenzelm@67917: \<^url>\https://www.infoq.com/presentations/Null-References-The-Billion-Dollar-Mistake-Tony-Hoare\.\ wenzelm@67917: wenzelm@67918: \<^item> Type \bool\ is the type of the truth values \<^verbatim>\true\ and \<^verbatim>\false\; it wenzelm@67917: corresponds to \<^verbatim>\Boolean\ in Scala. wenzelm@67917: wenzelm@67917: \<^item> Types \int\, \long\, \double\ are specific versions of the generic wenzelm@67917: \number\ type, corresponding to \<^verbatim>\Int\, \<^verbatim>\Long\, \<^verbatim>\Double\ in Scala, but wenzelm@67917: \<^verbatim>\Long\ is limited to 53 bit precision.\<^footnote>\Implementations of JSON typically wenzelm@67917: standardize \number\ to \<^verbatim>\Double\, which can absorb \<^verbatim>\Int\ faithfully, but wenzelm@67917: not all of \<^verbatim>\Long\.\ wenzelm@67917: wenzelm@67917: \<^item> Type \string\ represents Unicode text; it corresponds to type \<^verbatim>\String\ in wenzelm@67917: Scala. wenzelm@67917: wenzelm@67917: \<^item> Type \[t]\ is the array (or list) type over \t\; it corresponds to wenzelm@67917: \<^verbatim>\List[t]\ in Scala. The list type is co-variant as usual (i.e.\ monotonic wenzelm@67918: wrt. the subtype relation). wenzelm@67917: wenzelm@67917: \<^item> Object types describe the possible content of JSON records, with field wenzelm@67917: names and types. A question mark after a field name means that it is wenzelm@67918: optional. In Scala this could refer to an explicit type \<^verbatim>\Option[t]\, e.g.\ wenzelm@67918: \{a: int, b?: string}\ corresponding to a Scala case class with arguments wenzelm@67918: \<^verbatim>\a: Int\, \<^verbatim>\b: Option[String]\. wenzelm@67917: wenzelm@67917: Alternatively, optional fields can have a default value. If nothing else is wenzelm@67918: specified, a standard ``empty value'' is used for each type, i.e.\ \<^verbatim>\0\ for wenzelm@67918: the number types, \<^verbatim>\false\ for \bool\, or the empty string, array, object wenzelm@67918: etc. wenzelm@67917: wenzelm@67917: Object types are \<^emph>\permissive\ in the sense that only the specified field wenzelm@67917: names need to conform to the given types, but unspecified fields may be wenzelm@67917: present as well. wenzelm@67917: wenzelm@67917: \<^item> The type expression \t\<^sub>1 \ t\<^sub>2\ only works for two object types with wenzelm@67917: disjoint field names: it is the concatenation of the respective @{syntax wenzelm@67917: field_type} specifications taken together. For example: \{task: string} \ wenzelm@67917: {ok: bool}\ is the equivalent to \{task: string, ok: bool}\. wenzelm@67917: wenzelm@67917: \<^item> The type expression \t\<^sub>1 | t\<^sub>2\ is the disjoint union of two types, either wenzelm@67917: one of the two cases may occur. wenzelm@67917: wenzelm@67918: \<^item> Parentheses \(t)\ merely group type expressions syntactically. wenzelm@67918: wenzelm@67917: wenzelm@67917: These types correspond to JSON values in an obvious manner, which is not wenzelm@67917: further described here. For example, the JSON array \<^verbatim>\[1, 2, 3]\ conforms to wenzelm@67917: types \[int]\, \[long]\, \[double]\, \[any]\, \any\. wenzelm@67917: wenzelm@67917: Note that JSON objects require field names to be quoted, but the type wenzelm@67918: language omits quotes for clarity. Thus the object \<^verbatim>\{"a": 42, "b": "xyz"}\ wenzelm@67918: conforms to the type \{a: int, b: string}\, for example. wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67918: The absence of an argument or result is represented by the Scala type wenzelm@67918: \<^verbatim>\Unit\: it is written as empty text in the message \argument\ wenzelm@67917: (\secref{sec:input-output-messages}). This is not part of the JSON language. wenzelm@67917: wenzelm@67917: Server replies have name tags like \<^verbatim>\OK\, \<^verbatim>\ERROR\: these are used literally wenzelm@67917: together with type specifications to indicate the particular name with the wenzelm@67917: type of its argument, e.g.\ \<^verbatim>\OK\~\[string]\ for a regular result that is a wenzelm@67917: list (JSON array) of strings. wenzelm@67917: wenzelm@67917: \<^bigskip> wenzelm@67918: Here are some common type definitions, for use in particular specifications wenzelm@67918: of command arguments and results. wenzelm@67917: wenzelm@67917: \<^item> \<^bold>\type\~\position = {line?: int, offset?: int, end_offset?: int, file?: wenzelm@67918: string, id?: long}\ describes a source position within Isabelle text. Only wenzelm@67918: the \line\ and \file\ fields make immediate sense to external programs. wenzelm@67917: Detailed \offset\ and \end_offset\ positions are counted according to wenzelm@67917: Isabelle symbols, see @{ML_type Symbol.symbol} in Isabelle/ML @{cite wenzelm@67918: "isabelle-implementation"}. The position \id\ belongs to the representation wenzelm@67918: of command transactions in the Isabelle/PIDE protocol: it normally does not wenzelm@67918: occur in externalized positions. wenzelm@67917: wenzelm@67917: \<^item> \<^bold>\type\~\message = {kind?: string, message: string, pos?: position}\ where wenzelm@67917: the \kind\ provides some hint about the role and importance of the message. wenzelm@67917: The main message kinds are \<^verbatim>\writeln\ (for regular output), \<^verbatim>\warning\, and wenzelm@67917: \<^verbatim>\error\. The table \<^verbatim>\Markup.messages\ in Isabelle/Scala defines further wenzelm@67917: message kinds for more specific applications. wenzelm@67917: wenzelm@67917: \<^item> \<^bold>\type\~\error_message = {kind:\~\<^verbatim>\"error"\\, message: string}\ refers to wenzelm@67917: error messages in particular. These occur routinely with \<^verbatim>\ERROR\ or wenzelm@67917: \<^verbatim>\FAILED\ replies, but also as initial command syntax errors (which are wenzelm@67917: omitted in the command specifications below). wenzelm@67917: wenzelm@67917: \<^item> \<^bold>\type\~\theory_progress = {kind:\~\<^verbatim>\"writeln"\\, message: string, theory: wenzelm@67917: string, session: string}\ reports formal progress in loading theories (e.g.\ wenzelm@67917: when building a session image). Apart from a regular output message, it also wenzelm@67918: reveals the formal theory name (e.g.\ \<^verbatim>\"HOL.Nat"\) and session name (e.g.\ wenzelm@67918: \<^verbatim>\"HOL"\). Note that some rare theory names lack a proper session prefix, wenzelm@67918: e.g. theory \<^verbatim>\"Main"\ in session \<^verbatim>\"HOL"\. wenzelm@67917: wenzelm@67917: \<^item> \<^bold>\type\~\timing = {elapsed: double, cpu: double, gc: double}\ refers to wenzelm@67917: common Isabelle timing information in seconds, usually with a precision of wenzelm@67917: three digits after the point (whole milliseconds). wenzelm@67917: wenzelm@67917: \<^item> \<^bold>\type\~\uuid = string\ refers to a Universally Unique Identifier (UUID) wenzelm@67917: as plain text.\<^footnote>\See \<^url>\https://www.ietf.org/rfc/rfc4122.txt\ and wenzelm@67917: \<^url>\https://docs.oracle.com/javase/8/docs/api/java/util/UUID.html\.\ Such wenzelm@67917: identifiers are created as private random numbers of the server and only wenzelm@67917: revealed to the client that creates a certain resource (e.g.\ task or wenzelm@67917: session). A client may disclose this information for use in a different wenzelm@67918: client connection: this allows to share sessions between multiple wenzelm@67917: connections. wenzelm@67917: wenzelm@67917: Client commands need to provide syntactically wellformed UUIDs: this is wenzelm@67917: trivial to achieve by using only identifiers that have been produced by the wenzelm@67917: server beforehand. wenzelm@67917: wenzelm@67917: \<^item> \<^bold>\type\~\task = {task: uuid}\ identifies a newly created asynchronous task wenzelm@67917: and thus allows the client to control it by the \<^verbatim>\cancel\ command. The same wenzelm@67918: task identification is included in all messages produced by this task. wenzelm@67918: wenzelm@67918: \<^item> \<^bold>\type\ \node_status = {ok: bool, total: int, unprocessed: int, running: wenzelm@67918: int, warned: int, failed: int, finished: int, consolidated: bool}\ wenzelm@67918: represents a formal theory node status of the PIDE document model. Fields wenzelm@67918: \total\, \unprocessed\, \running\, \warned\, \failed\, \finished\ account wenzelm@67918: for individual commands within a theory node; \ok\ is an abstraction for wenzelm@67918: \failed = 0\. The \consolidated\ flag indicates whether the outermost theory wenzelm@67918: command structure has finished (or failed) and the final \<^theory_text>\end\ command has wenzelm@67918: been checked. wenzelm@67918: \ wenzelm@67918: wenzelm@67918: wenzelm@67918: section \Server commands and results\ wenzelm@67918: wenzelm@67918: text \ wenzelm@67918: Here follows an overview of particular Isabelle server commands with their wenzelm@67918: results, which are usually represented as JSON values with types according wenzelm@67918: to \secref{sec:json-types}. The general format of input and output messages wenzelm@67918: is described in \secref{sec:input-output-messages}. The relevant wenzelm@67918: Isabelle/Scala source files are: wenzelm@67918: wenzelm@67918: \<^medskip> wenzelm@67918: \begin{tabular}{l} wenzelm@67918: \<^file>\$ISABELLE_HOME/src/Pure/Tools/server_commands.scala\ \\ wenzelm@67918: \<^file>\$ISABELLE_HOME/src/Pure/Tools/server.scala\ \\ wenzelm@67918: \<^file>\$ISABELLE_HOME/src/Pure/General/json.scala\ \\ wenzelm@67918: \end{tabular} wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsection \Command \<^verbatim>\help\\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67917: \begin{tabular}{ll} wenzelm@67917: regular result: & \<^verbatim>\OK\ \[string]\ \\ wenzelm@67917: \end{tabular} wenzelm@67917: \<^medskip> wenzelm@67917: wenzelm@67918: The \<^verbatim>\help\ command has no argument and returns the list of server command wenzelm@67918: names. This is occasionally useful for interactive experimentation (see also wenzelm@67918: @{tool client} in \secref{sec:tool-client}). wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsection \Command \<^verbatim>\echo\\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67917: \begin{tabular}{ll} wenzelm@67917: argument: & \any\ \\ wenzelm@67917: regular result: & \<^verbatim>\OK\ \any\ \\ wenzelm@67917: \end{tabular} wenzelm@67917: \<^medskip> wenzelm@67917: wenzelm@67917: The \<^verbatim>\echo\ command is the identity function: it returns its argument as wenzelm@67917: regular result. This is occasionally useful for testing and interactive wenzelm@67917: experimentation (see also @{tool client} in \secref{sec:tool-client}). wenzelm@67917: wenzelm@67918: The Scala type of \<^verbatim>\echo\ is actually more general than given above: wenzelm@67918: \<^verbatim>\Unit\, \<^verbatim>\XML.Elem\, \<^verbatim>\JSON.T\ work uniformly. Note that \<^verbatim>\XML.Elem\ might wenzelm@67917: be difficult to type on the console in its YXML syntax wenzelm@67917: (\secref{sec:yxml-vs-xml}). wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsection \Command \<^verbatim>\shutdown\\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67917: \begin{tabular}{ll} wenzelm@67917: regular result: & \<^verbatim>\OK\ \\ wenzelm@67917: \end{tabular} wenzelm@67917: \<^medskip> wenzelm@67917: wenzelm@67917: The \<^verbatim>\shutdown\ command has no argument and result value. It forces a wenzelm@67917: shutdown of the connected server process, stopping all open sessions and wenzelm@67917: closing the server socket. This may disrupt pending commands on other wenzelm@67917: connections! wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67918: The command-line invocation \<^verbatim>\isabelle server -x\ opens a server connection wenzelm@67918: and issues a \<^verbatim>\shutdown\ command (see also \secref{sec:tool-server}). wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsection \Command \<^verbatim>\cancel\\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67917: \begin{tabular}{ll} wenzelm@67917: argument: & \uuid\ \\ wenzelm@67917: regular result: & \<^verbatim>\OK\ \\ wenzelm@67917: \end{tabular} wenzelm@67917: \<^medskip> wenzelm@67917: wenzelm@67918: The command invocation \<^verbatim>\cancel "\\uuid\\<^verbatim>\"\ attempts to cancel the wenzelm@67918: specified task. wenzelm@67917: wenzelm@67917: Cancellation is merely a hint that the client prefers an ongoing process to wenzelm@67917: be stopped. The command always succeeds formally, but it may get ignored by wenzelm@67918: a task that is still running; it might also refer to a non-existing or wenzelm@67918: no-longer existing task (without producing an error). wenzelm@67917: wenzelm@67917: Successful cancellation typically leads to an asynchronous failure of type wenzelm@67918: \<^verbatim>\FAILED {\\task: uuid, message:\~\<^verbatim>\"Interrupt"}\. A different message is wenzelm@67918: also possible, depending how the task handles the event. wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsection \Command \<^verbatim>\session_build\ \label{sec:command-session-build}\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67917: \begin{tabular}{lll} wenzelm@67917: argument: & \session_build_args\ \\ wenzelm@67917: immediate result: & \<^verbatim>\OK\ \task\ \\ wenzelm@67917: notifications: & \<^verbatim>\NOTE\ \task \ (theory_progress | message)\ \\ wenzelm@67917: regular result: & \<^verbatim>\FINISHED\ \task \ session_build_results\ \\ wenzelm@67917: error result: & \<^verbatim>\FAILED\ \task \ error_message \ session_build_results\ \\[2ex] wenzelm@67917: \end{tabular} wenzelm@67917: wenzelm@67917: \begin{tabular}{lll} wenzelm@67917: \<^bold>\type\ \session_build_args =\ \\ wenzelm@67917: \quad\{session: string,\ \\ wenzelm@67917: \quad~~\preferences?: string,\ & \<^bold>\default:\ server preferences \\ wenzelm@67917: \quad~~\options?: [string],\ \\ wenzelm@67917: \quad~~\dirs?: [string],\ \\ wenzelm@67917: \quad~~\ancestor_session?: string,\ \\ wenzelm@67917: \quad~~\all_known?: bool,\ \\ wenzelm@67917: \quad~~\focus_session?: bool,\ \\ wenzelm@67917: \quad~~\required_session?: bool,\ \\ wenzelm@67917: \quad~~\system_mode?: bool,\ \\ wenzelm@67917: \quad~~\verbose?: bool}\ \\[2ex] wenzelm@67917: wenzelm@67917: \<^bold>\type\ \session_build_result =\ \\ wenzelm@67917: \quad\{session: string,\ \\ wenzelm@67917: \quad~~\ok: bool,\ \\ wenzelm@67917: \quad~~\return_code: int,\ \\ wenzelm@67917: \quad~~\timeout: bool,\ \\ wenzelm@67918: \quad~~\timing: timing}\ \\[2ex] wenzelm@67918: wenzelm@67918: \<^bold>\type\ \session_build_results =\ \\ wenzelm@67918: \quad\{ok: bool,\ \\ wenzelm@67918: \quad~~\return_code: int,\ \\ wenzelm@67918: \quad~~\sessions: [session_build_result]}\ \\ wenzelm@67917: \end{tabular} wenzelm@67917: \ wenzelm@67917: wenzelm@67917: text \ wenzelm@67918: The \<^verbatim>\session_build\ command prepares a session image for interactive use of wenzelm@67918: theories. This is a limited version of command-line tool @{tool build} wenzelm@67918: (\secref{sec:tool-build}), with specific options to request a formal context wenzelm@67918: for an interactive PIDE session. wenzelm@67917: wenzelm@67917: The build process is asynchronous, with notifications that inform about the wenzelm@67918: progress of loaded theories. Some further informative messages are output as wenzelm@67918: well. wenzelm@67917: wenzelm@67917: Coordination of independent build processes is at the discretion of the wenzelm@67918: client (or end-user), just as for @{tool build} and @{tool jedit}. There is wenzelm@67918: no built-in coordination of conflicting builds with overlapping hierarchies wenzelm@67918: of session images. In the worst case, a session image produced by one task wenzelm@67918: may get overwritten by another task! wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsubsection \Arguments\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67918: The \session\ field is mandatory. It specifies the target session name: wenzelm@67918: either directly (default) or indirectly (if \required_session\ is \<^verbatim>\true\). wenzelm@67918: The build process will produce all required ancestor images for the wenzelm@67918: specified target. wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67917: The environment of Isabelle system options is determined from \preferences\ wenzelm@67917: that are augmented by \options\, which is a list individual updates of the wenzelm@67917: form the \name\\<^verbatim>\=\\value\ or \name\ (the latter abbreviates wenzelm@67917: \name\\<^verbatim>\=true\); see also command-line option \<^verbatim>\-o\ for @{tool build}. The wenzelm@67917: preferences are loaded from the file wenzelm@67917: \<^path>\$ISABELLE_HOME_USER/etc/preferences\ by default, but the client may wenzelm@67917: provide alternative contents for it (as text, not a file-name). This could wenzelm@67917: be relevant in situations where client and server run in different wenzelm@67917: operating-system contexts. wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67918: The \dirs\ field specifies additional directories for session ROOT and ROOTS wenzelm@67918: files (\secref{sec:session-root}). This augments the name space of available wenzelm@67918: sessions; see also option \<^verbatim>\-d\ in @{tool build}. wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67917: The \ancestor_session\ field specifies an alternative image as starting wenzelm@67917: point for the target image. The default is to use the parent session wenzelm@67917: according to the ROOT entry; see also option \<^verbatim>\-A\ in @{tool jedit}. This wenzelm@67917: can lead to a more light-weight build process, by skipping intermediate wenzelm@67917: session images of the hierarchy that are not used later on. wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67918: The \all_known\ field set to \<^verbatim>\true\ includes all sessions from reachable wenzelm@67918: ROOT entries into the name space of theories. This is relevant for proper wenzelm@67918: session-qualified names, instead of referring to theory file names. The wenzelm@67918: option enables the \<^verbatim>\use_theories\ command wenzelm@67917: (\secref{sec:command-use-theories}) to refer to arbitrary theories from wenzelm@67918: other sessions, but considerable time is required to explore all accessible wenzelm@67918: session directories and theory dependencies. wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67918: The \focus_session\ field set to \<^verbatim>\true\ focuses on the target session: wenzelm@67917: the accessible name space of theories is restricted to sessions that are wenzelm@67918: connected to it in the imports graph. wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67917: The \required_session\ field set to \<^verbatim>\true\ indicates that the target image wenzelm@67917: should not be the \session\, but its parent (or ancestor according to wenzelm@67917: \ancestor_session\). Thus it prepares a session context where theories from wenzelm@67917: the \session\ itself can be loaded. wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67918: The \system_mode\ field set to \<^verbatim>\true\ stores resulting session images and wenzelm@67917: log files in @{path "$ISABELLE_HOME/heaps"} instead of the default location wenzelm@67917: @{setting ISABELLE_OUTPUT} (which is normally in @{setting wenzelm@67918: ISABELLE_HOME_USER}). See also option \<^verbatim>\-s\ in @{tool build}. wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67918: The \verbose\ field set to \<^verbatim>\true\ yields extra verbosity. The effect is wenzelm@67917: similar to option \<^verbatim>\-v\ in @{tool build}. wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsubsection \Intermediate output\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67917: The asynchronous notifications of command \<^verbatim>\session_build\ mainly serve as wenzelm@67917: progress indicator: the output resembles that of the session build window of wenzelm@67917: Isabelle/jEdit after startup @{cite "isabelle-jedit"}. wenzelm@67917: wenzelm@67917: For the client it is usually sufficient to print the messages in plain text, wenzelm@67918: but note that \theory_progress\ also reveals formal \theory\ and wenzelm@67918: \session\ names directly. wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsubsection \Results\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67918: The overall \session_build_results\ contain both a summary and an entry wenzelm@67918: \session_build_result\ for each session in the build hierarchy. The result wenzelm@67918: is always provided, independently of overall success (\<^verbatim>\FINISHED\ task) or wenzelm@67918: failure (\<^verbatim>\FAILED\ task). wenzelm@67917: wenzelm@67917: The \ok\ field tells abstractly, whether all required session builds came wenzelm@67918: out as \ok\, i.e.\ with zero \return_code\. A non-zero \return_code\ wenzelm@67918: indicates an error according to usual POSIX conventions for process exit. wenzelm@67918: wenzelm@67918: The individual \session_build_result\ entries provide extra fields: wenzelm@67917: wenzelm@67918: \<^item> \timeout\ tells if the build process was aborted after running too long, wenzelm@67918: wenzelm@67918: \<^item> \timing\ gives the overall process timing in the usual Isabelle format wenzelm@67918: with elapsed, CPU, GC time. wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsubsection \Examples\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67917: Build of a session image from the Isabelle distribution: wenzelm@67917: @{verbatim [display] \session_build {"session": "HOL-Word"}\} wenzelm@67917: wenzelm@67917: Build a session image from the Archive of Formal Proofs: wenzelm@67917: @{verbatim [display] \session_build {"session": "Coinductive", "dirs": ["$AFP_BASE/thys"]}\} wenzelm@67917: wenzelm@67917: Build of a session image of \<^verbatim>\HOL-Number_Theory\ directly on top of \<^verbatim>\HOL\: wenzelm@67917: @{verbatim [display] \session_build {"session": "HOL-Number_Theory", "ancestor_session": "HOL"}\} wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsection \Command \<^verbatim>\session_start\ \label{sec:command-session-start}\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67917: \begin{tabular}{lll} wenzelm@67917: argument: & \session_build_args \ {print_mode?: [string]}\ \\ wenzelm@67917: immediate result: & \<^verbatim>\OK\ \task\ \\ wenzelm@67917: notifications: & \<^verbatim>\NOTE\ \task \ (theory_progress | message)\ \\ wenzelm@67917: regular result: & \<^verbatim>\FINISHED\ \task \ session_start_result\ \\ wenzelm@67917: error result: & \<^verbatim>\FAILED\ \task \ error_message\ \\[2ex] wenzelm@67917: \end{tabular} wenzelm@67917: wenzelm@67917: \begin{tabular}{l} wenzelm@67917: \<^bold>\type\ \session_start_result = {session: string, session_id: uuid}\ wenzelm@67917: \end{tabular} wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67917: The \<^verbatim>\session_start\ command starts a new Isabelle/PIDE session with wenzelm@67917: underlying Isabelle/ML process, based on a session image that it produces on wenzelm@67918: demand using \<^verbatim>\session_build\. Thus it accepts all \session_build_args\ and wenzelm@67918: produces similar notifications, but the detailed \session_build_results\ are wenzelm@67918: omitted. wenzelm@67917: wenzelm@67917: The session build and startup process is asynchronous: when the task is wenzelm@67918: finished, the session remains active for commands, until a \<^verbatim>\session_stop\ wenzelm@67918: or \<^verbatim>\shutdown\ command is sent to the server. wenzelm@67917: wenzelm@67917: Sessions are independent of client connections: it is possible to start a wenzelm@67917: session and later apply \<^verbatim>\use_theories\ on different connections, as long as wenzelm@67918: the internal session identifier is known: shared theory imports will be used wenzelm@67918: only once (and persist until purged explicitly). wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsubsection \Arguments\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67918: Most arguments are shared with \<^verbatim>\session_build\ wenzelm@67917: (\secref{sec:command-session-build}). wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67917: The \print_mode\ field adds identifiers of print modes to be made active for wenzelm@67917: this session. For example, \<^verbatim>\"print_mode": ["ASCII"]\ prefers ASCII wenzelm@67917: replacement syntax over mathematical Isabelle symbols. See also option \<^verbatim>\-m\ wenzelm@67917: in @{tool process} (\secref{sec:tool-process}). wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsubsection \Results\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67917: The \session\ field provides the active session name for clarity. wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67917: The \session_id\ field provides the internal identification of the session wenzelm@67917: object within the sever process. It can remain active as long as the server wenzelm@67917: is running, independently of the current client connection. wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsection \Examples\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67917: Start a default Isabelle/HOL session: wenzelm@67917: @{verbatim [display] \session_start {"session": "HOL"}\} wenzelm@67918: wenzelm@67918: Start a session from the Archive of Formal Proofs: wenzelm@67918: @{verbatim [display] \session_start {"session": "Coinductive", "dirs": ["$AFP_BASE/thys"]}\} wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsection \Command \<^verbatim>\session_stop\\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67917: \begin{tabular}{ll} wenzelm@67917: argument: & \{session_id?: uuid}\ \\ wenzelm@67917: immediate result: & \<^verbatim>\OK\ \task\ \\ wenzelm@67917: regular result: & \<^verbatim>\FINISHED\ \task \ session_stop_result\ \\ wenzelm@67917: error result: & \<^verbatim>\FAILED\ \task \ error_message \ session_stop_result\ \\[2ex] wenzelm@67917: \end{tabular} wenzelm@67917: wenzelm@67917: \begin{tabular}{l} wenzelm@67917: \<^bold>\type\ \session_stop_result = {ok: bool, return_code: int}\ wenzelm@67917: \end{tabular} wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67918: The \<^verbatim>\session_stop\ command forces a shutdown of the identified PIDE wenzelm@67918: session. This asynchronous tasks usually finishes quickly. Failure only wenzelm@67918: happens in unusual situations, according to the return code of the wenzelm@67917: underlying Isabelle/ML process. wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsubsection \Arguments\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67918: The \session_id\ field provides the UUID originally created by the server wenzelm@67918: for this session. wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsubsection \Results\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67918: The \ok\ field tells abstractly, whether the Isabelle/ML process has wenzelm@67918: terminated properly. wenzelm@67918: wenzelm@67918: The \return_code\ field expresses this information according to usual POSIX wenzelm@67918: conventions for process exit. wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsection \Command \<^verbatim>\use_theories\ \label{sec:command-use-theories}\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67917: \begin{tabular}{ll} wenzelm@67917: argument: & \use_theories_arguments\ \\ wenzelm@67917: regular result: & \<^verbatim>\OK\ \use_theories_results\ \\ wenzelm@67917: \end{tabular} wenzelm@67917: wenzelm@67917: \begin{tabular}{ll} wenzelm@67917: \<^bold>\type\ \theory_name = string | {name: string, pos: position}\ \\ wenzelm@67917: \end{tabular} wenzelm@67917: wenzelm@67917: \begin{tabular}{ll} wenzelm@67917: \<^bold>\type\ \use_theories_arguments =\ \\ wenzelm@67917: \quad\{session_id: uuid,\ \\ wenzelm@67917: \quad~~\theories: [theory_name],\ \\ wenzelm@67917: \quad~~\qualifier?: string,\ \\ wenzelm@67917: \quad~~\master_dir?: string,\ \\ wenzelm@67917: \quad~~\pretty_margin?: double\ & \<^bold>\default:\ \<^verbatim>\76\ \\ wenzelm@67917: \quad~~\unicode_symbols?: bool}\ \\[2ex] wenzelm@67918: \end{tabular} wenzelm@67917: wenzelm@67918: \begin{tabular}{ll} wenzelm@67917: \<^bold>\type\ \node_result =\ \\ wenzelm@67917: \quad\{node_name: string,\ \\ wenzelm@67917: \quad~~\theory: string,\ \\ wenzelm@67917: \quad~~\status: node_status,\ \\ wenzelm@67917: \quad~~\messages: [message]}\ \\[2ex] wenzelm@67917: wenzelm@67917: \<^bold>\type\ \use_theories_results =\ \\ wenzelm@67917: \quad\{ok: bool,\ \\ wenzelm@67917: \quad~~\errors: [message],\ \\ wenzelm@67917: \quad~~\nodes: [node_result]}\ \\[2ex] wenzelm@67917: \end{tabular} wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67918: The \<^verbatim>\use_theories\ command updates the identified session by adding the wenzelm@67918: current version of theory files to it, while dependencies are resolved wenzelm@67918: implicitly. The command succeeds eventually, when all theories have been wenzelm@67918: \<^emph>\consolidated\ in the sense the formal \node_status\ wenzelm@67918: (\secref{sec:json-types}): the outermost command structure has finished (or wenzelm@67918: failed) and the final \<^theory_text>\end\ command of each theory has been checked. wenzelm@67917: wenzelm@67918: Already used theories persist in the session until purged explicitly. This wenzelm@67918: also means that repeated invocations of \<^verbatim>\use_theories\ are idempotent: it wenzelm@67918: could make sense to do that with different values for \pretty_margin\ or wenzelm@67918: \unicode_symbols\ to get different formatting for \errors\ or \messages\. wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsubsection \Arguments\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67917: The \session_id\ is the identifier provided by the server, when the session wenzelm@67917: was created (possibly on a different client connection). wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67918: The \theories\ field specifies theory names as in theory \<^theory_text>\imports\ or in wenzelm@67918: ROOT \<^bold>\theories\. An explicit source position for these theory name wenzelm@67918: specifications may be given, which is occasionally useful for precise error wenzelm@67918: locations. wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67917: The \qualifier\ field provides an alternative session qualifier for theories wenzelm@67917: that are not formally recognized as a member of a particular session. The wenzelm@67917: default is \<^verbatim>\Draft\ as in Isabelle/jEdit. There is rarely a need to change wenzelm@67917: that, as theory nodes are already uniquely identified by their physical wenzelm@67918: file-system location. This already allows reuse of theory base names with wenzelm@67918: the same session qualifier --- as long as these theories are not used wenzelm@67918: together (e.g.\ in \<^theory_text>\imports\). wenzelm@67917: wenzelm@67918: \<^medskip> The \master_dir\ field explicit specifies the formal master directory of wenzelm@67918: the imported theory. Normally this is determined implicitly from the parent wenzelm@67918: directory of the theory file. wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67917: The \pretty_margin\ field specifies the line width for pretty-printing. The wenzelm@67918: default is suitable for classic console output. Formatting happens at the wenzelm@67918: end of \<^verbatim>\use_theories\, when all prover messages are exported to the client. wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67918: The \unicode_symbols\ field set to \<^verbatim>\true\ renders message output for direct wenzelm@67918: output on a Unicode capable channel, ideally with the Isabelle fonts as in wenzelm@67918: Isabelle/jEdit. The default is to keep the symbolic representation of wenzelm@67918: Isabelle text, e.g.\ \<^verbatim>\\\ instead of its rendering as \\\. This means the wenzelm@67918: client needs to perform its own rendering before presenting it to the wenzelm@67918: end-user. wenzelm@67917: \ wenzelm@67917: wenzelm@67918: wenzelm@67917: subsubsection \Results\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67917: The \ok\ field indicates overall success of processing the specified wenzelm@67917: theories with all their dependencies. wenzelm@67917: wenzelm@67918: When \ok\ is \<^verbatim>\false\, the \errors\ field lists all errors cumulatively wenzelm@67918: (including imported theories). The messages contain position information for wenzelm@67918: the original theory nodes. wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67918: The \nodes\ field provides detailed information about each imported theory wenzelm@67918: node. The individual fields are as follows: wenzelm@67917: wenzelm@67917: \<^item> \node_name\: the physical name for the theory node, based on its wenzelm@67918: file-system location; wenzelm@67917: wenzelm@67918: \<^item> \theory\: the logical theory name, e.g.\ \<^verbatim>\HOL-Library.AList\; wenzelm@67917: wenzelm@67917: \<^item> \status\: the overall node status, e.g.\ see the visualization in the wenzelm@67917: \Theories\ panel of Isabelle/jEdit @{cite "isabelle-jedit"}; wenzelm@67917: wenzelm@67918: \<^item> \messages\: the main bulk of prover messages produced in this theory wenzelm@67917: (\<^verbatim>\writeln\, \<^verbatim>\warning\, \<^verbatim>\error\, etc.). wenzelm@67917: \ wenzelm@67917: wenzelm@67917: wenzelm@67917: subsubsection \Examples\ wenzelm@67917: wenzelm@67917: text \ wenzelm@67917: Process some example theory from the Isabelle distribution, within the wenzelm@67917: context of an already started session for Isabelle/HOL (see also wenzelm@67917: \secref{sec:command-session-start}): wenzelm@67917: @{verbatim [display] \use_theories {"session_id": ..., "theories": ["~~/src/HOL/Isar_Examples/Drinker"]}\} wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67917: Process some example theories that import other theories via wenzelm@67917: session-qualified theory names: wenzelm@67917: wenzelm@67917: @{verbatim [display] \session_start {"session": "HOL", "all_known": true} wenzelm@67917: use_theories {"session_id": ..., "theories": ["~~/src/HOL/Unix/Unix"]} wenzelm@67917: session_stop {"session_id": ...}\} wenzelm@67917: wenzelm@67917: The option \all_known\ has been used here to the full name space of wenzelm@67917: session-qualified theory names accessible in this session. This is also the wenzelm@67917: default in Isabelle/jEdit, although it requires significant startup time. wenzelm@67917: wenzelm@67917: \<^medskip> wenzelm@67917: Process some example theories in the context of their (single) parent wenzelm@67917: session: wenzelm@67917: wenzelm@67917: @{verbatim [display] \session_start {"session": "HOL-Library"} wenzelm@67917: use_theories {"session_id": ..., "theories": ["~~/src/HOL/Unix/Unix"]} wenzelm@67917: session_stop {"session_id": ...}\} wenzelm@67904: \ wenzelm@67904: wenzelm@67904: end