more documentation;
authorwenzelm
Wed, 21 Mar 2018 21:50:28 +0100
changeset 67917 d13b2dd20f5e
parent 67916 a72f01c63262
child 67918 7dc204623770
more documentation;
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 @@
 \<close>
 
 
-subsection \<open>Input and output messages\<close>
+subsection \<open>Input and output messages \label{sec:input-output-messages}\<close>
 
 text \<open>
   Server input and output messages have a uniform format as follows:
@@ -346,9 +346,677 @@
   asynchronous messages properly.
 
   The synchronous command \<^verbatim>\<open>cancel\<close>~\<open>id\<close> tells the specified task to terminate
-  prematurely: usually causing a \<^verbatim>\<open>FAILED\<close> result with error message
-  \<^verbatim>\<open>Interrupt\<close>, but this is not guaranteed: the cancel event may come too
-  late or the task may just ignore it.
+  prematurely: usually causing a \<^verbatim>\<open>FAILED\<close> result, but this is not guaranteed:
+  the cancel event may come too late or the task may just ignore it.
+\<close>
+
+
+section \<open>Server commands and results\<close>
+
+text \<open>
+  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>\<open>$ISABELLE_HOME/src/Pure/Tools/server_commands.scala\<close> \\
+  \<^file>\<open>$ISABELLE_HOME/src/Pure/Tools/server.scala\<close>
+  \end{tabular}
+\<close>
+
+
+subsection \<open>Types for JSON values\<close>
+
+text \<open>
+  In order to specify concrete JSON types for command arguments and result
+  messages, the following type definition language shall be used:
+
+  \<^rail>\<open>
+    @{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} '\<oplus>' @{syntax type} |
+      @{syntax type} '|' @{syntax type}
+    ;
+    @{syntax field_type}: @{syntax name} ('?'?) ':' @{syntax type}
+  \<close>
+
+  This is a simplified version of interfaces in
+  TypeScript.\<^footnote>\<open>\<^url>\<open>https://www.typescriptlang.org/docs/handbook/interfaces.html\<close>\<close>
+  The meaning of these types is specified wrt. the implementation in
+  Isabelle/Scala as follows.
+
+  \<^item> A \<open>name\<close> 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 \<open>value\<close> in JSON notation represents the singleton type of the given
+  item. For example, the string \<^verbatim>\<open>"error"\<close> can be used is the type for a slot
+  that is guaranteed to contain that constant.
+
+  \<^item> Type \<open>any\<close> is the super type of all other types: it is an untyped slot in
+  the specification and corresponds to \<^verbatim>\<open>Any\<close> or \<^verbatim>\<open>JSON.T\<close> in Isabelle/Scala.
+
+  \<^item> Type \<open>null\<close> is the type of the improper value \<open>null\<close>; it corresponds to
+  type \<^verbatim>\<open>Null\<close> in Scala and is normally not used in Isabelle/Scala.\<^footnote>\<open>See also
+  ``Null References: The Billion Dollar Mistake'' by Tony Hoare
+  \<^url>\<open>https://www.infoq.com/presentations/Null-References-The-Billion-Dollar-Mistake-Tony-Hoare\<close>.\<close>
+
+  \<^item> Type \<open>bool\<close> is the type of the truth values \<open>true\<close> and \<open>false\<close>; it
+  corresponds to \<^verbatim>\<open>Boolean\<close> in Scala.
+
+  \<^item> Types \<open>int\<close>, \<open>long\<close>, \<open>double\<close> are specific versions of the generic
+  \<open>number\<close> type, corresponding to \<^verbatim>\<open>Int\<close>, \<^verbatim>\<open>Long\<close>, \<^verbatim>\<open>Double\<close> in Scala, but
+  \<^verbatim>\<open>Long\<close> is limited to 53 bit precision.\<^footnote>\<open>Implementations of JSON typically
+  standardize \<open>number\<close> to \<^verbatim>\<open>Double\<close>, which can absorb \<^verbatim>\<open>Int\<close> faithfully, but
+  not all of \<^verbatim>\<open>Long\<close>.\<close>
+
+  \<^item> Type \<open>string\<close> represents Unicode text; it corresponds to type \<^verbatim>\<open>String\<close> in
+  Scala.
+
+  \<^item> Type \<open>[t]\<close> is the array (or list) type over \<open>t\<close>; it corresponds to
+  \<^verbatim>\<open>List[t]\<close> 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>\<open>Option[t]\<close>. For
+  example, \<open>{a: int, b?: string}\<close> corresponds to a Scala case class with
+  arguments \<^verbatim>\<open>a: Int\<close>, \<^verbatim>\<open>b: Option[String]\<close>.
+
+  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>\<open>0\<close> for the number types, \<^verbatim>\<open>false\<close> for \<open>bool\<close>, or the empty string,
+  array, object etc.
+
+  Object types are \<^emph>\<open>permissive\<close> 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 \<open>t\<^sub>1 \<oplus> t\<^sub>2\<close> 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: \<open>{task: string} \<oplus>
+  {ok: bool}\<close> is the equivalent to \<open>{task: string, ok: bool}\<close>.
+
+  \<^item> The type expression \<open>t\<^sub>1 | t\<^sub>2\<close> 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>\<open>[1, 2, 3]\<close> conforms to
+  types \<open>[int]\<close>, \<open>[long]\<close>, \<open>[double]\<close>, \<open>[any]\<close>, \<open>any\<close>.
+
+  Note that JSON objects require field names to be quoted, but the type
+  language omits quotes for clarity. Thus \<^verbatim>\<open>{"a": 42, "b": "xyz"}\<close> conforms to
+  the type \<open>{a: int, b: string}\<close>, for example.
+
+  \<^medskip>
+  The absence of an argument or result is represented by type \<^verbatim>\<open>Unit\<close> in
+  Isabelle/Scala: it is written as empty text in the message \<open>argument\<close>
+  (\secref{sec:input-output-messages}). This is not part of the JSON language.
+
+  Server replies have name tags like \<^verbatim>\<open>OK\<close>, \<^verbatim>\<open>ERROR\<close>: these are used literally
+  together with type specifications to indicate the particular name with the
+  type of its argument, e.g.\ \<^verbatim>\<open>OK\<close>~\<open>[string]\<close> 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>\<open>type\<close>~\<open>position = {line?: int, offset?: int, end_offset?: int, file?:
+  string, id?: long}\<close> describes a source position within Isabelle source text.
+  Only the \<open>line\<close> and \<open>file\<close> fields make immediate sense to external programs.
+  Detailed \<open>offset\<close> and \<open>end_offset\<close> positions are counted according to
+  Isabelle symbols, see @{ML_type Symbol.symbol} in Isabelle/ML @{cite
+  "isabelle-implementation"}. The position \<open>id\<close> belongs to the internal
+  representation of command transactions in the Isabelle/PIDE protocol.
+
+  \<^item> \<^bold>\<open>type\<close>~\<open>message = {kind?: string, message: string, pos?: position}\<close> where
+  the \<open>kind\<close> provides some hint about the role and importance of the message.
+  The main message kinds are \<^verbatim>\<open>writeln\<close> (for regular output), \<^verbatim>\<open>warning\<close>, and
+  \<^verbatim>\<open>error\<close>. The table \<^verbatim>\<open>Markup.messages\<close> in Isabelle/Scala defines further
+  message kinds for more specific applications.
+
+  \<^item> \<^bold>\<open>type\<close>~\<open>error_message = {kind:\<close>~\<^verbatim>\<open>"error"\<close>\<open>, message: string}\<close> refers to
+  error messages in particular. These occur routinely with \<^verbatim>\<open>ERROR\<close> or
+  \<^verbatim>\<open>FAILED\<close> replies, but also as initial command syntax errors (which are
+  omitted in the command specifications below).
+
+  \<^item> \<^bold>\<open>type\<close>~\<open>theory_progress = {kind:\<close>~\<^verbatim>\<open>"writeln"\<close>\<open>, message: string, theory:
+  string, session: string}\<close> 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>\<open>HOL.Nat\<close>) and session name (e.g.\
+  \<^verbatim>\<open>HOL\<close>). Note that some rare theory names lack a proper session prefix, e.g.
+  theory \<^verbatim>\<open>Main\<close> in session \<^verbatim>\<open>HOL\<close>.
+
+  \<^item> \<^bold>\<open>type\<close>~\<open>timing = {elapsed: double, cpu: double, gc: double}\<close> refers to
+  common Isabelle timing information in seconds, usually with a precision of
+  three digits after the point (whole milliseconds).
+
+  \<^item> \<^bold>\<open>type\<close>~\<open>uuid = string\<close> refers to a Universally Unique Identifier (UUID)
+  as plain text.\<^footnote>\<open>See \<^url>\<open>https://www.ietf.org/rfc/rfc4122.txt\<close> and
+  \<^url>\<open>https://docs.oracle.com/javase/8/docs/api/java/util/UUID.html\<close>.\<close> 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>\<open>type\<close>~\<open>task = {task: uuid}\<close> identifies a newly created asynchronous task
+  and thus allows the client to control it by the \<^verbatim>\<open>cancel\<close> command. The same
+  task identification is included in messages and the final result produced by
+  this task.
+\<close>
+
+
+subsection \<open>Command \<^verbatim>\<open>help\<close>\<close>
+
+text \<open>
+  \begin{tabular}{ll}
+  regular result: & \<^verbatim>\<open>OK\<close> \<open>[string]\<close> \\
+  \end{tabular}
+  \<^medskip>
+
+  The \<^verbatim>\<open>help\<close> 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}).
+\<close>
+
+
+subsection \<open>Command \<^verbatim>\<open>echo\<close>\<close>
+
+text \<open>
+  \begin{tabular}{ll}
+  argument: & \<open>any\<close> \\
+  regular result: & \<^verbatim>\<open>OK\<close> \<open>any\<close> \\
+  \end{tabular}
+  \<^medskip>
+
+  The \<^verbatim>\<open>echo\<close> 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>\<open>echo\<close> is actually more general than given above: \<^verbatim>\<open>Unit\<close>,
+  \<^verbatim>\<open>XML.Elem\<close>, \<^verbatim>\<open>JSON.T\<close> are supported uniformly. Note that \<^verbatim>\<open>XML.Elem\<close> might
+  be difficult to type on the console in its YXML syntax
+  (\secref{sec:yxml-vs-xml}).
+\<close>
+
+
+subsection \<open>Command \<^verbatim>\<open>shutdown\<close>\<close>
+
+text \<open>
+  \begin{tabular}{ll}
+  regular result: & \<^verbatim>\<open>OK\<close> \\
+  \end{tabular}
+  \<^medskip>
+
+  The \<^verbatim>\<open>shutdown\<close> 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>\<open>isabelle server -x\<close> opens a server
+  connection and issues a \<^verbatim>\<open>shutdown\<close> command (see also
+  \secref{sec:tool-server}).
+\<close>
+
+
+subsection \<open>Command \<^verbatim>\<open>cancel\<close>\<close>
+
+text \<open>
+  \begin{tabular}{ll}
+  argument: & \<open>uuid\<close> \\
+  regular result: & \<^verbatim>\<open>OK\<close> \\
+  \end{tabular}
+  \<^medskip>
+
+  The command invocation \<^verbatim>\<open>cancel "\<close>\<open>id\<close>\<^verbatim>\<open>"\<close> 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>\<open>FAILED {\<close>\<open>task: string, message: \<close>\<^verbatim>\<open>"Interrupt"}\<close> --- or a slightly
+  different message, depending how the task handles the event.
+\<close>
+
+
+subsection \<open>Command \<^verbatim>\<open>session_build\<close> \label{sec:command-session-build}\<close>
+
+text \<open>
+  \begin{tabular}{lll}
+  argument: & \<open>session_build_args\<close> \\
+  immediate result: & \<^verbatim>\<open>OK\<close> \<open>task\<close> \\
+  notifications: & \<^verbatim>\<open>NOTE\<close> \<open>task \<oplus> (theory_progress | message)\<close> \\
+  regular result: & \<^verbatim>\<open>FINISHED\<close> \<open>task \<oplus> session_build_results\<close> \\
+  error result: & \<^verbatim>\<open>FAILED\<close> \<open>task \<oplus> error_message \<oplus> session_build_results\<close> \\[2ex]
+  \end{tabular}
+
+  \begin{tabular}{lll}
+  \<^bold>\<open>type\<close> \<open>session_build_args =\<close> \\
+  \quad\<open>{session: string,\<close> \\
+  \quad~~\<open>preferences?: string,\<close> & \<^bold>\<open>default:\<close> server preferences \\
+  \quad~~\<open>options?: [string],\<close> \\
+  \quad~~\<open>dirs?: [string],\<close> \\
+  \quad~~\<open>ancestor_session?: string,\<close> \\
+  \quad~~\<open>all_known?: bool,\<close> \\
+  \quad~~\<open>focus_session?: bool,\<close> \\
+  \quad~~\<open>required_session?: bool,\<close> \\
+  \quad~~\<open>system_mode?: bool,\<close> \\
+  \quad~~\<open>verbose?: bool}\<close> \\[2ex]
+
+  \<^bold>\<open>type\<close> \<open>session_build_results =\<close> \\
+  \quad\<open>{ok: bool,\<close> \\
+  \quad~~\<open>return_code: int,\<close> \\
+  \quad~~\<open>sessions: [session_build_result]}\<close> \\[2ex]
+
+  \<^bold>\<open>type\<close> \<open>session_build_result =\<close> \\
+  \quad\<open>{session: string,\<close> \\
+  \quad~~\<open>ok: bool,\<close> \\
+  \quad~~\<open>return_code: int,\<close> \\
+  \quad~~\<open>timeout: bool,\<close> \\
+  \quad~~\<open>timing: timing}\<close> \\
+  \end{tabular}
+\<close>
+
+text \<open>
+  The \<^verbatim>\<open>session_build\<close> 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.
+\<close>
+
+
+subsubsection \<open>Arguments\<close>
+
+text \<open>
+  The field \<open>session\<close> is mandatory. It specifies the target session name:
+  either directly (default) or indirectly (if \<open>required_session\<close> is
+  \<^verbatim>\<open>true\<close>).
+
+  \<^medskip>
+  The environment of Isabelle system options is determined from \<open>preferences\<close>
+  that are augmented by \<open>options\<close>, which is a list individual updates of the
+  form the \<open>name\<close>\<^verbatim>\<open>=\<close>\<open>value\<close> or \<open>name\<close> (the latter abbreviates
+  \<open>name\<close>\<^verbatim>\<open>=true\<close>); see also command-line option \<^verbatim>\<open>-o\<close> for @{tool build}. The
+  preferences are loaded from the file
+  \<^path>\<open>$ISABELLE_HOME_USER/etc/preferences\<close> 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 \<open>dirs\<close> field specifies additional directories for session ROOT files
+  (\secref{sec:session-root}). This augments the name space of available
+  sessions; see also option \<^verbatim>\<open>-d\<close> in @{tool build} or @{tool jedit}.
+
+  \<^medskip>
+  The \<open>ancestor_session\<close> 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>\<open>-A\<close> 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 \<open>all_known\<close> field set to \<^verbatim>\<open>true\<close> 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>\<open>use_theories\<close> 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 \<open>focus_session\<close> field set to \<^verbatim>\<open>true\<close> to focuses on the target session:
+  the accessible name space of theories is restricted to sessions that are
+  connected to it.
+
+  \<^medskip>
+  The \<open>required_session\<close> field set to \<^verbatim>\<open>true\<close> indicates that the target image
+  should not be the \<open>session\<close>, but its parent (or ancestor according to
+  \<open>ancestor_session\<close>). Thus it prepares a session context where theories from
+  the \<open>session\<close> itself can be loaded.
+
+  \<^medskip>
+  The \<open>system_mode\<close> field may be set to \<^verbatim>\<open>true\<close> 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>\<open>-s\<close>
+  in @{tool build} and @{tool jedit}.
+
+  \<^medskip>
+  The \<open>verbose\<close> field may be set to \<^verbatim>\<open>true\<close> for extra verbosity. The effect is
+  similar to option \<^verbatim>\<open>-v\<close> in @{tool build}.
+\<close>
+
+
+subsubsection \<open>Intermediate output\<close>
+
+text \<open>
+  The asynchronous notifications of command \<^verbatim>\<open>session_build\<close> 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 \<open>theory_progress\<close> also reveals the internal \<open>theory\<close> and
+  \<open>session\<close> names.
+\<close>
+
+
+subsubsection \<open>Results\<close>
+
+text \<open>
+  The overall \<open>session_build_results\<close> contain both a summary and and entry
+  \<open>session_build_result\<close> for each session in the build hierarchy, leading up
+  to the specified target session. The result is always provided,
+  independently of overall success (\<^verbatim>\<open>FINISHED\<close> task) or failure (\<^verbatim>\<open>FAILED\<close>
+  task).
+
+  The \<open>ok\<close> field tells abstractly, whether all required session builds came
+  out as \<open>ok\<close>, i.e.\ \<open>return_code\<close>. A non-zero \<open>return_code\<close> indicates an
+  error according to usual POSIX conventions for process exit.
+
+  For individual \<open>session_build_result\<close> entries, there are additional fields:
+  \<open>timeout\<close> to tell if the build process was aborted after running too long,
+  and \<open>timing\<close> for the overall process timing in the usual Isabelle format
+  with elapsed, CPU, gc time.
+\<close>
+
+
+subsubsection \<open>Examples\<close>
+
+text \<open>
+  Build of a session image from the Isabelle distribution:
+  @{verbatim [display] \<open>session_build {"session": "HOL-Word"}\<close>}
+
+  Build a session image from the Archive of Formal Proofs:
+  @{verbatim [display] \<open>session_build {"session": "Coinductive", "dirs": ["$AFP_BASE/thys"]}\<close>}
+
+  Build of a session image of \<^verbatim>\<open>HOL-Number_Theory\<close> directly on top of \<^verbatim>\<open>HOL\<close>:
+  @{verbatim [display] \<open>session_build {"session": "HOL-Number_Theory", "ancestor_session": "HOL"}\<close>}
+\<close>
+
+
+subsection \<open>Command \<^verbatim>\<open>session_start\<close> \label{sec:command-session-start}\<close>
+
+text \<open>
+  \begin{tabular}{lll}
+  argument: & \<open>session_build_args \<oplus> {print_mode?: [string]}\<close> \\
+  immediate result: & \<^verbatim>\<open>OK\<close> \<open>task\<close> \\
+  notifications: & \<^verbatim>\<open>NOTE\<close> \<open>task \<oplus> (theory_progress | message)\<close> \\
+  regular result: & \<^verbatim>\<open>FINISHED\<close> \<open>task \<oplus> session_start_result\<close> \\
+  error result: & \<^verbatim>\<open>FAILED\<close> \<open>task \<oplus> error_message\<close> \\[2ex]
+  \end{tabular}
+
+  \begin{tabular}{l}
+  \<^bold>\<open>type\<close> \<open>session_start_result = {session: string, session_id: uuid}\<close>
+  \end{tabular}
+
+  \<^medskip>
+  The \<^verbatim>\<open>session_start\<close> 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>\<open>session_build\<close>. Thus it accepts all
+  \<open>session_build_args\<close> and produces similar notifications, but the detailed
+  \<open>session_build_results\<close> are omitted.
+
+  The session build and startup process is asynchronous: when the task is
+  finished, the session remains active for commands such as \<^verbatim>\<open>use_theories\<close>
+  (\secref{sec:command-use-theories}), until a \<^verbatim>\<open>session_stop\<close> or \<^verbatim>\<open>shutdown\<close>
+  command is sent to the server.
+
+  Sessions are independent of client connections: it is possible to start a
+  session and later apply \<^verbatim>\<open>use_theories\<close> on different connections, as long as
+  the internal session identifier is known.
+\<close>
+
+
+subsubsection \<open>Arguments\<close>
+
+text \<open>
+  Most of the arguments are the same as for \<^verbatim>\<open>session_build\<close>
+  (\secref{sec:command-session-build}).
+
+  \<^medskip>
+  The \<open>print_mode\<close> field adds identifiers of print modes to be made active for
+  this session. For example, \<^verbatim>\<open>"print_mode": ["ASCII"]\<close> prefers ASCII
+  replacement syntax over mathematical Isabelle symbols. See also option \<^verbatim>\<open>-m\<close>
+  in @{tool process} (\secref{sec:tool-process}).
+\<close>
+
+
+subsubsection \<open>Results\<close>
+
+text \<open>
+  The \<open>session\<close> field provides the active session name for clarity.
+
+  \<^medskip>
+  The \<open>session_id\<close> 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.
+\<close>
+
+
+subsection \<open>Examples\<close>
+
+text \<open>
+  Start a default Isabelle/HOL session:
+  @{verbatim [display] \<open>session_start {"session": "HOL"}\<close>}
+\<close>
+
+
+subsection \<open>Command \<^verbatim>\<open>session_stop\<close>\<close>
+
+text \<open>
+  \begin{tabular}{ll}
+  argument: & \<open>{session_id?: uuid}\<close> \\
+  immediate result: & \<^verbatim>\<open>OK\<close> \<open>task\<close> \\
+  regular result: & \<^verbatim>\<open>FINISHED\<close> \<open>task \<oplus> session_stop_result\<close> \\
+  error result: & \<^verbatim>\<open>FAILED\<close> \<open>task \<oplus> error_message \<oplus> session_stop_result\<close> \\[2ex]
+  \end{tabular}
+
+  \begin{tabular}{l}
+  \<^bold>\<open>type\<close> \<open>session_stop_result = {ok: bool, return_code: int}\<close>
+  \end{tabular}
+
+  \<^medskip>
+  The \<^verbatim>\<open>session_stop\<close> 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.
+\<close>
+
+
+subsubsection \<open>Arguments\<close>
+
+text \<open>
+  The \<open>session_id\<close> field is the UUID originally created by the server for the
+  intended session.
+\<close>
+
+
+subsubsection \<open>Results\<close>
+
+text \<open>
+  The \<open>ok\<close> field tells abstractly, whether the Isabelle/ML process terminated
+  properly. The \<open>return_code\<close> expresses this information according to usual
+  POSIX conventions for process exit.
+\<close>
+
+
+subsection \<open>Command \<^verbatim>\<open>use_theories\<close> \label{sec:command-use-theories}\<close>
+
+text \<open>
+  \begin{tabular}{ll}
+  argument: & \<open>use_theories_arguments\<close> \\
+  regular result: & \<^verbatim>\<open>OK\<close> \<open>use_theories_results\<close> \\
+  \end{tabular}
+
+  \begin{tabular}{ll}
+  \<^bold>\<open>type\<close> \<open>theory_name = string | {name: string, pos: position}\<close> \\
+  \end{tabular}
+
+  \begin{tabular}{ll}
+  \<^bold>\<open>type\<close> \<open>use_theories_arguments =\<close> \\
+  \quad\<open>{session_id: uuid,\<close> \\
+  \quad~~\<open>theories: [theory_name],\<close> \\
+  \quad~~\<open>qualifier?: string,\<close> \\
+  \quad~~\<open>master_dir?: string,\<close> \\
+  \quad~~\<open>pretty_margin?: double\<close> & \<^bold>\<open>default:\<close> \<^verbatim>\<open>76\<close> \\
+  \quad~~\<open>unicode_symbols?: bool}\<close> \\[2ex]
+
+  \<^bold>\<open>type\<close> \<open>node_status =\<close> \\
+  \quad\<open>{ok: bool,\<close> \\
+  \quad~~\<open>total: int,\<close> \\
+  \quad~~\<open>unprocessed: int,\<close> \\
+  \quad~~\<open>running: int,\<close> \\
+  \quad~~\<open>warned: int,\<close> \\
+  \quad~~\<open>failed: int,\<close> \\
+  \quad~~\<open>finished: int,\<close> \\
+  \quad~~\<open>consolidated: bool}\<close> \\[2ex]
+
+  \<^bold>\<open>type\<close> \<open>node_result =\<close> \\
+  \quad\<open>{node_name: string,\<close> \\
+  \quad~~\<open>theory: string,\<close> \\
+  \quad~~\<open>status: node_status,\<close> \\
+  \quad~~\<open>messages: [message]}\<close> \\[2ex]
+
+  \<^bold>\<open>type\<close> \<open>use_theories_results =\<close> \\
+  \quad\<open>{ok: bool,\<close> \\
+  \quad~~\<open>errors: [message],\<close> \\
+  \quad~~\<open>nodes: [node_result]}\<close> \\[2ex]
+  \end{tabular}
+
+  \<^medskip>
+
+  The \<^verbatim>\<open>use_theories\<close> 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>\<open>consolidated\<close> in the sense that the outermost command structure has
+  finished (or failed) and the final \<^theory_text>\<open>end\<close> command of each theory has been
+  checked.
+\<close>
+
+
+subsubsection \<open>Arguments\<close>
+
+text \<open>
+  The \<open>session_id\<close> is the identifier provided by the server, when the session
+  was created (possibly on a different client connection).
+
+  \<^medskip>
+  The \<open>theories\<close> field specifies theory names as in \<^theory_text>\<open>theory imports\<close> or in
+  session ROOT \<^verbatim>\<open>theories\<close>: an explicit source position for these theory name
+  specifications may be given as well, which is occasionally useful for
+  precise error locations.
+
+  \<^medskip>
+  The \<open>qualifier\<close> field provides an alternative session qualifier for theories
+  that are not formally recognized as a member of a particular session. The
+  default is \<^verbatim>\<open>Draft\<close> 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 \<open>master_dir\<close> 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 \<open>pretty_margin\<close> field specifies the line width for pretty-printing. The
+  default is for classic console output. Formatting happens at the end of
+  \<^verbatim>\<open>use_theories\<close>, when all prover messages produced are exported to the
+  client.
+
+  \<^medskip>
+  The \<open>unicode_symbols\<close> field set to \<^verbatim>\<open>true\<close> 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>\<open>\<forall>\<close> instead
+  of its rendering as \<open>\<forall>\<close>. This means the client needs to perform its own
+  rendering to present it to the end-user.
+\<close>
+
+subsubsection \<open>Results\<close>
+
+text \<open>
+  The \<open>ok\<close> field indicates overall success of processing the specified
+  theories with all their dependencies.
+
+  When \<open>ok\<close> is \<^verbatim>\<open>false\<close>, the \<open>errors\<close> field lists all errors cumulatively,
+  including position information for the theory nodes where the error happened
+  (this includes imported theories).
+
+  \<^medskip>
+  The \<open>nodes\<close> field provides more detailed information about each imported
+  theory node. Here the individual fields are as follows:
+
+  \<^item> \<open>node_name\<close>: the physical name for the theory node, based on its
+  file-system location;\<^footnote>\<open>Clients may recognize the originally provided
+  file-names after careful normalization in the file-system of the server.\<close>
+
+  \<^item> \<open>theory\<close>: the logical theory name, e.g.\ \<^verbatim>\<open>HOL-Library.AList\<close> or
+  \<^verbatim>\<open>Draft.Test\<close>;
+
+  \<^item> \<open>status\<close>: the overall node status, e.g.\ see the visualization in the
+  \<open>Theories\<close> panel of Isabelle/jEdit @{cite "isabelle-jedit"};
+
+  \<^item> \<open>messages\<close>: the main bulk of prover messages produced in this theory note
+  (\<^verbatim>\<open>writeln\<close>, \<^verbatim>\<open>warning\<close>, \<^verbatim>\<open>error\<close>, etc.).
+\<close>
+
+
+subsubsection \<open>Examples\<close>
+
+text \<open>
+  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] \<open>use_theories {"session_id": ..., "theories": ["~~/src/HOL/Isar_Examples/Drinker"]}\<close>}
+
+  \<^medskip>
+  Process some example theories that import other theories via
+  session-qualified theory names:
+
+  @{verbatim [display] \<open>session_start {"session": "HOL", "all_known": true}
+use_theories {"session_id": ..., "theories": ["~~/src/HOL/Unix/Unix"]}
+session_stop {"session_id": ...}\<close>}
+
+  The option \<open>all_known\<close> 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] \<open>session_start {"session": "HOL-Library"}
+use_theories {"session_id": ..., "theories": ["~~/src/HOL/Unix/Unix"]}
+session_stop {"session_id": ...}\<close>}
 \<close>
 
 end