merged
authorwenzelm
Mon, 05 Jul 2010 10:47:39 +0200
changeset 37707 764d57a3a28d
parent 37706 c63649d8d75b (current diff)
parent 37690 b16231572c61 (diff)
child 37708 694815d76240
merged
--- a/src/Pure/Concurrent/future.ML	Mon Jul 05 10:42:27 2010 +0200
+++ b/src/Pure/Concurrent/future.ML	Mon Jul 05 10:47:39 2010 +0200
@@ -59,6 +59,7 @@
   val cancel_group: group -> unit
   val cancel: 'a future -> unit
   val shutdown: unit -> unit
+  val report: (unit -> 'a) -> 'a
 end;
 
 structure Future: FUTURE =
@@ -284,7 +285,7 @@
       if forall (Thread.isActive o #1) (! workers) then ()
       else
         let
-          val  (alive, dead) = List.partition (Thread.isActive o #1) (! workers);
+          val (alive, dead) = List.partition (Thread.isActive o #1) (! workers);
           val _ = workers := alive;
         in
           Multithreading.tracing 0 (fn () =>
@@ -523,6 +524,16 @@
   else ();
 
 
+(* report markup *)
+
+fun report e =
+  let
+    val _ = Output.status (Markup.markup Markup.forked "");
+    val x = e ();  (*sic -- report "joined" only for success*)
+    val _ = Output.status (Markup.markup Markup.joined "");
+  in x end;
+
+
 (*final declarations of this structure!*)
 val map = map_future;
 
--- a/src/Pure/Isar/toplevel.ML	Mon Jul 05 10:42:27 2010 +0200
+++ b/src/Pure/Isar/toplevel.ML	Mon Jul 05 10:47:39 2010 +0200
@@ -86,9 +86,9 @@
   val error_msg: transition -> exn * string -> unit
   val add_hook: (transition -> state -> state -> unit) -> unit
   val transition: bool -> transition -> state -> (state * (exn * string) option) option
+  val run_command: string -> transition -> state -> state option
   val commit_exit: Position.T -> transition
   val command: transition -> state -> state
-  val run_command: string -> transition -> state -> state option
   val excursion: (transition * transition list) list -> (transition * state) list lazy
 end;
 
@@ -561,6 +561,12 @@
 fun status tr m =
   setmp_thread_position tr (fn () => Output.status (Markup.markup m "")) ();
 
+fun async_state (tr as Transition {print, ...}) st =
+  if print then
+    ignore (Future.fork_group (Task_Queue.new_group (Future.worker_group ()))
+      (fn () => Future.report (setmp_thread_position tr (fn () => print_state false st))))
+  else ();
+
 fun error_msg tr exn_info =
   setmp_thread_position tr (fn () =>
     Output.error_msg (ML_Compiler.exn_message (EXCURSION_FAIL exn_info))) ();
@@ -614,6 +620,22 @@
 end;
 
 
+(* managed execution *)
+
+fun run_command thy_name (tr as Transition {print, ...}) st =
+  (case
+      (case init_of tr of
+        SOME name => Exn.capture (fn () => Thy_Load.check_name thy_name name) ()
+      | NONE => Exn.Result ()) of
+    Exn.Result () =>
+      (case transition false tr st of
+        SOME (st', NONE) => (status tr Markup.finished; async_state tr st'; SOME st')
+      | SOME (_, SOME (exn as Exn.Interrupt, _)) => reraise exn
+      | SOME (_, SOME exn_info) => (error_msg tr exn_info; status tr Markup.failed; NONE)
+      | NONE => (error_msg tr (TERMINATE, at_command tr); status tr Markup.failed; NONE))
+  | Exn.Exn exn => (error_msg tr (exn, at_command tr); status tr Markup.failed; NONE));
+
+
 (* commit final exit *)
 
 fun commit_exit pos =
@@ -635,19 +657,6 @@
   let val st' = command tr st
   in (st', st') end;
 
-fun run_command thy_name tr st =
-  (case
-      (case init_of tr of
-        SOME name => Exn.capture (fn () => Thy_Load.check_name thy_name name) ()
-      | NONE => Exn.Result ()) of
-    Exn.Result () =>
-      (case transition true tr st of
-        SOME (st', NONE) => (status tr Markup.finished; SOME st')
-      | SOME (_, SOME (exn as Exn.Interrupt, _)) => reraise exn
-      | SOME (_, SOME exn_info) => (error_msg tr exn_info; status tr Markup.failed; NONE)
-      | NONE => (error_msg tr (TERMINATE, at_command tr); status tr Markup.failed; NONE))
-  | Exn.Exn exn => (error_msg tr (exn, at_command tr); status tr Markup.failed; NONE));
-
 
 (* excursion of units, consisting of commands with proof *)
 
--- a/src/Pure/PIDE/document.scala	Mon Jul 05 10:42:27 2010 +0200
+++ b/src/Pure/PIDE/document.scala	Mon Jul 05 10:47:39 2010 +0200
@@ -14,13 +14,13 @@
 {
   /* command start positions */
 
-  def command_starts(commands: Linear_Set[Command]): Iterator[(Command, Int)] =
+  def command_starts(commands: Iterator[Command], offset: Int = 0): Iterator[(Command, Int)] =
   {
-    var offset = 0
-    for (cmd <- commands.iterator) yield {
-      val start = offset
-      offset += cmd.length
-      (cmd, start)
+    var i = offset
+    for (command <- commands) yield {
+      val start = i
+      i += command.length
+      (command, start)
     }
   }
 
@@ -60,9 +60,10 @@
     {
       eds match {
         case e :: es =>
-          command_starts(commands).find {   // FIXME relative search!
+          command_starts(commands.iterator).find {
             case (cmd, cmd_start) =>
-              e.can_edit(cmd.source, cmd_start) || e.is_insert && e.start == cmd_start + cmd.length
+              e.can_edit(cmd.source, cmd_start) ||
+                e.is_insert && e.start == cmd_start + cmd.length
           } match {
             case Some((cmd, cmd_start)) if e.can_edit(cmd.source, cmd_start) =>
               val (rest, text) = e.edit(cmd.source, cmd_start)
@@ -144,7 +145,7 @@
 {
   /* command ranges */
 
-  def command_starts: Iterator[(Command, Int)] = Document.command_starts(commands)
+  def command_starts: Iterator[(Command, Int)] = Document.command_starts(commands.iterator)
 
   def command_start(cmd: Command): Option[Int] =
     command_starts.find(_._1 == cmd).map(_._2)
--- a/src/Pure/Syntax/parser.ML	Mon Jul 05 10:42:27 2010 +0200
+++ b/src/Pure/Syntax/parser.ML	Mon Jul 05 10:47:39 2010 +0200
@@ -8,9 +8,9 @@
 sig
   type gram
   val empty_gram: gram
-  val extend_gram: gram -> Syn_Ext.xprod list -> gram
+  val extend_gram: Syn_Ext.xprod list -> gram -> gram
   val make_gram: Syn_Ext.xprod list -> gram
-  val merge_grams: gram -> gram -> gram
+  val merge_gram: gram * gram -> gram
   val pretty_gram: gram -> Pretty.T list
   datatype parsetree =
     Node of string * parsetree list |
@@ -23,19 +23,16 @@
 structure Parser: PARSER =
 struct
 
-open Lexicon Syn_Ext;
-
-
 (** datatype gram **)
 
 type nt_tag = int;              (*production for the NTs are stored in an array
                                   so we can identify NTs by their index*)
 
-datatype symb = Terminal of token
+datatype symb = Terminal of Lexicon.token
               | Nonterminal of nt_tag * int;              (*(tag, precedence)*)
 
-type nt_gram = ((nt_tag list * token list) *
-                (token option * (symb list * string * int) list) list);
+type nt_gram = ((nt_tag list * Lexicon.token list) *
+                (Lexicon.token option * (symb list * string * int) list) list);
                                      (*(([dependent_nts], [start_tokens]),
                                         [(start_token, [(rhs, name, prio)])])*)
                               (*depent_nts is a list of all NTs whose lookahead
@@ -53,8 +50,8 @@
                          lambda productions are stored as normal productions
                          and also as an entry in "lambdas"*)
 
-val UnknownStart = eof;       (*productions for which no starting token is
-                                known yet are associated with this token*)
+val UnknownStart = Lexicon.eof;       (*productions for which no starting token is
+                                        known yet are associated with this token*)
 
 (* get all NTs that are connected with a list of NTs
    (used for expanding chain list)*)
@@ -125,7 +122,7 @@
       (*get all known starting tokens for a nonterminal*)
       fun starts_for_nt nt = snd (fst (Array.sub (prods, nt)));
 
-      val token_union = uncurry (union matching_tokens);
+      val token_union = uncurry (union Lexicon.matching_tokens);
 
       (*update prods, lookaheads, and lambdas according to new lambda NTs*)
       val (added_starts', lambdas') =
@@ -158,7 +155,7 @@
                         val nt_dependencies' = union (op =) nts nt_dependencies;
 
                         (*associate production with new starting tokens*)
-                        fun copy ([]: token option list) nt_prods = nt_prods
+                        fun copy ([]: Lexicon.token option list) nt_prods = nt_prods
                           | copy (tk :: tks) nt_prods =
                             let val old_prods = these (AList.lookup (op =) nt_prods tk);
 
@@ -259,7 +256,7 @@
               let
                 val ((old_nts, old_tks), nt_prods) = Array.sub (prods, nt);
 
-                val new_tks = subtract matching_tokens old_tks start_tks;
+                val new_tks = subtract Lexicon.matching_tokens old_tks start_tks;
 
                 (*store new production*)
                 fun store [] prods is_new =
@@ -278,7 +275,7 @@
                           else (new_prod :: tk_prods, true);
 
                         val prods' = prods
-                          |> is_new' ? AList.update (op =) (tk: token option, tk_prods');
+                          |> is_new' ? AList.update (op =) (tk: Lexicon.token option, tk_prods');
                     in store tks prods' (is_new orelse is_new') end;
 
                 val (nt_prods', prod_count', changed) =
@@ -329,10 +326,10 @@
                                    andalso member (op =) new_tks (the tk);
 
                       (*associate production with new starting tokens*)
-                      fun copy ([]: token list) nt_prods = nt_prods
+                      fun copy ([]: Lexicon.token list) nt_prods = nt_prods
                         | copy (tk :: tks) nt_prods =
                           let
-                            val tk_prods = (these o AList.lookup (op =) nt_prods) (SOME tk);
+                            val tk_prods = these (AList.lookup (op =) nt_prods (SOME tk));
 
                             val tk_prods' =
                               if not lambda then p :: tk_prods
@@ -359,7 +356,7 @@
                       val (lookahead as (old_nts, old_tks), nt_prods) =
                         Array.sub (prods, nt);
 
-                      val tk_prods = (these o AList.lookup (op =) nt_prods) key;
+                      val tk_prods = these (AList.lookup (op =) nt_prods key);
 
                       (*associate productions with new lookahead tokens*)
                       val (tk_prods', nt_prods') =
@@ -370,7 +367,7 @@
                         |> (key = SOME UnknownStart) ? AList.update (op =) (key, tk_prods')
 
                       val added_tks =
-                        subtract matching_tokens old_tks new_tks;
+                        subtract Lexicon.matching_tokens old_tks new_tks;
                     in if null added_tks then
                          (Array.update (prods, nt, (lookahead, nt_prods'));
                           process_nts nts added)
@@ -381,7 +378,7 @@
                     end;
 
                 val ((dependent, _), _) = Array.sub (prods, changed_nt);
-              in add_starts (starts @ (process_nts dependent [])) end;
+              in add_starts (starts @ process_nts dependent []) end;
         in add_starts added_starts' end;
   in add_prods prods chains' lambdas' prod_count ps end;
 
@@ -394,8 +391,8 @@
 
     val nt_name = the o Inttab.lookup (Inttab.make (map swap (Symtab.dest tags)));
 
-    fun pretty_symb (Terminal (Token (Literal, s, _))) = Pretty.quote (Pretty.str s)
-      | pretty_symb (Terminal tok) = Pretty.str (str_of_token tok)
+    fun pretty_symb (Terminal (Lexicon.Token (Lexicon.Literal, s, _))) = Pretty.quote (Pretty.str s)
+      | pretty_symb (Terminal tok) = Pretty.str (Lexicon.str_of_token tok)
       | pretty_symb (Nonterminal (tag, p)) =
           Pretty.str (nt_name tag ^ "[" ^ signed_string_of_int p ^ "]");
 
@@ -422,7 +419,6 @@
 
 (** Operations on gramars **)
 
-(*The mother of all grammars*)
 val empty_gram = Gram {nt_count = 0, prod_count = 0,
                        tags = Symtab.empty, chains = [], lambdas = [],
                        prods = Array.array (0, (([], []), []))};
@@ -439,75 +435,75 @@
 
 
 (*Add productions to a grammar*)
-fun extend_gram gram [] = gram
-  | extend_gram (Gram {nt_count, prod_count, tags, chains, lambdas, prods})
-                xprods =
-  let
-    (*Get tag for existing nonterminal or create a new one*)
-    fun get_tag nt_count tags nt =
-      case Symtab.lookup tags nt of
-        SOME tag => (nt_count, tags, tag)
-      | NONE => (nt_count+1, Symtab.update_new (nt, nt_count) tags,
-                 nt_count);
+fun extend_gram [] gram = gram
+  | extend_gram xprods (Gram {nt_count, prod_count, tags, chains, lambdas, prods}) =
+    let
+      (*Get tag for existing nonterminal or create a new one*)
+      fun get_tag nt_count tags nt =
+        case Symtab.lookup tags nt of
+          SOME tag => (nt_count, tags, tag)
+        | NONE => (nt_count+1, Symtab.update_new (nt, nt_count) tags,
+                   nt_count);
 
-    (*Convert symbols to the form used by the parser;
-      delimiters and predefined terms are stored as terminals,
-      nonterminals are converted to integer tags*)
-    fun symb_of [] nt_count tags result = (nt_count, tags, rev result)
-      | symb_of ((Delim s) :: ss) nt_count tags result =
-          symb_of ss nt_count tags (Terminal (Token (Literal, s, Position.no_range)) :: result)
-      | symb_of ((Argument (s, p)) :: ss) nt_count tags result =
-          let
-            val (nt_count', tags', new_symb) =
-              case predef_term s of
-                NONE =>
-                  let val (nt_count', tags', s_tag) = get_tag nt_count tags s;
-                  in (nt_count', tags', Nonterminal (s_tag, p)) end
-              | SOME tk => (nt_count, tags, Terminal tk);
-          in symb_of ss nt_count' tags' (new_symb :: result) end
-      | symb_of (_ :: ss) nt_count tags result =
-          symb_of ss nt_count tags result;
+      (*Convert symbols to the form used by the parser;
+        delimiters and predefined terms are stored as terminals,
+        nonterminals are converted to integer tags*)
+      fun symb_of [] nt_count tags result = (nt_count, tags, rev result)
+        | symb_of ((Syn_Ext.Delim s) :: ss) nt_count tags result =
+            symb_of ss nt_count tags
+              (Terminal (Lexicon.Token (Lexicon.Literal, s, Position.no_range)) :: result)
+        | symb_of ((Syn_Ext.Argument (s, p)) :: ss) nt_count tags result =
+            let
+              val (nt_count', tags', new_symb) =
+                case Lexicon.predef_term s of
+                  NONE =>
+                    let val (nt_count', tags', s_tag) = get_tag nt_count tags s;
+                    in (nt_count', tags', Nonterminal (s_tag, p)) end
+                | SOME tk => (nt_count, tags, Terminal tk);
+            in symb_of ss nt_count' tags' (new_symb :: result) end
+        | symb_of (_ :: ss) nt_count tags result =
+            symb_of ss nt_count tags result;
 
-    (*Convert list of productions by invoking symb_of for each of them*)
-    fun prod_of [] nt_count prod_count tags result =
-          (nt_count, prod_count, tags, result)
-      | prod_of ((XProd (lhs, xsymbs, const, pri)) :: ps)
-                nt_count prod_count tags result =
-        let val (nt_count', tags', lhs_tag) = get_tag nt_count tags lhs;
+      (*Convert list of productions by invoking symb_of for each of them*)
+      fun prod_of [] nt_count prod_count tags result =
+            (nt_count, prod_count, tags, result)
+        | prod_of ((Syn_Ext.XProd (lhs, xsymbs, const, pri)) :: ps)
+                  nt_count prod_count tags result =
+          let val (nt_count', tags', lhs_tag) = get_tag nt_count tags lhs;
 
-            val (nt_count'', tags'', prods) =
-              symb_of xsymbs nt_count' tags' [];
-        in prod_of ps nt_count'' (prod_count+1) tags''
-                   ((lhs_tag, (prods, const, pri)) :: result)
-        end;
+              val (nt_count'', tags'', prods) =
+                symb_of xsymbs nt_count' tags' [];
+          in prod_of ps nt_count'' (prod_count+1) tags''
+                     ((lhs_tag, (prods, const, pri)) :: result)
+          end;
 
-    val (nt_count', prod_count', tags', xprods') =
-      prod_of xprods nt_count prod_count tags [];
+      val (nt_count', prod_count', tags', xprods') =
+        prod_of xprods nt_count prod_count tags [];
 
-    (*Copy array containing productions of old grammar;
-      this has to be done to preserve the old grammar while being able
-      to change the array's content*)
-    val prods' =
-      let fun get_prod i = if i < nt_count then Array.sub (prods, i)
-                           else (([], []), []);
-      in Array.tabulate (nt_count', get_prod) end;
+      (*Copy array containing productions of old grammar;
+        this has to be done to preserve the old grammar while being able
+        to change the array's content*)
+      val prods' =
+        let fun get_prod i = if i < nt_count then Array.sub (prods, i)
+                             else (([], []), []);
+        in Array.tabulate (nt_count', get_prod) end;
 
-    val fromto_chains = inverse_chains chains [];
+      val fromto_chains = inverse_chains chains [];
 
-    (*Add new productions to old ones*)
-    val (fromto_chains', lambdas', _) =
-      add_prods prods' fromto_chains lambdas NONE xprods';
+      (*Add new productions to old ones*)
+      val (fromto_chains', lambdas', _) =
+        add_prods prods' fromto_chains lambdas NONE xprods';
 
-    val chains' = inverse_chains fromto_chains' [];
-  in Gram {nt_count = nt_count', prod_count = prod_count', tags = tags',
-           chains = chains', lambdas = lambdas', prods = prods'}
-  end;
+      val chains' = inverse_chains fromto_chains' [];
+    in Gram {nt_count = nt_count', prod_count = prod_count', tags = tags',
+             chains = chains', lambdas = lambdas', prods = prods'}
+    end;
 
-val make_gram = extend_gram empty_gram;
+fun make_gram xprods = extend_gram xprods empty_gram;
 
 
 (*Merge two grammars*)
-fun merge_grams gram_a gram_b =
+fun merge_gram (gram_a, gram_b) =
   let
     (*find out which grammar is bigger*)
     val (Gram {nt_count = nt_count1, prod_count = prod_count1, tags = tags1,
@@ -604,7 +600,7 @@
 
 datatype parsetree =
   Node of string * parsetree list |
-  Tip of token;
+  Tip of Lexicon.token;
 
 type state =
   nt_tag * int *                (*identification and production precedence*)
@@ -675,7 +671,7 @@
 
 
 fun movedot_term (A, j, ts, Terminal a :: sa, id, i) c =
-  if valued_token c then
+  if Lexicon.valued_token c then
     (A, j, ts @ [Tip c], sa, id, i)
   else (A, j, ts, sa, id, i);
 
@@ -695,105 +691,105 @@
 (*get all productions of a NT and NTs chained to it which can
   be started by specified token*)
 fun prods_for prods chains include_none tk nts =
-let
-    fun token_assoc (list, key) =
-      let fun assoc [] result = result
-            | assoc ((keyi, pi) :: pairs) result =
-                if is_some keyi andalso matching_tokens (the keyi, key)
-                   orelse include_none andalso is_none keyi then
-                  assoc pairs (pi @ result)
-                else assoc pairs result;
-      in assoc list [] end;
+  let
+      fun token_assoc (list, key) =
+        let fun assoc [] result = result
+              | assoc ((keyi, pi) :: pairs) result =
+                  if is_some keyi andalso Lexicon.matching_tokens (the keyi, key)
+                     orelse include_none andalso is_none keyi then
+                    assoc pairs (pi @ result)
+                  else assoc pairs result;
+        in assoc list [] end;
 
-    fun get_prods [] result = result
-      | get_prods (nt :: nts) result =
-        let val nt_prods = snd (Array.sub (prods, nt));
-        in get_prods nts ((token_assoc (nt_prods, tk)) @ result) end;
-in get_prods (connected_with chains nts nts) [] end;
+      fun get_prods [] result = result
+        | get_prods (nt :: nts) result =
+          let val nt_prods = snd (Array.sub (prods, nt));
+          in get_prods nts ((token_assoc (nt_prods, tk)) @ result) end;
+  in get_prods (connected_with chains nts nts) [] end;
 
 
 fun PROCESSS warned prods chains Estate i c states =
-let
-fun all_prods_for nt = prods_for prods chains true c [nt];
+  let
+    fun all_prods_for nt = prods_for prods chains true c [nt];
 
-fun processS used [] (Si, Sii) = (Si, Sii)
-  | processS used (S :: States) (Si, Sii) =
-      (case S of
-        (_, _, _, Nonterminal (nt, minPrec) :: _, _, _) =>
-          let                                       (*predictor operation*)
-            val (used', new_states) =
-              (case AList.lookup (op =) used nt of
-                SOME (usedPrec, l) =>       (*nonterminal has been processed*)
-                  if usedPrec <= minPrec then
-                                      (*wanted precedence has been processed*)
-                    (used, movedot_lambda S l)
-                  else            (*wanted precedence hasn't been parsed yet*)
-                    let
-                      val tk_prods = all_prods_for nt;
+    fun processS used [] (Si, Sii) = (Si, Sii)
+      | processS used (S :: States) (Si, Sii) =
+          (case S of
+            (_, _, _, Nonterminal (nt, minPrec) :: _, _, _) =>
+              let                                       (*predictor operation*)
+                val (used', new_states) =
+                  (case AList.lookup (op =) used nt of
+                    SOME (usedPrec, l) =>       (*nonterminal has been processed*)
+                      if usedPrec <= minPrec then
+                                          (*wanted precedence has been processed*)
+                        (used, movedot_lambda S l)
+                      else            (*wanted precedence hasn't been parsed yet*)
+                        let
+                          val tk_prods = all_prods_for nt;
 
-                      val States' = mkStates i minPrec nt
-                                      (getRHS' minPrec usedPrec tk_prods);
-                    in (update_prec (nt, minPrec) used,
-                        movedot_lambda S l @ States')
-                    end
+                          val States' = mkStates i minPrec nt
+                                          (getRHS' minPrec usedPrec tk_prods);
+                        in (update_prec (nt, minPrec) used,
+                            movedot_lambda S l @ States')
+                        end
 
-              | NONE =>           (*nonterminal is parsed for the first time*)
-                  let val tk_prods = all_prods_for nt;
-                      val States' = mkStates i minPrec nt
-                                      (getRHS minPrec tk_prods);
-                  in ((nt, (minPrec, [])) :: used, States') end);
+                  | NONE =>           (*nonterminal is parsed for the first time*)
+                      let val tk_prods = all_prods_for nt;
+                          val States' = mkStates i minPrec nt
+                                          (getRHS minPrec tk_prods);
+                      in ((nt, (minPrec, [])) :: used, States') end);
 
-            val dummy =
-              if not (!warned) andalso
-                 length (new_states @ States) > (!branching_level) then
-                (warning "Currently parsed expression could be extremely ambiguous.";
-                 warned := true)
-              else ();
-          in
-            processS used' (new_states @ States) (S :: Si, Sii)
-          end
-      | (_, _, _, Terminal a :: _, _, _) =>               (*scanner operation*)
-          processS used States
-            (S :: Si,
-              if matching_tokens (a, c) then movedot_term S c :: Sii else Sii)
-      | (A, prec, ts, [], id, j) =>                   (*completer operation*)
-          let val tt = if id = "" then ts else [Node (id, ts)] in
-            if j = i then                             (*lambda production?*)
-              let
-                val (used', O) = update_trees used (A, (tt, prec));
+                val dummy =
+                  if not (!warned) andalso
+                     length (new_states @ States) > (!branching_level) then
+                    (warning "Currently parsed expression could be extremely ambiguous.";
+                     warned := true)
+                  else ();
               in
-                case O of
-                  NONE =>
-                    let val Slist = getS A prec Si;
-                        val States' = map (movedot_nonterm tt) Slist;
-                    in processS used' (States' @ States) (S :: Si, Sii) end
-                | SOME n =>
-                    if n >= prec then processS used' States (S :: Si, Sii)
-                    else
-                      let val Slist = getS' A prec n Si;
-                          val States' = map (movedot_nonterm tt) Slist;
-                      in processS used' (States' @ States) (S :: Si, Sii) end
+                processS used' (new_states @ States) (S :: Si, Sii)
               end
-            else
-              let val Slist = getStates Estate i j A prec
-              in processS used (map (movedot_nonterm tt) Slist @ States)
-                          (S :: Si, Sii)
-              end
-          end)
-in processS [] states ([], []) end;
+          | (_, _, _, Terminal a :: _, _, _) =>               (*scanner operation*)
+              processS used States
+                (S :: Si,
+                  if Lexicon.matching_tokens (a, c) then movedot_term S c :: Sii else Sii)
+          | (A, prec, ts, [], id, j) =>                   (*completer operation*)
+              let val tt = if id = "" then ts else [Node (id, ts)] in
+                if j = i then                             (*lambda production?*)
+                  let
+                    val (used', O) = update_trees used (A, (tt, prec));
+                  in
+                    case O of
+                      NONE =>
+                        let val Slist = getS A prec Si;
+                            val States' = map (movedot_nonterm tt) Slist;
+                        in processS used' (States' @ States) (S :: Si, Sii) end
+                    | SOME n =>
+                        if n >= prec then processS used' States (S :: Si, Sii)
+                        else
+                          let val Slist = getS' A prec n Si;
+                              val States' = map (movedot_nonterm tt) Slist;
+                          in processS used' (States' @ States) (S :: Si, Sii) end
+                  end
+                else
+                  let val Slist = getStates Estate i j A prec
+                  in processS used (map (movedot_nonterm tt) Slist @ States)
+                              (S :: Si, Sii)
+                  end
+              end)
+  in processS [] states ([], []) end;
 
 
 fun produce warned prods tags chains stateset i indata prev_token =
   (case Array.sub (stateset, i) of
     [] =>
       let
-        val toks = if is_eof prev_token then indata else prev_token :: indata;
-        val pos = Position.str_of (pos_of_token prev_token);
+        val toks = if Lexicon.is_eof prev_token then indata else prev_token :: indata;
+        val pos = Position.str_of (Lexicon.pos_of_token prev_token);
       in
         if null toks then error ("Inner syntax error: unexpected end of input" ^ pos)
         else error (Pretty.string_of (Pretty.block
           (Pretty.str ("Inner syntax error" ^ pos ^ " at \"") ::
-            Pretty.breaks (map (Pretty.str o str_of_token) (#1 (split_last toks))) @
+            Pretty.breaks (map (Pretty.str o Lexicon.str_of_token) (#1 (split_last toks))) @
             [Pretty.str "\""])))
       end
   | s =>
@@ -807,21 +803,20 @@
        end));
 
 
-fun get_trees l = map_filter (fn (_, _, [pt], _, _, _) => SOME pt | _ => NONE)
-                            l;
+fun get_trees l = map_filter (fn (_, _, [pt], _, _, _) => SOME pt | _ => NONE) l;
 
 fun earley prods tags chains startsymbol indata =
   let
-    val start_tag = case Symtab.lookup tags startsymbol of
-                       SOME tag => tag
-                     | NONE   => error ("parse: Unknown startsymbol " ^
-                                        quote startsymbol);
-    val S0 = [(~1, 0, [], [Nonterminal (start_tag, 0), Terminal eof], "", 0)];
+    val start_tag =
+      (case Symtab.lookup tags startsymbol of
+        SOME tag => tag
+      | NONE => error ("Inner syntax: unknown startsymbol " ^ quote startsymbol));
+    val S0 = [(~1, 0, [], [Nonterminal (start_tag, 0), Terminal Lexicon.eof], "", 0)];
     val s = length indata + 1;
     val Estate = Array.array (s, []);
   in
     Array.update (Estate, 0, S0);
-    get_trees (produce (Unsynchronized.ref false) prods tags chains Estate 0 indata eof)
+    get_trees (produce (Unsynchronized.ref false) prods tags chains Estate 0 indata Lexicon.eof)
   end;
 
 
@@ -830,10 +825,10 @@
     val end_pos =
       (case try List.last toks of
         NONE => Position.none
-      | SOME (Token (_, _, (_, end_pos))) => end_pos);
+      | SOME (Lexicon.Token (_, _, (_, end_pos))) => end_pos);
     val r =
-      (case earley prods tags chains start (toks @ [mk_eof end_pos]) of
-        [] => sys_error "parse: no parse trees"
+      (case earley prods tags chains start (toks @ [Lexicon.mk_eof end_pos]) of
+        [] => raise Fail "no parse trees"
       | pts => pts);
   in r end;
 
@@ -842,7 +837,8 @@
   let
     fun freeze a = map_range (curry Array.sub a) (Array.length a);
     val prods = maps snd (maps snd (freeze (#prods gram)));
-    fun guess (SOME ([Nonterminal (_, k), Terminal (Token (Literal, s, _)), Nonterminal (_, l)], _, j)) =
+    fun guess (SOME ([Nonterminal (_, k),
+            Terminal (Lexicon.Token (Lexicon.Literal, s, _)), Nonterminal (_, l)], _, j)) =
           if k = j andalso l = j + 1 then SOME (s, true, false, j)
           else if k = j + 1 then if l = j then SOME (s, false, true, j)
             else if l = j + 1 then SOME (s, false, false, j)
--- a/src/Pure/Syntax/syntax.ML	Mon Jul 05 10:42:27 2010 +0200
+++ b/src/Pure/Syntax/syntax.ML	Mon Jul 05 10:47:39 2010 +0200
@@ -297,7 +297,7 @@
     Syntax
     ({input = new_xprods @ input,
       lexicon = fold Scan.extend_lexicon (Syn_Ext.delims_of new_xprods) lexicon,
-      gram = Parser.extend_gram gram new_xprods,
+      gram = Parser.extend_gram new_xprods gram,
       consts = Library.merge (op =) (consts1, filter_out Lexicon.is_marked consts2),
       prmodes = insert (op =) mode (Library.merge (op =) (prmodes1, prmodes2)),
       parse_ast_trtab =
@@ -362,7 +362,7 @@
     Syntax
     ({input = Library.merge (op =) (input1, input2),
       lexicon = Scan.merge_lexicons (lexicon1, lexicon2),
-      gram = Parser.merge_grams gram1 gram2,
+      gram = Parser.merge_gram (gram1, gram2),
       consts = sort_distinct string_ord (consts1 @ consts2),
       prmodes = Library.merge (op =) (prmodes1, prmodes2),
       parse_ast_trtab =
--- a/src/Pure/System/isabelle_process.ML	Mon Jul 05 10:42:27 2010 +0200
+++ b/src/Pure/System/isabelle_process.ML	Mon Jul 05 10:47:39 2010 +0200
@@ -91,6 +91,7 @@
  (Unsynchronized.change print_mode
     (fold (update op =) [isabelle_processN, Keyword.keyword_status_reportN, Pretty.symbolicN]);
   setup_channels out |> init_message;
+  quick_and_dirty := true;  (* FIXME !? *)
   Keyword.report ();
   Output.status (Markup.markup Markup.ready "");
   Isar.toplevel_loop {init = true, welcome = false, sync = true, secure = true});
--- a/src/Pure/System/isabelle_process.scala	Mon Jul 05 10:42:27 2010 +0200
+++ b/src/Pure/System/isabelle_process.scala	Mon Jul 05 10:47:39 2010 +0200
@@ -19,89 +19,55 @@
 {
   /* results */
 
-  object Kind extends Enumeration {
-    //{{{ values and codes
-    // internal system notification
-    val SYSTEM = Value("SYSTEM")
-    // Posix channels/events
-    val STDIN = Value("STDIN")
-    val STDOUT = Value("STDOUT")
-    val SIGNAL = Value("SIGNAL")
-    val EXIT = Value("EXIT")
-    // Isabelle messages
-    val INIT = Value("INIT")
-    val STATUS = Value("STATUS")
-    val WRITELN = Value("WRITELN")
-    val TRACING = Value("TRACING")
-    val WARNING = Value("WARNING")
-    val ERROR = Value("ERROR")
-    val DEBUG = Value("DEBUG")
-    // messages codes
-    val code = Map(
-      ('A' : Int) -> Kind.INIT,
-      ('B' : Int) -> Kind.STATUS,
-      ('C' : Int) -> Kind.WRITELN,
-      ('D' : Int) -> Kind.TRACING,
-      ('E' : Int) -> Kind.WARNING,
-      ('F' : Int) -> Kind.ERROR,
-      ('G' : Int) -> Kind.DEBUG,
-      ('0' : Int) -> Kind.SYSTEM,
-      ('1' : Int) -> Kind.STDIN,
-      ('2' : Int) -> Kind.STDOUT,
-      ('3' : Int) -> Kind.SIGNAL,
-      ('4' : Int) -> Kind.EXIT)
+  object Kind {
     // message markup
     val markup = Map(
-      Kind.INIT -> Markup.INIT,
-      Kind.STATUS -> Markup.STATUS,
-      Kind.WRITELN -> Markup.WRITELN,
-      Kind.TRACING -> Markup.TRACING,
-      Kind.WARNING -> Markup.WARNING,
-      Kind.ERROR -> Markup.ERROR,
-      Kind.DEBUG -> Markup.DEBUG,
-      Kind.SYSTEM -> Markup.SYSTEM,
-      Kind.STDIN -> Markup.STDIN,
-      Kind.STDOUT -> Markup.STDOUT,
-      Kind.SIGNAL -> Markup.SIGNAL,
-      Kind.EXIT -> Markup.EXIT)
-    //}}}
-    def is_raw(kind: Value) =
-      kind == STDOUT
-    def is_control(kind: Value) =
-      kind == SYSTEM ||
-      kind == SIGNAL ||
-      kind == EXIT
-    def is_system(kind: Value) =
-      kind == SYSTEM ||
-      kind == STDIN ||
-      kind == SIGNAL ||
-      kind == EXIT ||
-      kind == STATUS
+      ('A' : Int) -> Markup.INIT,
+      ('B' : Int) -> Markup.STATUS,
+      ('C' : Int) -> Markup.WRITELN,
+      ('D' : Int) -> Markup.TRACING,
+      ('E' : Int) -> Markup.WARNING,
+      ('F' : Int) -> Markup.ERROR,
+      ('G' : Int) -> Markup.DEBUG)
+    def is_raw(kind: String) =
+      kind == Markup.STDOUT
+    def is_control(kind: String) =
+      kind == Markup.SYSTEM ||
+      kind == Markup.SIGNAL ||
+      kind == Markup.EXIT
+    def is_system(kind: String) =
+      kind == Markup.SYSTEM ||
+      kind == Markup.STDIN ||
+      kind == Markup.SIGNAL ||
+      kind == Markup.EXIT ||
+      kind == Markup.STATUS
   }
 
-  class Result(val kind: Kind.Value, val props: List[(String, String)], val body: List[XML.Tree])
+  class Result(val message: XML.Elem)
   {
-    def message = XML.Elem(Kind.markup(kind), props, body)
+    def kind = message.name
+    def body = message.body
+
+    def is_raw = Kind.is_raw(kind)
+    def is_control = Kind.is_control(kind)
+    def is_system = Kind.is_system(kind)
+    def is_status = kind == Markup.STATUS
+    def is_ready = is_status && body == List(XML.Elem(Markup.READY, Nil, Nil))
 
     override def toString: String =
     {
       val res =
-        if (kind == Kind.STATUS) body.map(_.toString).mkString
-        else Pretty.string_of(body)
+        if (is_status) message.body.map(_.toString).mkString
+        else Pretty.string_of(message.body)
+      val props = message.attributes
       if (props.isEmpty)
         kind.toString + " [[" + res + "]]"
       else
         kind.toString + " " +
           (for ((x, y) <- props) yield x + "=" + y).mkString("{", ",", "}") + " [[" + res + "]]"
     }
-    def is_raw = Kind.is_raw(kind)
-    def is_control = Kind.is_control(kind)
-    def is_system = Kind.is_system(kind)
 
-    def is_ready = kind == Kind.STATUS && body == List(XML.Elem(Markup.READY, Nil, Nil))
-
-    def cache(c: XML.Cache): Result =
-      new Result(kind, c.cache_props(props), c.cache_trees(body))
+    def cache(c: XML.Cache): Result = new Result(c.cache_tree(message).asInstanceOf[XML.Elem])
   }
 }
 
@@ -127,15 +93,15 @@
 
   /* results */
 
-  private def put_result(kind: Kind.Value, props: List[(String, String)], body: List[XML.Tree])
+  private def put_result(kind: String, props: List[(String, String)], body: List[XML.Tree])
   {
-    if (kind == Kind.INIT) {
+    if (kind == Markup.INIT) {
       for ((Markup.PID, p) <- props) pid = p
     }
-    receiver ! new Result(kind, props, body)
+    receiver ! new Result(XML.Elem(kind, props, body))
   }
 
-  private def put_result(kind: Kind.Value, text: String)
+  private def put_result(kind: String, text: String)
   {
     put_result(kind, Nil, List(XML.Text(system.symbols.decode(text))))
   }
@@ -145,13 +111,13 @@
 
   def interrupt() = synchronized {
     if (proc == null) error("Cannot interrupt Isabelle: no process")
-    if (pid == null) put_result(Kind.SYSTEM, "Cannot interrupt: unknown pid")
+    if (pid == null) put_result(Markup.SYSTEM, "Cannot interrupt: unknown pid")
     else {
       try {
         if (system.execute(true, "kill", "-INT", pid).waitFor == 0)
-          put_result(Kind.SIGNAL, "INT")
+          put_result(Markup.SIGNAL, "INT")
         else
-          put_result(Kind.SYSTEM, "Cannot interrupt: kill command failed")
+          put_result(Markup.SYSTEM, "Cannot interrupt: kill command failed")
       }
       catch { case e: IOException => error("Cannot interrupt Isabelle: " + e.getMessage) }
     }
@@ -162,7 +128,7 @@
     else {
       try_close()
       Thread.sleep(500)  // FIXME property!?
-      put_result(Kind.SIGNAL, "KILL")
+      put_result(Markup.SIGNAL, "KILL")
       proc.destroy
       proc = null
       pid = null
@@ -222,17 +188,17 @@
             finished = true
           }
           else {
-            put_result(Kind.STDIN, s)
+            put_result(Markup.STDIN, s)
             writer.write(s)
             writer.flush
           }
           //}}}
         }
         catch {
-          case e: IOException => put_result(Kind.SYSTEM, "Stdin thread: " + e.getMessage)
+          case e: IOException => put_result(Markup.SYSTEM, "Stdin thread: " + e.getMessage)
         }
       }
-      put_result(Kind.SYSTEM, "Stdin thread terminated")
+      put_result(Markup.SYSTEM, "Stdin thread terminated")
     }
   }
 
@@ -256,7 +222,7 @@
             else done = true
           }
           if (result.length > 0) {
-            put_result(Kind.STDOUT, result.toString)
+            put_result(Markup.STDOUT, result.toString)
             result.length = 0
           }
           else {
@@ -267,10 +233,10 @@
           //}}}
         }
         catch {
-          case e: IOException => put_result(Kind.SYSTEM, "Stdout thread: " + e.getMessage)
+          case e: IOException => put_result(Markup.SYSTEM, "Stdout thread: " + e.getMessage)
         }
       }
-      put_result(Kind.SYSTEM, "Stdout thread terminated")
+      put_result(Markup.SYSTEM, "Stdout thread terminated")
     }
   }
 
@@ -332,8 +298,8 @@
             val body = read_chunk()
             header match {
               case List(XML.Elem(name, props, Nil))
-                  if name.size == 1 && Kind.code.isDefinedAt(name(0)) =>
-                put_result(Kind.code(name(0)), props, body)
+                  if name.size == 1 && Kind.markup.isDefinedAt(name(0)) =>
+                put_result(Kind.markup(name(0)), props, body)
               case _ => throw new Protocol_Error("bad header: " + header.toString)
             }
           }
@@ -341,15 +307,15 @@
         }
         catch {
           case e: IOException =>
-            put_result(Kind.SYSTEM, "Cannot read message:\n" + e.getMessage)
+            put_result(Markup.SYSTEM, "Cannot read message:\n" + e.getMessage)
           case e: Protocol_Error =>
-            put_result(Kind.SYSTEM, "Malformed message:\n" + e.getMessage)
+            put_result(Markup.SYSTEM, "Malformed message:\n" + e.getMessage)
         }
       } while (c != -1)
       stream.close
       try_close()
 
-      put_result(Kind.SYSTEM, "Message thread terminated")
+      put_result(Markup.SYSTEM, "Message thread terminated")
     }
   }
 
@@ -392,8 +358,8 @@
       override def run() = {
         val rc = proc.waitFor()
         Thread.sleep(300)  // FIXME property!?
-        put_result(Kind.SYSTEM, "Exit thread terminated")
-        put_result(Kind.EXIT, rc.toString)
+        put_result(Markup.SYSTEM, "Exit thread terminated")
+        put_result(Markup.EXIT, rc.toString)
         rm_fifo()
       }
     }.start
--- a/src/Pure/System/session.scala	Mon Jul 05 10:42:27 2010 +0200
+++ b/src/Pure/System/session.scala	Mon Jul 05 10:47:39 2010 +0200
@@ -110,14 +110,14 @@
     {
       raw_results.event(result)
 
-      val target_id: Option[Session.Entity_ID] = Position.get_id(result.props)
+      val target_id: Option[Session.Entity_ID] = Position.get_id(result.message.attributes)
       val target: Option[Session.Entity] =
         target_id match {
           case None => None
           case Some(id) => lookup_entity(id)
         }
       if (target.isDefined) target.get.consume(result.message, indicate_command_change)
-      else if (result.kind == Isabelle_Process.Kind.STATUS) {
+      else if (result.is_status) {
         // global status message
         result.body match {
 
@@ -145,7 +145,7 @@
           case _ => if (!result.is_ready) bad_result(result)
         }
       }
-      else if (result.kind == Isabelle_Process.Kind.EXIT)
+      else if (result.kind == Markup.EXIT)
         prover = null
       else if (result.is_raw)
         raw_output.event(result)
@@ -176,7 +176,7 @@
     {
       receiveWithin(timeout) {
         case result: Isabelle_Process.Result
-          if result.kind == Isabelle_Process.Kind.INIT =>
+          if result.kind == Markup.INIT =>
           while (receive {
             case result: Isabelle_Process.Result =>
               handle_result(result); !result.is_ready
@@ -184,7 +184,7 @@
           None
 
         case result: Isabelle_Process.Result
-          if result.kind == Isabelle_Process.Kind.EXIT =>
+          if result.kind == Markup.EXIT =>
           Some(startup_error())
 
         case TIMEOUT =>  // FIXME clarify
--- a/src/Pure/goal.ML	Mon Jul 05 10:42:27 2010 +0200
+++ b/src/Pure/goal.ML	Mon Jul 05 10:47:39 2010 +0200
@@ -104,12 +104,7 @@
 
 (* parallel proofs *)
 
-fun fork e = Future.fork_pri ~1 (fn () =>
-  let
-    val _ = Output.status (Markup.markup Markup.forked "");
-    val x = e ();  (*sic*)
-    val _ = Output.status (Markup.markup Markup.joined "");
-  in x end);
+fun fork e = Future.fork_pri ~1 (fn () => Future.report e);
 
 val parallel_proofs = Unsynchronized.ref 1;
 
--- a/src/Pure/library.scala	Mon Jul 05 10:42:27 2010 +0200
+++ b/src/Pure/library.scala	Mon Jul 05 10:47:39 2010 +0200
@@ -129,11 +129,12 @@
 
   def timeit[A](message: String)(e: => A) =
   {
-    val start = System.currentTimeMillis()
+    val start = System.nanoTime()
     val result = Exn.capture(e)
-    val stop = System.currentTimeMillis()
+    val stop = System.nanoTime()
     System.err.println(
-      (if (message.isEmpty) "" else message + ": ") + (stop - start) + "ms elapsed time")
+      (if (message == null || message.isEmpty) "" else message + ": ") +
+        ((stop - start).toDouble / 1000000) + "ms elapsed time")
     Exn.release(result)
   }
 }
--- a/src/Tools/jEdit/src/jedit/document_model.scala	Mon Jul 05 10:42:27 2010 +0200
+++ b/src/Tools/jEdit/src/jedit/document_model.scala	Mon Jul 05 10:47:39 2010 +0200
@@ -95,14 +95,6 @@
   def to_current(doc: Document, offset: Int): Int =
     (offset /: changes_from(doc)) ((i, change) => change after i)
 
-  def lines_of_command(doc: Document, cmd: Command): (Int, Int) =
-  {
-    val start = doc.command_start(cmd).get  // FIXME total?
-    val stop = start + cmd.length
-    (buffer.getLineOfOffset(to_current(doc, start)),
-     buffer.getLineOfOffset(to_current(doc, stop)))
-  }
-
 
   /* text edits */
 
--- a/src/Tools/jEdit/src/jedit/document_view.scala	Mon Jul 05 10:42:27 2010 +0200
+++ b/src/Tools/jEdit/src/jedit/document_view.scala	Mon Jul 05 10:47:39 2010 +0200
@@ -104,13 +104,10 @@
       react {
         case Command_Set(changed) =>
           Swing_Thread.now {
+            // FIXME cover doc states as well!!?
             val document = model.recent_document()
-            // FIXME cover doc states as well!!?
-            for (command <- changed if document.commands.contains(command)) {
-              update_syntax(document, command)
-              invalidate_line(document, command)
-              overview.repaint()
-            }
+            if (changed.exists(document.commands.contains))
+              full_repaint(document, changed)
           }
 
         case bad => System.err.println("command_change_actor: ignoring bad message " + bad)
@@ -118,34 +115,82 @@
     }
   }
 
+  def full_repaint(document: Document, changed: Set[Command])
+  {
+    Swing_Thread.assert()
+
+    val buffer = model.buffer
+    var visible_change = false
+
+    for ((command, start) <- document.command_starts) {
+      if (changed(command)) {
+        val stop = start + command.length
+        val line1 = buffer.getLineOfOffset(model.to_current(document, start))
+        val line2 = buffer.getLineOfOffset(model.to_current(document, stop))
+        if (line2 >= text_area.getFirstLine &&
+            line1 <= text_area.getFirstLine + text_area.getVisibleLines)
+          visible_change = true
+        text_area.invalidateLineRange(line1, line2)
+      }
+    }
+    if (visible_change) model.buffer.propertiesChanged()
+
+    overview.repaint()  // FIXME paint here!?
+  }
+
 
   /* text_area_extension */
 
   private val text_area_extension = new TextAreaExtension
   {
-    override def paintValidLine(gfx: Graphics2D,
-      screen_line: Int, physical_line: Int, start: Int, end: Int, y: Int)
+    override def paintScreenLineRange(gfx: Graphics2D,
+      first_line: Int, last_line: Int, physical_lines: Array[Int],
+      start: Array[Int], end: Array[Int], y0: Int, line_height: Int)
     {
-      val document = model.recent_document()
-      def from_current(pos: Int) = model.from_current(document, pos)
-      def to_current(pos: Int) = model.to_current(document, pos)
+      Swing_Thread.now {
+        val document = model.recent_document()
+        def from_current(pos: Int) = model.from_current(document, pos)
+        def to_current(pos: Int) = model.to_current(document, pos)
 
-      val line_end = model.visible_line_end(start, end)
-      val line_height = text_area.getPainter.getFontMetrics.getHeight
+        val command_range: Iterable[(Command, Int)] =
+        {
+          val range = document.command_range(from_current(start(0)))
+          if (range.hasNext) {
+            val (cmd0, start0) = range.next
+            new Iterable[(Command, Int)] {
+              def iterator = Document.command_starts(document.commands.iterator(cmd0), start0)
+            }
+          }
+          else Iterable.empty
+        }
 
-      val saved_color = gfx.getColor
-      try {
-        for ((command, command_start) <-
-          document.command_range(from_current(start), from_current(line_end)) if !command.is_ignored)
-        {
-          val p = text_area.offsetToXY(start max to_current(command_start))
-          val q = text_area.offsetToXY(line_end min to_current(command_start + command.length))
-          assert(p.y == q.y)
-          gfx.setColor(Document_View.choose_color(document, command))
-          gfx.fillRect(p.x, y, q.x - p.x, line_height)
+        val saved_color = gfx.getColor
+        try {
+          var y = y0
+          for (i <- 0 until physical_lines.length) {
+            if (physical_lines(i) != -1) {
+              val line_start = start(i)
+              val line_end = model.visible_line_end(line_start, end(i))
+
+              val a = from_current(line_start)
+              val b = from_current(line_end)
+              val cmds = command_range.iterator.
+                dropWhile { case (cmd, c) => c + cmd.length <= a } .
+                takeWhile { case (_, c) => c < b }
+
+              for ((command, command_start) <- cmds if !command.is_ignored) {
+                val p = text_area.offsetToXY(line_start max to_current(command_start))
+                val q = text_area.offsetToXY(line_end min to_current(command_start + command.length))
+                assert(p.y == q.y)
+                gfx.setColor(Document_View.choose_color(document, command))
+                gfx.fillRect(p.x, y, q.x - p.x, line_height)
+              }
+            }
+            y += line_height
+          }
         }
+        finally { gfx.setColor(saved_color) }
       }
-      finally { gfx.setColor(saved_color) }
     }
 
     override def getToolTipText(x: Int, y: Int): String =
@@ -186,30 +231,6 @@
   }
 
 
-  /* (re)painting */
-
-  private val update_delay = Swing_Thread.delay_first(500) { model.buffer.propertiesChanged() }
-  // FIXME update_delay property
-
-  private def update_syntax(document: Document, cmd: Command)
-  {
-    val (line1, line2) = model.lines_of_command(document, cmd)
-    if (line2 >= text_area.getFirstLine &&
-      line1 <= text_area.getFirstLine + text_area.getVisibleLines)
-        update_delay()
-  }
-
-  private def invalidate_line(document: Document, cmd: Command) =
-  {
-    val (start, stop) = model.lines_of_command(document, cmd)
-    text_area.invalidateLineRange(start, stop)
-  }
-
-  private def invalidate_all() =
-    text_area.invalidateLineRange(text_area.getFirstPhysicalLine,
-      text_area.getLastPhysicalLine)
-
-
   /* overview of command status left of scrollbar */
 
   private val overview = new JPanel(new BorderLayout)
@@ -252,9 +273,9 @@
       val buffer = model.buffer
       val document = model.recent_document()
       def to_current(pos: Int) = model.to_current(document, pos)
-      val saved_color = gfx.getColor
+      val saved_color = gfx.getColor  // FIXME needed!?
       try {
-        for ((command, start) <- document.command_range(0) if !command.is_ignored) {
+        for ((command, start) <- document.command_starts if !command.is_ignored) {
           val line1 = buffer.getLineOfOffset(to_current(start))
           val line2 = buffer.getLineOfOffset(to_current(start + command.length)) + 1
           val y = line_to_y(line1)