more explicit / functional ML version of document model;
authorwenzelm
Sun, 15 Aug 2010 18:41:23 +0200
changeset 38418 9a7af64d71bb
parent 38417 b8922ae21111
child 38419 f9dc924e54f8
more explicit / functional ML version of document model; tuned;
src/Pure/PIDE/document.ML
src/Pure/PIDE/document.scala
src/Pure/ROOT.ML
src/Pure/System/isar_document.ML
src/Pure/System/isar_document.scala
--- a/src/Pure/PIDE/document.ML	Sun Aug 15 14:18:52 2010 +0200
+++ b/src/Pure/PIDE/document.ML	Sun Aug 15 18:41:23 2010 +0200
@@ -2,7 +2,7 @@
     Author:     Makarius
 
 Document as collection of named nodes, each consisting of an editable
-list of commands.
+list of commands, associated with asynchronous execution process.
 *)
 
 signature DOCUMENT =
@@ -12,9 +12,15 @@
   type command_id = id
   type exec_id = id
   val no_id: id
+  val create_id: unit -> id
   val parse_id: string -> id
   val print_id: id -> string
   type edit = string * ((command_id * command_id option) list) option
+  type state
+  val init_state: state
+  val define_command: command_id -> string -> state -> state
+  val edit: version_id -> version_id -> edit list -> state -> (command_id * exec_id) list * state
+  val execute: version_id -> state -> state
 end;
 
 structure Document: DOCUMENT =
@@ -29,15 +35,273 @@
 
 val no_id = 0;
 
+local
+  val id_count = Synchronized.var "id" 0;
+in
+  fun create_id () =
+    Synchronized.change_result id_count
+      (fn i =>
+        let val i' = i + 1
+        in (i', i') end);
+end;
+
 val parse_id = Markup.parse_int;
 val print_id = Markup.print_int;
 
+fun err_dup kind id = error ("Duplicate " ^ kind ^ ": " ^ print_id id);
+fun err_undef kind id = error ("Undefined " ^ kind ^ ": " ^ print_id id);
 
-(* edits *)
+
+
+(** document structure **)
+
+abstype entry = Entry of {next: command_id option, exec: exec_id option}
+  and node = Node of entry Inttab.table  (*unique entries indexed by command_id, start with no_id*)
+  and version = Version of node Graph.T  (*development graph wrt. static imports*)
+with
+
+
+(* command entries *)
+
+fun make_entry next exec = Entry {next = next, exec = exec};
+
+fun the_entry (Node entries) (id: command_id) =
+  (case Inttab.lookup entries id of
+    NONE => err_undef "command entry" id
+  | SOME (Entry entry) => entry);
+
+fun put_entry (id: command_id, entry: entry) (Node entries) =
+  Node (Inttab.update (id, entry) entries);
+
+fun put_entry_exec (id: command_id) exec node =
+  let val {next, ...} = the_entry node id
+  in put_entry (id, make_entry next exec) node end;
+
+fun reset_entry_exec id = put_entry_exec id NONE;
+fun set_entry_exec (id, exec_id) = put_entry_exec id (SOME exec_id);
+
+
+(* iterate entries *)
+
+fun fold_entries id0 f (node as Node entries) =
+  let
+    fun apply NONE x = x
+      | apply (SOME id) x =
+          let val entry = the_entry node id
+          in apply (#next entry) (f (id, entry) x) end;
+  in if Inttab.defined entries id0 then apply (SOME id0) else I end;
+
+fun first_entry P node =
+  let
+    fun first _ NONE = NONE
+      | first prev (SOME id) =
+          let val entry = the_entry node id
+          in if P (id, entry) then SOME (prev, id, entry) else first (SOME id) (#next entry) end;
+  in first NONE (SOME no_id) end;
+
+
+(* modify entries *)
+
+fun insert_after (id: command_id) (id2: command_id) node =
+  let val {next, exec} = the_entry node id in
+    node
+    |> put_entry (id, make_entry (SOME id2) exec)
+    |> put_entry (id2, make_entry next NONE)
+  end;
+
+fun delete_after (id: command_id) node =
+  let val {next, exec} = the_entry node id in
+    (case next of
+      NONE => error ("No next entry to delete: " ^ print_id id)
+    | SOME id2 =>
+        node |>
+          (case #next (the_entry node id2) of
+            NONE => put_entry (id, make_entry NONE exec)
+          | SOME id3 => put_entry (id, make_entry (SOME id3) exec) #> reset_entry_exec id3))
+  end;
+
+
+(* node edits *)
 
 type edit =
   string *  (*node name*)
   ((command_id * command_id option) list) option;  (*NONE: remove, SOME: insert/remove commands*)
 
+val empty_node = Node (Inttab.make [(no_id, make_entry NONE (SOME no_id))]);
+
+fun edit_node (id, SOME id2) = insert_after id id2
+  | edit_node (id, NONE) = delete_after id;
+
+
+(* version operations *)
+
+fun nodes_of (Version nodes) = nodes;
+val node_names_of = Graph.keys o nodes_of;
+
+fun edit_nodes (name, SOME edits) (Version nodes) =
+      Version (nodes
+        |> Graph.default_node (name, empty_node)
+        |> Graph.map_node name (fold edit_node edits))
+  | edit_nodes (name, NONE) (Version nodes) = Version (Graph.del_node name nodes);
+
+val empty_version = Version Graph.empty;
+
+fun the_node version name =
+  Graph.get_node (nodes_of version) name handle Graph.UNDEF _ => empty_node;
+
+fun put_node name node (Version nodes) =
+  Version (Graph.map_node name (K node) nodes);  (* FIXME Graph.UNDEF !? *)
+
 end;
 
+
+
+(** global state -- document structure and execution process **)
+
+abstype state = State of
+ {versions: version Inttab.table,                   (*version_id -> document content*)
+  commands: Toplevel.transition Inttab.table,       (*command_id -> transition function*)
+  execs: Toplevel.state option lazy Inttab.table,   (*exec_id -> execution process*)
+  execution: unit future list}                      (*global execution process*)
+with
+
+fun make_state (versions, commands, execs, execution) =
+  State {versions = versions, commands = commands, execs = execs, execution = execution};
+
+fun map_state f (State {versions, commands, execs, execution}) =
+  make_state (f (versions, commands, execs, execution));
+
+val init_state =
+  make_state (Inttab.make [(no_id, empty_version)],
+    Inttab.make [(no_id, Toplevel.empty)],
+    Inttab.make [(no_id, Lazy.value (SOME Toplevel.toplevel))],
+    []);
+
+
+(* document versions *)
+
+fun define_version (id: version_id) version =
+  map_state (fn (versions, commands, execs, execution) =>
+    let val versions' = Inttab.update_new (id, version) versions
+      handle Inttab.DUP dup => err_dup "document version" dup
+    in (versions', commands, execs, execution) end);
+
+fun the_version (State {versions, ...}) (id: version_id) =
+  (case Inttab.lookup versions id of
+    NONE => err_undef "document version" id
+  | SOME version => version);
+
+
+(* commands *)
+
+fun define_command (id: command_id) text =
+  map_state (fn (versions, commands, execs, execution) =>
+    let
+      val id_string = print_id id;
+      val tr =
+        Position.setmp_thread_data (Position.id_only id_string)
+          (fn () => Outer_Syntax.prepare_command (Position.id id_string) text) ();
+      val commands' =
+        Inttab.update_new (id, Toplevel.put_id id_string tr) commands
+          handle Inttab.DUP dup => err_dup "command" dup;
+    in (versions, commands', execs, execution) end);
+
+fun the_command (State {commands, ...}) (id: command_id) =
+  (case Inttab.lookup commands id of
+    NONE => err_undef "command" id
+  | SOME tr => tr);
+
+
+(* command executions *)
+
+fun define_exec (id: exec_id) exec =
+  map_state (fn (versions, commands, execs, execution) =>
+    let val execs' = Inttab.update_new (id, exec) execs
+      handle Inttab.DUP dup => err_dup "command execution" dup
+    in (versions, commands, execs', execution) end);
+
+fun the_exec (State {execs, ...}) (id: exec_id) =
+  (case Inttab.lookup execs id of
+    NONE => err_undef "command execution" id
+  | SOME exec => exec);
+
+end;
+
+
+
+(** editing **)
+
+(* edit *)
+
+local
+
+fun is_changed node' (id, {next = _, exec}) =
+  (case try (the_entry node') id of
+    NONE => true
+  | SOME {next = _, exec = exec'} => exec' <> exec);
+
+fun new_exec name (id: command_id) (exec_id, updates, state) =
+  let
+    val exec = the_exec state exec_id;
+    val exec_id' = create_id ();
+    val tr = Toplevel.put_id (print_id exec_id') (the_command state id);
+    val exec' =
+      Lazy.lazy (fn () =>
+        (case Lazy.force exec of
+          NONE => NONE
+        | SOME st => Toplevel.run_command name tr st));
+    val state' = define_exec exec_id' exec' state;
+  in (exec_id', (id, exec_id') :: updates, state') end;
+
+in
+
+fun edit (old_id: version_id) (new_id: version_id) edits state =
+  let
+    val old_version = the_version state old_id;
+    val new_version = fold edit_nodes edits old_version;
+
+    fun update_node name (rev_updates, version, st) =
+      let val node = the_node version name in
+        (case first_entry (is_changed (the_node old_version name)) node of
+          NONE => (rev_updates, version, st)
+        | SOME (prev, id, _) =>
+            let
+              val prev_exec = the (#exec (the_entry node (the prev)));
+              val (_, rev_upds, st') =
+                fold_entries id (new_exec name o #1) node (prev_exec, [], st);
+              val node' = fold set_entry_exec rev_upds node;
+            in (rev_upds @ rev_updates, put_node name node' version, st') end)
+      end;
+
+    (* FIXME proper node deps *)
+    val (rev_updates, new_version', state') =
+      fold update_node (node_names_of new_version) ([], new_version, state);
+    val state'' = define_version new_id new_version' state';
+  in (rev rev_updates, state'') end;
+
+end;
+
+
+(* execute *)
+
+fun execute version_id state =
+  state |> map_state (fn (versions, commands, execs, execution) =>
+    let
+      val version = the_version state version_id;
+
+      fun force_exec NONE = ()
+        | force_exec (SOME exec_id) = ignore (Lazy.force (the_exec state exec_id));
+
+      val _ = List.app Future.cancel execution;
+      fun await_cancellation () = uninterruptible (fn _ => Future.join_results) execution;
+
+      val execution' = (* FIXME proper node deps *)
+        node_names_of version |> map (fn name =>
+          Future.fork_pri 1 (fn () =>
+            (await_cancellation ();
+              fold_entries no_id (fn (_, {exec, ...}) => fn () => force_exec exec)
+                (the_node version name) ())));
+    in (versions, commands, execs, execution') end);
+
+end;
+
--- a/src/Pure/PIDE/document.scala	Sun Aug 15 14:18:52 2010 +0200
+++ b/src/Pure/PIDE/document.scala	Sun Aug 15 18:41:23 2010 +0200
@@ -2,7 +2,7 @@
     Author:     Makarius
 
 Document as collection of named nodes, each consisting of an editable
-list of commands.
+list of commands, associated with asynchronous execution process.
 */
 
 package isabelle
@@ -124,7 +124,7 @@
 
 
 
-  /** global state -- accumulated prover results **/
+  /** global state -- document structure and execution process **/
 
   object State
   {
--- a/src/Pure/ROOT.ML	Sun Aug 15 14:18:52 2010 +0200
+++ b/src/Pure/ROOT.ML	Sun Aug 15 18:41:23 2010 +0200
@@ -236,9 +236,9 @@
 use "Thy/term_style.ML";
 use "Thy/thy_output.ML";
 use "Thy/thy_syntax.ML";
-use "PIDE/document.ML";
 use "old_goals.ML";
 use "Isar/outer_syntax.ML";
+use "PIDE/document.ML";
 use "Thy/thy_info.ML";
 
 (*theory and proof operations*)
--- a/src/Pure/System/isar_document.ML	Sun Aug 15 14:18:52 2010 +0200
+++ b/src/Pure/System/isar_document.ML	Sun Aug 15 18:41:23 2010 +0200
@@ -1,295 +1,52 @@
 (*  Title:      Pure/System/isar_document.ML
     Author:     Makarius
 
-Interactive Isar documents, which are structured as follows:
-
-  - history: tree of documents (i.e. changes without merge)
-  - document: graph of nodes (cf. theory files)
-  - node: linear set of commands, potentially with proof structure
+Protocol commands for interactive Isar documents.
 *)
 
 structure Isar_Document: sig end =
 struct
 
-(* unique identifiers *)
-
-local
-  val id_count = Synchronized.var "id" 0;
-in
-  fun create_id () =
-    Synchronized.change_result id_count
-      (fn i =>
-        let val i' = i + 1
-        in (i', i') end);
-end;
-
-fun err_dup kind id = error ("Duplicate " ^ kind ^ ": " ^ Document.print_id id);
-fun err_undef kind id = error ("Undefined " ^ kind ^ ": " ^ Document.print_id id);
-
-
-(** document versions **)
-
-datatype entry = Entry of {next: Document.command_id option, exec: Document.exec_id option};
-type node = entry Inttab.table; (*unique command entries indexed by command_id, start with no_id*)
-type version = node Graph.T;   (*development graph via static imports*)
-
-
-(* command entries *)
-
-fun make_entry next exec = Entry {next = next, exec = exec};
+(* global document state *)
 
-fun the_entry (node: node) (id: Document.command_id) =
-  (case Inttab.lookup node id of
-    NONE => err_undef "command entry" id
-  | SOME (Entry entry) => entry);
-
-fun put_entry (id: Document.command_id, entry: entry) = Inttab.update (id, entry);
-
-fun put_entry_exec (id: Document.command_id) exec (node: node) =
-  let val {next, ...} = the_entry node id
-  in put_entry (id, make_entry next exec) node end;
-
-fun reset_entry_exec id = put_entry_exec id NONE;
-fun set_entry_exec (id, exec_id) = put_entry_exec id (SOME exec_id);
-
-
-(* iterate entries *)
-
-fun fold_entries id0 f (node: node) =
-  let
-    fun apply NONE x = x
-      | apply (SOME id) x =
-          let val entry = the_entry node id
-          in apply (#next entry) (f (id, entry) x) end;
-  in if Inttab.defined node id0 then apply (SOME id0) else I end;
-
-fun first_entry P (node: node) =
-  let
-    fun first _ NONE = NONE
-      | first prev (SOME id) =
-          let val entry = the_entry node id
-          in if P (id, entry) then SOME (prev, id, entry) else first (SOME id) (#next entry) end;
-  in first NONE (SOME Document.no_id) end;
+local val global_state = Unsynchronized.ref Document.init_state in
 
-
-(* modify entries *)
-
-fun insert_after (id: Document.command_id) (id2: Document.command_id) (node: node) =
-  let val {next, exec} = the_entry node id in
-    node
-    |> put_entry (id, make_entry (SOME id2) exec)
-    |> put_entry (id2, make_entry next NONE)
-  end;
-
-fun delete_after (id: Document.command_id) (node: node) =
-  let val {next, exec} = the_entry node id in
-    (case next of
-      NONE => error ("No next entry to delete: " ^ Document.print_id id)
-    | SOME id2 =>
-        node |>
-          (case #next (the_entry node id2) of
-            NONE => put_entry (id, make_entry NONE exec)
-          | SOME id3 => put_entry (id, make_entry (SOME id3) exec) #> reset_entry_exec id3))
-  end;
-
-
-(* node operations *)
-
-val empty_node: node = Inttab.make [(Document.no_id, make_entry NONE (SOME Document.no_id))];
-
-fun the_node (version: version) name =
-  Graph.get_node version name handle Graph.UNDEF _ => empty_node;
+fun change_state f =
+  NAMED_CRITICAL "Isar_Document" (fn () => Unsynchronized.change global_state f);
 
-fun edit_node (id, SOME id2) = insert_after id id2
-  | edit_node (id, NONE) = delete_after id;
-
-fun edit_nodes (name, SOME edits) =
-      Graph.default_node (name, empty_node) #>
-      Graph.map_node name (fold edit_node edits)
-  | edit_nodes (name, NONE) = Graph.del_node name;
-
-
-
-(** global configuration **)
-
-(* command executions *)
-
-local
-
-val global_execs =
-  Unsynchronized.ref (Inttab.make [(Document.no_id, Lazy.value (SOME Toplevel.toplevel))]);
-
-in
-
-fun define_exec (id: Document.exec_id) exec =
-  NAMED_CRITICAL "Isar" (fn () =>
-    Unsynchronized.change global_execs (Inttab.update_new (id, exec))
-      handle Inttab.DUP dup => err_dup "exec" dup);
-
-fun the_exec (id: Document.exec_id) =
-  (case Inttab.lookup (! global_execs) id of
-    NONE => err_undef "exec" id
-  | SOME exec => exec);
+fun current_state () = ! global_state;
 
 end;
 
 
-(* commands *)
-
-local
-
-val global_commands = Unsynchronized.ref (Inttab.make [(Document.no_id, Toplevel.empty)]);
-
-in
+(* define command *)
 
-fun define_command (id: Document.command_id) text =
-  let
-    val id_string = Document.print_id id;
-    val tr =
-      Position.setmp_thread_data (Position.id_only id_string) (fn () =>
-        Outer_Syntax.prepare_command (Position.id id_string) text) ();
-  in
-    NAMED_CRITICAL "Isar" (fn () =>
-      Unsynchronized.change global_commands (Inttab.update_new (id, Toplevel.put_id id_string tr))
-        handle Inttab.DUP dup => err_dup "command" dup)
-  end;
-
-fun the_command (id: Document.command_id) =
-  (case Inttab.lookup (! global_commands) id of
-    NONE => err_undef "command" id
-  | SOME tr => tr);
-
-end;
+val _ =
+  Isabelle_Process.add_command "Isar_Document.define_command"
+    (fn [id, text] => change_state (Document.define_command (Document.parse_id id) text));
 
 
-(* document versions *)
-
-local
-
-val global_versions = Unsynchronized.ref (Inttab.make [(Document.no_id, Graph.empty: version)]);
-
-in
+(* edit document version *)
 
-fun define_version (id: Document.version_id) version =
-  NAMED_CRITICAL "Isar" (fn () =>
-    Unsynchronized.change global_versions (Inttab.update_new (id, version))
-      handle Inttab.DUP dup => err_dup "version" dup);
+val _ =
+  Isabelle_Process.add_command "Isar_Document.edit_version"
+    (fn [old_id_string, new_id_string, edits_yxml] => change_state (fn state =>
+      let
+        val old_id = Document.parse_id old_id_string;
+        val new_id = Document.parse_id new_id_string;
+        val edits =
+          XML_Data.dest_list (XML_Data.dest_pair XML_Data.dest_string
+            (XML_Data.dest_option (XML_Data.dest_list
+                (XML_Data.dest_pair XML_Data.dest_int
+                  (XML_Data.dest_option XML_Data.dest_int))))) (YXML.parse_body edits_yxml);
 
-fun the_version (id: Document.version_id) =
-  (case Inttab.lookup (! global_versions) id of
-    NONE => err_undef "version" id
-  | SOME version => version);
+      val (updates, state') = Document.edit old_id new_id edits state;
+      val _ =
+        implode (map (fn (id, exec_id) => Markup.markup (Markup.edit id exec_id) "") updates)
+        |> Markup.markup (Markup.assign new_id)
+        |> Output.status;
+      val state'' = Document.execute new_id state';
+    in state'' end));
 
 end;
 
-
-
-(** document editing **)
-
-(* execution *)
-
-local
-
-val execution: unit future list Unsynchronized.ref = Unsynchronized.ref [];
-
-fun force_exec NONE = ()
-  | force_exec (SOME exec_id) = ignore (Lazy.force (the_exec exec_id));
-
-in
-
-fun execute version =
-  NAMED_CRITICAL "Isar" (fn () =>
-    let
-      val old_execution = ! execution;
-      val _ = List.app Future.cancel old_execution;
-      fun await_cancellation () = uninterruptible (K Future.join_results) old_execution;
-      (* FIXME proper node deps *)
-      val new_execution = Graph.keys version |> map (fn name =>
-        Future.fork_pri 1 (fn () =>
-          let
-            val _ = await_cancellation ();
-            val exec =
-              fold_entries Document.no_id (fn (_, {exec, ...}) => fn () => force_exec exec)
-                (the_node version name);
-          in exec () end));
-    in execution := new_execution end);
-
-end;
-
-
-(* editing *)
-
-local
-
-fun is_changed node' (id, {next = _, exec}) =
-  (case try (the_entry node') id of
-    NONE => true
-  | SOME {next = _, exec = exec'} => exec' <> exec);
-
-fun new_exec name (id: Document.command_id) (exec_id, updates) =
-  let
-    val exec = the_exec exec_id;
-    val exec_id' = create_id ();
-    val tr = Toplevel.put_id (Document.print_id exec_id') (the_command id);
-    val exec' =
-      Lazy.lazy (fn () =>
-        (case Lazy.force exec of
-          NONE => NONE
-        | SOME st => Toplevel.run_command name tr st));
-    val _ = define_exec exec_id' exec';
-  in (exec_id', (id, exec_id') :: updates) end;
-
-fun updates_status new_id updates =
-  implode (map (fn (id, exec_id) => Markup.markup (Markup.edit id exec_id) "") updates)
-  |> Markup.markup (Markup.assign new_id)
-  |> Output.status;
-
-in
-
-fun edit_document (old_id: Document.version_id) (new_id: Document.version_id) edits =
-  NAMED_CRITICAL "Isar" (fn () =>
-    let
-      val old_version = the_version old_id;
-      val new_version = fold edit_nodes edits old_version;
-
-      fun update_node name node =
-        (case first_entry (is_changed (the_node old_version name)) node of
-          NONE => ([], node)
-        | SOME (prev, id, _) =>
-            let
-              val prev_exec_id = the (#exec (the_entry node (the prev)));
-              val (_, updates) = fold_entries id (new_exec name o #1) node (prev_exec_id, []);
-              val node' = fold set_entry_exec updates node;
-            in (rev updates, node') end);
-
-      (* FIXME proper node deps *)
-      val (updatess, new_version') =
-        (Graph.keys new_version, new_version)
-          |-> fold_map (fn name => Graph.map_node_yield name (update_node name));
-
-      val _ = define_version new_id new_version';
-      val _ = updates_status new_id (flat updatess);
-      val _ = execute new_version';
-    in () end);
-
-end;
-
-
-
-(** Isabelle process commands **)
-
-val _ =
-  Isabelle_Process.add_command "Isar_Document.define_command"
-    (fn [id, text] => define_command (Document.parse_id id) text);
-
-val _ =
-  Isabelle_Process.add_command "Isar_Document.edit_version"
-    (fn [old_id, new_id, edits] =>
-      edit_document (Document.parse_id old_id) (Document.parse_id new_id)
-        (XML_Data.dest_list (XML_Data.dest_pair XML_Data.dest_string
-            (XML_Data.dest_option (XML_Data.dest_list
-                (XML_Data.dest_pair XML_Data.dest_int
-                  (XML_Data.dest_option XML_Data.dest_int))))) (YXML.parse_body edits)));
-
-end;
-
--- a/src/Pure/System/isar_document.scala	Sun Aug 15 14:18:52 2010 +0200
+++ b/src/Pure/System/isar_document.scala	Sun Aug 15 18:41:23 2010 +0200
@@ -1,7 +1,7 @@
 /*  Title:      Pure/System/isar_document.scala
     Author:     Makarius
 
-Interactive Isar documents.
+Protocol commands for interactive Isar documents.
 */
 
 package isabelle